merged
authorwenzelm
Wed, 16 Dec 2015 17:30:30 +0100
changeset 61855 32a530a5c54c
parent 61849 f8741f200f91 (current diff)
parent 61854 38b049cd3aad (diff)
child 61856 4b1b85f38944
child 61862 e2a9e46ac0fb
merged
--- a/src/Doc/Eisbach/Manual.thy	Tue Dec 15 14:41:47 2015 +0000
+++ b/src/Doc/Eisbach/Manual.thy	Wed Dec 16 17:30:30 2015 +0100
@@ -931,7 +931,7 @@
 
     attribute_setup get_split_rule =
       \<open>Args.term >> (fn t =>
-        Thm.rule_attribute (fn context => fn _ =>
+        Thm.rule_attribute [] (fn context => fn _ =>
           (case get_split_rule (Context.proof_of context) t of
             SOME thm => thm
           | NONE => Drule.dummy_thm)))\<close>
--- a/src/Doc/Implementation/Eq.thy	Tue Dec 15 14:41:47 2015 +0000
+++ b/src/Doc/Implementation/Eq.thy	Wed Dec 16 17:30:30 2015 +0100
@@ -6,52 +6,52 @@
 
 chapter \<open>Equational reasoning\<close>
 
-text \<open>Equality is one of the most fundamental concepts of
-  mathematics.  The Isabelle/Pure logic (\chref{ch:logic}) provides a
-  builtin relation \<open>\<equiv> :: \<alpha> \<Rightarrow> \<alpha> \<Rightarrow> prop\<close> that expresses equality
-  of arbitrary terms (or propositions) at the framework level, as
-  expressed by certain basic inference rules (\secref{sec:eq-rules}).
+text \<open>
+  Equality is one of the most fundamental concepts of mathematics. The
+  Isabelle/Pure logic (\chref{ch:logic}) provides a builtin relation \<open>\<equiv> :: \<alpha> \<Rightarrow>
+  \<alpha> \<Rightarrow> prop\<close> that expresses equality of arbitrary terms (or propositions) at
+  the framework level, as expressed by certain basic inference rules
+  (\secref{sec:eq-rules}).
 
-  Equational reasoning means to replace equals by equals, using
-  reflexivity and transitivity to form chains of replacement steps,
-  and congruence rules to access sub-structures.  Conversions
-  (\secref{sec:conv}) provide a convenient framework to compose basic
-  equational steps to build specific equational reasoning tools.
+  Equational reasoning means to replace equals by equals, using reflexivity
+  and transitivity to form chains of replacement steps, and congruence rules
+  to access sub-structures. Conversions (\secref{sec:conv}) provide a
+  convenient framework to compose basic equational steps to build specific
+  equational reasoning tools.
 
-  Higher-order matching is able to provide suitable instantiations for
-  giving equality rules, which leads to the versatile concept of
-  \<open>\<lambda>\<close>-term rewriting (\secref{sec:rewriting}).  Internally
-  this is based on the general-purpose Simplifier engine of Isabelle,
-  which is more specific and more efficient than plain conversions.
+  Higher-order matching is able to provide suitable instantiations for giving
+  equality rules, which leads to the versatile concept of \<open>\<lambda>\<close>-term rewriting
+  (\secref{sec:rewriting}). Internally this is based on the general-purpose
+  Simplifier engine of Isabelle, which is more specific and more efficient
+  than plain conversions.
 
-  Object-logics usually introduce specific notions of equality or
-  equivalence, and relate it with the Pure equality.  This enables to
-  re-use the Pure tools for equational reasoning for particular
-  object-logic connectives as well.
+  Object-logics usually introduce specific notions of equality or equivalence,
+  and relate it with the Pure equality. This enables to re-use the Pure tools
+  for equational reasoning for particular object-logic connectives as well.
 \<close>
 
 
 section \<open>Basic equality rules \label{sec:eq-rules}\<close>
 
-text \<open>Isabelle/Pure uses \<open>\<equiv>\<close> for equality of arbitrary
-  terms, which includes equivalence of propositions of the logical
-  framework.  The conceptual axiomatization of the constant \<open>\<equiv>
-  :: \<alpha> \<Rightarrow> \<alpha> \<Rightarrow> prop\<close> is given in \figref{fig:pure-equality}.  The
-  inference kernel presents slightly different equality rules, which
-  may be understood as derived rules from this minimal axiomatization.
-  The Pure theory also provides some theorems that express the same
-  reasoning schemes as theorems that can be composed like object-level
+text \<open>
+  Isabelle/Pure uses \<open>\<equiv>\<close> for equality of arbitrary terms, which includes
+  equivalence of propositions of the logical framework. The conceptual
+  axiomatization of the constant \<open>\<equiv> :: \<alpha> \<Rightarrow> \<alpha> \<Rightarrow> prop\<close> is given in
+  \figref{fig:pure-equality}. The inference kernel presents slightly different
+  equality rules, which may be understood as derived rules from this minimal
+  axiomatization. The Pure theory also provides some theorems that express the
+  same reasoning schemes as theorems that can be composed like object-level
   rules as explained in \secref{sec:obj-rules}.
 
-  For example, @{ML Thm.symmetric} as Pure inference is an ML function
-  that maps a theorem \<open>th\<close> stating \<open>t \<equiv> u\<close> to one
-  stating \<open>u \<equiv> t\<close>.  In contrast, @{thm [source]
-  Pure.symmetric} as Pure theorem expresses the same reasoning in
-  declarative form.  If used like \<open>th [THEN Pure.symmetric]\<close>
-  in Isar source notation, it achieves a similar effect as the ML
-  inference function, although the rule attribute @{attribute THEN} or
-  ML operator @{ML "op RS"} involve the full machinery of higher-order
-  unification (modulo \<open>\<beta>\<eta>\<close>-conversion) and lifting of \<open>\<And>/\<Longrightarrow>\<close> contexts.\<close>
+  For example, @{ML Thm.symmetric} as Pure inference is an ML function that
+  maps a theorem \<open>th\<close> stating \<open>t \<equiv> u\<close> to one stating \<open>u \<equiv> t\<close>. In contrast,
+  @{thm [source] Pure.symmetric} as Pure theorem expresses the same reasoning
+  in declarative form. If used like \<open>th [THEN Pure.symmetric]\<close> in Isar source
+  notation, it achieves a similar effect as the ML inference function,
+  although the rule attribute @{attribute THEN} or ML operator @{ML "op RS"}
+  involve the full machinery of higher-order unification (modulo
+  \<open>\<beta>\<eta>\<close>-conversion) and lifting of \<open>\<And>/\<Longrightarrow>\<close> contexts.
+\<close>
 
 text %mlref \<open>
   \begin{mldecls}
@@ -64,11 +64,11 @@
   @{index_ML Thm.equal_elim: "thm -> thm -> thm"} \\
   \end{mldecls}
 
-  See also @{file "~~/src/Pure/thm.ML" } for further description of
-  these inference rules, and a few more for primitive \<open>\<beta>\<close> and
-  \<open>\<eta>\<close> conversions.  Note that \<open>\<alpha>\<close> conversion is
-  implicit due to the representation of terms with de-Bruijn indices
-  (\secref{sec:terms}).\<close>
+  See also @{file "~~/src/Pure/thm.ML" } for further description of these
+  inference rules, and a few more for primitive \<open>\<beta>\<close> and \<open>\<eta>\<close> conversions. Note
+  that \<open>\<alpha>\<close> conversion is implicit due to the representation of terms with
+  de-Bruijn indices (\secref{sec:terms}).
+\<close>
 
 
 section \<open>Conversions \label{sec:conv}\<close>
@@ -76,19 +76,19 @@
 text \<open>
   %FIXME
 
-  The classic article that introduces the concept of conversion (for
-  Cambridge LCF) is @{cite "paulson:1983"}.
+  The classic article that introduces the concept of conversion (for Cambridge
+  LCF) is @{cite "paulson:1983"}.
 \<close>
 
 
 section \<open>Rewriting \label{sec:rewriting}\<close>
 
-text \<open>Rewriting normalizes a given term (theorem or goal) by
-  replacing instances of given equalities \<open>t \<equiv> u\<close> in subterms.
-  Rewriting continues until no rewrites are applicable to any subterm.
-  This may be used to unfold simple definitions of the form \<open>f
-  x\<^sub>1 \<dots> x\<^sub>n \<equiv> u\<close>, but is slightly more general than that.
-\<close>
+text \<open>
+  Rewriting normalizes a given term (theorem or goal) by replacing instances
+  of given equalities \<open>t \<equiv> u\<close> in subterms. Rewriting continues until no
+  rewrites are applicable to any subterm. This may be used to unfold simple
+  definitions of the form \<open>f x\<^sub>1 \<dots> x\<^sub>n \<equiv> u\<close>, but is slightly more general than
+  that. \<close>
 
 text %mlref \<open>
   \begin{mldecls}
@@ -99,24 +99,24 @@
   @{index_ML fold_goals_tac: "Proof.context -> thm list -> tactic"} \\
   \end{mldecls}
 
-  \<^descr> @{ML rewrite_rule}~\<open>ctxt rules thm\<close> rewrites the whole
-  theorem by the given rules.
+  \<^descr> @{ML rewrite_rule}~\<open>ctxt rules thm\<close> rewrites the whole theorem by the
+  given rules.
 
-  \<^descr> @{ML rewrite_goals_rule}~\<open>ctxt rules thm\<close> rewrites the
-  outer premises of the given theorem.  Interpreting the same as a
-  goal state (\secref{sec:tactical-goals}) it means to rewrite all
-  subgoals (in the same manner as @{ML rewrite_goals_tac}).
+  \<^descr> @{ML rewrite_goals_rule}~\<open>ctxt rules thm\<close> rewrites the outer premises of
+  the given theorem. Interpreting the same as a goal state
+  (\secref{sec:tactical-goals}) it means to rewrite all subgoals (in the same
+  manner as @{ML rewrite_goals_tac}).
 
-  \<^descr> @{ML rewrite_goal_tac}~\<open>ctxt rules i\<close> rewrites subgoal
-  \<open>i\<close> by the given rewrite rules.
+  \<^descr> @{ML rewrite_goal_tac}~\<open>ctxt rules i\<close> rewrites subgoal \<open>i\<close> by the given
+  rewrite rules.
 
-  \<^descr> @{ML rewrite_goals_tac}~\<open>ctxt rules\<close> rewrites all subgoals
-  by the given rewrite rules.
+  \<^descr> @{ML rewrite_goals_tac}~\<open>ctxt rules\<close> rewrites all subgoals by the given
+  rewrite rules.
 
-  \<^descr> @{ML fold_goals_tac}~\<open>ctxt rules\<close> essentially uses @{ML
-  rewrite_goals_tac} with the symmetric form of each member of \<open>rules\<close>, re-ordered to fold longer expression first.  This supports
-  to idea to fold primitive definitions that appear in expended form
-  in the proof state.
+  \<^descr> @{ML fold_goals_tac}~\<open>ctxt rules\<close> essentially uses @{ML rewrite_goals_tac}
+  with the symmetric form of each member of \<open>rules\<close>, re-ordered to fold longer
+  expression first. This supports to idea to fold primitive definitions that
+  appear in expended form in the proof state.
 \<close>
 
 end
--- a/src/Doc/Implementation/Integration.thy	Tue Dec 15 14:41:47 2015 +0000
+++ b/src/Doc/Implementation/Integration.thy	Wed Dec 16 17:30:30 2015 +0100
@@ -9,8 +9,8 @@
 section \<open>Isar toplevel \label{sec:isar-toplevel}\<close>
 
 text \<open>
-  The Isar \<^emph>\<open>toplevel state\<close> represents the outermost configuration that
-  is transformed by a sequence of transitions (commands) within a theory body.
+  The Isar \<^emph>\<open>toplevel state\<close> represents the outermost configuration that is
+  transformed by a sequence of transitions (commands) within a theory body.
   This is a pure value with pure functions acting on it in a timeless and
   stateless manner. Historically, the sequence of transitions was wrapped up
   as sequential command loop, such that commands are applied one-by-one. In
@@ -23,15 +23,16 @@
 subsection \<open>Toplevel state\<close>
 
 text \<open>
-  The toplevel state is a disjoint sum of empty \<open>toplevel\<close>, or \<open>theory\<close>, or \<open>proof\<close>. The initial toplevel is empty; a theory is
-  commenced by a @{command theory} header; within a theory we may use theory
-  commands such as @{command definition}, or state a @{command theorem} to be
-  proven. A proof state accepts a rich collection of Isar proof commands for
-  structured proof composition, or unstructured proof scripts. When the proof
-  is concluded we get back to the (local) theory, which is then updated by
-  defining the resulting fact. Further theory declarations or theorem
-  statements with proofs may follow, until we eventually conclude the theory
-  development by issuing @{command end} to get back to the empty toplevel.
+  The toplevel state is a disjoint sum of empty \<open>toplevel\<close>, or \<open>theory\<close>, or
+  \<open>proof\<close>. The initial toplevel is empty; a theory is commenced by a @{command
+  theory} header; within a theory we may use theory commands such as @{command
+  definition}, or state a @{command theorem} to be proven. A proof state
+  accepts a rich collection of Isar proof commands for structured proof
+  composition, or unstructured proof scripts. When the proof is concluded we
+  get back to the (local) theory, which is then updated by defining the
+  resulting fact. Further theory declarations or theorem statements with
+  proofs may follow, until we eventually conclude the theory development by
+  issuing @{command end} to get back to the empty toplevel.
 \<close>
 
 text %mlref \<open>
@@ -43,23 +44,21 @@
   @{index_ML Toplevel.proof_of: "Toplevel.state -> Proof.state"} \\
   \end{mldecls}
 
-  \<^descr> Type @{ML_type Toplevel.state} represents Isar toplevel
-  states, which are normally manipulated through the concept of
-  toplevel transitions only (\secref{sec:toplevel-transition}).
+  \<^descr> Type @{ML_type Toplevel.state} represents Isar toplevel states, which are
+  normally manipulated through the concept of toplevel transitions only
+  (\secref{sec:toplevel-transition}).
 
-  \<^descr> @{ML Toplevel.UNDEF} is raised for undefined toplevel
-  operations.  Many operations work only partially for certain cases,
-  since @{ML_type Toplevel.state} is a sum type.
+  \<^descr> @{ML Toplevel.UNDEF} is raised for undefined toplevel operations. Many
+  operations work only partially for certain cases, since @{ML_type
+  Toplevel.state} is a sum type.
 
-  \<^descr> @{ML Toplevel.is_toplevel}~\<open>state\<close> checks for an empty
-  toplevel state.
+  \<^descr> @{ML Toplevel.is_toplevel}~\<open>state\<close> checks for an empty toplevel state.
 
-  \<^descr> @{ML Toplevel.theory_of}~\<open>state\<close> selects the
-  background theory of \<open>state\<close>, it raises @{ML Toplevel.UNDEF}
-  for an empty toplevel state.
+  \<^descr> @{ML Toplevel.theory_of}~\<open>state\<close> selects the background theory of \<open>state\<close>,
+  it raises @{ML Toplevel.UNDEF} for an empty toplevel state.
 
-  \<^descr> @{ML Toplevel.proof_of}~\<open>state\<close> selects the Isar proof
-  state if available, otherwise it raises an error.
+  \<^descr> @{ML Toplevel.proof_of}~\<open>state\<close> selects the Isar proof state if available,
+  otherwise it raises an error.
 \<close>
 
 text %mlantiq \<open>
@@ -67,11 +66,11 @@
   @{ML_antiquotation_def "Isar.state"} & : & \<open>ML_antiquotation\<close> \\
   \end{matharray}
 
-  \<^descr> \<open>@{Isar.state}\<close> refers to Isar toplevel state at that
-  point --- as abstract value.
+  \<^descr> \<open>@{Isar.state}\<close> refers to Isar toplevel state at that point --- as
+  abstract value.
 
-  This only works for diagnostic ML commands, such as @{command
-  ML_val} or @{command ML_command}.
+  This only works for diagnostic ML commands, such as @{command ML_val} or
+  @{command ML_command}.
 \<close>
 
 
@@ -82,12 +81,11 @@
   state, with additional information for diagnostics and error reporting:
   there are fields for command name, source position, and other meta-data.
 
-  The operational part is represented as the sequential union of a
-  list of partial functions, which are tried in turn until the first
-  one succeeds.  This acts like an outer case-expression for various
-  alternative state transitions.  For example, \isakeyword{qed} works
-  differently for a local proofs vs.\ the global ending of an outermost
-  proof.
+  The operational part is represented as the sequential union of a list of
+  partial functions, which are tried in turn until the first one succeeds.
+  This acts like an outer case-expression for various alternative state
+  transitions. For example, \isakeyword{qed} works differently for a local
+  proofs vs.\ the global ending of an outermost proof.
 
   Transitions are composed via transition transformers. Internally, Isar
   commands are put together from an empty transition extended by name and
@@ -112,34 +110,32 @@
   Toplevel.transition -> Toplevel.transition"} \\
   \end{mldecls}
 
-  \<^descr> @{ML Toplevel.keep}~\<open>tr\<close> adjoins a diagnostic
-  function.
+  \<^descr> @{ML Toplevel.keep}~\<open>tr\<close> adjoins a diagnostic function.
+
+  \<^descr> @{ML Toplevel.theory}~\<open>tr\<close> adjoins a theory transformer.
 
-  \<^descr> @{ML Toplevel.theory}~\<open>tr\<close> adjoins a theory
-  transformer.
+  \<^descr> @{ML Toplevel.theory_to_proof}~\<open>tr\<close> adjoins a global goal function, which
+  turns a theory into a proof state. The theory may be changed before entering
+  the proof; the generic Isar goal setup includes an \<^verbatim>\<open>after_qed\<close> argument
+  that specifies how to apply the proven result to the enclosing context, when
+  the proof is finished.
 
-  \<^descr> @{ML Toplevel.theory_to_proof}~\<open>tr\<close> adjoins a global
-  goal function, which turns a theory into a proof state.  The theory
-  may be changed before entering the proof; the generic Isar goal
-  setup includes an \<^verbatim>\<open>after_qed\<close> argument that specifies how to
-  apply the proven result to the enclosing context, when the proof
-  is finished.
+  \<^descr> @{ML Toplevel.proof}~\<open>tr\<close> adjoins a deterministic proof command, with a
+  singleton result.
 
-  \<^descr> @{ML Toplevel.proof}~\<open>tr\<close> adjoins a deterministic
-  proof command, with a singleton result.
+  \<^descr> @{ML Toplevel.proofs}~\<open>tr\<close> adjoins a general proof command, with zero or
+  more result states (represented as a lazy list).
 
-  \<^descr> @{ML Toplevel.proofs}~\<open>tr\<close> adjoins a general proof
-  command, with zero or more result states (represented as a lazy
-  list).
-
-  \<^descr> @{ML Toplevel.end_proof}~\<open>tr\<close> adjoins a concluding
-  proof command, that returns the resulting theory, after applying the
-  resulting facts to the target context.
+  \<^descr> @{ML Toplevel.end_proof}~\<open>tr\<close> adjoins a concluding proof command, that
+  returns the resulting theory, after applying the resulting facts to the
+  target context.
 \<close>
 
-text %mlex \<open>The file @{"file" "~~/src/HOL/ex/Commands.thy"} shows some example
-Isar command definitions, with the all-important theory header declarations
-for outer syntax keywords.\<close>
+text %mlex \<open>
+  The file @{"file" "~~/src/HOL/ex/Commands.thy"} shows some example Isar
+  command definitions, with the all-important theory header declarations for
+  outer syntax keywords.
+\<close>
 
 
 section \<open>Theory loader database\<close>
@@ -149,8 +145,8 @@
   a collection of theories as a directed acyclic graph. A theory may refer to
   other theories as @{keyword "imports"}, or to auxiliary files via special
   \<^emph>\<open>load commands\<close> (e.g.\ @{command ML_file}). For each theory, the base
-  directory of its own theory file is called \<^emph>\<open>master directory\<close>: this is
-  used as the relative location to refer to other files from that theory.
+  directory of its own theory file is called \<^emph>\<open>master directory\<close>: this is used
+  as the relative location to refer to other files from that theory.
 \<close>
 
 text %mlref \<open>
@@ -162,28 +158,27 @@
   @{index_ML Thy_Info.register_thy: "theory -> unit"} \\
   \end{mldecls}
 
-  \<^descr> @{ML use_thy}~\<open>A\<close> ensures that theory \<open>A\<close> is fully
-  up-to-date wrt.\ the external file store; outdated ancestors are reloaded on
-  demand.
+  \<^descr> @{ML use_thy}~\<open>A\<close> ensures that theory \<open>A\<close> is fully up-to-date wrt.\ the
+  external file store; outdated ancestors are reloaded on demand.
 
-  \<^descr> @{ML use_thys} is similar to @{ML use_thy}, but handles several
-  theories simultaneously. Thus it acts like processing the import header of a
-  theory, without performing the merge of the result. By loading a whole
-  sub-graph of theories, the intrinsic parallelism can be exploited by the
-  system to speedup loading.
+  \<^descr> @{ML use_thys} is similar to @{ML use_thy}, but handles several theories
+  simultaneously. Thus it acts like processing the import header of a theory,
+  without performing the merge of the result. By loading a whole sub-graph of
+  theories, the intrinsic parallelism can be exploited by the system to
+  speedup loading.
 
   This variant is used by default in @{tool build} @{cite "isabelle-system"}.
 
-  \<^descr> @{ML Thy_Info.get_theory}~\<open>A\<close> retrieves the theory value
-  presently associated with name \<open>A\<close>. Note that the result might be
-  outdated wrt.\ the file-system content.
+  \<^descr> @{ML Thy_Info.get_theory}~\<open>A\<close> retrieves the theory value presently
+  associated with name \<open>A\<close>. Note that the result might be outdated wrt.\ the
+  file-system content.
 
-  \<^descr> @{ML Thy_Info.remove_thy}~\<open>A\<close> deletes theory \<open>A\<close> and all
-  descendants from the theory database.
+  \<^descr> @{ML Thy_Info.remove_thy}~\<open>A\<close> deletes theory \<open>A\<close> and all descendants from
+  the theory database.
 
-  \<^descr> @{ML Thy_Info.register_thy}~\<open>text thy\<close> registers an existing
-  theory value with the theory loader database and updates source version
-  information according to the file store.
+  \<^descr> @{ML Thy_Info.register_thy}~\<open>text thy\<close> registers an existing theory value
+  with the theory loader database and updates source version information
+  according to the file store.
 \<close>
 
 end
--- a/src/Doc/Implementation/Isar.thy	Tue Dec 15 14:41:47 2015 +0000
+++ b/src/Doc/Implementation/Isar.thy	Wed Dec 16 17:30:30 2015 +0100
@@ -6,60 +6,57 @@
 
 chapter \<open>Isar language elements\<close>
 
-text \<open>The Isar proof language (see also
-  @{cite \<open>\S2\<close> "isabelle-isar-ref"}) consists of three main categories of
-  language elements:
+text \<open>
+  The Isar proof language (see also @{cite \<open>\S2\<close> "isabelle-isar-ref"})
+  consists of three main categories of language elements:
 
-  \<^enum> Proof \<^emph>\<open>commands\<close> define the primary language of
-  transactions of the underlying Isar/VM interpreter.  Typical
-  examples are @{command "fix"}, @{command "assume"}, @{command
-  "show"}, @{command "proof"}, and @{command "qed"}.
+  \<^enum> Proof \<^emph>\<open>commands\<close> define the primary language of transactions of the
+  underlying Isar/VM interpreter. Typical examples are @{command "fix"},
+  @{command "assume"}, @{command "show"}, @{command "proof"}, and @{command
+  "qed"}.
 
-  Composing proof commands according to the rules of the Isar/VM leads
-  to expressions of structured proof text, such that both the machine
-  and the human reader can give it a meaning as formal reasoning.
+  Composing proof commands according to the rules of the Isar/VM leads to
+  expressions of structured proof text, such that both the machine and the
+  human reader can give it a meaning as formal reasoning.
 
-  \<^enum> Proof \<^emph>\<open>methods\<close> define a secondary language of mixed
-  forward-backward refinement steps involving facts and goals.
-  Typical examples are @{method rule}, @{method unfold}, and @{method
-  simp}.
+  \<^enum> Proof \<^emph>\<open>methods\<close> define a secondary language of mixed forward-backward
+  refinement steps involving facts and goals. Typical examples are @{method
+  rule}, @{method unfold}, and @{method simp}.
 
-  Methods can occur in certain well-defined parts of the Isar proof
-  language, say as arguments to @{command "proof"}, @{command "qed"},
-  or @{command "by"}.
+  Methods can occur in certain well-defined parts of the Isar proof language,
+  say as arguments to @{command "proof"}, @{command "qed"}, or @{command
+  "by"}.
 
-  \<^enum> \<^emph>\<open>Attributes\<close> define a tertiary language of small
-  annotations to theorems being defined or referenced.  Attributes can
-  modify both the context and the theorem.
+  \<^enum> \<^emph>\<open>Attributes\<close> define a tertiary language of small annotations to theorems
+  being defined or referenced. Attributes can modify both the context and the
+  theorem.
 
-  Typical examples are @{attribute intro} (which affects the context),
-  and @{attribute symmetric} (which affects the theorem).
+  Typical examples are @{attribute intro} (which affects the context), and
+  @{attribute symmetric} (which affects the theorem).
 \<close>
 
 
 section \<open>Proof commands\<close>
 
-text \<open>A \<^emph>\<open>proof command\<close> is state transition of the Isar/VM
-  proof interpreter.
+text \<open>
+  A \<^emph>\<open>proof command\<close> is state transition of the Isar/VM proof interpreter.
 
-  In principle, Isar proof commands could be defined in user-space as
-  well.  The system is built like that in the first place: one part of
-  the commands are primitive, the other part is defined as derived
-  elements.  Adding to the genuine structured proof language requires
-  profound understanding of the Isar/VM machinery, though, so this is
-  beyond the scope of this manual.
+  In principle, Isar proof commands could be defined in user-space as well.
+  The system is built like that in the first place: one part of the commands
+  are primitive, the other part is defined as derived elements. Adding to the
+  genuine structured proof language requires profound understanding of the
+  Isar/VM machinery, though, so this is beyond the scope of this manual.
 
-  What can be done realistically is to define some diagnostic commands
-  that inspect the general state of the Isar/VM, and report some
-  feedback to the user.  Typically this involves checking of the
-  linguistic \<^emph>\<open>mode\<close> of a proof state, or peeking at the pending
-  goals (if available).
+  What can be done realistically is to define some diagnostic commands that
+  inspect the general state of the Isar/VM, and report some feedback to the
+  user. Typically this involves checking of the linguistic \<^emph>\<open>mode\<close> of a proof
+  state, or peeking at the pending goals (if available).
 
-  Another common application is to define a toplevel command that
-  poses a problem to the user as Isar proof state and processes the
-  final result relatively to the context.  Thus a proof can be
-  incorporated into the context of some user-space tool, without
-  modifying the Isar proof language itself.\<close>
+  Another common application is to define a toplevel command that poses a
+  problem to the user as Isar proof state and processes the final result
+  relatively to the context. Thus a proof can be incorporated into the context
+  of some user-space tool, without modifying the Isar proof language itself.
+\<close>
 
 text %mlref \<open>
   \begin{mldecls}
@@ -77,61 +74,56 @@
   (term * term list) list list -> Proof.context -> Proof.state"} \\
   \end{mldecls}
 
-  \<^descr> Type @{ML_type Proof.state} represents Isar proof states.
-  This is a block-structured configuration with proof context,
-  linguistic mode, and optional goal.  The latter consists of goal
-  context, goal facts (``\<open>using\<close>''), and tactical goal state
-  (see \secref{sec:tactical-goals}).
+  \<^descr> Type @{ML_type Proof.state} represents Isar proof states. This is a
+  block-structured configuration with proof context, linguistic mode, and
+  optional goal. The latter consists of goal context, goal facts
+  (``\<open>using\<close>''), and tactical goal state (see \secref{sec:tactical-goals}).
 
-  The general idea is that the facts shall contribute to the
-  refinement of some parts of the tactical goal --- how exactly is
-  defined by the proof method that is applied in that situation.
+  The general idea is that the facts shall contribute to the refinement of
+  some parts of the tactical goal --- how exactly is defined by the proof
+  method that is applied in that situation.
 
   \<^descr> @{ML Proof.assert_forward}, @{ML Proof.assert_chain}, @{ML
-  Proof.assert_backward} are partial identity functions that fail
-  unless a certain linguistic mode is active, namely ``\<open>proof(state)\<close>'', ``\<open>proof(chain)\<close>'', ``\<open>proof(prove)\<close>'', respectively (using the terminology of
-  @{cite "isabelle-isar-ref"}).
+  Proof.assert_backward} are partial identity functions that fail unless a
+  certain linguistic mode is active, namely ``\<open>proof(state)\<close>'',
+  ``\<open>proof(chain)\<close>'', ``\<open>proof(prove)\<close>'', respectively (using the terminology
+  of @{cite "isabelle-isar-ref"}).
 
-  It is advisable study the implementations of existing proof commands
-  for suitable modes to be asserted.
+  It is advisable study the implementations of existing proof commands for
+  suitable modes to be asserted.
 
-  \<^descr> @{ML Proof.simple_goal}~\<open>state\<close> returns the structured
-  Isar goal (if available) in the form seen by ``simple'' methods
-  (like @{method simp} or @{method blast}).  The Isar goal facts are
-  already inserted as premises into the subgoals, which are presented
-  individually as in @{ML Proof.goal}.
+  \<^descr> @{ML Proof.simple_goal}~\<open>state\<close> returns the structured Isar goal (if
+  available) in the form seen by ``simple'' methods (like @{method simp} or
+  @{method blast}). The Isar goal facts are already inserted as premises into
+  the subgoals, which are presented individually as in @{ML Proof.goal}.
 
-  \<^descr> @{ML Proof.goal}~\<open>state\<close> returns the structured Isar
-  goal (if available) in the form seen by regular methods (like
-  @{method rule}).  The auxiliary internal encoding of Pure
-  conjunctions is split into individual subgoals as usual.
+  \<^descr> @{ML Proof.goal}~\<open>state\<close> returns the structured Isar goal (if available)
+  in the form seen by regular methods (like @{method rule}). The auxiliary
+  internal encoding of Pure conjunctions is split into individual subgoals as
+  usual.
 
-  \<^descr> @{ML Proof.raw_goal}~\<open>state\<close> returns the structured
-  Isar goal (if available) in the raw internal form seen by ``raw''
-  methods (like @{method induct}).  This form is rarely appropriate
-  for diagnostic tools; @{ML Proof.simple_goal} or @{ML Proof.goal}
-  should be used in most situations.
+  \<^descr> @{ML Proof.raw_goal}~\<open>state\<close> returns the structured Isar goal (if
+  available) in the raw internal form seen by ``raw'' methods (like @{method
+  induct}). This form is rarely appropriate for diagnostic tools; @{ML
+  Proof.simple_goal} or @{ML Proof.goal} should be used in most situations.
 
-  \<^descr> @{ML Proof.theorem}~\<open>before_qed after_qed statement ctxt\<close>
-  initializes a toplevel Isar proof state within a given context.
+  \<^descr> @{ML Proof.theorem}~\<open>before_qed after_qed statement ctxt\<close> initializes a
+  toplevel Isar proof state within a given context.
 
-  The optional \<open>before_qed\<close> method is applied at the end of
-  the proof, just before extracting the result (this feature is rarely
-  used).
+  The optional \<open>before_qed\<close> method is applied at the end of the proof, just
+  before extracting the result (this feature is rarely used).
 
-  The \<open>after_qed\<close> continuation receives the extracted result
-  in order to apply it to the final context in a suitable way (e.g.\
-  storing named facts).  Note that at this generic level the target
-  context is specified as @{ML_type Proof.context}, but the usual
-  wrapping of toplevel proofs into command transactions will provide a
-  @{ML_type local_theory} here (\chref{ch:local-theory}).  This
-  affects the way how results are stored.
+  The \<open>after_qed\<close> continuation receives the extracted result in order to apply
+  it to the final context in a suitable way (e.g.\ storing named facts). Note
+  that at this generic level the target context is specified as @{ML_type
+  Proof.context}, but the usual wrapping of toplevel proofs into command
+  transactions will provide a @{ML_type local_theory} here
+  (\chref{ch:local-theory}). This affects the way how results are stored.
 
-  The \<open>statement\<close> is given as a nested list of terms, each
-  associated with optional @{keyword "is"} patterns as usual in the
-  Isar source language.  The original nested list structure over terms
-  is turned into one over theorems when \<open>after_qed\<close> is
-  invoked.
+  The \<open>statement\<close> is given as a nested list of terms, each associated with
+  optional @{keyword "is"} patterns as usual in the Isar source language. The
+  original nested list structure over terms is turned into one over theorems
+  when \<open>after_qed\<close> is invoked.
 \<close>
 
 
@@ -140,15 +132,16 @@
   @{ML_antiquotation_def "Isar.goal"} & : & \<open>ML_antiquotation\<close> \\
   \end{matharray}
 
-  \<^descr> \<open>@{Isar.goal}\<close> refers to the regular goal state (if
-  available) of the current proof state managed by the Isar toplevel
-  --- as abstract value.
+  \<^descr> \<open>@{Isar.goal}\<close> refers to the regular goal state (if available) of the
+  current proof state managed by the Isar toplevel --- as abstract value.
 
-  This only works for diagnostic ML commands, such as @{command
-  ML_val} or @{command ML_command}.
+  This only works for diagnostic ML commands, such as @{command ML_val} or
+  @{command ML_command}.
 \<close>
 
-text %mlex \<open>The following example peeks at a certain goal configuration.\<close>
+text %mlex \<open>
+  The following example peeks at a certain goal configuration.
+\<close>
 
 notepad
 begin
@@ -159,8 +152,10 @@
     sorry
 end
 
-text \<open>Here we see 3 individual subgoals in the same way as regular
-  proof methods would do.\<close>
+text \<open>
+  Here we see 3 individual subgoals in the same way as regular proof methods
+  would do.
+\<close>
 
 
 section \<open>Proof methods\<close>
@@ -172,42 +167,38 @@
   normal circumstances, the goal context remains unchanged, but it is also
   possible to declare named extensions of the proof context (\<^emph>\<open>cases\<close>).
 
-  This means a proof method is like a structurally enhanced tactic
-  (cf.\ \secref{sec:tactics}).  The well-formedness conditions for
-  tactics need to hold for methods accordingly, with the following
-  additions.
+  This means a proof method is like a structurally enhanced tactic (cf.\
+  \secref{sec:tactics}). The well-formedness conditions for tactics need to
+  hold for methods accordingly, with the following additions.
 
-  \<^item> Goal addressing is further limited either to operate
-  uniformly on \<^emph>\<open>all\<close> subgoals, or specifically on the
-  \<^emph>\<open>first\<close> subgoal.
+  \<^item> Goal addressing is further limited either to operate uniformly on \<^emph>\<open>all\<close>
+  subgoals, or specifically on the \<^emph>\<open>first\<close> subgoal.
 
-  Exception: old-style tactic emulations that are embedded into the
-  method space, e.g.\ @{method rule_tac}.
+  Exception: old-style tactic emulations that are embedded into the method
+  space, e.g.\ @{method rule_tac}.
 
-  \<^item> A non-trivial method always needs to make progress: an
-  identical follow-up goal state has to be avoided.\<^footnote>\<open>This
-  enables the user to write method expressions like \<open>meth\<^sup>+\<close>
-  without looping, while the trivial do-nothing case can be recovered
-  via \<open>meth\<^sup>?\<close>.\<close>
+  \<^item> A non-trivial method always needs to make progress: an identical follow-up
+  goal state has to be avoided.\<^footnote>\<open>This enables the user to write method
+  expressions like \<open>meth\<^sup>+\<close> without looping, while the trivial do-nothing case
+  can be recovered via \<open>meth\<^sup>?\<close>.\<close>
 
-  Exception: trivial stuttering steps, such as ``@{method -}'' or
-  @{method succeed}.
+  Exception: trivial stuttering steps, such as ``@{method -}'' or @{method
+  succeed}.
 
-  \<^item> Goal facts passed to the method must not be ignored.  If there
-  is no sensible use of facts outside the goal state, facts should be
-  inserted into the subgoals that are addressed by the method.
+  \<^item> Goal facts passed to the method must not be ignored. If there is no
+  sensible use of facts outside the goal state, facts should be inserted into
+  the subgoals that are addressed by the method.
 
 
   \<^medskip>
-  Syntactically, the language of proof methods appears as
-  arguments to Isar commands like @{command "by"} or @{command apply}.
-  User-space additions are reasonably easy by plugging suitable
-  method-valued parser functions into the framework, using the
-  @{command method_setup} command, for example.
+  Syntactically, the language of proof methods appears as arguments to Isar
+  commands like @{command "by"} or @{command apply}. User-space additions are
+  reasonably easy by plugging suitable method-valued parser functions into the
+  framework, using the @{command method_setup} command, for example.
 
   To get a better idea about the range of possibilities, consider the
-  following Isar proof schemes.  This is the general form of
-  structured proof text:
+  following Isar proof schemes. This is the general form of structured proof
+  text:
 
   \<^medskip>
   \begin{tabular}{l}
@@ -218,11 +209,12 @@
   \end{tabular}
   \<^medskip>
 
-  The goal configuration consists of \<open>facts\<^sub>1\<close> and
-  \<open>facts\<^sub>2\<close> appended in that order, and various \<open>props\<close> being claimed.  The \<open>initial_method\<close> is invoked
-  with facts and goals together and refines the problem to something
-  that is handled recursively in the proof \<open>body\<close>.  The \<open>terminal_method\<close> has another chance to finish any remaining
-  subgoals, but it does not see the facts of the initial step.
+  The goal configuration consists of \<open>facts\<^sub>1\<close> and \<open>facts\<^sub>2\<close> appended in that
+  order, and various \<open>props\<close> being claimed. The \<open>initial_method\<close> is invoked
+  with facts and goals together and refines the problem to something that is
+  handled recursively in the proof \<open>body\<close>. The \<open>terminal_method\<close> has another
+  chance to finish any remaining subgoals, but it does not see the facts of
+  the initial step.
 
   \<^medskip>
   This pattern illustrates unstructured proof scripts:
@@ -237,51 +229,48 @@
   \end{tabular}
   \<^medskip>
 
-  The \<open>method\<^sub>1\<close> operates on the original claim while
-  using \<open>facts\<^sub>1\<close>.  Since the @{command apply} command
-  structurally resets the facts, the \<open>method\<^sub>2\<close> will
-  operate on the remaining goal state without facts.  The \<open>method\<^sub>3\<close> will see again a collection of \<open>facts\<^sub>3\<close> that has been inserted into the script explicitly.
+  The \<open>method\<^sub>1\<close> operates on the original claim while using \<open>facts\<^sub>1\<close>. Since
+  the @{command apply} command structurally resets the facts, the \<open>method\<^sub>2\<close>
+  will operate on the remaining goal state without facts. The \<open>method\<^sub>3\<close> will
+  see again a collection of \<open>facts\<^sub>3\<close> that has been inserted into the script
+  explicitly.
 
   \<^medskip>
-  Empirically, any Isar proof method can be categorized as
-  follows.
+  Empirically, any Isar proof method can be categorized as follows.
 
-  \<^enum> \<^emph>\<open>Special method with cases\<close> with named context additions
-  associated with the follow-up goal state.
+    \<^enum> \<^emph>\<open>Special method with cases\<close> with named context additions associated
+    with the follow-up goal state.
 
-  Example: @{method "induct"}, which is also a ``raw'' method since it
-  operates on the internal representation of simultaneous claims as
-  Pure conjunction (\secref{sec:logic-aux}), instead of separate
-  subgoals (\secref{sec:tactical-goals}).
+    Example: @{method "induct"}, which is also a ``raw'' method since it
+    operates on the internal representation of simultaneous claims as Pure
+    conjunction (\secref{sec:logic-aux}), instead of separate subgoals
+    (\secref{sec:tactical-goals}).
 
-  \<^enum> \<^emph>\<open>Structured method\<close> with strong emphasis on facts outside
-  the goal state.
+    \<^enum> \<^emph>\<open>Structured method\<close> with strong emphasis on facts outside the goal
+    state.
 
-  Example: @{method "rule"}, which captures the key ideas behind
-  structured reasoning in Isar in its purest form.
+    Example: @{method "rule"}, which captures the key ideas behind structured
+    reasoning in Isar in its purest form.
 
-  \<^enum> \<^emph>\<open>Simple method\<close> with weaker emphasis on facts, which are
-  inserted into subgoals to emulate old-style tactical ``premises''.
-
-  Examples: @{method "simp"}, @{method "blast"}, @{method "auto"}.
+    \<^enum> \<^emph>\<open>Simple method\<close> with weaker emphasis on facts, which are inserted into
+    subgoals to emulate old-style tactical ``premises''.
 
-  \<^enum> \<^emph>\<open>Old-style tactic emulation\<close> with detailed numeric goal
-  addressing and explicit references to entities of the internal goal
-  state (which are otherwise invisible from proper Isar proof text).
-  The naming convention \<open>foo_tac\<close> makes this special
-  non-standard status clear.
+    Examples: @{method "simp"}, @{method "blast"}, @{method "auto"}.
 
-  Example: @{method "rule_tac"}.
+    \<^enum> \<^emph>\<open>Old-style tactic emulation\<close> with detailed numeric goal addressing and
+    explicit references to entities of the internal goal state (which are
+    otherwise invisible from proper Isar proof text). The naming convention
+    \<open>foo_tac\<close> makes this special non-standard status clear.
 
+    Example: @{method "rule_tac"}.
 
   When implementing proof methods, it is advisable to study existing
-  implementations carefully and imitate the typical ``boiler plate''
-  for context-sensitive parsing and further combinators to wrap-up
-  tactic expressions as methods.\<^footnote>\<open>Aliases or abbreviations of
-  the standard method combinators should be avoided.  Note that from
-  Isabelle99 until Isabelle2009 the system did provide various odd
-  combinations of method syntax wrappers that made applications more
-  complicated than necessary.\<close>
+  implementations carefully and imitate the typical ``boiler plate'' for
+  context-sensitive parsing and further combinators to wrap-up tactic
+  expressions as methods.\<^footnote>\<open>Aliases or abbreviations of the standard method
+  combinators should be avoided. Note that from Isabelle99 until Isabelle2009
+  the system did provide various odd combinations of method syntax wrappers
+  that made applications more complicated than necessary.\<close>
 \<close>
 
 text %mlref \<open>
@@ -296,8 +285,7 @@
   string -> theory -> theory"} \\
   \end{mldecls}
 
-  \<^descr> Type @{ML_type Proof.method} represents proof methods as
-  abstract type.
+  \<^descr> Type @{ML_type Proof.method} represents proof methods as abstract type.
 
   \<^descr> @{ML CONTEXT_METHOD}~\<open>(fn facts => context_tactic)\<close> wraps \<open>context_tactic\<close>
   depending on goal facts as a general proof method that may change the proof
@@ -305,36 +293,36 @@
   Proof_Context.update_cases}, which is wrapped up as combinator @{index_ML
   CONTEXT_CASES} for convenience.
 
-  \<^descr> @{ML METHOD}~\<open>(fn facts => tactic)\<close> wraps \<open>tactic\<close> depending on goal facts as regular proof method; the goal
-  context is passed via method syntax.
+  \<^descr> @{ML METHOD}~\<open>(fn facts => tactic)\<close> wraps \<open>tactic\<close> depending on goal facts
+  as regular proof method; the goal context is passed via method syntax.
 
-  \<^descr> @{ML SIMPLE_METHOD}~\<open>tactic\<close> wraps a tactic that
-  addresses all subgoals uniformly as simple proof method.  Goal facts
-  are already inserted into all subgoals before \<open>tactic\<close> is
-  applied.
+  \<^descr> @{ML SIMPLE_METHOD}~\<open>tactic\<close> wraps a tactic that addresses all subgoals
+  uniformly as simple proof method. Goal facts are already inserted into all
+  subgoals before \<open>tactic\<close> is applied.
 
-  \<^descr> @{ML SIMPLE_METHOD'}~\<open>tactic\<close> wraps a tactic that
-  addresses a specific subgoal as simple proof method that operates on
-  subgoal 1.  Goal facts are inserted into the subgoal then the \<open>tactic\<close> is applied.
+  \<^descr> @{ML SIMPLE_METHOD'}~\<open>tactic\<close> wraps a tactic that addresses a specific
+  subgoal as simple proof method that operates on subgoal 1. Goal facts are
+  inserted into the subgoal then the \<open>tactic\<close> is applied.
 
-  \<^descr> @{ML Method.insert_tac}~\<open>ctxt facts i\<close> inserts \<open>facts\<close> into subgoal \<open>i\<close>.  This is convenient to reproduce
-  part of the @{ML SIMPLE_METHOD} or @{ML SIMPLE_METHOD'} wrapping
-  within regular @{ML METHOD}, for example.
+  \<^descr> @{ML Method.insert_tac}~\<open>ctxt facts i\<close> inserts \<open>facts\<close> into subgoal \<open>i\<close>.
+  This is convenient to reproduce part of the @{ML SIMPLE_METHOD} or @{ML
+  SIMPLE_METHOD'} wrapping within regular @{ML METHOD}, for example.
 
-  \<^descr> @{ML Method.setup}~\<open>name parser description\<close> provides
-  the functionality of the Isar command @{command method_setup} as ML
-  function.
+  \<^descr> @{ML Method.setup}~\<open>name parser description\<close> provides the functionality of
+  the Isar command @{command method_setup} as ML function.
 \<close>
 
-text %mlex \<open>See also @{command method_setup} in
-  @{cite "isabelle-isar-ref"} which includes some abstract examples.
+text %mlex \<open>
+  See also @{command method_setup} in @{cite "isabelle-isar-ref"} which
+  includes some abstract examples.
 
   \<^medskip>
-  The following toy examples illustrate how the goal facts
-  and state are passed to proof methods.  The predefined proof method
-  called ``@{method tactic}'' wraps ML source of type @{ML_type
-  tactic} (abstracted over @{ML_text facts}).  This allows immediate
-  experimentation without parsing of concrete syntax.\<close>
+  The following toy examples illustrate how the goal facts and state are
+  passed to proof methods. The predefined proof method called ``@{method
+  tactic}'' wraps ML source of type @{ML_type tactic} (abstracted over
+  @{ML_text facts}). This allows immediate experimentation without parsing of
+  concrete syntax.
+\<close>
 
 notepad
 begin
@@ -357,8 +345,9 @@
 
 text \<open>
   \<^medskip>
-  The next example implements a method that simplifies
-  the first subgoal by rewrite rules that are given as arguments.\<close>
+  The next example implements a method that simplifies the first subgoal by
+  rewrite rules that are given as arguments.
+\<close>
 
 method_setup my_simp =
   \<open>Attrib.thms >> (fn thms => fn ctxt =>
@@ -367,21 +356,21 @@
         (put_simpset HOL_basic_ss ctxt addsimps thms) i)))\<close>
   "rewrite subgoal by given rules"
 
-text \<open>The concrete syntax wrapping of @{command method_setup} always
-  passes-through the proof context at the end of parsing, but it is
-  not used in this example.
+text \<open>
+  The concrete syntax wrapping of @{command method_setup} always
+  passes-through the proof context at the end of parsing, but it is not used
+  in this example.
 
-  The @{ML Attrib.thms} parser produces a list of theorems from the
-  usual Isar syntax involving attribute expressions etc.\ (syntax
-  category @{syntax thmrefs}) @{cite "isabelle-isar-ref"}.  The resulting
-  @{ML_text thms} are added to @{ML HOL_basic_ss} which already
-  contains the basic Simplifier setup for HOL.
+  The @{ML Attrib.thms} parser produces a list of theorems from the usual Isar
+  syntax involving attribute expressions etc.\ (syntax category @{syntax
+  thmrefs}) @{cite "isabelle-isar-ref"}. The resulting @{ML_text thms} are
+  added to @{ML HOL_basic_ss} which already contains the basic Simplifier
+  setup for HOL.
 
-  The tactic @{ML asm_full_simp_tac} is the one that is also used in
-  method @{method simp} by default.  The extra wrapping by the @{ML
-  CHANGED} tactical ensures progress of simplification: identical goal
-  states are filtered out explicitly to make the raw tactic conform to
-  standard Isar method behaviour.
+  The tactic @{ML asm_full_simp_tac} is the one that is also used in method
+  @{method simp} by default. The extra wrapping by the @{ML CHANGED} tactical
+  ensures progress of simplification: identical goal states are filtered out
+  explicitly to make the raw tactic conform to standard Isar method behaviour.
 
   \<^medskip>
   Method @{method my_simp} can be used in Isar proofs like this:
@@ -395,8 +384,9 @@
   have "a = c" by (my_simp a b)
 end
 
-text \<open>Here is a similar method that operates on all subgoals,
-  instead of just the first one.\<close>
+text \<open>
+  Here is a similar method that operates on all subgoals, instead of just the
+  first one.\<close>
 
 method_setup my_simp_all =
   \<open>Attrib.thms >> (fn thms => fn ctxt =>
@@ -416,15 +406,18 @@
 
 text \<open>
   \<^medskip>
-  Apart from explicit arguments, common proof methods
-  typically work with a default configuration provided by the context. As a
-  shortcut to rule management we use a cheap solution via the @{command
-  named_theorems} command to declare a dynamic fact in the context.\<close>
+  Apart from explicit arguments, common proof methods typically work with a
+  default configuration provided by the context. As a shortcut to rule
+  management we use a cheap solution via the @{command named_theorems} command
+  to declare a dynamic fact in the context.
+\<close>
 
 named_theorems my_simp
 
-text \<open>The proof method is now defined as before, but we append the
-  explicit arguments and the rules from the context.\<close>
+text \<open>
+  The proof method is now defined as before, but we append the explicit
+  arguments and the rules from the context.
+\<close>
 
 method_setup my_simp' =
   \<open>Attrib.thms >> (fn thms => fn ctxt =>
@@ -440,8 +433,7 @@
 
 text \<open>
   \<^medskip>
-  Method @{method my_simp'} can be used in Isar proofs
-  like this:
+  Method @{method my_simp'} can be used in Isar proofs like this:
 \<close>
 
 notepad
@@ -454,84 +446,88 @@
 
 text \<open>
   \<^medskip>
-  The @{method my_simp} variants defined above are
-  ``simple'' methods, i.e.\ the goal facts are merely inserted as goal
-  premises by the @{ML SIMPLE_METHOD'} or @{ML SIMPLE_METHOD} wrapper.
-  For proof methods that are similar to the standard collection of
-  @{method simp}, @{method blast}, @{method fast}, @{method auto}
-  there is little more that can be done.
+  The @{method my_simp} variants defined above are ``simple'' methods, i.e.\
+  the goal facts are merely inserted as goal premises by the @{ML
+  SIMPLE_METHOD'} or @{ML SIMPLE_METHOD} wrapper. For proof methods that are
+  similar to the standard collection of @{method simp}, @{method blast},
+  @{method fast}, @{method auto} there is little more that can be done.
 
-  Note that using the primary goal facts in the same manner as the
-  method arguments obtained via concrete syntax or the context does
-  not meet the requirement of ``strong emphasis on facts'' of regular
-  proof methods, because rewrite rules as used above can be easily
-  ignored.  A proof text ``@{command using}~\<open>foo\<close>~@{command
-  "by"}~\<open>my_simp\<close>'' where \<open>foo\<close> is not used would
-  deceive the reader.
+  Note that using the primary goal facts in the same manner as the method
+  arguments obtained via concrete syntax or the context does not meet the
+  requirement of ``strong emphasis on facts'' of regular proof methods,
+  because rewrite rules as used above can be easily ignored. A proof text
+  ``@{command using}~\<open>foo\<close>~@{command "by"}~\<open>my_simp\<close>'' where \<open>foo\<close> is not used
+  would deceive the reader.
 
   \<^medskip>
-  The technical treatment of rules from the context requires
-  further attention.  Above we rebuild a fresh @{ML_type simpset} from
-  the arguments and \<^emph>\<open>all\<close> rules retrieved from the context on
-  every invocation of the method.  This does not scale to really large
-  collections of rules, which easily emerges in the context of a big
-  theory library, for example.
+  The technical treatment of rules from the context requires further
+  attention. Above we rebuild a fresh @{ML_type simpset} from the arguments
+  and \<^emph>\<open>all\<close> rules retrieved from the context on every invocation of the
+  method. This does not scale to really large collections of rules, which
+  easily emerges in the context of a big theory library, for example.
 
   This is an inherent limitation of the simplistic rule management via
-  @{command named_theorems}, because it lacks tool-specific
-  storage and retrieval.  More realistic applications require
-  efficient index-structures that organize theorems in a customized
-  manner, such as a discrimination net that is indexed by the
-  left-hand sides of rewrite rules.  For variations on the Simplifier,
-  re-use of the existing type @{ML_type simpset} is adequate, but
-  scalability would require it be maintained statically within the
-  context data, not dynamically on each tool invocation.\<close>
+  @{command named_theorems}, because it lacks tool-specific storage and
+  retrieval. More realistic applications require efficient index-structures
+  that organize theorems in a customized manner, such as a discrimination net
+  that is indexed by the left-hand sides of rewrite rules. For variations on
+  the Simplifier, re-use of the existing type @{ML_type simpset} is adequate,
+  but scalability would require it be maintained statically within the context
+  data, not dynamically on each tool invocation.
+\<close>
 
 
 section \<open>Attributes \label{sec:attributes}\<close>
 
-text \<open>An \<^emph>\<open>attribute\<close> is a function \<open>context \<times> thm \<rightarrow>
-  context \<times> thm\<close>, which means both a (generic) context and a theorem
-  can be modified simultaneously.  In practice this mixed form is very
-  rare, instead attributes are presented either as \<^emph>\<open>declaration
-  attribute:\<close> \<open>thm \<rightarrow> context \<rightarrow> context\<close> or \<^emph>\<open>rule
+text \<open>
+  An \<^emph>\<open>attribute\<close> is a function \<open>context \<times> thm \<rightarrow> context \<times> thm\<close>, which means
+  both a (generic) context and a theorem can be modified simultaneously. In
+  practice this mixed form is very rare, instead attributes are presented
+  either as \<^emph>\<open>declaration attribute:\<close> \<open>thm \<rightarrow> context \<rightarrow> context\<close> or \<^emph>\<open>rule
   attribute:\<close> \<open>context \<rightarrow> thm \<rightarrow> thm\<close>.
 
-  Attributes can have additional arguments via concrete syntax.  There
-  is a collection of context-sensitive parsers for various logical
-  entities (types, terms, theorems).  These already take care of
-  applying morphisms to the arguments when attribute expressions are
-  moved into a different context (see also \secref{sec:morphisms}).
+  Attributes can have additional arguments via concrete syntax. There is a
+  collection of context-sensitive parsers for various logical entities (types,
+  terms, theorems). These already take care of applying morphisms to the
+  arguments when attribute expressions are moved into a different context (see
+  also \secref{sec:morphisms}).
 
-  When implementing declaration attributes, it is important to operate
-  exactly on the variant of the generic context that is provided by
-  the system, which is either global theory context or local proof
-  context.  In particular, the background theory of a local context
-  must not be modified in this situation!\<close>
+  When implementing declaration attributes, it is important to operate exactly
+  on the variant of the generic context that is provided by the system, which
+  is either global theory context or local proof context. In particular, the
+  background theory of a local context must not be modified in this
+  situation!
+\<close>
 
 text %mlref \<open>
   \begin{mldecls}
   @{index_ML_type attribute} \\
-  @{index_ML Thm.rule_attribute: "(Context.generic -> thm -> thm) -> attribute"} \\
+  @{index_ML Thm.rule_attribute: "thm list ->
+  (Context.generic -> thm -> thm) -> attribute"} \\
   @{index_ML Thm.declaration_attribute: "
   (thm -> Context.generic -> Context.generic) -> attribute"} \\
   @{index_ML Attrib.setup: "binding -> attribute context_parser ->
   string -> theory -> theory"} \\
   \end{mldecls}
 
-  \<^descr> Type @{ML_type attribute} represents attributes as concrete
-  type alias.
+  \<^descr> Type @{ML_type attribute} represents attributes as concrete type alias.
 
-  \<^descr> @{ML Thm.rule_attribute}~\<open>(fn context => rule)\<close> wraps
-  a context-dependent rule (mapping on @{ML_type thm}) as attribute.
+  \<^descr> @{ML Thm.rule_attribute}~\<open>thms (fn context => rule)\<close> wraps a
+  context-dependent rule (mapping on @{ML_type thm}) as attribute.
+
+  The \<open>thms\<close> are additional parameters: when forming an abstract closure, the
+  system may provide dummy facts that are propagated according to strict
+  evaluation discipline. In that case, \<open>rule\<close> is bypassed.
 
-  \<^descr> @{ML Thm.declaration_attribute}~\<open>(fn thm => decl)\<close>
-  wraps a theorem-dependent declaration (mapping on @{ML_type
-  Context.generic}) as attribute.
+  \<^descr> @{ML Thm.declaration_attribute}~\<open>(fn thm => decl)\<close> wraps a
+  theorem-dependent declaration (mapping on @{ML_type Context.generic}) as
+  attribute.
 
-  \<^descr> @{ML Attrib.setup}~\<open>name parser description\<close> provides
-  the functionality of the Isar command @{command attribute_setup} as
-  ML function.
+  When forming an abstract closure, the system may provide a dummy fact as
+  \<open>thm\<close>. In that case, \<open>decl\<close> is bypassed.
+
+  \<^descr> @{ML Attrib.setup}~\<open>name parser description\<close> provides the functionality of
+  the Isar command @{command attribute_setup} as ML function.
 \<close>
 
 text %mlantiq \<open>
@@ -543,15 +539,16 @@
   @@{ML_antiquotation attributes} attributes
   \<close>}
 
-  \<^descr> \<open>@{attributes [\<dots>]}\<close> embeds attribute source
-  representation into the ML text, which is particularly useful with
-  declarations like @{ML Local_Theory.note}.  Attribute names are
-  internalized at compile time, but the source is unevaluated.  This
-  means attributes with formal arguments (types, terms, theorems) may
-  be subject to odd effects of dynamic scoping!
+  \<^descr> \<open>@{attributes [\<dots>]}\<close> embeds attribute source representation into the ML
+  text, which is particularly useful with declarations like @{ML
+  Local_Theory.note}. Attribute names are internalized at compile time, but
+  the source is unevaluated. This means attributes with formal arguments
+  (types, terms, theorems) may be subject to odd effects of dynamic scoping!
 \<close>
 
-text %mlex \<open>See also @{command attribute_setup} in
-  @{cite "isabelle-isar-ref"} which includes some abstract examples.\<close>
+text %mlex \<open>
+  See also @{command attribute_setup} in @{cite "isabelle-isar-ref"} which
+  includes some abstract examples.
+\<close>
 
 end
--- a/src/Doc/Implementation/Local_Theory.thy	Tue Dec 15 14:41:47 2015 +0000
+++ b/src/Doc/Implementation/Local_Theory.thy	Wed Dec 16 17:30:30 2015 +0100
@@ -7,49 +7,45 @@
 chapter \<open>Local theory specifications \label{ch:local-theory}\<close>
 
 text \<open>
-  A \<^emph>\<open>local theory\<close> combines aspects of both theory and proof
-  context (cf.\ \secref{sec:context}), such that definitional
-  specifications may be given relatively to parameters and
-  assumptions.  A local theory is represented as a regular proof
-  context, augmented by administrative data about the \<^emph>\<open>target
+  A \<^emph>\<open>local theory\<close> combines aspects of both theory and proof context (cf.\
+  \secref{sec:context}), such that definitional specifications may be given
+  relatively to parameters and assumptions. A local theory is represented as a
+  regular proof context, augmented by administrative data about the \<^emph>\<open>target
   context\<close>.
 
-  The target is usually derived from the background theory by adding
-  local \<open>\<FIX>\<close> and \<open>\<ASSUME>\<close> elements, plus
-  suitable modifications of non-logical context data (e.g.\ a special
-  type-checking discipline).  Once initialized, the target is ready to
-  absorb definitional primitives: \<open>\<DEFINE>\<close> for terms and
-  \<open>\<NOTE>\<close> for theorems.  Such definitions may get
-  transformed in a target-specific way, but the programming interface
-  hides such details.
+  The target is usually derived from the background theory by adding local
+  \<open>\<FIX>\<close> and \<open>\<ASSUME>\<close> elements, plus suitable modifications of
+  non-logical context data (e.g.\ a special type-checking discipline). Once
+  initialized, the target is ready to absorb definitional primitives:
+  \<open>\<DEFINE>\<close> for terms and \<open>\<NOTE>\<close> for theorems. Such definitions may get
+  transformed in a target-specific way, but the programming interface hides
+  such details.
 
   Isabelle/Pure provides target mechanisms for locales, type-classes,
-  type-class instantiations, and general overloading.  In principle,
-  users can implement new targets as well, but this rather arcane
-  discipline is beyond the scope of this manual.  In contrast,
-  implementing derived definitional packages to be used within a local
-  theory context is quite easy: the interfaces are even simpler and
-  more abstract than the underlying primitives for raw theories.
+  type-class instantiations, and general overloading. In principle, users can
+  implement new targets as well, but this rather arcane discipline is beyond
+  the scope of this manual. In contrast, implementing derived definitional
+  packages to be used within a local theory context is quite easy: the
+  interfaces are even simpler and more abstract than the underlying primitives
+  for raw theories.
 
-  Many definitional packages for local theories are available in
-  Isabelle.  Although a few old packages only work for global
-  theories, the standard way of implementing definitional packages in
-  Isabelle is via the local theory interface.
+  Many definitional packages for local theories are available in Isabelle.
+  Although a few old packages only work for global theories, the standard way
+  of implementing definitional packages in Isabelle is via the local theory
+  interface.
 \<close>
 
 
 section \<open>Definitional elements\<close>
 
 text \<open>
-  There are separate elements \<open>\<DEFINE> c \<equiv> t\<close> for terms, and
-  \<open>\<NOTE> b = thm\<close> for theorems.  Types are treated
-  implicitly, according to Hindley-Milner discipline (cf.\
-  \secref{sec:variables}).  These definitional primitives essentially
-  act like \<open>let\<close>-bindings within a local context that may
-  already contain earlier \<open>let\<close>-bindings and some initial
-  \<open>\<lambda>\<close>-bindings.  Thus we gain \<^emph>\<open>dependent definitions\<close>
-  that are relative to an initial axiomatic context.  The following
-  diagram illustrates this idea of axiomatic elements versus
+  There are separate elements \<open>\<DEFINE> c \<equiv> t\<close> for terms, and \<open>\<NOTE> b =
+  thm\<close> for theorems. Types are treated implicitly, according to Hindley-Milner
+  discipline (cf.\ \secref{sec:variables}). These definitional primitives
+  essentially act like \<open>let\<close>-bindings within a local context that may already
+  contain earlier \<open>let\<close>-bindings and some initial \<open>\<lambda>\<close>-bindings. Thus we gain
+  \<^emph>\<open>dependent definitions\<close> that are relative to an initial axiomatic context.
+  The following diagram illustrates this idea of axiomatic elements versus
   definitional elements:
 
   \begin{center}
@@ -64,34 +60,33 @@
   \end{tabular}
   \end{center}
 
-  A user package merely needs to produce suitable \<open>\<DEFINE>\<close>
-  and \<open>\<NOTE>\<close> elements according to the application.  For
-  example, a package for inductive definitions might first \<open>\<DEFINE>\<close> a certain predicate as some fixed-point construction,
-  then \<open>\<NOTE>\<close> a proven result about monotonicity of the
+  A user package merely needs to produce suitable \<open>\<DEFINE>\<close> and \<open>\<NOTE>\<close>
+  elements according to the application. For example, a package for inductive
+  definitions might first \<open>\<DEFINE>\<close> a certain predicate as some fixed-point
+  construction, then \<open>\<NOTE>\<close> a proven result about monotonicity of the
   functor involved here, and then produce further derived concepts via
   additional \<open>\<DEFINE>\<close> and \<open>\<NOTE>\<close> elements.
 
-  The cumulative sequence of \<open>\<DEFINE>\<close> and \<open>\<NOTE>\<close>
-  produced at package runtime is managed by the local theory
-  infrastructure by means of an \<^emph>\<open>auxiliary context\<close>.  Thus the
-  system holds up the impression of working within a fully abstract
-  situation with hypothetical entities: \<open>\<DEFINE> c \<equiv> t\<close>
-  always results in a literal fact \<open>\<^BG>c \<equiv> t\<^EN>\<close>, where
-  \<open>c\<close> is a fixed variable \<open>c\<close>.  The details about
-  global constants, name spaces etc. are handled internally.
+  The cumulative sequence of \<open>\<DEFINE>\<close> and \<open>\<NOTE>\<close> produced at package
+  runtime is managed by the local theory infrastructure by means of an
+  \<^emph>\<open>auxiliary context\<close>. Thus the system holds up the impression of working
+  within a fully abstract situation with hypothetical entities: \<open>\<DEFINE> c \<equiv>
+  t\<close> always results in a literal fact \<open>\<^BG>c \<equiv> t\<^EN>\<close>, where \<open>c\<close> is a
+  fixed variable \<open>c\<close>. The details about global constants, name spaces etc. are
+  handled internally.
 
-  So the general structure of a local theory is a sandwich of three
-  layers:
+  So the general structure of a local theory is a sandwich of three layers:
 
   \begin{center}
   \framebox{\quad auxiliary context \quad\framebox{\quad target context \quad\framebox{\quad background theory\quad}}}
   \end{center}
 
-  When a definitional package is finished, the auxiliary context is
-  reset to the target context.  The target now holds definitions for
-  terms and theorems that stem from the hypothetical \<open>\<DEFINE>\<close> and \<open>\<NOTE>\<close> elements, transformed by the
-  particular target policy (see @{cite \<open>\S4--5\<close> "Haftmann-Wenzel:2009"}
-  for details).\<close>
+  When a definitional package is finished, the auxiliary context is reset to
+  the target context. The target now holds definitions for terms and theorems
+  that stem from the hypothetical \<open>\<DEFINE>\<close> and \<open>\<NOTE>\<close> elements,
+  transformed by the particular target policy (see @{cite \<open>\S4--5\<close>
+  "Haftmann-Wenzel:2009"} for details).
+\<close>
 
 text %mlref \<open>
   \begin{mldecls}
@@ -103,50 +98,45 @@
     local_theory -> (string * thm list) * local_theory"} \\
   \end{mldecls}
 
-  \<^descr> Type @{ML_type local_theory} represents local theories.
-  Although this is merely an alias for @{ML_type Proof.context}, it is
-  semantically a subtype of the same: a @{ML_type local_theory} holds
-  target information as special context data.  Subtyping means that
-  any value \<open>lthy:\<close>~@{ML_type local_theory} can be also used
-  with operations on expecting a regular \<open>ctxt:\<close>~@{ML_type
+  \<^descr> Type @{ML_type local_theory} represents local theories. Although this is
+  merely an alias for @{ML_type Proof.context}, it is semantically a subtype
+  of the same: a @{ML_type local_theory} holds target information as special
+  context data. Subtyping means that any value \<open>lthy:\<close>~@{ML_type local_theory}
+  can be also used with operations on expecting a regular \<open>ctxt:\<close>~@{ML_type
   Proof.context}.
 
-  \<^descr> @{ML Named_Target.init}~\<open>before_exit name thy\<close>
-  initializes a local theory derived from the given background theory.
-  An empty name refers to a \<^emph>\<open>global theory\<close> context, and a
-  non-empty name refers to a @{command locale} or @{command class}
-  context (a fully-qualified internal name is expected here).  This is
-  useful for experimentation --- normally the Isar toplevel already
+  \<^descr> @{ML Named_Target.init}~\<open>before_exit name thy\<close> initializes a local theory
+  derived from the given background theory. An empty name refers to a \<^emph>\<open>global
+  theory\<close> context, and a non-empty name refers to a @{command locale} or
+  @{command class} context (a fully-qualified internal name is expected here).
+  This is useful for experimentation --- normally the Isar toplevel already
   takes care to initialize the local theory context.
 
-  \<^descr> @{ML Local_Theory.define}~\<open>((b, mx), (a, rhs))
-  lthy\<close> defines a local entity according to the specification that is
-  given relatively to the current \<open>lthy\<close> context.  In
-  particular the term of the RHS may refer to earlier local entities
-  from the auxiliary context, or hypothetical parameters from the
-  target context.  The result is the newly defined term (which is
-  always a fixed variable with exactly the same name as specified for
-  the LHS), together with an equational theorem that states the
-  definition as a hypothetical fact.
+  \<^descr> @{ML Local_Theory.define}~\<open>((b, mx), (a, rhs)) lthy\<close> defines a local
+  entity according to the specification that is given relatively to the
+  current \<open>lthy\<close> context. In particular the term of the RHS may refer to
+  earlier local entities from the auxiliary context, or hypothetical
+  parameters from the target context. The result is the newly defined term
+  (which is always a fixed variable with exactly the same name as specified
+  for the LHS), together with an equational theorem that states the definition
+  as a hypothetical fact.
 
-  Unless an explicit name binding is given for the RHS, the resulting
-  fact will be called \<open>b_def\<close>.  Any given attributes are
-  applied to that same fact --- immediately in the auxiliary context
-  \<^emph>\<open>and\<close> in any transformed versions stemming from target-specific
-  policies or any later interpretations of results from the target
-  context (think of @{command locale} and @{command interpretation},
-  for example).  This means that attributes should be usually plain
-  declarations such as @{attribute simp}, while non-trivial rules like
+  Unless an explicit name binding is given for the RHS, the resulting fact
+  will be called \<open>b_def\<close>. Any given attributes are applied to that same fact
+  --- immediately in the auxiliary context \<^emph>\<open>and\<close> in any transformed versions
+  stemming from target-specific policies or any later interpretations of
+  results from the target context (think of @{command locale} and @{command
+  interpretation}, for example). This means that attributes should be usually
+  plain declarations such as @{attribute simp}, while non-trivial rules like
   @{attribute simplified} are better avoided.
 
-  \<^descr> @{ML Local_Theory.note}~\<open>(a, ths) lthy\<close> is
-  analogous to @{ML Local_Theory.define}, but defines facts instead of
-  terms.  There is also a slightly more general variant @{ML
-  Local_Theory.notes} that defines several facts (with attribute
-  expressions) simultaneously.
+  \<^descr> @{ML Local_Theory.note}~\<open>(a, ths) lthy\<close> is analogous to @{ML
+  Local_Theory.define}, but defines facts instead of terms. There is also a
+  slightly more general variant @{ML Local_Theory.notes} that defines several
+  facts (with attribute expressions) simultaneously.
 
-  This is essentially the internal version of the @{command lemmas}
-  command, or @{command declare} if an empty name binding is given.
+  This is essentially the internal version of the @{command lemmas} command,
+  or @{command declare} if an empty name binding is given.
 \<close>
 
 
--- a/src/Doc/Implementation/Logic.thy	Tue Dec 15 14:41:47 2015 +0000
+++ b/src/Doc/Implementation/Logic.thy	Wed Dec 16 17:30:30 2015 +0100
@@ -7,110 +7,97 @@
 chapter \<open>Primitive logic \label{ch:logic}\<close>
 
 text \<open>
-  The logical foundations of Isabelle/Isar are that of the Pure logic,
-  which has been introduced as a Natural Deduction framework in
-  @{cite paulson700}.  This is essentially the same logic as ``\<open>\<lambda>HOL\<close>'' in the more abstract setting of Pure Type Systems (PTS)
-  @{cite "Barendregt-Geuvers:2001"}, although there are some key
-  differences in the specific treatment of simple types in
-  Isabelle/Pure.
+  The logical foundations of Isabelle/Isar are that of the Pure logic, which
+  has been introduced as a Natural Deduction framework in @{cite paulson700}.
+  This is essentially the same logic as ``\<open>\<lambda>HOL\<close>'' in the more abstract
+  setting of Pure Type Systems (PTS) @{cite "Barendregt-Geuvers:2001"},
+  although there are some key differences in the specific treatment of simple
+  types in Isabelle/Pure.
 
-  Following type-theoretic parlance, the Pure logic consists of three
-  levels of \<open>\<lambda>\<close>-calculus with corresponding arrows, \<open>\<Rightarrow>\<close> for syntactic function space (terms depending on terms), \<open>\<And>\<close> for universal quantification (proofs depending on terms), and
-  \<open>\<Longrightarrow>\<close> for implication (proofs depending on proofs).
+  Following type-theoretic parlance, the Pure logic consists of three levels
+  of \<open>\<lambda>\<close>-calculus with corresponding arrows, \<open>\<Rightarrow>\<close> for syntactic function space
+  (terms depending on terms), \<open>\<And>\<close> for universal quantification (proofs
+  depending on terms), and \<open>\<Longrightarrow>\<close> for implication (proofs depending on proofs).
 
   Derivations are relative to a logical theory, which declares type
-  constructors, constants, and axioms.  Theory declarations support
-  schematic polymorphism, which is strictly speaking outside the
-  logic.\<^footnote>\<open>This is the deeper logical reason, why the theory
-  context \<open>\<Theta>\<close> is separate from the proof context \<open>\<Gamma>\<close>
-  of the core calculus: type constructors, term constants, and facts
-  (proof constants) may involve arbitrary type schemes, but the type
-  of a locally fixed term parameter is also fixed!\<close>
+  constructors, constants, and axioms. Theory declarations support schematic
+  polymorphism, which is strictly speaking outside the logic.\<^footnote>\<open>This is the
+  deeper logical reason, why the theory context \<open>\<Theta>\<close> is separate from the proof
+  context \<open>\<Gamma>\<close> of the core calculus: type constructors, term constants, and
+  facts (proof constants) may involve arbitrary type schemes, but the type of
+  a locally fixed term parameter is also fixed!\<close>
 \<close>
 
 
 section \<open>Types \label{sec:types}\<close>
 
 text \<open>
-  The language of types is an uninterpreted order-sorted first-order
-  algebra; types are qualified by ordered type classes.
+  The language of types is an uninterpreted order-sorted first-order algebra;
+  types are qualified by ordered type classes.
 
   \<^medskip>
-  A \<^emph>\<open>type class\<close> is an abstract syntactic entity
-  declared in the theory context.  The \<^emph>\<open>subclass relation\<close> \<open>c\<^sub>1 \<subseteq> c\<^sub>2\<close> is specified by stating an acyclic
-  generating relation; the transitive closure is maintained
-  internally.  The resulting relation is an ordering: reflexive,
-  transitive, and antisymmetric.
+  A \<^emph>\<open>type class\<close> is an abstract syntactic entity declared in the theory
+  context. The \<^emph>\<open>subclass relation\<close> \<open>c\<^sub>1 \<subseteq> c\<^sub>2\<close> is specified by stating an
+  acyclic generating relation; the transitive closure is maintained
+  internally. The resulting relation is an ordering: reflexive, transitive,
+  and antisymmetric.
 
-  A \<^emph>\<open>sort\<close> is a list of type classes written as \<open>s = {c\<^sub>1,
-  \<dots>, c\<^sub>m}\<close>, it represents symbolic intersection.  Notationally, the
-  curly braces are omitted for singleton intersections, i.e.\ any
-  class \<open>c\<close> may be read as a sort \<open>{c}\<close>.  The ordering
-  on type classes is extended to sorts according to the meaning of
-  intersections: \<open>{c\<^sub>1, \<dots> c\<^sub>m} \<subseteq> {d\<^sub>1, \<dots>, d\<^sub>n}\<close> iff \<open>\<forall>j. \<exists>i. c\<^sub>i \<subseteq> d\<^sub>j\<close>.  The empty intersection \<open>{}\<close> refers to
-  the universal sort, which is the largest element wrt.\ the sort
-  order.  Thus \<open>{}\<close> represents the ``full sort'', not the
-  empty one!  The intersection of all (finitely many) classes declared
-  in the current theory is the least element wrt.\ the sort ordering.
+  A \<^emph>\<open>sort\<close> is a list of type classes written as \<open>s = {c\<^sub>1, \<dots>, c\<^sub>m}\<close>, it
+  represents symbolic intersection. Notationally, the curly braces are omitted
+  for singleton intersections, i.e.\ any class \<open>c\<close> may be read as a sort
+  \<open>{c}\<close>. The ordering on type classes is extended to sorts according to the
+  meaning of intersections: \<open>{c\<^sub>1, \<dots> c\<^sub>m} \<subseteq> {d\<^sub>1, \<dots>, d\<^sub>n}\<close> iff \<open>\<forall>j. \<exists>i. c\<^sub>i \<subseteq>
+  d\<^sub>j\<close>. The empty intersection \<open>{}\<close> refers to the universal sort, which is the
+  largest element wrt.\ the sort order. Thus \<open>{}\<close> represents the ``full
+  sort'', not the empty one! The intersection of all (finitely many) classes
+  declared in the current theory is the least element wrt.\ the sort ordering.
 
   \<^medskip>
-  A \<^emph>\<open>fixed type variable\<close> is a pair of a basic name
-  (starting with a \<open>'\<close> character) and a sort constraint, e.g.\
-  \<open>('a, s)\<close> which is usually printed as \<open>\<alpha>\<^sub>s\<close>.
-  A \<^emph>\<open>schematic type variable\<close> is a pair of an indexname and a
-  sort constraint, e.g.\ \<open>(('a, 0), s)\<close> which is usually
-  printed as \<open>?\<alpha>\<^sub>s\<close>.
+  A \<^emph>\<open>fixed type variable\<close> is a pair of a basic name (starting with a \<open>'\<close>
+  character) and a sort constraint, e.g.\ \<open>('a, s)\<close> which is usually printed
+  as \<open>\<alpha>\<^sub>s\<close>. A \<^emph>\<open>schematic type variable\<close> is a pair of an indexname and a sort
+  constraint, e.g.\ \<open>(('a, 0), s)\<close> which is usually printed as \<open>?\<alpha>\<^sub>s\<close>.
 
-  Note that \<^emph>\<open>all\<close> syntactic components contribute to the identity
-  of type variables: basic name, index, and sort constraint.  The core
-  logic handles type variables with the same name but different sorts
-  as different, although the type-inference layer (which is outside
-  the core) rejects anything like that.
+  Note that \<^emph>\<open>all\<close> syntactic components contribute to the identity of type
+  variables: basic name, index, and sort constraint. The core logic handles
+  type variables with the same name but different sorts as different, although
+  the type-inference layer (which is outside the core) rejects anything like
+  that.
 
-  A \<^emph>\<open>type constructor\<close> \<open>\<kappa>\<close> is a \<open>k\<close>-ary operator
-  on types declared in the theory.  Type constructor application is
-  written postfix as \<open>(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>k)\<kappa>\<close>.  For
-  \<open>k = 0\<close> the argument tuple is omitted, e.g.\ \<open>prop\<close>
-  instead of \<open>()prop\<close>.  For \<open>k = 1\<close> the parentheses
-  are omitted, e.g.\ \<open>\<alpha> list\<close> instead of \<open>(\<alpha>)list\<close>.
-  Further notation is provided for specific constructors, notably the
-  right-associative infix \<open>\<alpha> \<Rightarrow> \<beta>\<close> instead of \<open>(\<alpha>,
-  \<beta>)fun\<close>.
+  A \<^emph>\<open>type constructor\<close> \<open>\<kappa>\<close> is a \<open>k\<close>-ary operator on types declared in the
+  theory. Type constructor application is written postfix as \<open>(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>k)\<kappa>\<close>.
+  For \<open>k = 0\<close> the argument tuple is omitted, e.g.\ \<open>prop\<close> instead of \<open>()prop\<close>.
+  For \<open>k = 1\<close> the parentheses are omitted, e.g.\ \<open>\<alpha> list\<close> instead of
+  \<open>(\<alpha>)list\<close>. Further notation is provided for specific constructors, notably
+  the right-associative infix \<open>\<alpha> \<Rightarrow> \<beta>\<close> instead of \<open>(\<alpha>, \<beta>)fun\<close>.
   
-  The logical category \<^emph>\<open>type\<close> is defined inductively over type
-  variables and type constructors as follows: \<open>\<tau> = \<alpha>\<^sub>s | ?\<alpha>\<^sub>s |
-  (\<tau>\<^sub>1, \<dots>, \<tau>\<^sub>k)\<kappa>\<close>.
+  The logical category \<^emph>\<open>type\<close> is defined inductively over type variables and
+  type constructors as follows: \<open>\<tau> = \<alpha>\<^sub>s | ?\<alpha>\<^sub>s | (\<tau>\<^sub>1, \<dots>, \<tau>\<^sub>k)\<kappa>\<close>.
 
-  A \<^emph>\<open>type abbreviation\<close> is a syntactic definition \<open>(\<^vec>\<alpha>)\<kappa> = \<tau>\<close> of an arbitrary type expression \<open>\<tau>\<close> over
-  variables \<open>\<^vec>\<alpha>\<close>.  Type abbreviations appear as type
-  constructors in the syntax, but are expanded before entering the
-  logical core.
+  A \<^emph>\<open>type abbreviation\<close> is a syntactic definition \<open>(\<^vec>\<alpha>)\<kappa> = \<tau>\<close> of an
+  arbitrary type expression \<open>\<tau>\<close> over variables \<open>\<^vec>\<alpha>\<close>. Type abbreviations
+  appear as type constructors in the syntax, but are expanded before entering
+  the logical core.
 
-  A \<^emph>\<open>type arity\<close> declares the image behavior of a type
-  constructor wrt.\ the algebra of sorts: \<open>\<kappa> :: (s\<^sub>1, \<dots>,
-  s\<^sub>k)s\<close> means that \<open>(\<tau>\<^sub>1, \<dots>, \<tau>\<^sub>k)\<kappa>\<close> is
-  of sort \<open>s\<close> if every argument type \<open>\<tau>\<^sub>i\<close> is
-  of sort \<open>s\<^sub>i\<close>.  Arity declarations are implicitly
-  completed, i.e.\ \<open>\<kappa> :: (\<^vec>s)c\<close> entails \<open>\<kappa> ::
+  A \<^emph>\<open>type arity\<close> declares the image behavior of a type constructor wrt.\ the
+  algebra of sorts: \<open>\<kappa> :: (s\<^sub>1, \<dots>, s\<^sub>k)s\<close> means that \<open>(\<tau>\<^sub>1, \<dots>, \<tau>\<^sub>k)\<kappa>\<close> is of
+  sort \<open>s\<close> if every argument type \<open>\<tau>\<^sub>i\<close> is of sort \<open>s\<^sub>i\<close>. Arity declarations
+  are implicitly completed, i.e.\ \<open>\<kappa> :: (\<^vec>s)c\<close> entails \<open>\<kappa> ::
   (\<^vec>s)c'\<close> for any \<open>c' \<supseteq> c\<close>.
 
   \<^medskip>
-  The sort algebra is always maintained as \<^emph>\<open>coregular\<close>,
-  which means that type arities are consistent with the subclass
-  relation: for any type constructor \<open>\<kappa>\<close>, and classes \<open>c\<^sub>1 \<subseteq> c\<^sub>2\<close>, and arities \<open>\<kappa> ::
-  (\<^vec>s\<^sub>1)c\<^sub>1\<close> and \<open>\<kappa> ::
-  (\<^vec>s\<^sub>2)c\<^sub>2\<close> holds \<open>\<^vec>s\<^sub>1 \<subseteq>
-  \<^vec>s\<^sub>2\<close> component-wise.
+  The sort algebra is always maintained as \<^emph>\<open>coregular\<close>, which means that type
+  arities are consistent with the subclass relation: for any type constructor
+  \<open>\<kappa>\<close>, and classes \<open>c\<^sub>1 \<subseteq> c\<^sub>2\<close>, and arities \<open>\<kappa> :: (\<^vec>s\<^sub>1)c\<^sub>1\<close> and \<open>\<kappa> ::
+  (\<^vec>s\<^sub>2)c\<^sub>2\<close> holds \<open>\<^vec>s\<^sub>1 \<subseteq> \<^vec>s\<^sub>2\<close> component-wise.
 
   The key property of a coregular order-sorted algebra is that sort
   constraints can be solved in a most general fashion: for each type
-  constructor \<open>\<kappa>\<close> and sort \<open>s\<close> there is a most general
-  vector of argument sorts \<open>(s\<^sub>1, \<dots>, s\<^sub>k)\<close> such
-  that a type scheme \<open>(\<alpha>\<^bsub>s\<^sub>1\<^esub>, \<dots>,
-  \<alpha>\<^bsub>s\<^sub>k\<^esub>)\<kappa>\<close> is of sort \<open>s\<close>.
-  Consequently, type unification has most general solutions (modulo
-  equivalence of sorts), so type-inference produces primary types as
-  expected @{cite "nipkow-prehofer"}.
+  constructor \<open>\<kappa>\<close> and sort \<open>s\<close> there is a most general vector of argument
+  sorts \<open>(s\<^sub>1, \<dots>, s\<^sub>k)\<close> such that a type scheme \<open>(\<alpha>\<^bsub>s\<^sub>1\<^esub>, \<dots>, \<alpha>\<^bsub>s\<^sub>k\<^esub>)\<kappa>\<close> is of
+  sort \<open>s\<close>. Consequently, type unification has most general solutions (modulo
+  equivalence of sorts), so type-inference produces primary types as expected
+  @{cite "nipkow-prehofer"}.
 \<close>
 
 text %mlref \<open>
@@ -135,48 +122,42 @@
 
   \<^descr> Type @{ML_type class} represents type classes.
 
-  \<^descr> Type @{ML_type sort} represents sorts, i.e.\ finite
-  intersections of classes.  The empty list @{ML "[]: sort"} refers to
-  the empty class intersection, i.e.\ the ``full sort''.
+  \<^descr> Type @{ML_type sort} represents sorts, i.e.\ finite intersections of
+  classes. The empty list @{ML "[]: sort"} refers to the empty class
+  intersection, i.e.\ the ``full sort''.
 
-  \<^descr> Type @{ML_type arity} represents type arities.  A triple
-  \<open>(\<kappa>, \<^vec>s, s) : arity\<close> represents \<open>\<kappa> ::
-  (\<^vec>s)s\<close> as described above.
+  \<^descr> Type @{ML_type arity} represents type arities. A triple \<open>(\<kappa>, \<^vec>s, s)
+  : arity\<close> represents \<open>\<kappa> :: (\<^vec>s)s\<close> as described above.
 
-  \<^descr> Type @{ML_type typ} represents types; this is a datatype with
-  constructors @{ML TFree}, @{ML TVar}, @{ML Type}.
+  \<^descr> Type @{ML_type typ} represents types; this is a datatype with constructors
+  @{ML TFree}, @{ML TVar}, @{ML Type}.
 
-  \<^descr> @{ML Term.map_atyps}~\<open>f \<tau>\<close> applies the mapping \<open>f\<close> to all atomic types (@{ML TFree}, @{ML TVar}) occurring in
-  \<open>\<tau>\<close>.
+  \<^descr> @{ML Term.map_atyps}~\<open>f \<tau>\<close> applies the mapping \<open>f\<close> to all atomic types
+  (@{ML TFree}, @{ML TVar}) occurring in \<open>\<tau>\<close>.
 
-  \<^descr> @{ML Term.fold_atyps}~\<open>f \<tau>\<close> iterates the operation
-  \<open>f\<close> over all occurrences of atomic types (@{ML TFree}, @{ML
-  TVar}) in \<open>\<tau>\<close>; the type structure is traversed from left to
-  right.
+  \<^descr> @{ML Term.fold_atyps}~\<open>f \<tau>\<close> iterates the operation \<open>f\<close> over all
+  occurrences of atomic types (@{ML TFree}, @{ML TVar}) in \<open>\<tau>\<close>; the type
+  structure is traversed from left to right.
 
-  \<^descr> @{ML Sign.subsort}~\<open>thy (s\<^sub>1, s\<^sub>2)\<close>
-  tests the subsort relation \<open>s\<^sub>1 \<subseteq> s\<^sub>2\<close>.
+  \<^descr> @{ML Sign.subsort}~\<open>thy (s\<^sub>1, s\<^sub>2)\<close> tests the subsort relation \<open>s\<^sub>1 \<subseteq>
+  s\<^sub>2\<close>.
 
-  \<^descr> @{ML Sign.of_sort}~\<open>thy (\<tau>, s)\<close> tests whether type
-  \<open>\<tau>\<close> is of sort \<open>s\<close>.
+  \<^descr> @{ML Sign.of_sort}~\<open>thy (\<tau>, s)\<close> tests whether type \<open>\<tau>\<close> is of sort \<open>s\<close>.
 
-  \<^descr> @{ML Sign.add_type}~\<open>ctxt (\<kappa>, k, mx)\<close> declares a
-  new type constructors \<open>\<kappa>\<close> with \<open>k\<close> arguments and
-  optional mixfix syntax.
+  \<^descr> @{ML Sign.add_type}~\<open>ctxt (\<kappa>, k, mx)\<close> declares a new type constructors \<open>\<kappa>\<close>
+  with \<open>k\<close> arguments and optional mixfix syntax.
 
-  \<^descr> @{ML Sign.add_type_abbrev}~\<open>ctxt (\<kappa>, \<^vec>\<alpha>, \<tau>)\<close>
-  defines a new type abbreviation \<open>(\<^vec>\<alpha>)\<kappa> = \<tau>\<close>.
+  \<^descr> @{ML Sign.add_type_abbrev}~\<open>ctxt (\<kappa>, \<^vec>\<alpha>, \<tau>)\<close> defines a new type
+  abbreviation \<open>(\<^vec>\<alpha>)\<kappa> = \<tau>\<close>.
 
-  \<^descr> @{ML Sign.primitive_class}~\<open>(c, [c\<^sub>1, \<dots>,
-  c\<^sub>n])\<close> declares a new class \<open>c\<close>, together with class
-  relations \<open>c \<subseteq> c\<^sub>i\<close>, for \<open>i = 1, \<dots>, n\<close>.
+  \<^descr> @{ML Sign.primitive_class}~\<open>(c, [c\<^sub>1, \<dots>, c\<^sub>n])\<close> declares a new class \<open>c\<close>,
+  together with class relations \<open>c \<subseteq> c\<^sub>i\<close>, for \<open>i = 1, \<dots>, n\<close>.
 
-  \<^descr> @{ML Sign.primitive_classrel}~\<open>(c\<^sub>1,
-  c\<^sub>2)\<close> declares the class relation \<open>c\<^sub>1 \<subseteq>
-  c\<^sub>2\<close>.
+  \<^descr> @{ML Sign.primitive_classrel}~\<open>(c\<^sub>1, c\<^sub>2)\<close> declares the class relation
+  \<open>c\<^sub>1 \<subseteq> c\<^sub>2\<close>.
 
-  \<^descr> @{ML Sign.primitive_arity}~\<open>(\<kappa>, \<^vec>s, s)\<close> declares
-  the arity \<open>\<kappa> :: (\<^vec>s)s\<close>.
+  \<^descr> @{ML Sign.primitive_arity}~\<open>(\<kappa>, \<^vec>s, s)\<close> declares the arity \<open>\<kappa> ::
+  (\<^vec>s)s\<close>.
 \<close>
 
 text %mlantiq \<open>
@@ -201,92 +182,84 @@
   @@{ML_antiquotation typ} type
   \<close>}
 
-  \<^descr> \<open>@{class c}\<close> inlines the internalized class \<open>c\<close> --- as @{ML_type string} literal.
-
-  \<^descr> \<open>@{sort s}\<close> inlines the internalized sort \<open>s\<close>
-  --- as @{ML_type "string list"} literal.
-
-  \<^descr> \<open>@{type_name c}\<close> inlines the internalized type
-  constructor \<open>c\<close> --- as @{ML_type string} literal.
-
-  \<^descr> \<open>@{type_abbrev c}\<close> inlines the internalized type
-  abbreviation \<open>c\<close> --- as @{ML_type string} literal.
-
-  \<^descr> \<open>@{nonterminal c}\<close> inlines the internalized syntactic
-  type~/ grammar nonterminal \<open>c\<close> --- as @{ML_type string}
+  \<^descr> \<open>@{class c}\<close> inlines the internalized class \<open>c\<close> --- as @{ML_type string}
   literal.
 
-  \<^descr> \<open>@{typ \<tau>}\<close> inlines the internalized type \<open>\<tau>\<close>
-  --- as constructor term for datatype @{ML_type typ}.
+  \<^descr> \<open>@{sort s}\<close> inlines the internalized sort \<open>s\<close> --- as @{ML_type "string
+  list"} literal.
+
+  \<^descr> \<open>@{type_name c}\<close> inlines the internalized type constructor \<open>c\<close> --- as
+  @{ML_type string} literal.
+
+  \<^descr> \<open>@{type_abbrev c}\<close> inlines the internalized type abbreviation \<open>c\<close> --- as
+  @{ML_type string} literal.
+
+  \<^descr> \<open>@{nonterminal c}\<close> inlines the internalized syntactic type~/ grammar
+  nonterminal \<open>c\<close> --- as @{ML_type string} literal.
+
+  \<^descr> \<open>@{typ \<tau>}\<close> inlines the internalized type \<open>\<tau>\<close> --- as constructor term for
+  datatype @{ML_type typ}.
 \<close>
 
 
 section \<open>Terms \label{sec:terms}\<close>
 
 text \<open>
-  The language of terms is that of simply-typed \<open>\<lambda>\<close>-calculus
-  with de-Bruijn indices for bound variables (cf.\ @{cite debruijn72}
-  or @{cite "paulson-ml2"}), with the types being determined by the
-  corresponding binders.  In contrast, free variables and constants
-  have an explicit name and type in each occurrence.
+  The language of terms is that of simply-typed \<open>\<lambda>\<close>-calculus with de-Bruijn
+  indices for bound variables (cf.\ @{cite debruijn72} or @{cite
+  "paulson-ml2"}), with the types being determined by the corresponding
+  binders. In contrast, free variables and constants have an explicit name and
+  type in each occurrence.
 
   \<^medskip>
-  A \<^emph>\<open>bound variable\<close> is a natural number \<open>b\<close>,
-  which accounts for the number of intermediate binders between the
-  variable occurrence in the body and its binding position.  For
-  example, the de-Bruijn term \<open>\<lambda>\<^bsub>bool\<^esub>. \<lambda>\<^bsub>bool\<^esub>. 1 \<and> 0\<close> would
-  correspond to \<open>\<lambda>x\<^bsub>bool\<^esub>. \<lambda>y\<^bsub>bool\<^esub>. x \<and> y\<close> in a named
-  representation.  Note that a bound variable may be represented by
-  different de-Bruijn indices at different occurrences, depending on
-  the nesting of abstractions.
+  A \<^emph>\<open>bound variable\<close> is a natural number \<open>b\<close>, which accounts for the number
+  of intermediate binders between the variable occurrence in the body and its
+  binding position. For example, the de-Bruijn term \<open>\<lambda>\<^bsub>bool\<^esub>. \<lambda>\<^bsub>bool\<^esub>. 1 \<and> 0\<close>
+  would correspond to \<open>\<lambda>x\<^bsub>bool\<^esub>. \<lambda>y\<^bsub>bool\<^esub>. x \<and> y\<close> in a named representation.
+  Note that a bound variable may be represented by different de-Bruijn indices
+  at different occurrences, depending on the nesting of abstractions.
 
-  A \<^emph>\<open>loose variable\<close> is a bound variable that is outside the
-  scope of local binders.  The types (and names) for loose variables
-  can be managed as a separate context, that is maintained as a stack
-  of hypothetical binders.  The core logic operates on closed terms,
-  without any loose variables.
+  A \<^emph>\<open>loose variable\<close> is a bound variable that is outside the scope of local
+  binders. The types (and names) for loose variables can be managed as a
+  separate context, that is maintained as a stack of hypothetical binders. The
+  core logic operates on closed terms, without any loose variables.
 
-  A \<^emph>\<open>fixed variable\<close> is a pair of a basic name and a type, e.g.\
-  \<open>(x, \<tau>)\<close> which is usually printed \<open>x\<^sub>\<tau>\<close> here.  A
-  \<^emph>\<open>schematic variable\<close> is a pair of an indexname and a type,
-  e.g.\ \<open>((x, 0), \<tau>)\<close> which is likewise printed as \<open>?x\<^sub>\<tau>\<close>.
+  A \<^emph>\<open>fixed variable\<close> is a pair of a basic name and a type, e.g.\ \<open>(x, \<tau>)\<close>
+  which is usually printed \<open>x\<^sub>\<tau>\<close> here. A \<^emph>\<open>schematic variable\<close> is a pair of an
+  indexname and a type, e.g.\ \<open>((x, 0), \<tau>)\<close> which is likewise printed as
+  \<open>?x\<^sub>\<tau>\<close>.
 
   \<^medskip>
-  A \<^emph>\<open>constant\<close> is a pair of a basic name and a type,
-  e.g.\ \<open>(c, \<tau>)\<close> which is usually printed as \<open>c\<^sub>\<tau>\<close>
-  here.  Constants are declared in the context as polymorphic families
-  \<open>c :: \<sigma>\<close>, meaning that all substitution instances \<open>c\<^sub>\<tau>\<close> for \<open>\<tau> = \<sigma>\<vartheta>\<close> are valid.
+  A \<^emph>\<open>constant\<close> is a pair of a basic name and a type, e.g.\ \<open>(c, \<tau>)\<close> which is
+  usually printed as \<open>c\<^sub>\<tau>\<close> here. Constants are declared in the context as
+  polymorphic families \<open>c :: \<sigma>\<close>, meaning that all substitution instances \<open>c\<^sub>\<tau>\<close>
+  for \<open>\<tau> = \<sigma>\<vartheta>\<close> are valid.
 
-  The vector of \<^emph>\<open>type arguments\<close> of constant \<open>c\<^sub>\<tau>\<close> wrt.\
-  the declaration \<open>c :: \<sigma>\<close> is defined as the codomain of the
-  matcher \<open>\<vartheta> = {?\<alpha>\<^sub>1 \<mapsto> \<tau>\<^sub>1, \<dots>, ?\<alpha>\<^sub>n \<mapsto> \<tau>\<^sub>n}\<close> presented in
-  canonical order \<open>(\<tau>\<^sub>1, \<dots>, \<tau>\<^sub>n)\<close>, corresponding to the
-  left-to-right occurrences of the \<open>\<alpha>\<^sub>i\<close> in \<open>\<sigma>\<close>.
-  Within a given theory context, there is a one-to-one correspondence
-  between any constant \<open>c\<^sub>\<tau>\<close> and the application \<open>c(\<tau>\<^sub>1,
-  \<dots>, \<tau>\<^sub>n)\<close> of its type arguments.  For example, with \<open>plus :: \<alpha>
-  \<Rightarrow> \<alpha> \<Rightarrow> \<alpha>\<close>, the instance \<open>plus\<^bsub>nat \<Rightarrow> nat \<Rightarrow> nat\<^esub>\<close> corresponds to
+  The vector of \<^emph>\<open>type arguments\<close> of constant \<open>c\<^sub>\<tau>\<close> wrt.\ the declaration \<open>c
+  :: \<sigma>\<close> is defined as the codomain of the matcher \<open>\<vartheta> = {?\<alpha>\<^sub>1 \<mapsto> \<tau>\<^sub>1,
+  \<dots>, ?\<alpha>\<^sub>n \<mapsto> \<tau>\<^sub>n}\<close> presented in canonical order \<open>(\<tau>\<^sub>1, \<dots>, \<tau>\<^sub>n)\<close>, corresponding
+  to the left-to-right occurrences of the \<open>\<alpha>\<^sub>i\<close> in \<open>\<sigma>\<close>. Within a given theory
+  context, there is a one-to-one correspondence between any constant \<open>c\<^sub>\<tau>\<close> and
+  the application \<open>c(\<tau>\<^sub>1, \<dots>, \<tau>\<^sub>n)\<close> of its type arguments. For example, with
+  \<open>plus :: \<alpha> \<Rightarrow> \<alpha> \<Rightarrow> \<alpha>\<close>, the instance \<open>plus\<^bsub>nat \<Rightarrow> nat \<Rightarrow> nat\<^esub>\<close> corresponds to
   \<open>plus(nat)\<close>.
 
-  Constant declarations \<open>c :: \<sigma>\<close> may contain sort constraints
-  for type variables in \<open>\<sigma>\<close>.  These are observed by
-  type-inference as expected, but \<^emph>\<open>ignored\<close> by the core logic.
-  This means the primitive logic is able to reason with instances of
-  polymorphic constants that the user-level type-checker would reject
-  due to violation of type class restrictions.
+  Constant declarations \<open>c :: \<sigma>\<close> may contain sort constraints for type
+  variables in \<open>\<sigma>\<close>. These are observed by type-inference as expected, but
+  \<^emph>\<open>ignored\<close> by the core logic. This means the primitive logic is able to
+  reason with instances of polymorphic constants that the user-level
+  type-checker would reject due to violation of type class restrictions.
 
   \<^medskip>
-  An \<^emph>\<open>atomic term\<close> is either a variable or constant.
-  The logical category \<^emph>\<open>term\<close> is defined inductively over atomic
-  terms, with abstraction and application as follows: \<open>t = b |
-  x\<^sub>\<tau> | ?x\<^sub>\<tau> | c\<^sub>\<tau> | \<lambda>\<^sub>\<tau>. t | t\<^sub>1 t\<^sub>2\<close>.  Parsing and printing takes care of
-  converting between an external representation with named bound
-  variables.  Subsequently, we shall use the latter notation instead
-  of internal de-Bruijn representation.
+  An \<^emph>\<open>atomic term\<close> is either a variable or constant. The logical category
+  \<^emph>\<open>term\<close> is defined inductively over atomic terms, with abstraction and
+  application as follows: \<open>t = b | x\<^sub>\<tau> | ?x\<^sub>\<tau> | c\<^sub>\<tau> | \<lambda>\<^sub>\<tau>. t | t\<^sub>1 t\<^sub>2\<close>.
+  Parsing and printing takes care of converting between an external
+  representation with named bound variables. Subsequently, we shall use the
+  latter notation instead of internal de-Bruijn representation.
 
-  The inductive relation \<open>t :: \<tau>\<close> assigns a (unique) type to a
-  term according to the structure of atomic terms, abstractions, and
-  applications:
+  The inductive relation \<open>t :: \<tau>\<close> assigns a (unique) type to a term according
+  to the structure of atomic terms, abstractions, and applications:
   \[
   \infer{\<open>a\<^sub>\<tau> :: \<tau>\<close>}{}
   \qquad
@@ -296,47 +269,46 @@
   \]
   A \<^emph>\<open>well-typed term\<close> is a term that can be typed according to these rules.
 
-  Typing information can be omitted: type-inference is able to
-  reconstruct the most general type of a raw term, while assigning
-  most general types to all of its variables and constants.
-  Type-inference depends on a context of type constraints for fixed
-  variables, and declarations for polymorphic constants.
+  Typing information can be omitted: type-inference is able to reconstruct the
+  most general type of a raw term, while assigning most general types to all
+  of its variables and constants. Type-inference depends on a context of type
+  constraints for fixed variables, and declarations for polymorphic constants.
 
   The identity of atomic terms consists both of the name and the type
-  component.  This means that different variables \<open>x\<^bsub>\<tau>\<^sub>1\<^esub>\<close> and \<open>x\<^bsub>\<tau>\<^sub>2\<^esub>\<close> may become the same after
-  type instantiation.  Type-inference rejects variables of the same
-  name, but different types.  In contrast, mixed instances of
+  component. This means that different variables \<open>x\<^bsub>\<tau>\<^sub>1\<^esub>\<close> and \<open>x\<^bsub>\<tau>\<^sub>2\<^esub>\<close> may
+  become the same after type instantiation. Type-inference rejects variables
+  of the same name, but different types. In contrast, mixed instances of
   polymorphic constants occur routinely.
 
   \<^medskip>
-  The \<^emph>\<open>hidden polymorphism\<close> of a term \<open>t :: \<sigma>\<close>
-  is the set of type variables occurring in \<open>t\<close>, but not in
-  its type \<open>\<sigma>\<close>.  This means that the term implicitly depends
-  on type arguments that are not accounted in the result type, i.e.\
-  there are different type instances \<open>t\<vartheta> :: \<sigma>\<close> and
-  \<open>t\<vartheta>' :: \<sigma>\<close> with the same type.  This slightly
-  pathological situation notoriously demands additional care.
+  The \<^emph>\<open>hidden polymorphism\<close> of a term \<open>t :: \<sigma>\<close> is the set of type variables
+  occurring in \<open>t\<close>, but not in its type \<open>\<sigma>\<close>. This means that the term
+  implicitly depends on type arguments that are not accounted in the result
+  type, i.e.\ there are different type instances \<open>t\<vartheta> :: \<sigma>\<close> and
+  \<open>t\<vartheta>' :: \<sigma>\<close> with the same type. This slightly pathological
+  situation notoriously demands additional care.
 
   \<^medskip>
-  A \<^emph>\<open>term abbreviation\<close> is a syntactic definition \<open>c\<^sub>\<sigma> \<equiv> t\<close> of a closed term \<open>t\<close> of type \<open>\<sigma>\<close>,
-  without any hidden polymorphism.  A term abbreviation looks like a
-  constant in the syntax, but is expanded before entering the logical
-  core.  Abbreviations are usually reverted when printing terms, using
-  \<open>t \<rightarrow> c\<^sub>\<sigma>\<close> as rules for higher-order rewriting.
+  A \<^emph>\<open>term abbreviation\<close> is a syntactic definition \<open>c\<^sub>\<sigma> \<equiv> t\<close> of a closed term
+  \<open>t\<close> of type \<open>\<sigma>\<close>, without any hidden polymorphism. A term abbreviation looks
+  like a constant in the syntax, but is expanded before entering the logical
+  core. Abbreviations are usually reverted when printing terms, using \<open>t \<rightarrow>
+  c\<^sub>\<sigma>\<close> as rules for higher-order rewriting.
 
   \<^medskip>
-  Canonical operations on \<open>\<lambda>\<close>-terms include \<open>\<alpha>\<beta>\<eta>\<close>-conversion: \<open>\<alpha>\<close>-conversion refers to capture-free
-  renaming of bound variables; \<open>\<beta>\<close>-conversion contracts an
-  abstraction applied to an argument term, substituting the argument
-  in the body: \<open>(\<lambda>x. b)a\<close> becomes \<open>b[a/x]\<close>; \<open>\<eta>\<close>-conversion contracts vacuous application-abstraction: \<open>\<lambda>x. f x\<close> becomes \<open>f\<close>, provided that the bound variable
-  does not occur in \<open>f\<close>.
+  Canonical operations on \<open>\<lambda>\<close>-terms include \<open>\<alpha>\<beta>\<eta>\<close>-conversion: \<open>\<alpha>\<close>-conversion
+  refers to capture-free renaming of bound variables; \<open>\<beta>\<close>-conversion contracts
+  an abstraction applied to an argument term, substituting the argument in the
+  body: \<open>(\<lambda>x. b)a\<close> becomes \<open>b[a/x]\<close>; \<open>\<eta>\<close>-conversion contracts vacuous
+  application-abstraction: \<open>\<lambda>x. f x\<close> becomes \<open>f\<close>, provided that the bound
+  variable does not occur in \<open>f\<close>.
 
-  Terms are normally treated modulo \<open>\<alpha>\<close>-conversion, which is
-  implicit in the de-Bruijn representation.  Names for bound variables
-  in abstractions are maintained separately as (meaningless) comments,
-  mostly for parsing and printing.  Full \<open>\<alpha>\<beta>\<eta>\<close>-conversion is
-  commonplace in various standard operations (\secref{sec:obj-rules})
-  that are based on higher-order unification and matching.
+  Terms are normally treated modulo \<open>\<alpha>\<close>-conversion, which is implicit in the
+  de-Bruijn representation. Names for bound variables in abstractions are
+  maintained separately as (meaningless) comments, mostly for parsing and
+  printing. Full \<open>\<alpha>\<beta>\<eta>\<close>-conversion is commonplace in various standard
+  operations (\secref{sec:obj-rules}) that are based on higher-order
+  unification and matching.
 \<close>
 
 text %mlref \<open>
@@ -361,56 +333,52 @@
   @{index_ML Sign.const_instance: "theory -> string * typ list -> typ"} \\
   \end{mldecls}
 
-  \<^descr> Type @{ML_type term} represents de-Bruijn terms, with comments
-  in abstractions, and explicitly named free variables and constants;
-  this is a datatype with constructors @{index_ML Bound}, @{index_ML
-  Free}, @{index_ML Var}, @{index_ML Const}, @{index_ML Abs},
-  @{index_ML_op "$"}.
+  \<^descr> Type @{ML_type term} represents de-Bruijn terms, with comments in
+  abstractions, and explicitly named free variables and constants; this is a
+  datatype with constructors @{index_ML Bound}, @{index_ML Free}, @{index_ML
+  Var}, @{index_ML Const}, @{index_ML Abs}, @{index_ML_op "$"}.
 
-  \<^descr> \<open>t\<close>~@{ML_text aconv}~\<open>u\<close> checks \<open>\<alpha>\<close>-equivalence of two terms.  This is the basic equality relation
-  on type @{ML_type term}; raw datatype equality should only be used
-  for operations related to parsing or printing!
-
-  \<^descr> @{ML Term.map_types}~\<open>f t\<close> applies the mapping \<open>f\<close> to all types occurring in \<open>t\<close>.
+  \<^descr> \<open>t\<close>~@{ML_text aconv}~\<open>u\<close> checks \<open>\<alpha>\<close>-equivalence of two terms. This is the
+  basic equality relation on type @{ML_type term}; raw datatype equality
+  should only be used for operations related to parsing or printing!
 
-  \<^descr> @{ML Term.fold_types}~\<open>f t\<close> iterates the operation
-  \<open>f\<close> over all occurrences of types in \<open>t\<close>; the term
-  structure is traversed from left to right.
+  \<^descr> @{ML Term.map_types}~\<open>f t\<close> applies the mapping \<open>f\<close> to all types occurring
+  in \<open>t\<close>.
+
+  \<^descr> @{ML Term.fold_types}~\<open>f t\<close> iterates the operation \<open>f\<close> over all
+  occurrences of types in \<open>t\<close>; the term structure is traversed from left to
+  right.
 
-  \<^descr> @{ML Term.map_aterms}~\<open>f t\<close> applies the mapping \<open>f\<close> to all atomic terms (@{ML Bound}, @{ML Free}, @{ML Var}, @{ML
-  Const}) occurring in \<open>t\<close>.
+  \<^descr> @{ML Term.map_aterms}~\<open>f t\<close> applies the mapping \<open>f\<close> to all atomic terms
+  (@{ML Bound}, @{ML Free}, @{ML Var}, @{ML Const}) occurring in \<open>t\<close>.
 
-  \<^descr> @{ML Term.fold_aterms}~\<open>f t\<close> iterates the operation
-  \<open>f\<close> over all occurrences of atomic terms (@{ML Bound}, @{ML
-  Free}, @{ML Var}, @{ML Const}) in \<open>t\<close>; the term structure is
-  traversed from left to right.
+  \<^descr> @{ML Term.fold_aterms}~\<open>f t\<close> iterates the operation \<open>f\<close> over all
+  occurrences of atomic terms (@{ML Bound}, @{ML Free}, @{ML Var}, @{ML
+  Const}) in \<open>t\<close>; the term structure is traversed from left to right.
 
-  \<^descr> @{ML fastype_of}~\<open>t\<close> determines the type of a
-  well-typed term.  This operation is relatively slow, despite the
-  omission of any sanity checks.
+  \<^descr> @{ML fastype_of}~\<open>t\<close> determines the type of a well-typed term. This
+  operation is relatively slow, despite the omission of any sanity checks.
 
-  \<^descr> @{ML lambda}~\<open>a b\<close> produces an abstraction \<open>\<lambda>a. b\<close>, where occurrences of the atomic term \<open>a\<close> in the
-  body \<open>b\<close> are replaced by bound variables.
+  \<^descr> @{ML lambda}~\<open>a b\<close> produces an abstraction \<open>\<lambda>a. b\<close>, where occurrences of
+  the atomic term \<open>a\<close> in the body \<open>b\<close> are replaced by bound variables.
 
-  \<^descr> @{ML betapply}~\<open>(t, u)\<close> produces an application \<open>t u\<close>, with topmost \<open>\<beta>\<close>-conversion if \<open>t\<close> is an
-  abstraction.
+  \<^descr> @{ML betapply}~\<open>(t, u)\<close> produces an application \<open>t u\<close>, with topmost
+  \<open>\<beta>\<close>-conversion if \<open>t\<close> is an abstraction.
 
-  \<^descr> @{ML incr_boundvars}~\<open>j\<close> increments a term's dangling
-  bound variables by the offset \<open>j\<close>.  This is required when
-  moving a subterm into a context where it is enclosed by a different
-  number of abstractions.  Bound variables with a matching abstraction
-  are unaffected.
+  \<^descr> @{ML incr_boundvars}~\<open>j\<close> increments a term's dangling bound variables by
+  the offset \<open>j\<close>. This is required when moving a subterm into a context where
+  it is enclosed by a different number of abstractions. Bound variables with a
+  matching abstraction are unaffected.
 
-  \<^descr> @{ML Sign.declare_const}~\<open>ctxt ((c, \<sigma>), mx)\<close> declares
-  a new constant \<open>c :: \<sigma>\<close> with optional mixfix syntax.
+  \<^descr> @{ML Sign.declare_const}~\<open>ctxt ((c, \<sigma>), mx)\<close> declares a new constant \<open>c ::
+  \<sigma>\<close> with optional mixfix syntax.
 
-  \<^descr> @{ML Sign.add_abbrev}~\<open>print_mode (c, t)\<close>
-  introduces a new term abbreviation \<open>c \<equiv> t\<close>.
+  \<^descr> @{ML Sign.add_abbrev}~\<open>print_mode (c, t)\<close> introduces a new term
+  abbreviation \<open>c \<equiv> t\<close>.
 
-  \<^descr> @{ML Sign.const_typargs}~\<open>thy (c, \<tau>)\<close> and @{ML
-  Sign.const_instance}~\<open>thy (c, [\<tau>\<^sub>1, \<dots>, \<tau>\<^sub>n])\<close>
-  convert between two representations of polymorphic constants: full
-  type instance vs.\ compact type arguments form.
+  \<^descr> @{ML Sign.const_typargs}~\<open>thy (c, \<tau>)\<close> and @{ML Sign.const_instance}~\<open>thy
+  (c, [\<tau>\<^sub>1, \<dots>, \<tau>\<^sub>n])\<close> convert between two representations of polymorphic
+  constants: full type instance vs.\ compact type arguments form.
 \<close>
 
 text %mlantiq \<open>
@@ -433,33 +401,31 @@
   @@{ML_antiquotation prop} prop
   \<close>}
 
-  \<^descr> \<open>@{const_name c}\<close> inlines the internalized logical
-  constant name \<open>c\<close> --- as @{ML_type string} literal.
+  \<^descr> \<open>@{const_name c}\<close> inlines the internalized logical constant name \<open>c\<close> ---
+  as @{ML_type string} literal.
+
+  \<^descr> \<open>@{const_abbrev c}\<close> inlines the internalized abbreviated constant name \<open>c\<close>
+  --- as @{ML_type string} literal.
 
-  \<^descr> \<open>@{const_abbrev c}\<close> inlines the internalized
-  abbreviated constant name \<open>c\<close> --- as @{ML_type string}
-  literal.
+  \<^descr> \<open>@{const c(\<^vec>\<tau>)}\<close> inlines the internalized constant \<open>c\<close> with precise
+  type instantiation in the sense of @{ML Sign.const_instance} --- as @{ML
+  Const} constructor term for datatype @{ML_type term}.
 
-  \<^descr> \<open>@{const c(\<^vec>\<tau>)}\<close> inlines the internalized
-  constant \<open>c\<close> with precise type instantiation in the sense of
-  @{ML Sign.const_instance} --- as @{ML Const} constructor term for
+  \<^descr> \<open>@{term t}\<close> inlines the internalized term \<open>t\<close> --- as constructor term for
   datatype @{ML_type term}.
 
-  \<^descr> \<open>@{term t}\<close> inlines the internalized term \<open>t\<close>
-  --- as constructor term for datatype @{ML_type term}.
-
-  \<^descr> \<open>@{prop \<phi>}\<close> inlines the internalized proposition
-  \<open>\<phi>\<close> --- as constructor term for datatype @{ML_type term}.
+  \<^descr> \<open>@{prop \<phi>}\<close> inlines the internalized proposition \<open>\<phi>\<close> --- as constructor
+  term for datatype @{ML_type term}.
 \<close>
 
 
 section \<open>Theorems \label{sec:thms}\<close>
 
 text \<open>
-  A \<^emph>\<open>proposition\<close> is a well-typed term of type \<open>prop\<close>, a
-  \<^emph>\<open>theorem\<close> is a proven proposition (depending on a context of
-  hypotheses and the background theory).  Primitive inferences include
-  plain Natural Deduction rules for the primary connectives \<open>\<And>\<close> and \<open>\<Longrightarrow>\<close> of the framework.  There is also a builtin
+  A \<^emph>\<open>proposition\<close> is a well-typed term of type \<open>prop\<close>, a \<^emph>\<open>theorem\<close> is a
+  proven proposition (depending on a context of hypotheses and the background
+  theory). Primitive inferences include plain Natural Deduction rules for the
+  primary connectives \<open>\<And>\<close> and \<open>\<Longrightarrow>\<close> of the framework. There is also a builtin
   notion of equality/equivalence \<open>\<equiv>\<close>.
 \<close>
 
@@ -467,16 +433,14 @@
 subsection \<open>Primitive connectives and rules \label{sec:prim-rules}\<close>
 
 text \<open>
-  The theory \<open>Pure\<close> contains constant declarations for the
-  primitive connectives \<open>\<And>\<close>, \<open>\<Longrightarrow>\<close>, and \<open>\<equiv>\<close> of
-  the logical framework, see \figref{fig:pure-connectives}.  The
-  derivability judgment \<open>A\<^sub>1, \<dots>, A\<^sub>n \<turnstile> B\<close> is
-  defined inductively by the primitive inferences given in
-  \figref{fig:prim-rules}, with the global restriction that the
-  hypotheses must \<^emph>\<open>not\<close> contain any schematic variables.  The
-  builtin equality is conceptually axiomatized as shown in
-  \figref{fig:pure-equality}, although the implementation works
-  directly with derived inferences.
+  The theory \<open>Pure\<close> contains constant declarations for the primitive
+  connectives \<open>\<And>\<close>, \<open>\<Longrightarrow>\<close>, and \<open>\<equiv>\<close> of the logical framework, see
+  \figref{fig:pure-connectives}. The derivability judgment \<open>A\<^sub>1, \<dots>, A\<^sub>n \<turnstile> B\<close>
+  is defined inductively by the primitive inferences given in
+  \figref{fig:prim-rules}, with the global restriction that the hypotheses
+  must \<^emph>\<open>not\<close> contain any schematic variables. The builtin equality is
+  conceptually axiomatized as shown in \figref{fig:pure-equality}, although
+  the implementation works directly with derived inferences.
 
   \begin{figure}[htb]
   \begin{center}
@@ -523,26 +487,29 @@
   \end{center}
   \end{figure}
 
-  The introduction and elimination rules for \<open>\<And>\<close> and \<open>\<Longrightarrow>\<close> are analogous to formation of dependently typed \<open>\<lambda>\<close>-terms representing the underlying proof objects.  Proof terms
-  are irrelevant in the Pure logic, though; they cannot occur within
-  propositions.  The system provides a runtime option to record
+  The introduction and elimination rules for \<open>\<And>\<close> and \<open>\<Longrightarrow>\<close> are analogous to
+  formation of dependently typed \<open>\<lambda>\<close>-terms representing the underlying proof
+  objects. Proof terms are irrelevant in the Pure logic, though; they cannot
+  occur within propositions. The system provides a runtime option to record
   explicit proof terms for primitive inferences, see also
-  \secref{sec:proof-terms}.  Thus all three levels of \<open>\<lambda>\<close>-calculus become explicit: \<open>\<Rightarrow>\<close> for terms, and \<open>\<And>/\<Longrightarrow>\<close> for proofs (cf.\ @{cite "Berghofer-Nipkow:2000:TPHOL"}).
+  \secref{sec:proof-terms}. Thus all three levels of \<open>\<lambda>\<close>-calculus become
+  explicit: \<open>\<Rightarrow>\<close> for terms, and \<open>\<And>/\<Longrightarrow>\<close> for proofs (cf.\ @{cite
+  "Berghofer-Nipkow:2000:TPHOL"}).
 
-  Observe that locally fixed parameters (as in \<open>\<And>\<hyphen>intro\<close>) need not be recorded in the hypotheses, because
-  the simple syntactic types of Pure are always inhabitable.
-  ``Assumptions'' \<open>x :: \<tau>\<close> for type-membership are only
-  present as long as some \<open>x\<^sub>\<tau>\<close> occurs in the statement
-  body.\<^footnote>\<open>This is the key difference to ``\<open>\<lambda>HOL\<close>'' in
-  the PTS framework @{cite "Barendregt-Geuvers:2001"}, where hypotheses
-  \<open>x : A\<close> are treated uniformly for propositions and types.\<close>
+  Observe that locally fixed parameters (as in \<open>\<And>\<hyphen>intro\<close>) need not be recorded
+  in the hypotheses, because the simple syntactic types of Pure are always
+  inhabitable. ``Assumptions'' \<open>x :: \<tau>\<close> for type-membership are only present
+  as long as some \<open>x\<^sub>\<tau>\<close> occurs in the statement body.\<^footnote>\<open>This is the key
+  difference to ``\<open>\<lambda>HOL\<close>'' in the PTS framework @{cite
+  "Barendregt-Geuvers:2001"}, where hypotheses \<open>x : A\<close> are treated uniformly
+  for propositions and types.\<close>
 
   \<^medskip>
-  The axiomatization of a theory is implicitly closed by
-  forming all instances of type and term variables: \<open>\<turnstile>
-  A\<vartheta>\<close> holds for any substitution instance of an axiom
-  \<open>\<turnstile> A\<close>.  By pushing substitutions through derivations
-  inductively, we also get admissible \<open>generalize\<close> and \<open>instantiate\<close> rules as shown in \figref{fig:subst-rules}.
+  The axiomatization of a theory is implicitly closed by forming all instances
+  of type and term variables: \<open>\<turnstile> A\<vartheta>\<close> holds for any substitution
+  instance of an axiom \<open>\<turnstile> A\<close>. By pushing substitutions through derivations
+  inductively, we also get admissible \<open>generalize\<close> and \<open>instantiate\<close> rules as
+  shown in \figref{fig:subst-rules}.
 
   \begin{figure}[htb]
   \begin{center}
@@ -560,40 +527,39 @@
   \end{center}
   \end{figure}
 
-  Note that \<open>instantiate\<close> does not require an explicit
-  side-condition, because \<open>\<Gamma>\<close> may never contain schematic
-  variables.
+  Note that \<open>instantiate\<close> does not require an explicit side-condition, because
+  \<open>\<Gamma>\<close> may never contain schematic variables.
 
-  In principle, variables could be substituted in hypotheses as well,
-  but this would disrupt the monotonicity of reasoning: deriving
-  \<open>\<Gamma>\<vartheta> \<turnstile> B\<vartheta>\<close> from \<open>\<Gamma> \<turnstile> B\<close> is
-  correct, but \<open>\<Gamma>\<vartheta> \<supseteq> \<Gamma>\<close> does not necessarily hold:
-  the result belongs to a different proof context.
+  In principle, variables could be substituted in hypotheses as well, but this
+  would disrupt the monotonicity of reasoning: deriving \<open>\<Gamma>\<vartheta> \<turnstile>
+  B\<vartheta>\<close> from \<open>\<Gamma> \<turnstile> B\<close> is correct, but \<open>\<Gamma>\<vartheta> \<supseteq> \<Gamma>\<close> does not
+  necessarily hold: the result belongs to a different proof context.
 
   \<^medskip>
-  An \<^emph>\<open>oracle\<close> is a function that produces axioms on the
-  fly.  Logically, this is an instance of the \<open>axiom\<close> rule
-  (\figref{fig:prim-rules}), but there is an operational difference.
-  The system always records oracle invocations within derivations of
-  theorems by a unique tag.
+  An \<^emph>\<open>oracle\<close> is a function that produces axioms on the fly. Logically, this
+  is an instance of the \<open>axiom\<close> rule (\figref{fig:prim-rules}), but there is
+  an operational difference. The system always records oracle invocations
+  within derivations of theorems by a unique tag.
 
-  Axiomatizations should be limited to the bare minimum, typically as
-  part of the initial logical basis of an object-logic formalization.
-  Later on, theories are usually developed in a strictly definitional
-  fashion, by stating only certain equalities over new constants.
+  Axiomatizations should be limited to the bare minimum, typically as part of
+  the initial logical basis of an object-logic formalization. Later on,
+  theories are usually developed in a strictly definitional fashion, by
+  stating only certain equalities over new constants.
 
-  A \<^emph>\<open>simple definition\<close> consists of a constant declaration \<open>c :: \<sigma>\<close> together with an axiom \<open>\<turnstile> c \<equiv> t\<close>, where \<open>t
-  :: \<sigma>\<close> is a closed term without any hidden polymorphism.  The RHS
-  may depend on further defined constants, but not \<open>c\<close> itself.
-  Definitions of functions may be presented as \<open>c \<^vec>x \<equiv>
-  t\<close> instead of the puristic \<open>c \<equiv> \<lambda>\<^vec>x. t\<close>.
+  A \<^emph>\<open>simple definition\<close> consists of a constant declaration \<open>c :: \<sigma>\<close> together
+  with an axiom \<open>\<turnstile> c \<equiv> t\<close>, where \<open>t :: \<sigma>\<close> is a closed term without any hidden
+  polymorphism. The RHS may depend on further defined constants, but not \<open>c\<close>
+  itself. Definitions of functions may be presented as \<open>c \<^vec>x \<equiv> t\<close>
+  instead of the puristic \<open>c \<equiv> \<lambda>\<^vec>x. t\<close>.
 
-  An \<^emph>\<open>overloaded definition\<close> consists of a collection of axioms
-  for the same constant, with zero or one equations \<open>c((\<^vec>\<alpha>)\<kappa>) \<equiv> t\<close> for each type constructor \<open>\<kappa>\<close> (for
-  distinct variables \<open>\<^vec>\<alpha>\<close>).  The RHS may mention
-  previously defined constants as above, or arbitrary constants \<open>d(\<alpha>\<^sub>i)\<close> for some \<open>\<alpha>\<^sub>i\<close> projected from \<open>\<^vec>\<alpha>\<close>.  Thus overloaded definitions essentially work by
-  primitive recursion over the syntactic structure of a single type
-  argument.  See also @{cite \<open>\S4.3\<close> "Haftmann-Wenzel:2006:classes"}.
+  An \<^emph>\<open>overloaded definition\<close> consists of a collection of axioms for the same
+  constant, with zero or one equations \<open>c((\<^vec>\<alpha>)\<kappa>) \<equiv> t\<close> for each type
+  constructor \<open>\<kappa>\<close> (for distinct variables \<open>\<^vec>\<alpha>\<close>). The RHS may mention
+  previously defined constants as above, or arbitrary constants \<open>d(\<alpha>\<^sub>i)\<close> for
+  some \<open>\<alpha>\<^sub>i\<close> projected from \<open>\<^vec>\<alpha>\<close>. Thus overloaded definitions
+  essentially work by primitive recursion over the syntactic structure of a
+  single type argument. See also @{cite \<open>\S4.3\<close>
+  "Haftmann-Wenzel:2006:classes"}.
 \<close>
 
 text %mlref \<open>
@@ -635,101 +601,89 @@
   Defs.entry -> Defs.entry list -> theory -> theory"} \\
   \end{mldecls}
 
-  \<^descr> @{ML Thm.peek_status}~\<open>thm\<close> informs about the current
-  status of the derivation object behind the given theorem.  This is a
-  snapshot of a potentially ongoing (parallel) evaluation of proofs.
-  The three Boolean values indicate the following: \<^verbatim>\<open>oracle\<close>
-  if the finished part contains some oracle invocation; \<^verbatim>\<open>unfinished\<close>
-  if some future proofs are still pending; \<^verbatim>\<open>failed\<close> if some future
-  proof has failed, rendering the theorem invalid!
+  \<^descr> @{ML Thm.peek_status}~\<open>thm\<close> informs about the current status of the
+  derivation object behind the given theorem. This is a snapshot of a
+  potentially ongoing (parallel) evaluation of proofs. The three Boolean
+  values indicate the following: \<^verbatim>\<open>oracle\<close> if the finished part contains some
+  oracle invocation; \<^verbatim>\<open>unfinished\<close> if some future proofs are still pending;
+  \<^verbatim>\<open>failed\<close> if some future proof has failed, rendering the theorem invalid!
 
-  \<^descr> @{ML Logic.all}~\<open>a B\<close> produces a Pure quantification
-  \<open>\<And>a. B\<close>, where occurrences of the atomic term \<open>a\<close> in
-  the body proposition \<open>B\<close> are replaced by bound variables.
-  (See also @{ML lambda} on terms.)
+  \<^descr> @{ML Logic.all}~\<open>a B\<close> produces a Pure quantification \<open>\<And>a. B\<close>, where
+  occurrences of the atomic term \<open>a\<close> in the body proposition \<open>B\<close> are replaced
+  by bound variables. (See also @{ML lambda} on terms.)
 
-  \<^descr> @{ML Logic.mk_implies}~\<open>(A, B)\<close> produces a Pure
-  implication \<open>A \<Longrightarrow> B\<close>.
+  \<^descr> @{ML Logic.mk_implies}~\<open>(A, B)\<close> produces a Pure implication \<open>A \<Longrightarrow> B\<close>.
 
-  \<^descr> Types @{ML_type ctyp} and @{ML_type cterm} represent certified
-  types and terms, respectively.  These are abstract datatypes that
-  guarantee that its values have passed the full well-formedness (and
-  well-typedness) checks, relative to the declarations of type
-  constructors, constants etc.\ in the background theory.  The
-  abstract types @{ML_type ctyp} and @{ML_type cterm} are part of the
-  same inference kernel that is mainly responsible for @{ML_type thm}.
-  Thus syntactic operations on @{ML_type ctyp} and @{ML_type cterm}
-  are located in the @{ML_structure Thm} module, even though theorems are
-  not yet involved at that stage.
+  \<^descr> Types @{ML_type ctyp} and @{ML_type cterm} represent certified types and
+  terms, respectively. These are abstract datatypes that guarantee that its
+  values have passed the full well-formedness (and well-typedness) checks,
+  relative to the declarations of type constructors, constants etc.\ in the
+  background theory. The abstract types @{ML_type ctyp} and @{ML_type cterm}
+  are part of the same inference kernel that is mainly responsible for
+  @{ML_type thm}. Thus syntactic operations on @{ML_type ctyp} and @{ML_type
+  cterm} are located in the @{ML_structure Thm} module, even though theorems
+  are not yet involved at that stage.
 
-  \<^descr> @{ML Thm.ctyp_of}~\<open>ctxt \<tau>\<close> and @{ML
-  Thm.cterm_of}~\<open>ctxt t\<close> explicitly check types and terms,
-  respectively.  This also involves some basic normalizations, such
-  expansion of type and term abbreviations from the underlying
-  theory context.
-  Full re-certification is relatively slow and should be avoided in
-  tight reasoning loops.
+  \<^descr> @{ML Thm.ctyp_of}~\<open>ctxt \<tau>\<close> and @{ML Thm.cterm_of}~\<open>ctxt t\<close> explicitly
+  check types and terms, respectively. This also involves some basic
+  normalizations, such expansion of type and term abbreviations from the
+  underlying theory context. Full re-certification is relatively slow and
+  should be avoided in tight reasoning loops.
 
-  \<^descr> @{ML Thm.apply}, @{ML Thm.lambda}, @{ML Thm.all}, @{ML
-  Drule.mk_implies} etc.\ compose certified terms (or propositions)
-  incrementally.  This is equivalent to @{ML Thm.cterm_of} after
-  unchecked @{ML_op "$"}, @{ML lambda}, @{ML Logic.all}, @{ML
-  Logic.mk_implies} etc., but there can be a big difference in
-  performance when large existing entities are composed by a few extra
-  constructions on top.  There are separate operations to decompose
+  \<^descr> @{ML Thm.apply}, @{ML Thm.lambda}, @{ML Thm.all}, @{ML Drule.mk_implies}
+  etc.\ compose certified terms (or propositions) incrementally. This is
+  equivalent to @{ML Thm.cterm_of} after unchecked @{ML_op "$"}, @{ML lambda},
+  @{ML Logic.all}, @{ML Logic.mk_implies} etc., but there can be a big
+  difference in performance when large existing entities are composed by a few
+  extra constructions on top. There are separate operations to decompose
   certified terms and theorems to produce certified terms again.
 
-  \<^descr> Type @{ML_type thm} represents proven propositions.  This is
-  an abstract datatype that guarantees that its values have been
-  constructed by basic principles of the @{ML_structure Thm} module.
-  Every @{ML_type thm} value refers its background theory,
-  cf.\ \secref{sec:context-theory}.
+  \<^descr> Type @{ML_type thm} represents proven propositions. This is an abstract
+  datatype that guarantees that its values have been constructed by basic
+  principles of the @{ML_structure Thm} module. Every @{ML_type thm} value
+  refers its background theory, cf.\ \secref{sec:context-theory}.
 
-  \<^descr> @{ML Thm.transfer}~\<open>thy thm\<close> transfers the given
-  theorem to a \<^emph>\<open>larger\<close> theory, see also \secref{sec:context}.
-  This formal adjustment of the background context has no logical
-  significance, but is occasionally required for formal reasons, e.g.\
-  when theorems that are imported from more basic theories are used in
-  the current situation.
+  \<^descr> @{ML Thm.transfer}~\<open>thy thm\<close> transfers the given theorem to a \<^emph>\<open>larger\<close>
+  theory, see also \secref{sec:context}. This formal adjustment of the
+  background context has no logical significance, but is occasionally required
+  for formal reasons, e.g.\ when theorems that are imported from more basic
+  theories are used in the current situation.
 
-  \<^descr> @{ML Thm.assume}, @{ML Thm.forall_intr}, @{ML
-  Thm.forall_elim}, @{ML Thm.implies_intr}, and @{ML Thm.implies_elim}
-  correspond to the primitive inferences of \figref{fig:prim-rules}.
+  \<^descr> @{ML Thm.assume}, @{ML Thm.forall_intr}, @{ML Thm.forall_elim}, @{ML
+  Thm.implies_intr}, and @{ML Thm.implies_elim} correspond to the primitive
+  inferences of \figref{fig:prim-rules}.
 
-  \<^descr> @{ML Thm.generalize}~\<open>(\<^vec>\<alpha>, \<^vec>x)\<close>
-  corresponds to the \<open>generalize\<close> rules of
-  \figref{fig:subst-rules}.  Here collections of type and term
-  variables are generalized simultaneously, specified by the given
-  basic names.
+  \<^descr> @{ML Thm.generalize}~\<open>(\<^vec>\<alpha>, \<^vec>x)\<close> corresponds to the
+  \<open>generalize\<close> rules of \figref{fig:subst-rules}. Here collections of type and
+  term variables are generalized simultaneously, specified by the given basic
+  names.
 
-  \<^descr> @{ML Thm.instantiate}~\<open>(\<^vec>\<alpha>\<^sub>s,
-  \<^vec>x\<^sub>\<tau>)\<close> corresponds to the \<open>instantiate\<close> rules
-  of \figref{fig:subst-rules}.  Type variables are substituted before
-  term variables.  Note that the types in \<open>\<^vec>x\<^sub>\<tau>\<close>
-  refer to the instantiated versions.
+  \<^descr> @{ML Thm.instantiate}~\<open>(\<^vec>\<alpha>\<^sub>s, \<^vec>x\<^sub>\<tau>)\<close> corresponds to the
+  \<open>instantiate\<close> rules of \figref{fig:subst-rules}. Type variables are
+  substituted before term variables. Note that the types in \<open>\<^vec>x\<^sub>\<tau>\<close> refer
+  to the instantiated versions.
 
-  \<^descr> @{ML Thm.add_axiom}~\<open>ctxt (name, A)\<close> declares an
-  arbitrary proposition as axiom, and retrieves it as a theorem from
-  the resulting theory, cf.\ \<open>axiom\<close> in
-  \figref{fig:prim-rules}.  Note that the low-level representation in
-  the axiom table may differ slightly from the returned theorem.
+  \<^descr> @{ML Thm.add_axiom}~\<open>ctxt (name, A)\<close> declares an arbitrary proposition as
+  axiom, and retrieves it as a theorem from the resulting theory, cf.\ \<open>axiom\<close>
+  in \figref{fig:prim-rules}. Note that the low-level representation in the
+  axiom table may differ slightly from the returned theorem.
 
-  \<^descr> @{ML Thm.add_oracle}~\<open>(binding, oracle)\<close> produces a named
-  oracle rule, essentially generating arbitrary axioms on the fly,
-  cf.\ \<open>axiom\<close> in \figref{fig:prim-rules}.
+  \<^descr> @{ML Thm.add_oracle}~\<open>(binding, oracle)\<close> produces a named oracle rule,
+  essentially generating arbitrary axioms on the fly, cf.\ \<open>axiom\<close> in
+  \figref{fig:prim-rules}.
 
-  \<^descr> @{ML Thm.add_def}~\<open>ctxt unchecked overloaded (name, c
-  \<^vec>x \<equiv> t)\<close> states a definitional axiom for an existing constant
-  \<open>c\<close>.  Dependencies are recorded via @{ML Theory.add_deps},
-  unless the \<open>unchecked\<close> option is set.  Note that the
-  low-level representation in the axiom table may differ slightly from
-  the returned theorem.
+  \<^descr> @{ML Thm.add_def}~\<open>ctxt unchecked overloaded (name, c \<^vec>x \<equiv> t)\<close>
+  states a definitional axiom for an existing constant \<open>c\<close>. Dependencies are
+  recorded via @{ML Theory.add_deps}, unless the \<open>unchecked\<close> option is set.
+  Note that the low-level representation in the axiom table may differ
+  slightly from the returned theorem.
 
-  \<^descr> @{ML Theory.add_deps}~\<open>ctxt name c\<^sub>\<tau> \<^vec>d\<^sub>\<sigma>\<close>
-  declares dependencies of a named specification for constant \<open>c\<^sub>\<tau>\<close>, relative to existing specifications for constants \<open>\<^vec>d\<^sub>\<sigma>\<close>.  This also works for type constructors.
+  \<^descr> @{ML Theory.add_deps}~\<open>ctxt name c\<^sub>\<tau> \<^vec>d\<^sub>\<sigma>\<close> declares dependencies of
+  a named specification for constant \<open>c\<^sub>\<tau>\<close>, relative to existing
+  specifications for constants \<open>\<^vec>d\<^sub>\<sigma>\<close>. This also works for type
+  constructors.
 \<close>
 
-
 text %mlantiq \<open>
   \begin{matharray}{rcl}
   @{ML_antiquotation_def "ctyp"} & : & \<open>ML_antiquotation\<close> \\
@@ -755,46 +709,42 @@
     @'by' method method?
   \<close>}
 
-  \<^descr> \<open>@{ctyp \<tau>}\<close> produces a certified type wrt.\ the
-  current background theory --- as abstract value of type @{ML_type
-  ctyp}.
+  \<^descr> \<open>@{ctyp \<tau>}\<close> produces a certified type wrt.\ the current background theory
+  --- as abstract value of type @{ML_type ctyp}.
+
+  \<^descr> \<open>@{cterm t}\<close> and \<open>@{cprop \<phi>}\<close> produce a certified term wrt.\ the current
+  background theory --- as abstract value of type @{ML_type cterm}.
 
-  \<^descr> \<open>@{cterm t}\<close> and \<open>@{cprop \<phi>}\<close> produce a
-  certified term wrt.\ the current background theory --- as abstract
-  value of type @{ML_type cterm}.
+  \<^descr> \<open>@{thm a}\<close> produces a singleton fact --- as abstract value of type
+  @{ML_type thm}.
 
-  \<^descr> \<open>@{thm a}\<close> produces a singleton fact --- as abstract
-  value of type @{ML_type thm}.
-
-  \<^descr> \<open>@{thms a}\<close> produces a general fact --- as abstract
-  value of type @{ML_type "thm list"}.
+  \<^descr> \<open>@{thms a}\<close> produces a general fact --- as abstract value of type
+  @{ML_type "thm list"}.
 
-  \<^descr> \<open>@{lemma \<phi> by meth}\<close> produces a fact that is proven on
-  the spot according to the minimal proof, which imitates a terminal
-  Isar proof.  The result is an abstract value of type @{ML_type thm}
-  or @{ML_type "thm list"}, depending on the number of propositions
-  given here.
+  \<^descr> \<open>@{lemma \<phi> by meth}\<close> produces a fact that is proven on the spot according
+  to the minimal proof, which imitates a terminal Isar proof. The result is an
+  abstract value of type @{ML_type thm} or @{ML_type "thm list"}, depending on
+  the number of propositions given here.
 
-  The internal derivation object lacks a proper theorem name, but it
-  is formally closed, unless the \<open>(open)\<close> option is specified
-  (this may impact performance of applications with proof terms).
+  The internal derivation object lacks a proper theorem name, but it is
+  formally closed, unless the \<open>(open)\<close> option is specified (this may impact
+  performance of applications with proof terms).
 
-  Since ML antiquotations are always evaluated at compile-time, there
-  is no run-time overhead even for non-trivial proofs.  Nonetheless,
-  the justification is syntactically limited to a single @{command
-  "by"} step.  More complex Isar proofs should be done in regular
-  theory source, before compiling the corresponding ML text that uses
-  the result.
+  Since ML antiquotations are always evaluated at compile-time, there is no
+  run-time overhead even for non-trivial proofs. Nonetheless, the
+  justification is syntactically limited to a single @{command "by"} step.
+  More complex Isar proofs should be done in regular theory source, before
+  compiling the corresponding ML text that uses the result.
 \<close>
 
 
 subsection \<open>Auxiliary connectives \label{sec:logic-aux}\<close>
 
-text \<open>Theory \<open>Pure\<close> provides a few auxiliary connectives
-  that are defined on top of the primitive ones, see
-  \figref{fig:pure-aux}.  These special constants are useful in
-  certain internal encodings, and are normally not directly exposed to
-  the user.
+text \<open>
+  Theory \<open>Pure\<close> provides a few auxiliary connectives that are defined on top
+  of the primitive ones, see \figref{fig:pure-aux}. These special constants
+  are useful in certain internal encodings, and are normally not directly
+  exposed to the user.
 
   \begin{figure}[htb]
   \begin{center}
@@ -812,37 +762,32 @@
   \end{center}
   \end{figure}
 
-  The introduction \<open>A \<Longrightarrow> B \<Longrightarrow> A &&& B\<close>, and eliminations
-  (projections) \<open>A &&& B \<Longrightarrow> A\<close> and \<open>A &&& B \<Longrightarrow> B\<close> are
-  available as derived rules.  Conjunction allows to treat
-  simultaneous assumptions and conclusions uniformly, e.g.\ consider
-  \<open>A \<Longrightarrow> B \<Longrightarrow> C &&& D\<close>.  In particular, the goal mechanism
-  represents multiple claims as explicit conjunction internally, but
-  this is refined (via backwards introduction) into separate sub-goals
-  before the user commences the proof; the final result is projected
-  into a list of theorems using eliminations (cf.\
-  \secref{sec:tactical-goals}).
+  The introduction \<open>A \<Longrightarrow> B \<Longrightarrow> A &&& B\<close>, and eliminations (projections) \<open>A &&& B
+  \<Longrightarrow> A\<close> and \<open>A &&& B \<Longrightarrow> B\<close> are available as derived rules. Conjunction allows to
+  treat simultaneous assumptions and conclusions uniformly, e.g.\ consider \<open>A
+  \<Longrightarrow> B \<Longrightarrow> C &&& D\<close>. In particular, the goal mechanism represents multiple claims
+  as explicit conjunction internally, but this is refined (via backwards
+  introduction) into separate sub-goals before the user commences the proof;
+  the final result is projected into a list of theorems using eliminations
+  (cf.\ \secref{sec:tactical-goals}).
 
-  The \<open>prop\<close> marker (\<open>#\<close>) makes arbitrarily complex
-  propositions appear as atomic, without changing the meaning: \<open>\<Gamma> \<turnstile> A\<close> and \<open>\<Gamma> \<turnstile> #A\<close> are interchangeable.  See
-  \secref{sec:tactical-goals} for specific operations.
+  The \<open>prop\<close> marker (\<open>#\<close>) makes arbitrarily complex propositions appear as
+  atomic, without changing the meaning: \<open>\<Gamma> \<turnstile> A\<close> and \<open>\<Gamma> \<turnstile> #A\<close> are
+  interchangeable. See \secref{sec:tactical-goals} for specific operations.
 
-  The \<open>term\<close> marker turns any well-typed term into a derivable
-  proposition: \<open>\<turnstile> TERM t\<close> holds unconditionally.  Although
-  this is logically vacuous, it allows to treat terms and proofs
-  uniformly, similar to a type-theoretic framework.
+  The \<open>term\<close> marker turns any well-typed term into a derivable proposition: \<open>\<turnstile>
+  TERM t\<close> holds unconditionally. Although this is logically vacuous, it allows
+  to treat terms and proofs uniformly, similar to a type-theoretic framework.
 
-  The \<open>TYPE\<close> constructor is the canonical representative of
-  the unspecified type \<open>\<alpha> itself\<close>; it essentially injects the
-  language of types into that of terms.  There is specific notation
-  \<open>TYPE(\<tau>)\<close> for \<open>TYPE\<^bsub>\<tau> itself\<^esub>\<close>.
-  Although being devoid of any particular meaning, the term \<open>TYPE(\<tau>)\<close> accounts for the type \<open>\<tau>\<close> within the term
-  language.  In particular, \<open>TYPE(\<alpha>)\<close> may be used as formal
-  argument in primitive definitions, in order to circumvent hidden
-  polymorphism (cf.\ \secref{sec:terms}).  For example, \<open>c
-  TYPE(\<alpha>) \<equiv> A[\<alpha>]\<close> defines \<open>c :: \<alpha> itself \<Rightarrow> prop\<close> in terms of
-  a proposition \<open>A\<close> that depends on an additional type
-  argument, which is essentially a predicate on types.
+  The \<open>TYPE\<close> constructor is the canonical representative of the unspecified
+  type \<open>\<alpha> itself\<close>; it essentially injects the language of types into that of
+  terms. There is specific notation \<open>TYPE(\<tau>)\<close> for \<open>TYPE\<^bsub>\<tau> itself\<^esub>\<close>. Although
+  being devoid of any particular meaning, the term \<open>TYPE(\<tau>)\<close> accounts for the
+  type \<open>\<tau>\<close> within the term language. In particular, \<open>TYPE(\<alpha>)\<close> may be used as
+  formal argument in primitive definitions, in order to circumvent hidden
+  polymorphism (cf.\ \secref{sec:terms}). For example, \<open>c TYPE(\<alpha>) \<equiv> A[\<alpha>]\<close>
+  defines \<open>c :: \<alpha> itself \<Rightarrow> prop\<close> in terms of a proposition \<open>A\<close> that depends on
+  an additional type argument, which is essentially a predicate on types.
 \<close>
 
 text %mlref \<open>
@@ -857,8 +802,7 @@
 
   \<^descr> @{ML Conjunction.intr} derives \<open>A &&& B\<close> from \<open>A\<close> and \<open>B\<close>.
 
-  \<^descr> @{ML Conjunction.elim} derives \<open>A\<close> and \<open>B\<close>
-  from \<open>A &&& B\<close>.
+  \<^descr> @{ML Conjunction.elim} derives \<open>A\<close> and \<open>B\<close> from \<open>A &&& B\<close>.
 
   \<^descr> @{ML Drule.mk_term} derives \<open>TERM t\<close>.
 
@@ -866,35 +810,35 @@
 
   \<^descr> @{ML Logic.mk_type}~\<open>\<tau>\<close> produces the term \<open>TYPE(\<tau>)\<close>.
 
-  \<^descr> @{ML Logic.dest_type}~\<open>TYPE(\<tau>)\<close> recovers the type
-  \<open>\<tau>\<close>.
+  \<^descr> @{ML Logic.dest_type}~\<open>TYPE(\<tau>)\<close> recovers the type \<open>\<tau>\<close>.
 \<close>
 
 
 subsection \<open>Sort hypotheses\<close>
 
-text \<open>Type variables are decorated with sorts, as explained in
-  \secref{sec:types}.  This constrains type instantiation to certain
-  ranges of types: variable \<open>\<alpha>\<^sub>s\<close> may only be assigned to types
-  \<open>\<tau>\<close> that belong to sort \<open>s\<close>.  Within the logic, sort
-  constraints act like implicit preconditions on the result \<open>\<lparr>\<alpha>\<^sub>1 : s\<^sub>1\<rparr>, \<dots>, \<lparr>\<alpha>\<^sub>n : s\<^sub>n\<rparr>, \<Gamma> \<turnstile> \<phi>\<close> where the type variables \<open>\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>n\<close> cover the propositions \<open>\<Gamma>\<close>, \<open>\<phi>\<close>, as
-  well as the proof of \<open>\<Gamma> \<turnstile> \<phi>\<close>.
+text \<open>
+  Type variables are decorated with sorts, as explained in \secref{sec:types}.
+  This constrains type instantiation to certain ranges of types: variable
+  \<open>\<alpha>\<^sub>s\<close> may only be assigned to types \<open>\<tau>\<close> that belong to sort \<open>s\<close>. Within the
+  logic, sort constraints act like implicit preconditions on the result \<open>\<lparr>\<alpha>\<^sub>1
+  : s\<^sub>1\<rparr>, \<dots>, \<lparr>\<alpha>\<^sub>n : s\<^sub>n\<rparr>, \<Gamma> \<turnstile> \<phi>\<close> where the type variables \<open>\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>n\<close> cover
+  the propositions \<open>\<Gamma>\<close>, \<open>\<phi>\<close>, as well as the proof of \<open>\<Gamma> \<turnstile> \<phi>\<close>.
 
-  These \<^emph>\<open>sort hypothesis\<close> of a theorem are passed monotonically
-  through further derivations.  They are redundant, as long as the
-  statement of a theorem still contains the type variables that are
-  accounted here.  The logical significance of sort hypotheses is
-  limited to the boundary case where type variables disappear from the
-  proposition, e.g.\ \<open>\<lparr>\<alpha>\<^sub>s : s\<rparr> \<turnstile> \<phi>\<close>.  Since such dangling type
-  variables can be renamed arbitrarily without changing the
-  proposition \<open>\<phi>\<close>, the inference kernel maintains sort
-  hypotheses in anonymous form \<open>s \<turnstile> \<phi>\<close>.
+  These \<^emph>\<open>sort hypothesis\<close> of a theorem are passed monotonically through
+  further derivations. They are redundant, as long as the statement of a
+  theorem still contains the type variables that are accounted here. The
+  logical significance of sort hypotheses is limited to the boundary case
+  where type variables disappear from the proposition, e.g.\ \<open>\<lparr>\<alpha>\<^sub>s : s\<rparr> \<turnstile> \<phi>\<close>.
+  Since such dangling type variables can be renamed arbitrarily without
+  changing the proposition \<open>\<phi>\<close>, the inference kernel maintains sort hypotheses
+  in anonymous form \<open>s \<turnstile> \<phi>\<close>.
 
-  In most practical situations, such extra sort hypotheses may be
-  stripped in a final bookkeeping step, e.g.\ at the end of a proof:
-  they are typically left over from intermediate reasoning with type
-  classes that can be satisfied by some concrete type \<open>\<tau>\<close> of
-  sort \<open>s\<close> to replace the hypothetical type variable \<open>\<alpha>\<^sub>s\<close>.\<close>
+  In most practical situations, such extra sort hypotheses may be stripped in
+  a final bookkeeping step, e.g.\ at the end of a proof: they are typically
+  left over from intermediate reasoning with type classes that can be
+  satisfied by some concrete type \<open>\<tau>\<close> of sort \<open>s\<close> to replace the hypothetical
+  type variable \<open>\<alpha>\<^sub>s\<close>.
+\<close>
 
 text %mlref \<open>
   \begin{mldecls}
@@ -902,17 +846,18 @@
   @{index_ML Thm.strip_shyps: "thm -> thm"} \\
   \end{mldecls}
 
-  \<^descr> @{ML Thm.extra_shyps}~\<open>thm\<close> determines the extraneous
-  sort hypotheses of the given theorem, i.e.\ the sorts that are not
-  present within type variables of the statement.
+  \<^descr> @{ML Thm.extra_shyps}~\<open>thm\<close> determines the extraneous sort hypotheses of
+  the given theorem, i.e.\ the sorts that are not present within type
+  variables of the statement.
 
-  \<^descr> @{ML Thm.strip_shyps}~\<open>thm\<close> removes any extraneous
-  sort hypotheses that can be witnessed from the type signature.
+  \<^descr> @{ML Thm.strip_shyps}~\<open>thm\<close> removes any extraneous sort hypotheses that
+  can be witnessed from the type signature.
 \<close>
 
-text %mlex \<open>The following artificial example demonstrates the
-  derivation of @{prop False} with a pending sort hypothesis involving
-  a logically empty sort.\<close>
+text %mlex \<open>
+  The following artificial example demonstrates the derivation of @{prop
+  False} with a pending sort hypothesis involving a logically empty sort.
+\<close>
 
 class empty =
   assumes bad: "\<And>(x::'a) y. x \<noteq> y"
@@ -922,55 +867,54 @@
 
 ML_val \<open>@{assert} (Thm.extra_shyps @{thm false} = [@{sort empty}])\<close>
 
-text \<open>Thanks to the inference kernel managing sort hypothesis
-  according to their logical significance, this example is merely an
-  instance of \<^emph>\<open>ex falso quodlibet consequitur\<close> --- not a collapse
-  of the logical framework!\<close>
+text \<open>
+  Thanks to the inference kernel managing sort hypothesis according to their
+  logical significance, this example is merely an instance of \<^emph>\<open>ex falso
+  quodlibet consequitur\<close> --- not a collapse of the logical framework!
+\<close>
 
 
 section \<open>Object-level rules \label{sec:obj-rules}\<close>
 
 text \<open>
-  The primitive inferences covered so far mostly serve foundational
-  purposes.  User-level reasoning usually works via object-level rules
-  that are represented as theorems of Pure.  Composition of rules
-  involves \<^emph>\<open>backchaining\<close>, \<^emph>\<open>higher-order unification\<close> modulo
-  \<open>\<alpha>\<beta>\<eta>\<close>-conversion of \<open>\<lambda>\<close>-terms, and so-called
-  \<^emph>\<open>lifting\<close> of rules into a context of \<open>\<And>\<close> and \<open>\<Longrightarrow>\<close> connectives.  Thus the full power of higher-order Natural
-  Deduction in Isabelle/Pure becomes readily available.
+  The primitive inferences covered so far mostly serve foundational purposes.
+  User-level reasoning usually works via object-level rules that are
+  represented as theorems of Pure. Composition of rules involves
+  \<^emph>\<open>backchaining\<close>, \<^emph>\<open>higher-order unification\<close> modulo \<open>\<alpha>\<beta>\<eta>\<close>-conversion of
+  \<open>\<lambda>\<close>-terms, and so-called \<^emph>\<open>lifting\<close> of rules into a context of \<open>\<And>\<close> and \<open>\<Longrightarrow>\<close>
+  connectives. Thus the full power of higher-order Natural Deduction in
+  Isabelle/Pure becomes readily available.
 \<close>
 
 
 subsection \<open>Hereditary Harrop Formulae\<close>
 
 text \<open>
-  The idea of object-level rules is to model Natural Deduction
-  inferences in the style of Gentzen @{cite "Gentzen:1935"}, but we allow
-  arbitrary nesting similar to @{cite extensions91}.  The most basic
-  rule format is that of a \<^emph>\<open>Horn Clause\<close>:
+  The idea of object-level rules is to model Natural Deduction inferences in
+  the style of Gentzen @{cite "Gentzen:1935"}, but we allow arbitrary nesting
+  similar to @{cite extensions91}. The most basic rule format is that of a
+  \<^emph>\<open>Horn Clause\<close>:
   \[
   \infer{\<open>A\<close>}{\<open>A\<^sub>1\<close> & \<open>\<dots>\<close> & \<open>A\<^sub>n\<close>}
   \]
-  where \<open>A, A\<^sub>1, \<dots>, A\<^sub>n\<close> are atomic propositions
-  of the framework, usually of the form \<open>Trueprop B\<close>, where
-  \<open>B\<close> is a (compound) object-level statement.  This
-  object-level inference corresponds to an iterated implication in
-  Pure like this:
+  where \<open>A, A\<^sub>1, \<dots>, A\<^sub>n\<close> are atomic propositions of the framework, usually of
+  the form \<open>Trueprop B\<close>, where \<open>B\<close> is a (compound) object-level statement.
+  This object-level inference corresponds to an iterated implication in Pure
+  like this:
   \[
   \<open>A\<^sub>1 \<Longrightarrow> \<dots> A\<^sub>n \<Longrightarrow> A\<close>
   \]
-  As an example consider conjunction introduction: \<open>A \<Longrightarrow> B \<Longrightarrow> A \<and>
-  B\<close>.  Any parameters occurring in such rule statements are
-  conceptionally treated as arbitrary:
+  As an example consider conjunction introduction: \<open>A \<Longrightarrow> B \<Longrightarrow> A \<and> B\<close>. Any
+  parameters occurring in such rule statements are conceptionally treated as
+  arbitrary:
   \[
   \<open>\<And>x\<^sub>1 \<dots> x\<^sub>m. A\<^sub>1 x\<^sub>1 \<dots> x\<^sub>m \<Longrightarrow> \<dots> A\<^sub>n x\<^sub>1 \<dots> x\<^sub>m \<Longrightarrow> A x\<^sub>1 \<dots> x\<^sub>m\<close>
   \]
 
-  Nesting of rules means that the positions of \<open>A\<^sub>i\<close> may
-  again hold compound rules, not just atomic propositions.
-  Propositions of this format are called \<^emph>\<open>Hereditary Harrop
-  Formulae\<close> in the literature @{cite "Miller:1991"}.  Here we give an
-  inductive characterization as follows:
+  Nesting of rules means that the positions of \<open>A\<^sub>i\<close> may again hold compound
+  rules, not just atomic propositions. Propositions of this format are called
+  \<^emph>\<open>Hereditary Harrop Formulae\<close> in the literature @{cite "Miller:1991"}. Here
+  we give an inductive characterization as follows:
 
   \<^medskip>
   \begin{tabular}{ll}
@@ -980,29 +924,26 @@
   \end{tabular}
   \<^medskip>
 
-  Thus we essentially impose nesting levels on propositions formed
-  from \<open>\<And>\<close> and \<open>\<Longrightarrow>\<close>.  At each level there is a prefix
-  of parameters and compound premises, concluding an atomic
-  proposition.  Typical examples are \<open>\<longrightarrow>\<close>-introduction \<open>(A \<Longrightarrow> B) \<Longrightarrow> A \<longrightarrow> B\<close> or mathematical induction \<open>P 0 \<Longrightarrow> (\<And>n. P n
-  \<Longrightarrow> P (Suc n)) \<Longrightarrow> P n\<close>.  Even deeper nesting occurs in well-founded
-  induction \<open>(\<And>x. (\<And>y. y \<prec> x \<Longrightarrow> P y) \<Longrightarrow> P x) \<Longrightarrow> P x\<close>, but this
-  already marks the limit of rule complexity that is usually seen in
-  practice.
+  Thus we essentially impose nesting levels on propositions formed from \<open>\<And>\<close>
+  and \<open>\<Longrightarrow>\<close>. At each level there is a prefix of parameters and compound
+  premises, concluding an atomic proposition. Typical examples are
+  \<open>\<longrightarrow>\<close>-introduction \<open>(A \<Longrightarrow> B) \<Longrightarrow> A \<longrightarrow> B\<close> or mathematical induction \<open>P 0 \<Longrightarrow> (\<And>n. P n
+  \<Longrightarrow> P (Suc n)) \<Longrightarrow> P n\<close>. Even deeper nesting occurs in well-founded induction
+  \<open>(\<And>x. (\<And>y. y \<prec> x \<Longrightarrow> P y) \<Longrightarrow> P x) \<Longrightarrow> P x\<close>, but this already marks the limit of
+  rule complexity that is usually seen in practice.
 
   \<^medskip>
-  Regular user-level inferences in Isabelle/Pure always
-  maintain the following canonical form of results:
+  Regular user-level inferences in Isabelle/Pure always maintain the following
+  canonical form of results:
 
-  \<^item> Normalization by \<open>(A \<Longrightarrow> (\<And>x. B x)) \<equiv> (\<And>x. A \<Longrightarrow> B x)\<close>,
-  which is a theorem of Pure, means that quantifiers are pushed in
-  front of implication at each level of nesting.  The normal form is a
-  Hereditary Harrop Formula.
+  \<^item> Normalization by \<open>(A \<Longrightarrow> (\<And>x. B x)) \<equiv> (\<And>x. A \<Longrightarrow> B x)\<close>, which is a theorem of
+  Pure, means that quantifiers are pushed in front of implication at each
+  level of nesting. The normal form is a Hereditary Harrop Formula.
 
-  \<^item> The outermost prefix of parameters is represented via
-  schematic variables: instead of \<open>\<And>\<^vec>x. \<^vec>H \<^vec>x
-  \<Longrightarrow> A \<^vec>x\<close> we have \<open>\<^vec>H ?\<^vec>x \<Longrightarrow> A ?\<^vec>x\<close>.
-  Note that this representation looses information about the order of
-  parameters, and vacuous quantifiers vanish automatically.
+  \<^item> The outermost prefix of parameters is represented via schematic variables:
+  instead of \<open>\<And>\<^vec>x. \<^vec>H \<^vec>x \<Longrightarrow> A \<^vec>x\<close> we have \<open>\<^vec>H
+  ?\<^vec>x \<Longrightarrow> A ?\<^vec>x\<close>. Note that this representation looses information
+  about the order of parameters, and vacuous quantifiers vanish automatically.
 \<close>
 
 text %mlref \<open>
@@ -1010,43 +951,42 @@
   @{index_ML Simplifier.norm_hhf: "Proof.context -> thm -> thm"} \\
   \end{mldecls}
 
-  \<^descr> @{ML Simplifier.norm_hhf}~\<open>ctxt thm\<close> normalizes the given
-  theorem according to the canonical form specified above.  This is
-  occasionally helpful to repair some low-level tools that do not
-  handle Hereditary Harrop Formulae properly.
+  \<^descr> @{ML Simplifier.norm_hhf}~\<open>ctxt thm\<close> normalizes the given theorem
+  according to the canonical form specified above. This is occasionally
+  helpful to repair some low-level tools that do not handle Hereditary Harrop
+  Formulae properly.
 \<close>
 
 
 subsection \<open>Rule composition\<close>
 
 text \<open>
-  The rule calculus of Isabelle/Pure provides two main inferences:
-  @{inference resolution} (i.e.\ back-chaining of rules) and
-  @{inference assumption} (i.e.\ closing a branch), both modulo
-  higher-order unification.  There are also combined variants, notably
-  @{inference elim_resolution} and @{inference dest_resolution}.
+  The rule calculus of Isabelle/Pure provides two main inferences: @{inference
+  resolution} (i.e.\ back-chaining of rules) and @{inference assumption}
+  (i.e.\ closing a branch), both modulo higher-order unification. There are
+  also combined variants, notably @{inference elim_resolution} and @{inference
+  dest_resolution}.
 
-  To understand the all-important @{inference resolution} principle,
-  we first consider raw @{inference_def composition} (modulo
-  higher-order unification with substitution \<open>\<vartheta>\<close>):
+  To understand the all-important @{inference resolution} principle, we first
+  consider raw @{inference_def composition} (modulo higher-order unification
+  with substitution \<open>\<vartheta>\<close>):
   \[
   \infer[(@{inference_def composition})]{\<open>\<^vec>A\<vartheta> \<Longrightarrow> C\<vartheta>\<close>}
   {\<open>\<^vec>A \<Longrightarrow> B\<close> & \<open>B' \<Longrightarrow> C\<close> & \<open>B\<vartheta> = B'\<vartheta>\<close>}
   \]
-  Here the conclusion of the first rule is unified with the premise of
-  the second; the resulting rule instance inherits the premises of the
-  first and conclusion of the second.  Note that \<open>C\<close> can again
-  consist of iterated implications.  We can also permute the premises
-  of the second rule back-and-forth in order to compose with \<open>B'\<close> in any position (subsequently we shall always refer to
-  position 1 w.l.o.g.).
+  Here the conclusion of the first rule is unified with the premise of the
+  second; the resulting rule instance inherits the premises of the first and
+  conclusion of the second. Note that \<open>C\<close> can again consist of iterated
+  implications. We can also permute the premises of the second rule
+  back-and-forth in order to compose with \<open>B'\<close> in any position (subsequently
+  we shall always refer to position 1 w.l.o.g.).
 
-  In @{inference composition} the internal structure of the common
-  part \<open>B\<close> and \<open>B'\<close> is not taken into account.  For
-  proper @{inference resolution} we require \<open>B\<close> to be atomic,
-  and explicitly observe the structure \<open>\<And>\<^vec>x. \<^vec>H
-  \<^vec>x \<Longrightarrow> B' \<^vec>x\<close> of the premise of the second rule.  The
-  idea is to adapt the first rule by ``lifting'' it into this context,
-  by means of iterated application of the following inferences:
+  In @{inference composition} the internal structure of the common part \<open>B\<close>
+  and \<open>B'\<close> is not taken into account. For proper @{inference resolution} we
+  require \<open>B\<close> to be atomic, and explicitly observe the structure \<open>\<And>\<^vec>x.
+  \<^vec>H \<^vec>x \<Longrightarrow> B' \<^vec>x\<close> of the premise of the second rule. The idea
+  is to adapt the first rule by ``lifting'' it into this context, by means of
+  iterated application of the following inferences:
   \[
   \infer[(@{inference_def imp_lift})]{\<open>(\<^vec>H \<Longrightarrow> \<^vec>A) \<Longrightarrow> (\<^vec>H \<Longrightarrow> B)\<close>}{\<open>\<^vec>A \<Longrightarrow> B\<close>}
   \]
@@ -1065,10 +1005,10 @@
    \end{tabular}}
   \]
 
-  Continued resolution of rules allows to back-chain a problem towards
-  more and sub-problems.  Branches are closed either by resolving with
-  a rule of 0 premises, or by producing a ``short-circuit'' within a
-  solved situation (again modulo unification):
+  Continued resolution of rules allows to back-chain a problem towards more
+  and sub-problems. Branches are closed either by resolving with a rule of 0
+  premises, or by producing a ``short-circuit'' within a solved situation
+  (again modulo unification):
   \[
   \infer[(@{inference_def assumption})]{\<open>C\<vartheta>\<close>}
   {\<open>(\<And>\<^vec>x. \<^vec>H \<^vec>x \<Longrightarrow> A \<^vec>x) \<Longrightarrow> C\<close> & \<open>A\<vartheta> = H\<^sub>i\<vartheta>\<close>~~\text{(for some~\<open>i\<close>)}}
@@ -1089,133 +1029,125 @@
   @{index_ML_op "OF": "thm * thm list -> thm"} \\
   \end{mldecls}
 
-  \<^descr> \<open>rule\<^sub>1 RSN (i, rule\<^sub>2)\<close> resolves the conclusion of
-  \<open>rule\<^sub>1\<close> with the \<open>i\<close>-th premise of \<open>rule\<^sub>2\<close>,
-  according to the @{inference resolution} principle explained above.
-  Unless there is precisely one resolvent it raises exception @{ML
-  THM}.
+  \<^descr> \<open>rule\<^sub>1 RSN (i, rule\<^sub>2)\<close> resolves the conclusion of \<open>rule\<^sub>1\<close> with the
+  \<open>i\<close>-th premise of \<open>rule\<^sub>2\<close>, according to the @{inference resolution}
+  principle explained above. Unless there is precisely one resolvent it raises
+  exception @{ML THM}.
 
-  This corresponds to the rule attribute @{attribute THEN} in Isar
-  source language.
+  This corresponds to the rule attribute @{attribute THEN} in Isar source
+  language.
 
-  \<^descr> \<open>rule\<^sub>1 RS rule\<^sub>2\<close> abbreviates \<open>rule\<^sub>1 RSN (1,
-  rule\<^sub>2)\<close>.
+  \<^descr> \<open>rule\<^sub>1 RS rule\<^sub>2\<close> abbreviates \<open>rule\<^sub>1 RSN (1, rule\<^sub>2)\<close>.
 
-  \<^descr> \<open>rules\<^sub>1 RLN (i, rules\<^sub>2)\<close> joins lists of rules.  For
-  every \<open>rule\<^sub>1\<close> in \<open>rules\<^sub>1\<close> and \<open>rule\<^sub>2\<close> in
-  \<open>rules\<^sub>2\<close>, it resolves the conclusion of \<open>rule\<^sub>1\<close> with
-  the \<open>i\<close>-th premise of \<open>rule\<^sub>2\<close>, accumulating multiple
-  results in one big list.  Note that such strict enumerations of
-  higher-order unifications can be inefficient compared to the lazy
-  variant seen in elementary tactics like @{ML resolve_tac}.
+  \<^descr> \<open>rules\<^sub>1 RLN (i, rules\<^sub>2)\<close> joins lists of rules. For every \<open>rule\<^sub>1\<close> in
+  \<open>rules\<^sub>1\<close> and \<open>rule\<^sub>2\<close> in \<open>rules\<^sub>2\<close>, it resolves the conclusion of \<open>rule\<^sub>1\<close>
+  with the \<open>i\<close>-th premise of \<open>rule\<^sub>2\<close>, accumulating multiple results in one
+  big list. Note that such strict enumerations of higher-order unifications
+  can be inefficient compared to the lazy variant seen in elementary tactics
+  like @{ML resolve_tac}.
 
-  \<^descr> \<open>rules\<^sub>1 RL rules\<^sub>2\<close> abbreviates \<open>rules\<^sub>1 RLN (1,
-  rules\<^sub>2)\<close>.
+  \<^descr> \<open>rules\<^sub>1 RL rules\<^sub>2\<close> abbreviates \<open>rules\<^sub>1 RLN (1, rules\<^sub>2)\<close>.
 
-  \<^descr> \<open>[rule\<^sub>1, \<dots>, rule\<^sub>n] MRS rule\<close> resolves \<open>rule\<^sub>i\<close>
-  against premise \<open>i\<close> of \<open>rule\<close>, for \<open>i = n, \<dots>,
-  1\<close>.  By working from right to left, newly emerging premises are
-  concatenated in the result, without interfering.
+  \<^descr> \<open>[rule\<^sub>1, \<dots>, rule\<^sub>n] MRS rule\<close> resolves \<open>rule\<^sub>i\<close> against premise \<open>i\<close> of
+  \<open>rule\<close>, for \<open>i = n, \<dots>, 1\<close>. By working from right to left, newly emerging
+  premises are concatenated in the result, without interfering.
 
-  \<^descr> \<open>rule OF rules\<close> is an alternative notation for \<open>rules MRS rule\<close>, which makes rule composition look more like
-  function application.  Note that the argument \<open>rules\<close> need
-  not be atomic.
+  \<^descr> \<open>rule OF rules\<close> is an alternative notation for \<open>rules MRS rule\<close>, which
+  makes rule composition look more like function application. Note that the
+  argument \<open>rules\<close> need not be atomic.
 
-  This corresponds to the rule attribute @{attribute OF} in Isar
-  source language.
+  This corresponds to the rule attribute @{attribute OF} in Isar source
+  language.
 \<close>
 
 
 section \<open>Proof terms \label{sec:proof-terms}\<close>
 
-text \<open>The Isabelle/Pure inference kernel can record the proof of
-  each theorem as a proof term that contains all logical inferences in
-  detail.  Rule composition by resolution (\secref{sec:obj-rules}) and
-  type-class reasoning is broken down to primitive rules of the
-  logical framework.  The proof term can be inspected by a separate
-  proof-checker, for example.
+text \<open>
+  The Isabelle/Pure inference kernel can record the proof of each theorem as a
+  proof term that contains all logical inferences in detail. Rule composition
+  by resolution (\secref{sec:obj-rules}) and type-class reasoning is broken
+  down to primitive rules of the logical framework. The proof term can be
+  inspected by a separate proof-checker, for example.
 
-  According to the well-known \<^emph>\<open>Curry-Howard isomorphism\<close>, a proof
-  can be viewed as a \<open>\<lambda>\<close>-term. Following this idea, proofs in
-  Isabelle are internally represented by a datatype similar to the one
-  for terms described in \secref{sec:terms}.  On top of these
-  syntactic terms, two more layers of \<open>\<lambda>\<close>-calculus are added,
-  which correspond to \<open>\<And>x :: \<alpha>. B x\<close> and \<open>A \<Longrightarrow> B\<close>
-  according to the propositions-as-types principle.  The resulting
-  3-level \<open>\<lambda>\<close>-calculus resembles ``\<open>\<lambda>HOL\<close>'' in the
-  more abstract setting of Pure Type Systems (PTS)
-  @{cite "Barendregt-Geuvers:2001"}, if some fine points like schematic
-  polymorphism and type classes are ignored.
+  According to the well-known \<^emph>\<open>Curry-Howard isomorphism\<close>, a proof can be
+  viewed as a \<open>\<lambda>\<close>-term. Following this idea, proofs in Isabelle are internally
+  represented by a datatype similar to the one for terms described in
+  \secref{sec:terms}. On top of these syntactic terms, two more layers of
+  \<open>\<lambda>\<close>-calculus are added, which correspond to \<open>\<And>x :: \<alpha>. B x\<close> and \<open>A \<Longrightarrow> B\<close>
+  according to the propositions-as-types principle. The resulting 3-level
+  \<open>\<lambda>\<close>-calculus resembles ``\<open>\<lambda>HOL\<close>'' in the more abstract setting of Pure Type
+  Systems (PTS) @{cite "Barendregt-Geuvers:2001"}, if some fine points like
+  schematic polymorphism and type classes are ignored.
 
   \<^medskip>
-  \<^emph>\<open>Proof abstractions\<close> of the form \<open>\<^bold>\<lambda>x :: \<alpha>. prf\<close>
-  or \<open>\<^bold>\<lambda>p : A. prf\<close> correspond to introduction of \<open>\<And>\<close>/\<open>\<Longrightarrow>\<close>, and \<^emph>\<open>proof applications\<close> of the form \<open>p \<cdot> t\<close> or \<open>p \<bullet> q\<close> correspond to elimination of \<open>\<And>\<close>/\<open>\<Longrightarrow>\<close>.  Actual types \<open>\<alpha>\<close>, propositions \<open>A\<close>, and terms \<open>t\<close> might be suppressed and reconstructed
-  from the overall proof term.
+  \<^emph>\<open>Proof abstractions\<close> of the form \<open>\<^bold>\<lambda>x :: \<alpha>. prf\<close> or \<open>\<^bold>\<lambda>p : A. prf\<close>
+  correspond to introduction of \<open>\<And>\<close>/\<open>\<Longrightarrow>\<close>, and \<^emph>\<open>proof applications\<close> of the form
+  \<open>p \<cdot> t\<close> or \<open>p \<bullet> q\<close> correspond to elimination of \<open>\<And>\<close>/\<open>\<Longrightarrow>\<close>. Actual types \<open>\<alpha>\<close>,
+  propositions \<open>A\<close>, and terms \<open>t\<close> might be suppressed and reconstructed from
+  the overall proof term.
 
   \<^medskip>
-  Various atomic proofs indicate special situations within
-  the proof construction as follows.
+  Various atomic proofs indicate special situations within the proof
+  construction as follows.
 
-  A \<^emph>\<open>bound proof variable\<close> is a natural number \<open>b\<close> that
-  acts as de-Bruijn index for proof term abstractions.
+  A \<^emph>\<open>bound proof variable\<close> is a natural number \<open>b\<close> that acts as de-Bruijn
+  index for proof term abstractions.
 
-  A \<^emph>\<open>minimal proof\<close> ``\<open>?\<close>'' is a dummy proof term.  This
-  indicates some unrecorded part of the proof.
+  A \<^emph>\<open>minimal proof\<close> ``\<open>?\<close>'' is a dummy proof term. This indicates some
+  unrecorded part of the proof.
 
-  \<open>Hyp A\<close> refers to some pending hypothesis by giving its
-  proposition.  This indicates an open context of implicit hypotheses,
-  similar to loose bound variables or free variables within a term
-  (\secref{sec:terms}).
+  \<open>Hyp A\<close> refers to some pending hypothesis by giving its proposition. This
+  indicates an open context of implicit hypotheses, similar to loose bound
+  variables or free variables within a term (\secref{sec:terms}).
 
-  An \<^emph>\<open>axiom\<close> or \<^emph>\<open>oracle\<close> \<open>a : A[\<^vec>\<tau>]\<close> refers
-  some postulated \<open>proof constant\<close>, which is subject to
-  schematic polymorphism of theory content, and the particular type
-  instantiation may be given explicitly.  The vector of types \<open>\<^vec>\<tau>\<close> refers to the schematic type variables in the generic
+  An \<^emph>\<open>axiom\<close> or \<^emph>\<open>oracle\<close> \<open>a : A[\<^vec>\<tau>]\<close> refers some postulated \<open>proof
+  constant\<close>, which is subject to schematic polymorphism of theory content, and
+  the particular type instantiation may be given explicitly. The vector of
+  types \<open>\<^vec>\<tau>\<close> refers to the schematic type variables in the generic
   proposition \<open>A\<close> in canonical order.
 
-  A \<^emph>\<open>proof promise\<close> \<open>a : A[\<^vec>\<tau>]\<close> is a placeholder
-  for some proof of polymorphic proposition \<open>A\<close>, with explicit
-  type instantiation as given by the vector \<open>\<^vec>\<tau>\<close>, as
-  above.  Unlike axioms or oracles, proof promises may be
-  \<^emph>\<open>fulfilled\<close> eventually, by substituting \<open>a\<close> by some
-  particular proof \<open>q\<close> at the corresponding type instance.
-  This acts like Hindley-Milner \<open>let\<close>-polymorphism: a generic
-  local proof definition may get used at different type instances, and
-  is replaced by the concrete instance eventually.
+  A \<^emph>\<open>proof promise\<close> \<open>a : A[\<^vec>\<tau>]\<close> is a placeholder for some proof of
+  polymorphic proposition \<open>A\<close>, with explicit type instantiation as given by
+  the vector \<open>\<^vec>\<tau>\<close>, as above. Unlike axioms or oracles, proof promises
+  may be \<^emph>\<open>fulfilled\<close> eventually, by substituting \<open>a\<close> by some particular proof
+  \<open>q\<close> at the corresponding type instance. This acts like Hindley-Milner
+  \<open>let\<close>-polymorphism: a generic local proof definition may get used at
+  different type instances, and is replaced by the concrete instance
+  eventually.
 
-  A \<^emph>\<open>named theorem\<close> wraps up some concrete proof as a closed
-  formal entity, in the manner of constant definitions for proof
-  terms.  The \<^emph>\<open>proof body\<close> of such boxed theorems involves some
-  digest about oracles and promises occurring in the original proof.
-  This allows the inference kernel to manage this critical information
-  without the full overhead of explicit proof terms.
+  A \<^emph>\<open>named theorem\<close> wraps up some concrete proof as a closed formal entity,
+  in the manner of constant definitions for proof terms. The \<^emph>\<open>proof body\<close> of
+  such boxed theorems involves some digest about oracles and promises
+  occurring in the original proof. This allows the inference kernel to manage
+  this critical information without the full overhead of explicit proof terms.
 \<close>
 
 
 subsection \<open>Reconstructing and checking proof terms\<close>
 
-text \<open>Fully explicit proof terms can be large, but most of this
-  information is redundant and can be reconstructed from the context.
-  Therefore, the Isabelle/Pure inference kernel records only
-  \<^emph>\<open>implicit\<close> proof terms, by omitting all typing information in
-  terms, all term and type labels of proof abstractions, and some
-  argument terms of applications \<open>p \<cdot> t\<close> (if possible).
+text \<open>
+  Fully explicit proof terms can be large, but most of this information is
+  redundant and can be reconstructed from the context. Therefore, the
+  Isabelle/Pure inference kernel records only \<^emph>\<open>implicit\<close> proof terms, by
+  omitting all typing information in terms, all term and type labels of proof
+  abstractions, and some argument terms of applications \<open>p \<cdot> t\<close> (if possible).
 
-  There are separate operations to reconstruct the full proof term
-  later on, using \<^emph>\<open>higher-order pattern unification\<close>
-  @{cite "nipkow-patterns" and "Berghofer-Nipkow:2000:TPHOL"}.
+  There are separate operations to reconstruct the full proof term later on,
+  using \<^emph>\<open>higher-order pattern unification\<close> @{cite "nipkow-patterns" and
+  "Berghofer-Nipkow:2000:TPHOL"}.
 
-  The \<^emph>\<open>proof checker\<close> expects a fully reconstructed proof term,
-  and can turn it into a theorem by replaying its primitive inferences
-  within the kernel.\<close>
+  The \<^emph>\<open>proof checker\<close> expects a fully reconstructed proof term, and can turn
+  it into a theorem by replaying its primitive inferences within the kernel.
+\<close>
 
 
 subsection \<open>Concrete syntax of proof terms\<close>
 
-text \<open>The concrete syntax of proof terms is a slight extension of
-  the regular inner syntax of Isabelle/Pure @{cite "isabelle-isar-ref"}.
-  Its main syntactic category @{syntax (inner) proof} is defined as
-  follows:
+text \<open>
+  The concrete syntax of proof terms is a slight extension of the regular
+  inner syntax of Isabelle/Pure @{cite "isabelle-isar-ref"}. Its main
+  syntactic category @{syntax (inner) proof} is defined as follows:
 
   \begin{center}
   \begin{supertabular}{rclr}
@@ -1240,13 +1172,14 @@
   \end{supertabular}
   \end{center}
 
-  Implicit term arguments in partial proofs are indicated by ``\<open>_\<close>''.  Type arguments for theorems and axioms may be specified
-  using \<open>p \<cdot> TYPE(type)\<close> (they must appear before any other
-  term argument of a theorem or axiom, but may be omitted altogether).
+  Implicit term arguments in partial proofs are indicated by ``\<open>_\<close>''. Type
+  arguments for theorems and axioms may be specified using \<open>p \<cdot> TYPE(type)\<close>
+  (they must appear before any other term argument of a theorem or axiom, but
+  may be omitted altogether).
 
   \<^medskip>
-  There are separate read and print operations for proof
-  terms, in order to avoid conflicts with the regular term language.
+  There are separate read and print operations for proof terms, in order to
+  avoid conflicts with the regular term language.
 \<close>
 
 text %mlref \<open>
@@ -1263,65 +1196,60 @@
   @{index_ML Proof_Syntax.pretty_proof: "Proof.context -> proof -> Pretty.T"} \\
   \end{mldecls}
 
-  \<^descr> Type @{ML_type proof} represents proof terms; this is a
-  datatype with constructors @{index_ML Abst}, @{index_ML AbsP},
-  @{index_ML_op "%"}, @{index_ML_op "%%"}, @{index_ML PBound},
-  @{index_ML MinProof}, @{index_ML Hyp}, @{index_ML PAxm}, @{index_ML
-  Oracle}, @{index_ML Promise}, @{index_ML PThm} as explained above.
-  %FIXME OfClass (!?)
+  \<^descr> Type @{ML_type proof} represents proof terms; this is a datatype with
+  constructors @{index_ML Abst}, @{index_ML AbsP}, @{index_ML_op "%"},
+  @{index_ML_op "%%"}, @{index_ML PBound}, @{index_ML MinProof}, @{index_ML
+  Hyp}, @{index_ML PAxm}, @{index_ML Oracle}, @{index_ML Promise}, @{index_ML
+  PThm} as explained above. %FIXME OfClass (!?)
+
+  \<^descr> Type @{ML_type proof_body} represents the nested proof information of a
+  named theorem, consisting of a digest of oracles and named theorem over some
+  proof term. The digest only covers the directly visible part of the proof:
+  in order to get the full information, the implicit graph of nested theorems
+  needs to be traversed (e.g.\ using @{ML Proofterm.fold_body_thms}).
 
-  \<^descr> Type @{ML_type proof_body} represents the nested proof
-  information of a named theorem, consisting of a digest of oracles
-  and named theorem over some proof term.  The digest only covers the
-  directly visible part of the proof: in order to get the full
-  information, the implicit graph of nested theorems needs to be
-  traversed (e.g.\ using @{ML Proofterm.fold_body_thms}).
+  \<^descr> @{ML Thm.proof_of}~\<open>thm\<close> and @{ML Thm.proof_body_of}~\<open>thm\<close> produce the
+  proof term or proof body (with digest of oracles and theorems) from a given
+  theorem. Note that this involves a full join of internal futures that
+  fulfill pending proof promises, and thus disrupts the natural bottom-up
+  construction of proofs by introducing dynamic ad-hoc dependencies. Parallel
+  performance may suffer by inspecting proof terms at run-time.
 
-  \<^descr> @{ML Thm.proof_of}~\<open>thm\<close> and @{ML
-  Thm.proof_body_of}~\<open>thm\<close> produce the proof term or proof
-  body (with digest of oracles and theorems) from a given theorem.
-  Note that this involves a full join of internal futures that fulfill
-  pending proof promises, and thus disrupts the natural bottom-up
-  construction of proofs by introducing dynamic ad-hoc dependencies.
-  Parallel performance may suffer by inspecting proof terms at
-  run-time.
+  \<^descr> @{ML proofs} specifies the detail of proof recording within @{ML_type thm}
+  values produced by the inference kernel: @{ML 0} records only the names of
+  oracles, @{ML 1} records oracle names and propositions, @{ML 2} additionally
+  records full proof terms. Officially named theorems that contribute to a
+  result are recorded in any case.
 
-  \<^descr> @{ML proofs} specifies the detail of proof recording within
-  @{ML_type thm} values produced by the inference kernel: @{ML 0}
-  records only the names of oracles, @{ML 1} records oracle names and
-  propositions, @{ML 2} additionally records full proof terms.
-  Officially named theorems that contribute to a result are recorded
-  in any case.
+  \<^descr> @{ML Reconstruct.reconstruct_proof}~\<open>thy prop prf\<close> turns the implicit
+  proof term \<open>prf\<close> into a full proof of the given proposition.
 
-  \<^descr> @{ML Reconstruct.reconstruct_proof}~\<open>thy prop prf\<close>
-  turns the implicit proof term \<open>prf\<close> into a full proof of the
-  given proposition.
-
-  Reconstruction may fail if \<open>prf\<close> is not a proof of \<open>prop\<close>, or if it does not contain sufficient information for
-  reconstruction.  Failure may only happen for proofs that are
-  constructed manually, but not for those produced automatically by
-  the inference kernel.
+  Reconstruction may fail if \<open>prf\<close> is not a proof of \<open>prop\<close>, or if it does not
+  contain sufficient information for reconstruction. Failure may only happen
+  for proofs that are constructed manually, but not for those produced
+  automatically by the inference kernel.
 
-  \<^descr> @{ML Reconstruct.expand_proof}~\<open>thy [thm\<^sub>1, \<dots>, thm\<^sub>n]
-  prf\<close> expands and reconstructs the proofs of all specified theorems,
-  with the given (full) proof.  Theorems that are not unique specified
-  via their name may be disambiguated by giving their proposition.
+  \<^descr> @{ML Reconstruct.expand_proof}~\<open>thy [thm\<^sub>1, \<dots>, thm\<^sub>n] prf\<close> expands and
+  reconstructs the proofs of all specified theorems, with the given (full)
+  proof. Theorems that are not unique specified via their name may be
+  disambiguated by giving their proposition.
 
-  \<^descr> @{ML Proof_Checker.thm_of_proof}~\<open>thy prf\<close> turns the
-  given (full) proof into a theorem, by replaying it using only
-  primitive rules of the inference kernel.
+  \<^descr> @{ML Proof_Checker.thm_of_proof}~\<open>thy prf\<close> turns the given (full) proof
+  into a theorem, by replaying it using only primitive rules of the inference
+  kernel.
 
-  \<^descr> @{ML Proof_Syntax.read_proof}~\<open>thy b\<^sub>1 b\<^sub>2 s\<close> reads in a
-  proof term. The Boolean flags indicate the use of sort and type
-  information.  Usually, typing information is left implicit and is
-  inferred during proof reconstruction.  %FIXME eliminate flags!?
+  \<^descr> @{ML Proof_Syntax.read_proof}~\<open>thy b\<^sub>1 b\<^sub>2 s\<close> reads in a proof term. The
+  Boolean flags indicate the use of sort and type information. Usually, typing
+  information is left implicit and is inferred during proof reconstruction.
+  %FIXME eliminate flags!?
 
-  \<^descr> @{ML Proof_Syntax.pretty_proof}~\<open>ctxt prf\<close>
-  pretty-prints the given proof term.
+  \<^descr> @{ML Proof_Syntax.pretty_proof}~\<open>ctxt prf\<close> pretty-prints the given proof
+  term.
 \<close>
 
-text %mlex \<open>Detailed proof information of a theorem may be retrieved
-  as follows:\<close>
+text %mlex \<open>
+  Detailed proof information of a theorem may be retrieved as follows:
+\<close>
 
 lemma ex: "A \<and> B \<longrightarrow> B \<and> A"
 proof
@@ -1344,15 +1272,16 @@
       (fn (name, _, _) => insert (op =) name) [body] [];
 \<close>
 
-text \<open>The result refers to various basic facts of Isabelle/HOL:
-  @{thm [source] HOL.impI}, @{thm [source] HOL.conjE}, @{thm [source]
-  HOL.conjI} etc.  The combinator @{ML Proofterm.fold_body_thms}
-  recursively explores the graph of the proofs of all theorems being
-  used here.
+text \<open>
+  The result refers to various basic facts of Isabelle/HOL: @{thm [source]
+  HOL.impI}, @{thm [source] HOL.conjE}, @{thm [source] HOL.conjI} etc. The
+  combinator @{ML Proofterm.fold_body_thms} recursively explores the graph of
+  the proofs of all theorems being used here.
 
   \<^medskip>
-  Alternatively, we may produce a proof term manually, and
-  turn it into a theorem as follows:\<close>
+  Alternatively, we may produce a proof term manually, and turn it into a
+  theorem as follows:
+\<close>
 
 ML_val \<open>
   val thy = @{theory};
@@ -1371,9 +1300,8 @@
 
 text \<open>
   \<^medskip>
-  See also @{file "~~/src/HOL/Proofs/ex/XML_Data.thy"}
-  for further examples, with export and import of proof terms via
-  XML/ML data representation.
+  See also @{file "~~/src/HOL/Proofs/ex/XML_Data.thy"} for further examples,
+  with export and import of proof terms via XML/ML data representation.
 \<close>
 
 end
--- a/src/Doc/Implementation/ML.thy	Tue Dec 15 14:41:47 2015 +0000
+++ b/src/Doc/Implementation/ML.thy	Wed Dec 16 17:30:30 2015 +0100
@@ -6,75 +6,72 @@
 
 chapter \<open>Isabelle/ML\<close>
 
-text \<open>Isabelle/ML is best understood as a certain culture based on
-  Standard ML.  Thus it is not a new programming language, but a
-  certain way to use SML at an advanced level within the Isabelle
-  environment.  This covers a variety of aspects that are geared
-  towards an efficient and robust platform for applications of formal
-  logic with fully foundational proof construction --- according to
-  the well-known \<^emph>\<open>LCF principle\<close>.  There is specific
-  infrastructure with library modules to address the needs of this
-  difficult task.  For example, the raw parallel programming model of
-  Poly/ML is presented as considerably more abstract concept of
-  \<^emph>\<open>futures\<close>, which is then used to augment the inference
-  kernel, Isar theory and proof interpreter, and PIDE document management.
+text \<open>
+  Isabelle/ML is best understood as a certain culture based on Standard ML.
+  Thus it is not a new programming language, but a certain way to use SML at
+  an advanced level within the Isabelle environment. This covers a variety of
+  aspects that are geared towards an efficient and robust platform for
+  applications of formal logic with fully foundational proof construction ---
+  according to the well-known \<^emph>\<open>LCF principle\<close>. There is specific
+  infrastructure with library modules to address the needs of this difficult
+  task. For example, the raw parallel programming model of Poly/ML is
+  presented as considerably more abstract concept of \<^emph>\<open>futures\<close>, which is then
+  used to augment the inference kernel, Isar theory and proof interpreter, and
+  PIDE document management.
 
-  The main aspects of Isabelle/ML are introduced below.  These
-  first-hand explanations should help to understand how proper
-  Isabelle/ML is to be read and written, and to get access to the
-  wealth of experience that is expressed in the source text and its
-  history of changes.\<^footnote>\<open>See
-  @{url "http://isabelle.in.tum.de/repos/isabelle"} for the full
-  Mercurial history.  There are symbolic tags to refer to official
-  Isabelle releases, as opposed to arbitrary \<^emph>\<open>tip\<close> versions that
-  merely reflect snapshots that are never really up-to-date.\<close>\<close>
+  The main aspects of Isabelle/ML are introduced below. These first-hand
+  explanations should help to understand how proper Isabelle/ML is to be read
+  and written, and to get access to the wealth of experience that is expressed
+  in the source text and its history of changes.\<^footnote>\<open>See @{url
+  "http://isabelle.in.tum.de/repos/isabelle"} for the full Mercurial history.
+  There are symbolic tags to refer to official Isabelle releases, as opposed
+  to arbitrary \<^emph>\<open>tip\<close> versions that merely reflect snapshots that are never
+  really up-to-date.\<close>
+\<close>
 
 
 section \<open>Style and orthography\<close>
 
-text \<open>The sources of Isabelle/Isar are optimized for
-  \<^emph>\<open>readability\<close> and \<^emph>\<open>maintainability\<close>.  The main purpose is
-  to tell an informed reader what is really going on and how things
-  really work.  This is a non-trivial aim, but it is supported by a
-  certain style of writing Isabelle/ML that has emerged from long
-  years of system development.\<^footnote>\<open>See also the interesting style
-  guide for OCaml
-  @{url "http://caml.inria.fr/resources/doc/guides/guidelines.en.html"}
-  which shares many of our means and ends.\<close>
+text \<open>
+  The sources of Isabelle/Isar are optimized for \<^emph>\<open>readability\<close> and
+  \<^emph>\<open>maintainability\<close>. The main purpose is to tell an informed reader what is
+  really going on and how things really work. This is a non-trivial aim, but
+  it is supported by a certain style of writing Isabelle/ML that has emerged
+  from long years of system development.\<^footnote>\<open>See also the interesting style guide
+  for OCaml @{url
+  "http://caml.inria.fr/resources/doc/guides/guidelines.en.html"} which shares
+  many of our means and ends.\<close>
 
-  The main principle behind any coding style is \<^emph>\<open>consistency\<close>.
-  For a single author of a small program this merely means ``choose
-  your style and stick to it''.  A complex project like Isabelle, with
-  long years of development and different contributors, requires more
-  standardization.  A coding style that is changed every few years or
-  with every new contributor is no style at all, because consistency
-  is quickly lost.  Global consistency is hard to achieve, though.
-  Nonetheless, one should always strive at least for local consistency
-  of modules and sub-systems, without deviating from some general
-  principles how to write Isabelle/ML.
+  The main principle behind any coding style is \<^emph>\<open>consistency\<close>. For a single
+  author of a small program this merely means ``choose your style and stick to
+  it''. A complex project like Isabelle, with long years of development and
+  different contributors, requires more standardization. A coding style that
+  is changed every few years or with every new contributor is no style at all,
+  because consistency is quickly lost. Global consistency is hard to achieve,
+  though. Nonetheless, one should always strive at least for local consistency
+  of modules and sub-systems, without deviating from some general principles
+  how to write Isabelle/ML.
 
-  In a sense, good coding style is like an \<^emph>\<open>orthography\<close> for the
-  sources: it helps to read quickly over the text and see through the
-  main points, without getting distracted by accidental presentation
-  of free-style code.
+  In a sense, good coding style is like an \<^emph>\<open>orthography\<close> for the sources: it
+  helps to read quickly over the text and see through the main points, without
+  getting distracted by accidental presentation of free-style code.
 \<close>
 
 
 subsection \<open>Header and sectioning\<close>
 
-text \<open>Isabelle source files have a certain standardized header
-  format (with precise spacing) that follows ancient traditions
-  reaching back to the earliest versions of the system by Larry
-  Paulson.  See @{file "~~/src/Pure/thm.ML"}, for example.
+text \<open>
+  Isabelle source files have a certain standardized header format (with
+  precise spacing) that follows ancient traditions reaching back to the
+  earliest versions of the system by Larry Paulson. See @{file
+  "~~/src/Pure/thm.ML"}, for example.
 
-  The header includes at least \<^verbatim>\<open>Title\<close> and \<^verbatim>\<open>Author\<close> entries,
-  followed by a prose description of the purpose of
-  the module.  The latter can range from a single line to several
-  paragraphs of explanations.
+  The header includes at least \<^verbatim>\<open>Title\<close> and \<^verbatim>\<open>Author\<close> entries, followed by a
+  prose description of the purpose of the module. The latter can range from a
+  single line to several paragraphs of explanations.
 
-  The rest of the file is divided into sections, subsections,
-  subsubsections, paragraphs etc.\ using a simple layout via ML
-  comments as follows.
+  The rest of the file is divided into sections, subsections, subsubsections,
+  paragraphs etc.\ using a simple layout via ML comments as follows.
 
   @{verbatim [display]
 \<open>  (*** section ***)
@@ -90,26 +87,28 @@
     with more text
   *)\<close>}
 
-  As in regular typography, there is some extra space \<^emph>\<open>before\<close>
-  section headings that are adjacent to plain text, but not other headings
-  as in the example above.
+  As in regular typography, there is some extra space \<^emph>\<open>before\<close> section
+  headings that are adjacent to plain text, but not other headings as in the
+  example above.
 
   \<^medskip>
-  The precise wording of the prose text given in these
-  headings is chosen carefully to introduce the main theme of the
-  subsequent formal ML text.
+  The precise wording of the prose text given in these headings is chosen
+  carefully to introduce the main theme of the subsequent formal ML text.
 \<close>
 
 
 subsection \<open>Naming conventions\<close>
 
-text \<open>Since ML is the primary medium to express the meaning of the
-  source text, naming of ML entities requires special care.\<close>
+text \<open>
+  Since ML is the primary medium to express the meaning of the source text,
+  naming of ML entities requires special care.
+\<close>
 
 paragraph \<open>Notation.\<close>
-text \<open>A name consists of 1--3 \<^emph>\<open>words\<close> (rarely 4, but not more) that are
-  separated by underscore. There are three variants concerning upper or lower
-  case letters, which are used for certain ML categories as follows:
+text \<open>
+  A name consists of 1--3 \<^emph>\<open>words\<close> (rarely 4, but not more) that are separated
+  by underscore. There are three variants concerning upper or lower case
+  letters, which are used for certain ML categories as follows:
 
   \<^medskip>
   \begin{tabular}{lll}
@@ -120,38 +119,37 @@
   \end{tabular}
   \<^medskip>
 
-  For historical reasons, many capitalized names omit underscores,
-  e.g.\ old-style @{ML_text FooBar} instead of @{ML_text Foo_Bar}.
-  Genuine mixed-case names are \<^emph>\<open>not\<close> used, because clear division
-  of words is essential for readability.\<^footnote>\<open>Camel-case was
-  invented to workaround the lack of underscore in some early
-  non-ASCII character sets.  Later it became habitual in some language
-  communities that are now strong in numbers.\<close>
+  For historical reasons, many capitalized names omit underscores, e.g.\
+  old-style @{ML_text FooBar} instead of @{ML_text Foo_Bar}. Genuine
+  mixed-case names are \<^emph>\<open>not\<close> used, because clear division of words is
+  essential for readability.\<^footnote>\<open>Camel-case was invented to workaround the lack
+  of underscore in some early non-ASCII character sets. Later it became
+  habitual in some language communities that are now strong in numbers.\<close>
 
-  A single (capital) character does not count as ``word'' in this
-  respect: some Isabelle/ML names are suffixed by extra markers like
-  this: @{ML_text foo_barT}.
+  A single (capital) character does not count as ``word'' in this respect:
+  some Isabelle/ML names are suffixed by extra markers like this: @{ML_text
+  foo_barT}.
 
-  Name variants are produced by adding 1--3 primes, e.g.\ @{ML_text
-  foo'}, @{ML_text foo''}, or @{ML_text foo'''}, but not @{ML_text
-  foo''''} or more.  Decimal digits scale better to larger numbers,
-  e.g.\ @{ML_text foo0}, @{ML_text foo1}, @{ML_text foo42}.
+  Name variants are produced by adding 1--3 primes, e.g.\ @{ML_text foo'},
+  @{ML_text foo''}, or @{ML_text foo'''}, but not @{ML_text foo''''} or more.
+  Decimal digits scale better to larger numbers, e.g.\ @{ML_text foo0},
+  @{ML_text foo1}, @{ML_text foo42}.
 \<close>
 
 paragraph\<open>Scopes.\<close>
-text \<open>Apart from very basic library modules, ML structures are not ``opened'',
-  but names are referenced with explicit qualification, as in @{ML
+text \<open>
+  Apart from very basic library modules, ML structures are not ``opened'', but
+  names are referenced with explicit qualification, as in @{ML
   Syntax.string_of_term} for example. When devising names for structures and
   their components it is important to aim at eye-catching compositions of both
   parts, because this is how they are seen in the sources and documentation.
   For the same reasons, aliases of well-known library functions should be
   avoided.
 
-  Local names of function abstraction or case/let bindings are
-  typically shorter, sometimes using only rudiments of ``words'',
-  while still avoiding cryptic shorthands.  An auxiliary function
-  called @{ML_text helper}, @{ML_text aux}, or @{ML_text f} is
-  considered bad style.
+  Local names of function abstraction or case/let bindings are typically
+  shorter, sometimes using only rudiments of ``words'', while still avoiding
+  cryptic shorthands. An auxiliary function called @{ML_text helper},
+  @{ML_text aux}, or @{ML_text f} is considered bad style.
 
   Example:
 
@@ -184,28 +182,28 @@
 \<close>
 
 paragraph \<open>Specific conventions.\<close>
-text \<open>Here are some specific name forms that occur frequently in the sources.
+text \<open>
+  Here are some specific name forms that occur frequently in the sources.
 
-  \<^item> A function that maps @{ML_text foo} to @{ML_text bar} is
-  called @{ML_text foo_to_bar} or @{ML_text bar_of_foo} (never
-  @{ML_text foo2bar}, nor @{ML_text bar_from_foo}, nor @{ML_text
-  bar_for_foo}, nor @{ML_text bar4foo}).
+  \<^item> A function that maps @{ML_text foo} to @{ML_text bar} is called @{ML_text
+  foo_to_bar} or @{ML_text bar_of_foo} (never @{ML_text foo2bar}, nor
+  @{ML_text bar_from_foo}, nor @{ML_text bar_for_foo}, nor @{ML_text
+  bar4foo}).
 
-  \<^item> The name component @{ML_text legacy} means that the operation
-  is about to be discontinued soon.
+  \<^item> The name component @{ML_text legacy} means that the operation is about to
+  be discontinued soon.
 
-  \<^item> The name component @{ML_text global} means that this works
-  with the background theory instead of the regular local context
-  (\secref{sec:context}), sometimes for historical reasons, sometimes
-  due a genuine lack of locality of the concept involved, sometimes as
-  a fall-back for the lack of a proper context in the application
-  code.  Whenever there is a non-global variant available, the
-  application should be migrated to use it with a proper local
-  context.
+  \<^item> The name component @{ML_text global} means that this works with the
+  background theory instead of the regular local context
+  (\secref{sec:context}), sometimes for historical reasons, sometimes due a
+  genuine lack of locality of the concept involved, sometimes as a fall-back
+  for the lack of a proper context in the application code. Whenever there is
+  a non-global variant available, the application should be migrated to use it
+  with a proper local context.
 
-  \<^item> Variables of the main context types of the Isabelle/Isar
-  framework (\secref{sec:context} and \chref{ch:local-theory}) have
-  firm naming conventions as follows:
+  \<^item> Variables of the main context types of the Isabelle/Isar framework
+  (\secref{sec:context} and \chref{ch:local-theory}) have firm naming
+  conventions as follows:
 
     \<^item> theories are called @{ML_text thy}, rarely @{ML_text theory}
     (never @{ML_text thry})
@@ -219,54 +217,51 @@
     theories that are treated as proof context (which is a semantic
     super-type)
 
-  Variations with primed or decimal numbers are always possible, as
-  well as semantic prefixes like @{ML_text foo_thy} or @{ML_text
-  bar_ctxt}, but the base conventions above need to be preserved.
-  This allows to emphasize their data flow via plain regular
-  expressions in the text editor.
+  Variations with primed or decimal numbers are always possible, as well as
+  semantic prefixes like @{ML_text foo_thy} or @{ML_text bar_ctxt}, but the
+  base conventions above need to be preserved. This allows to emphasize their
+  data flow via plain regular expressions in the text editor.
 
-  \<^item> The main logical entities (\secref{ch:logic}) have established
-  naming convention as follows:
+  \<^item> The main logical entities (\secref{ch:logic}) have established naming
+  convention as follows:
 
     \<^item> sorts are called @{ML_text S}
 
-    \<^item> types are called @{ML_text T}, @{ML_text U}, or @{ML_text
-    ty} (never @{ML_text t})
+    \<^item> types are called @{ML_text T}, @{ML_text U}, or @{ML_text ty} (never
+    @{ML_text t})
 
-    \<^item> terms are called @{ML_text t}, @{ML_text u}, or @{ML_text
-    tm} (never @{ML_text trm})
+    \<^item> terms are called @{ML_text t}, @{ML_text u}, or @{ML_text tm} (never
+    @{ML_text trm})
 
-    \<^item> certified types are called @{ML_text cT}, rarely @{ML_text
-    T}, with variants as for types
+    \<^item> certified types are called @{ML_text cT}, rarely @{ML_text T}, with
+    variants as for types
 
-    \<^item> certified terms are called @{ML_text ct}, rarely @{ML_text
-    t}, with variants as for terms (never @{ML_text ctrm})
+    \<^item> certified terms are called @{ML_text ct}, rarely @{ML_text t}, with
+    variants as for terms (never @{ML_text ctrm})
 
     \<^item> theorems are called @{ML_text th}, or @{ML_text thm}
 
-  Proper semantic names override these conventions completely.  For
-  example, the left-hand side of an equation (as a term) can be called
-  @{ML_text lhs} (not @{ML_text lhs_tm}).  Or a term that is known
-  to be a variable can be called @{ML_text v} or @{ML_text x}.
+  Proper semantic names override these conventions completely. For example,
+  the left-hand side of an equation (as a term) can be called @{ML_text lhs}
+  (not @{ML_text lhs_tm}). Or a term that is known to be a variable can be
+  called @{ML_text v} or @{ML_text x}.
 
-  \<^item> Tactics (\secref{sec:tactics}) are sufficiently important to
-  have specific naming conventions.  The name of a basic tactic
-  definition always has a @{ML_text "_tac"} suffix, the subgoal index
-  (if applicable) is always called @{ML_text i}, and the goal state
-  (if made explicit) is usually called @{ML_text st} instead of the
-  somewhat misleading @{ML_text thm}.  Any other arguments are given
-  before the latter two, and the general context is given first.
-  Example:
+  \<^item> Tactics (\secref{sec:tactics}) are sufficiently important to have specific
+  naming conventions. The name of a basic tactic definition always has a
+  @{ML_text "_tac"} suffix, the subgoal index (if applicable) is always called
+  @{ML_text i}, and the goal state (if made explicit) is usually called
+  @{ML_text st} instead of the somewhat misleading @{ML_text thm}. Any other
+  arguments are given before the latter two, and the general context is given
+  first. Example:
 
   @{verbatim [display] \<open>  fun my_tac ctxt arg1 arg2 i st = ...\<close>}
 
-  Note that the goal state @{ML_text st} above is rarely made
-  explicit, if tactic combinators (tacticals) are used as usual.
+  Note that the goal state @{ML_text st} above is rarely made explicit, if
+  tactic combinators (tacticals) are used as usual.
 
   A tactic that requires a proof context needs to make that explicit as seen
-  in the \<^verbatim>\<open>ctxt\<close> argument above. Do not refer to the background
-  theory of \<^verbatim>\<open>st\<close> -- it is not a proper context, but merely a formal
-  certificate.
+  in the \<^verbatim>\<open>ctxt\<close> argument above. Do not refer to the background theory of
+  \<^verbatim>\<open>st\<close> -- it is not a proper context, but merely a formal certificate.
 \<close>
 
 
@@ -275,30 +270,34 @@
 text \<open>
   The general Isabelle/ML source layout imitates regular type-setting
   conventions, augmented by the requirements for deeply nested expressions
-  that are commonplace in functional programming.\<close>
+  that are commonplace in functional programming.
+\<close>
 
 paragraph \<open>Line length\<close>
-text \<open>is limited to 80 characters according to ancient standards, but we allow
-  as much as 100 characters (not more).\<^footnote>\<open>Readability requires to keep
-  the beginning of a line in view while watching its end. Modern wide-screen
+text \<open>
+  is limited to 80 characters according to ancient standards, but we allow as
+  much as 100 characters (not more).\<^footnote>\<open>Readability requires to keep the
+  beginning of a line in view while watching its end. Modern wide-screen
   displays do not change the way how the human brain works. Sources also need
   to be printable on plain paper with reasonable font-size.\<close> The extra 20
   characters acknowledge the space requirements due to qualified library
-  references in Isabelle/ML.\<close>
+  references in Isabelle/ML.
+\<close>
 
 paragraph \<open>White-space\<close>
-text \<open>is used to emphasize the structure of expressions, following mostly
-  standard conventions for mathematical typesetting, as can be seen in plain
-  {\TeX} or {\LaTeX}. This defines positioning of spaces for parentheses,
-  punctuation, and infixes as illustrated here:
+text \<open>
+  is used to emphasize the structure of expressions, following mostly standard
+  conventions for mathematical typesetting, as can be seen in plain {\TeX} or
+  {\LaTeX}. This defines positioning of spaces for parentheses, punctuation,
+  and infixes as illustrated here:
 
   @{verbatim [display]
 \<open>  val x = y + z * (a + b);
   val pair = (a, b);
   val record = {foo = 1, bar = 2};\<close>}
 
-  Lines are normally broken \<^emph>\<open>after\<close> an infix operator or
-  punctuation character.  For example:
+  Lines are normally broken \<^emph>\<open>after\<close> an infix operator or punctuation
+  character. For example:
 
   @{verbatim [display]
 \<open>
@@ -313,32 +312,32 @@
     c);
 \<close>}
 
-  Some special infixes (e.g.\ @{ML_text "|>"}) work better at the
-  start of the line, but punctuation is always at the end.
+  Some special infixes (e.g.\ @{ML_text "|>"}) work better at the start of the
+  line, but punctuation is always at the end.
 
-  Function application follows the tradition of \<open>\<lambda>\<close>-calculus,
-  not informal mathematics.  For example: @{ML_text "f a b"} for a
-  curried function, or @{ML_text "g (a, b)"} for a tupled function.
-  Note that the space between @{ML_text g} and the pair @{ML_text
-  "(a, b)"} follows the important principle of
-  \<^emph>\<open>compositionality\<close>: the layout of @{ML_text "g p"} does not
-  change when @{ML_text "p"} is refined to the concrete pair
-  @{ML_text "(a, b)"}.
+  Function application follows the tradition of \<open>\<lambda>\<close>-calculus, not informal
+  mathematics. For example: @{ML_text "f a b"} for a curried function, or
+  @{ML_text "g (a, b)"} for a tupled function. Note that the space between
+  @{ML_text g} and the pair @{ML_text "(a, b)"} follows the important
+  principle of \<^emph>\<open>compositionality\<close>: the layout of @{ML_text "g p"} does not
+  change when @{ML_text "p"} is refined to the concrete pair @{ML_text "(a,
+  b)"}.
 \<close>
 
 paragraph \<open>Indentation\<close>
-text \<open>uses plain spaces, never hard tabulators.\<^footnote>\<open>Tabulators were
-  invented to move the carriage of a type-writer to certain predefined
-  positions. In software they could be used as a primitive run-length
-  compression of consecutive spaces, but the precise result would depend on
-  non-standardized text editor configuration.\<close>
+text \<open>
+  uses plain spaces, never hard tabulators.\<^footnote>\<open>Tabulators were invented to move
+  the carriage of a type-writer to certain predefined positions. In software
+  they could be used as a primitive run-length compression of consecutive
+  spaces, but the precise result would depend on non-standardized text editor
+  configuration.\<close>
 
-  Each level of nesting is indented by 2 spaces, sometimes 1, very
-  rarely 4, never 8 or any other odd number.
+  Each level of nesting is indented by 2 spaces, sometimes 1, very rarely 4,
+  never 8 or any other odd number.
 
-  Indentation follows a simple logical format that only depends on the
-  nesting depth, not the accidental length of the text that initiates
-  a level of nesting.  Example:
+  Indentation follows a simple logical format that only depends on the nesting
+  depth, not the accidental length of the text that initiates a level of
+  nesting. Example:
 
   @{verbatim [display]
 \<open>  (* RIGHT *)
@@ -358,29 +357,27 @@
   else expr2_part1
        expr2_part2\<close>}
 
-  The second form has many problems: it assumes a fixed-width font
-  when viewing the sources, it uses more space on the line and thus
-  makes it hard to observe its strict length limit (working against
-  \<^emph>\<open>readability\<close>), it requires extra editing to adapt the layout
-  to changes of the initial text (working against
-  \<^emph>\<open>maintainability\<close>) etc.
+  The second form has many problems: it assumes a fixed-width font when
+  viewing the sources, it uses more space on the line and thus makes it hard
+  to observe its strict length limit (working against \<^emph>\<open>readability\<close>), it
+  requires extra editing to adapt the layout to changes of the initial text
+  (working against \<^emph>\<open>maintainability\<close>) etc.
 
   \<^medskip>
-  For similar reasons, any kind of two-dimensional or tabular
-  layouts, ASCII-art with lines or boxes of asterisks etc.\ should be
-  avoided.
+  For similar reasons, any kind of two-dimensional or tabular layouts,
+  ASCII-art with lines or boxes of asterisks etc.\ should be avoided.
 \<close>
 
 paragraph \<open>Complex expressions\<close>
-
-text \<open>that consist of multi-clausal function definitions, @{ML_text handle},
+text \<open>
+  that consist of multi-clausal function definitions, @{ML_text handle},
   @{ML_text case}, @{ML_text let} (and combinations) require special
   attention. The syntax of Standard ML is quite ambitious and admits a lot of
   variance that can distort the meaning of the text.
 
   Multiple clauses of @{ML_text fun}, @{ML_text fn}, @{ML_text handle},
-  @{ML_text case} get extra indentation to indicate the nesting
-  clearly.  Example:
+  @{ML_text case} get extra indentation to indicate the nesting clearly.
+  Example:
 
   @{verbatim [display]
 \<open>  (* RIGHT *)
@@ -398,10 +395,9 @@
     | foo p2 =
     expr2\<close>}
 
-  Body expressions consisting of @{ML_text case} or @{ML_text let}
-  require care to maintain compositionality, to prevent loss of
-  logical indentation where it is especially important to see the
-  structure of the text.  Example:
+  Body expressions consisting of @{ML_text case} or @{ML_text let} require
+  care to maintain compositionality, to prevent loss of logical indentation
+  where it is especially important to see the structure of the text. Example:
 
   @{verbatim [display]
 \<open>  (* RIGHT *)
@@ -430,13 +426,12 @@
       ...
     end\<close>}
 
-  Extra parentheses around @{ML_text case} expressions are optional,
-  but help to analyse the nesting based on character matching in the
-  text editor.
+  Extra parentheses around @{ML_text case} expressions are optional, but help
+  to analyse the nesting based on character matching in the text editor.
 
   \<^medskip>
-  There are two main exceptions to the overall principle of
-  compositionality in the layout of complex expressions.
+  There are two main exceptions to the overall principle of compositionality
+  in the layout of complex expressions.
 
   \<^enum> @{ML_text "if"} expressions are iterated as if ML had multi-branch
   conditionals, e.g.
@@ -448,10 +443,10 @@
   else if b2 then e2
   else e3\<close>}
 
-  \<^enum> @{ML_text fn} abstractions are often layed-out as if they
-  would lack any structure by themselves.  This traditional form is
-  motivated by the possibility to shift function arguments back and
-  forth wrt.\ additional combinators.  Example:
+  \<^enum> @{ML_text fn} abstractions are often layed-out as if they would lack any
+  structure by themselves. This traditional form is motivated by the
+  possibility to shift function arguments back and forth wrt.\ additional
+  combinators. Example:
 
   @{verbatim [display]
 \<open>  (* RIGHT *)
@@ -463,8 +458,8 @@
   @{ML_text y}, @{ML_text z} in a row.
 
 
-  Such weakly structured layout should be use with great care.  Here
-  are some counter-examples involving @{ML_text let} expressions:
+  Such weakly structured layout should be use with great care. Here are some
+  counter-examples involving @{ML_text let} expressions:
 
   @{verbatim [display]
 \<open>  (* WRONG *)
@@ -498,19 +493,19 @@
       ... end\<close>}
 
   \<^medskip>
-  In general the source layout is meant to emphasize the
-  structure of complex language expressions, not to pretend that SML
-  had a completely different syntax (say that of Haskell, Scala, Java).
+  In general the source layout is meant to emphasize the structure of complex
+  language expressions, not to pretend that SML had a completely different
+  syntax (say that of Haskell, Scala, Java).
 \<close>
 
 
 section \<open>ML embedded into Isabelle/Isar\<close>
 
-text \<open>ML and Isar are intertwined via an open-ended bootstrap
-  process that provides more and more programming facilities and
-  logical content in an alternating manner.  Bootstrapping starts from
-  the raw environment of existing implementations of Standard ML
-  (mainly Poly/ML, but also SML/NJ).
+text \<open>
+  ML and Isar are intertwined via an open-ended bootstrap process that
+  provides more and more programming facilities and logical content in an
+  alternating manner. Bootstrapping starts from the raw environment of
+  existing implementations of Standard ML (mainly Poly/ML, but also SML/NJ).
 
   Isabelle/Pure marks the point where the raw ML toplevel is superseded by
   Isabelle/ML within the Isar theory and proof language, with a uniform
@@ -521,11 +516,11 @@
   Object-logics like Isabelle/HOL are built within the Isabelle/ML/Isar
   environment by introducing suitable theories with associated ML modules,
   either inlined within \<^verbatim>\<open>.thy\<close> files, or as separate \<^verbatim>\<open>.ML\<close> files that are
-  loading from some theory. Thus Isabelle/HOL is defined
-  as a regular user-space application within the Isabelle framework. Further
-  add-on tools can be implemented in ML within the Isar context in the same
-  manner: ML is part of the standard repertoire of Isabelle, and there is no
-  distinction between ``users'' and ``developers'' in this respect.
+  loading from some theory. Thus Isabelle/HOL is defined as a regular
+  user-space application within the Isabelle framework. Further add-on tools
+  can be implemented in ML within the Isar context in the same manner: ML is
+  part of the standard repertoire of Isabelle, and there is no distinction
+  between ``users'' and ``developers'' in this respect.
 \<close>
 
 
@@ -543,10 +538,10 @@
   pending goal state via a given expression of type @{ML_type tactic}.
 \<close>
 
-text %mlex \<open>The following artificial example demonstrates some ML
-  toplevel declarations within the implicit Isar theory context.  This
-  is regular functional programming without referring to logical
-  entities yet.
+text %mlex \<open>
+  The following artificial example demonstrates some ML toplevel declarations
+  within the implicit Isar theory context. This is regular functional
+  programming without referring to logical entities yet.
 \<close>
 
 ML \<open>
@@ -554,24 +549,25 @@
     | factorial n = n * factorial (n - 1)
 \<close>
 
-text \<open>Here the ML environment is already managed by Isabelle, i.e.\
-  the @{ML factorial} function is not yet accessible in the preceding
-  paragraph, nor in a different theory that is independent from the
-  current one in the import hierarchy.
+text \<open>
+  Here the ML environment is already managed by Isabelle, i.e.\ the @{ML
+  factorial} function is not yet accessible in the preceding paragraph, nor in
+  a different theory that is independent from the current one in the import
+  hierarchy.
 
   Removing the above ML declaration from the source text will remove any trace
   of this definition, as expected. The Isabelle/ML toplevel environment is
-  managed in a \<^emph>\<open>stateless\<close> way: in contrast to the raw ML toplevel, there
-  are no global side-effects involved here.\<^footnote>\<open>Such a stateless
-  compilation environment is also a prerequisite for robust parallel
-  compilation within independent nodes of the implicit theory development
-  graph.\<close>
+  managed in a \<^emph>\<open>stateless\<close> way: in contrast to the raw ML toplevel, there are
+  no global side-effects involved here.\<^footnote>\<open>Such a stateless compilation
+  environment is also a prerequisite for robust parallel compilation within
+  independent nodes of the implicit theory development graph.\<close>
 
   \<^medskip>
-  The next example shows how to embed ML into Isar proofs, using
-  @{command_ref "ML_prf"} instead of @{command_ref "ML"}. As illustrated
-  below, the effect on the ML environment is local to the whole proof body,
-  but ignoring the block structure.\<close>
+  The next example shows how to embed ML into Isar proofs, using @{command_ref
+  "ML_prf"} instead of @{command_ref "ML"}. As illustrated below, the effect
+  on the ML environment is local to the whole proof body, but ignoring the
+  block structure.
+\<close>
 
 notepad
 begin
@@ -582,19 +578,19 @@
   ML_prf %"ML" \<open>val c = b + 1\<close>
 end
 
-text \<open>By side-stepping the normal scoping rules for Isar proof
-  blocks, embedded ML code can refer to the different contexts and
-  manipulate corresponding entities, e.g.\ export a fact from a block
-  context.
+text \<open>
+  By side-stepping the normal scoping rules for Isar proof blocks, embedded ML
+  code can refer to the different contexts and manipulate corresponding
+  entities, e.g.\ export a fact from a block context.
 
   \<^medskip>
-  Two further ML commands are useful in certain situations:
-  @{command_ref ML_val} and @{command_ref ML_command} are \<^emph>\<open>diagnostic\<close> in
-  the sense that there is no effect on the underlying environment, and can
-  thus be used anywhere. The examples below produce long strings of digits by
-  invoking @{ML factorial}: @{command ML_val} takes care of printing the ML
-  toplevel result, but @{command ML_command} is silent so we produce an
-  explicit output message.
+  Two further ML commands are useful in certain situations: @{command_ref
+  ML_val} and @{command_ref ML_command} are \<^emph>\<open>diagnostic\<close> in the sense that
+  there is no effect on the underlying environment, and can thus be used
+  anywhere. The examples below produce long strings of digits by invoking @{ML
+  factorial}: @{command ML_val} takes care of printing the ML toplevel result,
+  but @{command ML_command} is silent so we produce an explicit output
+  message.
 \<close>
 
 ML_val \<open>factorial 100\<close>
@@ -609,12 +605,13 @@
 
 subsection \<open>Compile-time context\<close>
 
-text \<open>Whenever the ML compiler is invoked within Isabelle/Isar, the
-  formal context is passed as a thread-local reference variable.  Thus
-  ML code may access the theory context during compilation, by reading
-  or writing the (local) theory under construction.  Note that such
-  direct access to the compile-time context is rare.  In practice it
-  is typically done via some derived ML functions instead.
+text \<open>
+  Whenever the ML compiler is invoked within Isabelle/Isar, the formal context
+  is passed as a thread-local reference variable. Thus ML code may access the
+  theory context during compilation, by reading or writing the (local) theory
+  under construction. Note that such direct access to the compile-time context
+  is rare. In practice it is typically done via some derived ML functions
+  instead.
 \<close>
 
 text %mlref \<open>
@@ -625,36 +622,35 @@
   @{index_ML ML_Thms.bind_thm: "string * thm -> unit"} \\
   \end{mldecls}
 
-  \<^descr> @{ML "ML_Context.the_generic_context ()"} refers to the theory
-  context of the ML toplevel --- at compile time.  ML code needs to
-  take care to refer to @{ML "ML_Context.the_generic_context ()"}
-  correctly.  Recall that evaluation of a function body is delayed
-  until actual run-time.
+    \<^descr> @{ML "ML_Context.the_generic_context ()"} refers to the theory context of
+    the ML toplevel --- at compile time. ML code needs to take care to refer to
+    @{ML "ML_Context.the_generic_context ()"} correctly. Recall that evaluation
+    of a function body is delayed until actual run-time.
 
-  \<^descr> @{ML "Context.>>"}~\<open>f\<close> applies context transformation
-  \<open>f\<close> to the implicit context of the ML toplevel.
+    \<^descr> @{ML "Context.>>"}~\<open>f\<close> applies context transformation \<open>f\<close> to the implicit
+    context of the ML toplevel.
 
-  \<^descr> @{ML ML_Thms.bind_thms}~\<open>(name, thms)\<close> stores a list of
-  theorems produced in ML both in the (global) theory context and the
-  ML toplevel, associating it with the provided name.
-
-  \<^descr> @{ML ML_Thms.bind_thm} is similar to @{ML ML_Thms.bind_thms} but
-  refers to a singleton fact.
+    \<^descr> @{ML ML_Thms.bind_thms}~\<open>(name, thms)\<close> stores a list of theorems produced
+    in ML both in the (global) theory context and the ML toplevel, associating
+    it with the provided name.
 
+    \<^descr> @{ML ML_Thms.bind_thm} is similar to @{ML ML_Thms.bind_thms} but refers to
+    a singleton fact.
 
-  It is important to note that the above functions are really
-  restricted to the compile time, even though the ML compiler is
-  invoked at run-time.  The majority of ML code either uses static
-  antiquotations (\secref{sec:ML-antiq}) or refers to the theory or
-  proof context at run-time, by explicit functional abstraction.
+  It is important to note that the above functions are really restricted to
+  the compile time, even though the ML compiler is invoked at run-time. The
+  majority of ML code either uses static antiquotations
+  (\secref{sec:ML-antiq}) or refers to the theory or proof context at
+  run-time, by explicit functional abstraction.
 \<close>
 
 
 subsection \<open>Antiquotations \label{sec:ML-antiq}\<close>
 
-text \<open>A very important consequence of embedding ML into Isar is the
-  concept of \<^emph>\<open>ML antiquotation\<close>.  The standard token language of
-  ML is augmented by special syntactic entities of the following form:
+text \<open>
+  A very important consequence of embedding ML into Isar is the concept of
+  \<^emph>\<open>ML antiquotation\<close>. The standard token language of ML is augmented by
+  special syntactic entities of the following form:
 
   @{rail \<open>
   @{syntax_def antiquote}: '@{' nameref args '}'
@@ -664,30 +660,29 @@
   defined in @{cite "isabelle-isar-ref"}.
 
   \<^medskip>
-  A regular antiquotation \<open>@{name args}\<close> processes
-  its arguments by the usual means of the Isar source language, and
-  produces corresponding ML source text, either as literal
-  \<^emph>\<open>inline\<close> text (e.g.\ \<open>@{term t}\<close>) or abstract
-  \<^emph>\<open>value\<close> (e.g. \<open>@{thm th}\<close>).  This pre-compilation
-  scheme allows to refer to formal entities in a robust manner, with
-  proper static scoping and with some degree of logical checking of
-  small portions of the code.
+  A regular antiquotation \<open>@{name args}\<close> processes its arguments by the usual
+  means of the Isar source language, and produces corresponding ML source
+  text, either as literal \<^emph>\<open>inline\<close> text (e.g.\ \<open>@{term t}\<close>) or abstract
+  \<^emph>\<open>value\<close> (e.g. \<open>@{thm th}\<close>). This pre-compilation scheme allows to refer to
+  formal entities in a robust manner, with proper static scoping and with some
+  degree of logical checking of small portions of the code.
 \<close>
 
 
 subsection \<open>Printing ML values\<close>
 
-text \<open>The ML compiler knows about the structure of values according
-  to their static type, and can print them in the manner of its
-  toplevel, although the details are non-portable.  The
-  antiquotations @{ML_antiquotation_def "make_string"} and
-  @{ML_antiquotation_def "print"} provide a quasi-portable way to
-  refer to this potential capability of the underlying ML system in
+text \<open>
+  The ML compiler knows about the structure of values according to their
+  static type, and can print them in the manner of its toplevel, although the
+  details are non-portable. The antiquotations @{ML_antiquotation_def
+  "make_string"} and @{ML_antiquotation_def "print"} provide a quasi-portable
+  way to refer to this potential capability of the underlying ML system in
   generic Isabelle/ML sources.
 
-  This is occasionally useful for diagnostic or demonstration
-  purposes.  Note that production-quality tools require proper
-  user-level error messages, avoiding raw ML values in the output.\<close>
+  This is occasionally useful for diagnostic or demonstration purposes. Note
+  that production-quality tools require proper user-level error messages,
+  avoiding raw ML values in the output.
+\<close>
 
 text %mlantiq \<open>
   \begin{matharray}{rcl}
@@ -701,19 +696,20 @@
   @@{ML_antiquotation print} @{syntax name}?
   \<close>}
 
-  \<^descr> \<open>@{make_string}\<close> inlines a function to print arbitrary values
-  similar to the ML toplevel. The result is compiler dependent and may fall
-  back on "?" in certain situations. The value of configuration option
-  @{attribute_ref ML_print_depth} determines further details of output.
+  \<^descr> \<open>@{make_string}\<close> inlines a function to print arbitrary values similar to
+  the ML toplevel. The result is compiler dependent and may fall back on "?"
+  in certain situations. The value of configuration option @{attribute_ref
+  ML_print_depth} determines further details of output.
 
-  \<^descr> \<open>@{print f}\<close> uses the ML function \<open>f: string ->
-  unit\<close> to output the result of \<open>@{make_string}\<close> above,
-  together with the source position of the antiquotation.  The default
-  output function is @{ML writeln}.
+  \<^descr> \<open>@{print f}\<close> uses the ML function \<open>f: string -> unit\<close> to output the result
+  of \<open>@{make_string}\<close> above, together with the source position of the
+  antiquotation. The default output function is @{ML writeln}.
 \<close>
 
-text %mlex \<open>The following artificial examples show how to produce
-  adhoc output of ML values for debugging purposes.\<close>
+text %mlex \<open>
+  The following artificial examples show how to produce adhoc output of ML
+  values for debugging purposes.
+\<close>
 
 ML_val \<open>
   val x = 42;
@@ -728,47 +724,44 @@
 
 section \<open>Canonical argument order \label{sec:canonical-argument-order}\<close>
 
-text \<open>Standard ML is a language in the tradition of \<open>\<lambda>\<close>-calculus and \<^emph>\<open>higher-order functional programming\<close>,
-  similar to OCaml, Haskell, or Isabelle/Pure and HOL as logical
-  languages.  Getting acquainted with the native style of representing
-  functions in that setting can save a lot of extra boiler-plate of
-  redundant shuffling of arguments, auxiliary abstractions etc.
+text \<open>
+  Standard ML is a language in the tradition of \<open>\<lambda>\<close>-calculus and
+  \<^emph>\<open>higher-order functional programming\<close>, similar to OCaml, Haskell, or
+  Isabelle/Pure and HOL as logical languages. Getting acquainted with the
+  native style of representing functions in that setting can save a lot of
+  extra boiler-plate of redundant shuffling of arguments, auxiliary
+  abstractions etc.
 
-  Functions are usually \<^emph>\<open>curried\<close>: the idea of turning arguments
-  of type \<open>\<tau>\<^sub>i\<close> (for \<open>i \<in> {1, \<dots> n}\<close>) into a result of
-  type \<open>\<tau>\<close> is represented by the iterated function space
-  \<open>\<tau>\<^sub>1 \<rightarrow> \<dots> \<rightarrow> \<tau>\<^sub>n \<rightarrow> \<tau>\<close>.  This is isomorphic to the well-known
-  encoding via tuples \<open>\<tau>\<^sub>1 \<times> \<dots> \<times> \<tau>\<^sub>n \<rightarrow> \<tau>\<close>, but the curried
-  version fits more smoothly into the basic calculus.\<^footnote>\<open>The
-  difference is even more significant in HOL, because the redundant
-  tuple structure needs to be accommodated extraneous proof steps.\<close>
+  Functions are usually \<^emph>\<open>curried\<close>: the idea of turning arguments of type
+  \<open>\<tau>\<^sub>i\<close> (for \<open>i \<in> {1, \<dots> n}\<close>) into a result of type \<open>\<tau>\<close> is represented by the
+  iterated function space \<open>\<tau>\<^sub>1 \<rightarrow> \<dots> \<rightarrow> \<tau>\<^sub>n \<rightarrow> \<tau>\<close>. This is isomorphic to the
+  well-known encoding via tuples \<open>\<tau>\<^sub>1 \<times> \<dots> \<times> \<tau>\<^sub>n \<rightarrow> \<tau>\<close>, but the curried version
+  fits more smoothly into the basic calculus.\<^footnote>\<open>The difference is even more
+  significant in HOL, because the redundant tuple structure needs to be
+  accommodated extraneous proof steps.\<close>
 
-  Currying gives some flexibility due to \<^emph>\<open>partial application\<close>.  A
-  function \<open>f: \<tau>\<^sub>1 \<rightarrow> \<tau>\<^sub>2 \<rightarrow> \<tau>\<close> can be applied to \<open>x: \<tau>\<^sub>1\<close>
-  and the remaining \<open>(f x): \<tau>\<^sub>2 \<rightarrow> \<tau>\<close> passed to another function
-  etc.  How well this works in practice depends on the order of
-  arguments.  In the worst case, arguments are arranged erratically,
-  and using a function in a certain situation always requires some
-  glue code.  Thus we would get exponentially many opportunities to
+  Currying gives some flexibility due to \<^emph>\<open>partial application\<close>. A function
+  \<open>f: \<tau>\<^sub>1 \<rightarrow> \<tau>\<^sub>2 \<rightarrow> \<tau>\<close> can be applied to \<open>x: \<tau>\<^sub>1\<close> and the remaining \<open>(f x): \<tau>\<^sub>2
+  \<rightarrow> \<tau>\<close> passed to another function etc. How well this works in practice depends
+  on the order of arguments. In the worst case, arguments are arranged
+  erratically, and using a function in a certain situation always requires
+  some glue code. Thus we would get exponentially many opportunities to
   decorate the code with meaningless permutations of arguments.
 
-  This can be avoided by \<^emph>\<open>canonical argument order\<close>, which
-  observes certain standard patterns and minimizes adhoc permutations
-  in their application.  In Isabelle/ML, large portions of text can be
-  written without auxiliary operations like \<open>swap: \<alpha> \<times> \<beta> \<rightarrow> \<beta> \<times>
-  \<alpha>\<close> or \<open>C: (\<alpha> \<rightarrow> \<beta> \<rightarrow> \<gamma>) \<rightarrow> (\<beta> \<rightarrow> \<alpha> \<rightarrow> \<gamma>)\<close> (the latter is not
-  present in the Isabelle/ML library).
+  This can be avoided by \<^emph>\<open>canonical argument order\<close>, which observes certain
+  standard patterns and minimizes adhoc permutations in their application. In
+  Isabelle/ML, large portions of text can be written without auxiliary
+  operations like \<open>swap: \<alpha> \<times> \<beta> \<rightarrow> \<beta> \<times> \<alpha>\<close> or \<open>C: (\<alpha> \<rightarrow> \<beta> \<rightarrow> \<gamma>) \<rightarrow> (\<beta> \<rightarrow> \<alpha> \<rightarrow> \<gamma>)\<close> (the
+  latter is not present in the Isabelle/ML library).
 
   \<^medskip>
-  The main idea is that arguments that vary less are moved
-  further to the left than those that vary more.  Two particularly
-  important categories of functions are \<^emph>\<open>selectors\<close> and
-  \<^emph>\<open>updates\<close>.
+  The main idea is that arguments that vary less are moved further to the left
+  than those that vary more. Two particularly important categories of
+  functions are \<^emph>\<open>selectors\<close> and \<^emph>\<open>updates\<close>.
 
-  The subsequent scheme is based on a hypothetical set-like container
-  of type \<open>\<beta>\<close> that manages elements of type \<open>\<alpha>\<close>.  Both
-  the names and types of the associated operations are canonical for
-  Isabelle/ML.
+  The subsequent scheme is based on a hypothetical set-like container of type
+  \<open>\<beta>\<close> that manages elements of type \<open>\<alpha>\<close>. Both the names and types of the
+  associated operations are canonical for Isabelle/ML.
 
   \begin{center}
   \begin{tabular}{ll}
@@ -778,35 +771,33 @@
   \end{tabular}
   \end{center}
 
-  Given a container \<open>B: \<beta>\<close>, the partially applied \<open>member B\<close> is a predicate over elements \<open>\<alpha> \<rightarrow> bool\<close>, and
-  thus represents the intended denotation directly.  It is customary
-  to pass the abstract predicate to further operations, not the
-  concrete container.  The argument order makes it easy to use other
-  combinators: \<open>forall (member B) list\<close> will check a list of
-  elements for membership in \<open>B\<close> etc. Often the explicit
-  \<open>list\<close> is pointless and can be contracted to \<open>forall
-  (member B)\<close> to get directly a predicate again.
+  Given a container \<open>B: \<beta>\<close>, the partially applied \<open>member B\<close> is a predicate
+  over elements \<open>\<alpha> \<rightarrow> bool\<close>, and thus represents the intended denotation
+  directly. It is customary to pass the abstract predicate to further
+  operations, not the concrete container. The argument order makes it easy to
+  use other combinators: \<open>forall (member B) list\<close> will check a list of
+  elements for membership in \<open>B\<close> etc. Often the explicit \<open>list\<close> is pointless
+  and can be contracted to \<open>forall (member B)\<close> to get directly a predicate
+  again.
 
-  In contrast, an update operation varies the container, so it moves
-  to the right: \<open>insert a\<close> is a function \<open>\<beta> \<rightarrow> \<beta>\<close> to
-  insert a value \<open>a\<close>.  These can be composed naturally as
-  \<open>insert c \<circ> insert b \<circ> insert a\<close>.  The slightly awkward
-  inversion of the composition order is due to conventional
-  mathematical notation, which can be easily amended as explained
-  below.
+  In contrast, an update operation varies the container, so it moves to the
+  right: \<open>insert a\<close> is a function \<open>\<beta> \<rightarrow> \<beta>\<close> to insert a value \<open>a\<close>. These can be
+  composed naturally as \<open>insert c \<circ> insert b \<circ> insert a\<close>. The slightly awkward
+  inversion of the composition order is due to conventional mathematical
+  notation, which can be easily amended as explained below.
 \<close>
 
 
 subsection \<open>Forward application and composition\<close>
 
-text \<open>Regular function application and infix notation works best for
-  relatively deeply structured expressions, e.g.\ \<open>h (f x y + g
-  z)\<close>.  The important special case of \<^emph>\<open>linear transformation\<close>
-  applies a cascade of functions \<open>f\<^sub>n (\<dots> (f\<^sub>1 x))\<close>.  This
-  becomes hard to read and maintain if the functions are themselves
-  given as complex expressions.  The notation can be significantly
-  improved by introducing \<^emph>\<open>forward\<close> versions of application and
-  composition as follows:
+text \<open>
+  Regular function application and infix notation works best for relatively
+  deeply structured expressions, e.g.\ \<open>h (f x y + g z)\<close>. The important
+  special case of \<^emph>\<open>linear transformation\<close> applies a cascade of functions \<open>f\<^sub>n
+  (\<dots> (f\<^sub>1 x))\<close>. This becomes hard to read and maintain if the functions are
+  themselves given as complex expressions. The notation can be significantly
+  improved by introducing \<^emph>\<open>forward\<close> versions of application and composition
+  as follows:
 
   \<^medskip>
   \begin{tabular}{lll}
@@ -815,13 +806,12 @@
   \end{tabular}
   \<^medskip>
 
-  This enables to write conveniently \<open>x |> f\<^sub>1 |> \<dots> |> f\<^sub>n\<close> or
-  \<open>f\<^sub>1 #> \<dots> #> f\<^sub>n\<close> for its functional abstraction over \<open>x\<close>.
+  This enables to write conveniently \<open>x |> f\<^sub>1 |> \<dots> |> f\<^sub>n\<close> or \<open>f\<^sub>1 #> \<dots> #>
+  f\<^sub>n\<close> for its functional abstraction over \<open>x\<close>.
 
   \<^medskip>
-  There is an additional set of combinators to accommodate
-  multiple results (via pairs) that are passed on as multiple
-  arguments (via currying).
+  There is an additional set of combinators to accommodate multiple results
+  (via pairs) that are passed on as multiple arguments (via currying).
 
   \<^medskip>
   \begin{tabular}{lll}
@@ -843,26 +833,27 @@
 
 subsection \<open>Canonical iteration\<close>
 
-text \<open>As explained above, a function \<open>f: \<alpha> \<rightarrow> \<beta> \<rightarrow> \<beta>\<close> can be
-  understood as update on a configuration of type \<open>\<beta>\<close>,
-  parameterized by an argument of type \<open>\<alpha>\<close>.  Given \<open>a: \<alpha>\<close>
-  the partial application \<open>(f a): \<beta> \<rightarrow> \<beta>\<close> operates
-  homogeneously on \<open>\<beta>\<close>.  This can be iterated naturally over a
-  list of parameters \<open>[a\<^sub>1, \<dots>, a\<^sub>n]\<close> as \<open>f a\<^sub>1 #> \<dots> #> f a\<^sub>n\<close>.
-  The latter expression is again a function \<open>\<beta> \<rightarrow> \<beta>\<close>.
-  It can be applied to an initial configuration \<open>b: \<beta>\<close> to
-  start the iteration over the given list of arguments: each \<open>a\<close> in \<open>a\<^sub>1, \<dots>, a\<^sub>n\<close> is applied consecutively by updating a
-  cumulative configuration.
+text \<open>
+  As explained above, a function \<open>f: \<alpha> \<rightarrow> \<beta> \<rightarrow> \<beta>\<close> can be understood as update on
+  a configuration of type \<open>\<beta>\<close>, parameterized by an argument of type \<open>\<alpha>\<close>. Given
+  \<open>a: \<alpha>\<close> the partial application \<open>(f a): \<beta> \<rightarrow> \<beta>\<close> operates homogeneously on \<open>\<beta>\<close>.
+  This can be iterated naturally over a list of parameters \<open>[a\<^sub>1, \<dots>, a\<^sub>n]\<close> as
+  \<open>f a\<^sub>1 #> \<dots> #> f a\<^sub>n\<close>. The latter expression is again a function \<open>\<beta> \<rightarrow> \<beta>\<close>. It
+  can be applied to an initial configuration \<open>b: \<beta>\<close> to start the iteration
+  over the given list of arguments: each \<open>a\<close> in \<open>a\<^sub>1, \<dots>, a\<^sub>n\<close> is applied
+  consecutively by updating a cumulative configuration.
 
-  The \<open>fold\<close> combinator in Isabelle/ML lifts a function \<open>f\<close> as above to its iterated version over a list of arguments.
-  Lifting can be repeated, e.g.\ \<open>(fold \<circ> fold) f\<close> iterates
-  over a list of lists as expected.
+  The \<open>fold\<close> combinator in Isabelle/ML lifts a function \<open>f\<close> as above to its
+  iterated version over a list of arguments. Lifting can be repeated, e.g.\
+  \<open>(fold \<circ> fold) f\<close> iterates over a list of lists as expected.
 
-  The variant \<open>fold_rev\<close> works inside-out over the list of
-  arguments, such that \<open>fold_rev f \<equiv> fold f \<circ> rev\<close> holds.
+  The variant \<open>fold_rev\<close> works inside-out over the list of arguments, such
+  that \<open>fold_rev f \<equiv> fold f \<circ> rev\<close> holds.
 
-  The \<open>fold_map\<close> combinator essentially performs \<open>fold\<close> and \<open>map\<close> simultaneously: each application of \<open>f\<close> produces an updated configuration together with a side-result;
-  the iteration collects all such side-results as a separate list.
+  The \<open>fold_map\<close> combinator essentially performs \<open>fold\<close> and \<open>map\<close>
+  simultaneously: each application of \<open>f\<close> produces an updated configuration
+  together with a side-result; the iteration collects all such side-results as
+  a separate list.
 \<close>
 
 text %mlref \<open>
@@ -872,30 +863,29 @@
   @{index_ML fold_map: "('a -> 'b -> 'c * 'b) -> 'a list -> 'b -> 'c list * 'b"} \\
   \end{mldecls}
 
-  \<^descr> @{ML fold}~\<open>f\<close> lifts the parametrized update function
-  \<open>f\<close> to a list of parameters.
+  \<^descr> @{ML fold}~\<open>f\<close> lifts the parametrized update function \<open>f\<close> to a list of
+  parameters.
 
-  \<^descr> @{ML fold_rev}~\<open>f\<close> is similar to @{ML fold}~\<open>f\<close>, but works inside-out, as if the list would be reversed.
+  \<^descr> @{ML fold_rev}~\<open>f\<close> is similar to @{ML fold}~\<open>f\<close>, but works inside-out, as
+  if the list would be reversed.
 
-  \<^descr> @{ML fold_map}~\<open>f\<close> lifts the parametrized update
-  function \<open>f\<close> (with side-result) to a list of parameters and
-  cumulative side-results.
+  \<^descr> @{ML fold_map}~\<open>f\<close> lifts the parametrized update function \<open>f\<close> (with
+  side-result) to a list of parameters and cumulative side-results.
 
 
   \begin{warn}
   The literature on functional programming provides a confusing multitude of
-  combinators called \<open>foldl\<close>, \<open>foldr\<close> etc. SML97 provides its
-  own variations as @{ML List.foldl} and @{ML List.foldr}, while the classic
-  Isabelle library also has the historic @{ML Library.foldl} and @{ML
-  Library.foldr}. To avoid unnecessary complication, all these historical
-  versions should be ignored, and the canonical @{ML fold} (or @{ML fold_rev})
-  used exclusively.
+  combinators called \<open>foldl\<close>, \<open>foldr\<close> etc. SML97 provides its own variations
+  as @{ML List.foldl} and @{ML List.foldr}, while the classic Isabelle library
+  also has the historic @{ML Library.foldl} and @{ML Library.foldr}. To avoid
+  unnecessary complication, all these historical versions should be ignored,
+  and the canonical @{ML fold} (or @{ML fold_rev}) used exclusively.
   \end{warn}
 \<close>
 
-text %mlex \<open>The following example shows how to fill a text buffer
-  incrementally by adding strings, either individually or from a given
-  list.
+text %mlex \<open>
+  The following example shows how to fill a text buffer incrementally by
+  adding strings, either individually or from a given list.
 \<close>
 
 ML_val \<open>
@@ -908,16 +898,15 @@
   @{assert} (s = "digits: 0123456789");
 \<close>
 
-text \<open>Note how @{ML "fold (Buffer.add o string_of_int)"} above saves
-  an extra @{ML "map"} over the given list.  This kind of peephole
-  optimization reduces both the code size and the tree structures in
-  memory (``deforestation''), but it requires some practice to read
-  and write fluently.
+text \<open>
+  Note how @{ML "fold (Buffer.add o string_of_int)"} above saves an extra @{ML
+  "map"} over the given list. This kind of peephole optimization reduces both
+  the code size and the tree structures in memory (``deforestation''), but it
+  requires some practice to read and write fluently.
 
   \<^medskip>
-  The next example elaborates the idea of canonical
-  iteration, demonstrating fast accumulation of tree content using a
-  text buffer.
+  The next example elaborates the idea of canonical iteration, demonstrating
+  fast accumulation of tree content using a text buffer.
 \<close>
 
 ML \<open>
@@ -939,51 +928,51 @@
     Buffer.empty |> add_content tree |> Buffer.content;
 \<close>
 
-text \<open>The slowness of @{ML slow_content} is due to the @{ML implode} of
-  the recursive results, because it copies previously produced strings
-  again and again.
+text \<open>
+  The slowness of @{ML slow_content} is due to the @{ML implode} of the
+  recursive results, because it copies previously produced strings again and
+  again.
 
-  The incremental @{ML add_content} avoids this by operating on a
-  buffer that is passed through in a linear fashion.  Using @{ML_text
-  "#>"} and contraction over the actual buffer argument saves some
-  additional boiler-plate.  Of course, the two @{ML "Buffer.add"}
-  invocations with concatenated strings could have been split into
-  smaller parts, but this would have obfuscated the source without
-  making a big difference in performance.  Here we have done some
-  peephole-optimization for the sake of readability.
+  The incremental @{ML add_content} avoids this by operating on a buffer that
+  is passed through in a linear fashion. Using @{ML_text "#>"} and contraction
+  over the actual buffer argument saves some additional boiler-plate. Of
+  course, the two @{ML "Buffer.add"} invocations with concatenated strings
+  could have been split into smaller parts, but this would have obfuscated the
+  source without making a big difference in performance. Here we have done
+  some peephole-optimization for the sake of readability.
 
-  Another benefit of @{ML add_content} is its ``open'' form as a
-  function on buffers that can be continued in further linear
-  transformations, folding etc.  Thus it is more compositional than
-  the naive @{ML slow_content}.  As realistic example, compare the
-  old-style @{ML "Term.maxidx_of_term: term -> int"} with the newer
-  @{ML "Term.maxidx_term: term -> int -> int"} in Isabelle/Pure.
+  Another benefit of @{ML add_content} is its ``open'' form as a function on
+  buffers that can be continued in further linear transformations, folding
+  etc. Thus it is more compositional than the naive @{ML slow_content}. As
+  realistic example, compare the old-style @{ML "Term.maxidx_of_term: term ->
+  int"} with the newer @{ML "Term.maxidx_term: term -> int -> int"} in
+  Isabelle/Pure.
 
-  Note that @{ML fast_content} above is only defined as example.  In
-  many practical situations, it is customary to provide the
-  incremental @{ML add_content} only and leave the initialization and
-  termination to the concrete application to the user.
+  Note that @{ML fast_content} above is only defined as example. In many
+  practical situations, it is customary to provide the incremental @{ML
+  add_content} only and leave the initialization and termination to the
+  concrete application to the user.
 \<close>
 
 
 section \<open>Message output channels \label{sec:message-channels}\<close>
 
-text \<open>Isabelle provides output channels for different kinds of
-  messages: regular output, high-volume tracing information, warnings,
-  and errors.
+text \<open>
+  Isabelle provides output channels for different kinds of messages: regular
+  output, high-volume tracing information, warnings, and errors.
 
-  Depending on the user interface involved, these messages may appear
-  in different text styles or colours.  The standard output for
-  batch sessions prefixes each line of warnings by \<^verbatim>\<open>###\<close> and errors by
-  \<^verbatim>\<open>***\<close>, but leaves anything else
-  unchanged.  The message body may contain further markup and formatting,
-  which is routinely used in the Prover IDE @{cite "isabelle-jedit"}.
+  Depending on the user interface involved, these messages may appear in
+  different text styles or colours. The standard output for batch sessions
+  prefixes each line of warnings by \<^verbatim>\<open>###\<close> and errors by \<^verbatim>\<open>***\<close>, but leaves
+  anything else unchanged. The message body may contain further markup and
+  formatting, which is routinely used in the Prover IDE @{cite
+  "isabelle-jedit"}.
 
-  Messages are associated with the transaction context of the running
-  Isar command.  This enables the front-end to manage commands and
-  resulting messages together.  For example, after deleting a command
-  from a given theory document version, the corresponding message
-  output can be retracted from the display.
+  Messages are associated with the transaction context of the running Isar
+  command. This enables the front-end to manage commands and resulting
+  messages together. For example, after deleting a command from a given theory
+  document version, the corresponding message output can be retracted from the
+  display.
 \<close>
 
 text %mlref \<open>
@@ -994,32 +983,29 @@
   @{index_ML error: "string -> 'a"} % FIXME Output.error_message (!?) \\
   \end{mldecls}
 
-  \<^descr> @{ML writeln}~\<open>text\<close> outputs \<open>text\<close> as regular
-  message.  This is the primary message output operation of Isabelle
-  and should be used by default.
+  \<^descr> @{ML writeln}~\<open>text\<close> outputs \<open>text\<close> as regular message. This is the
+  primary message output operation of Isabelle and should be used by default.
 
-  \<^descr> @{ML tracing}~\<open>text\<close> outputs \<open>text\<close> as special
-  tracing message, indicating potential high-volume output to the
-  front-end (hundreds or thousands of messages issued by a single
-  command).  The idea is to allow the user-interface to downgrade the
-  quality of message display to achieve higher throughput.
+  \<^descr> @{ML tracing}~\<open>text\<close> outputs \<open>text\<close> as special tracing message, indicating
+  potential high-volume output to the front-end (hundreds or thousands of
+  messages issued by a single command). The idea is to allow the
+  user-interface to downgrade the quality of message display to achieve higher
+  throughput.
 
-  Note that the user might have to take special actions to see tracing
-  output, e.g.\ switch to a different output window.  So this channel
-  should not be used for regular output.
+  Note that the user might have to take special actions to see tracing output,
+  e.g.\ switch to a different output window. So this channel should not be
+  used for regular output.
 
-  \<^descr> @{ML warning}~\<open>text\<close> outputs \<open>text\<close> as
-  warning, which typically means some extra emphasis on the front-end
-  side (color highlighting, icons, etc.).
+  \<^descr> @{ML warning}~\<open>text\<close> outputs \<open>text\<close> as warning, which typically means some
+  extra emphasis on the front-end side (color highlighting, icons, etc.).
 
-  \<^descr> @{ML error}~\<open>text\<close> raises exception @{ML ERROR}~\<open>text\<close> and thus lets the Isar toplevel print \<open>text\<close> on the
-  error channel, which typically means some extra emphasis on the
-  front-end side (color highlighting, icons, etc.).
+  \<^descr> @{ML error}~\<open>text\<close> raises exception @{ML ERROR}~\<open>text\<close> and thus lets the
+  Isar toplevel print \<open>text\<close> on the error channel, which typically means some
+  extra emphasis on the front-end side (color highlighting, icons, etc.).
 
   This assumes that the exception is not handled before the command
-  terminates.  Handling exception @{ML ERROR}~\<open>text\<close> is a
-  perfectly legal alternative: it means that the error is absorbed
-  without any message output.
+  terminates. Handling exception @{ML ERROR}~\<open>text\<close> is a perfectly legal
+  alternative: it means that the error is absorbed without any message output.
 
   \begin{warn}
   The actual error channel is accessed via @{ML Output.error_message}, but
@@ -1028,28 +1014,28 @@
 
 
   \begin{warn}
-  Regular Isabelle/ML code should output messages exclusively by the
-  official channels.  Using raw I/O on \<^emph>\<open>stdout\<close> or \<^emph>\<open>stderr\<close>
-  instead (e.g.\ via @{ML TextIO.output}) is apt to cause problems in
-  the presence of parallel and asynchronous processing of Isabelle
-  theories.  Such raw output might be displayed by the front-end in
-  some system console log, with a low chance that the user will ever
-  see it.  Moreover, as a genuine side-effect on global process
-  channels, there is no proper way to retract output when Isar command
+  Regular Isabelle/ML code should output messages exclusively by the official
+  channels. Using raw I/O on \<^emph>\<open>stdout\<close> or \<^emph>\<open>stderr\<close> instead (e.g.\ via @{ML
+  TextIO.output}) is apt to cause problems in the presence of parallel and
+  asynchronous processing of Isabelle theories. Such raw output might be
+  displayed by the front-end in some system console log, with a low chance
+  that the user will ever see it. Moreover, as a genuine side-effect on global
+  process channels, there is no proper way to retract output when Isar command
   transactions are reset by the system.
   \end{warn}
 
   \begin{warn}
-  The message channels should be used in a message-oriented manner.
-  This means that multi-line output that logically belongs together is
-  issued by a single invocation of @{ML writeln} etc.\ with the
-  functional concatenation of all message constituents.
+  The message channels should be used in a message-oriented manner. This means
+  that multi-line output that logically belongs together is issued by a single
+  invocation of @{ML writeln} etc.\ with the functional concatenation of all
+  message constituents.
   \end{warn}
 \<close>
 
-text %mlex \<open>The following example demonstrates a multi-line
-  warning.  Note that in some situations the user sees only the first
-  line, so the most important point should be made first.
+text %mlex \<open>
+  The following example demonstrates a multi-line warning. Note that in some
+  situations the user sees only the first line, so the most important point
+  should be made first.
 \<close>
 
 ML_command \<open>
@@ -1062,8 +1048,7 @@
 
 text \<open>
   \<^medskip>
-  An alternative is to make a paragraph of freely-floating words as
-  follows.
+  An alternative is to make a paragraph of freely-floating words as follows.
 \<close>
 
 ML_command \<open>
@@ -1083,60 +1068,60 @@
 
 section \<open>Exceptions \label{sec:exceptions}\<close>
 
-text \<open>The Standard ML semantics of strict functional evaluation
-  together with exceptions is rather well defined, but some delicate
-  points need to be observed to avoid that ML programs go wrong
-  despite static type-checking.  Exceptions in Isabelle/ML are
-  subsequently categorized as follows.\<close>
+text \<open>
+  The Standard ML semantics of strict functional evaluation together with
+  exceptions is rather well defined, but some delicate points need to be
+  observed to avoid that ML programs go wrong despite static type-checking.
+  Exceptions in Isabelle/ML are subsequently categorized as follows.
+\<close>
 
 paragraph \<open>Regular user errors.\<close>
-text \<open>These are meant to provide informative feedback about malformed input
-  etc.
+text \<open>
+  These are meant to provide informative feedback about malformed input etc.
+
+  The \<^emph>\<open>error\<close> function raises the corresponding @{ML ERROR} exception, with a
+  plain text message as argument. @{ML ERROR} exceptions can be handled
+  internally, in order to be ignored, turned into other exceptions, or
+  cascaded by appending messages. If the corresponding Isabelle/Isar command
+  terminates with an @{ML ERROR} exception state, the system will print the
+  result on the error channel (see \secref{sec:message-channels}).
 
-  The \<^emph>\<open>error\<close> function raises the corresponding @{ML ERROR}
-  exception, with a plain text message as argument.  @{ML ERROR}
-  exceptions can be handled internally, in order to be ignored, turned
-  into other exceptions, or cascaded by appending messages.  If the
-  corresponding Isabelle/Isar command terminates with an @{ML ERROR}
-  exception state, the system will print the result on the error
-  channel (see \secref{sec:message-channels}).
+  It is considered bad style to refer to internal function names or values in
+  ML source notation in user error messages. Do not use \<open>@{make_string}\<close> nor
+  \<open>@{here}\<close>!
 
-  It is considered bad style to refer to internal function names or
-  values in ML source notation in user error messages.  Do not use
-  \<open>@{make_string}\<close> nor \<open>@{here}\<close>!
-
-  Grammatical correctness of error messages can be improved by
-  \<^emph>\<open>omitting\<close> final punctuation: messages are often concatenated
-  or put into a larger context (e.g.\ augmented with source position).
-  Note that punctuation after formal entities (types, terms, theorems) is
-  particularly prone to user confusion.
+  Grammatical correctness of error messages can be improved by \<^emph>\<open>omitting\<close>
+  final punctuation: messages are often concatenated or put into a larger
+  context (e.g.\ augmented with source position). Note that punctuation after
+  formal entities (types, terms, theorems) is particularly prone to user
+  confusion.
 \<close>
 
 paragraph \<open>Program failures.\<close>
-text \<open>There is a handful of standard exceptions that indicate general failure
+text \<open>
+  There is a handful of standard exceptions that indicate general failure
   situations, or failures of core operations on logical entities (types,
   terms, theorems, theories, see \chref{ch:logic}).
 
-  These exceptions indicate a genuine breakdown of the program, so the
-  main purpose is to determine quickly what has happened where.
-  Traditionally, the (short) exception message would include the name
-  of an ML function, although this is no longer necessary, because the
-  ML runtime system attaches detailed source position stemming from the
-  corresponding @{ML_text raise} keyword.
+  These exceptions indicate a genuine breakdown of the program, so the main
+  purpose is to determine quickly what has happened where. Traditionally, the
+  (short) exception message would include the name of an ML function, although
+  this is no longer necessary, because the ML runtime system attaches detailed
+  source position stemming from the corresponding @{ML_text raise} keyword.
 
   \<^medskip>
-  User modules can always introduce their own custom
-  exceptions locally, e.g.\ to organize internal failures robustly
-  without overlapping with existing exceptions.  Exceptions that are
-  exposed in module signatures require extra care, though, and should
-  \<^emph>\<open>not\<close> be introduced by default.  Surprise by users of a module
-  can be often minimized by using plain user errors instead.
+  User modules can always introduce their own custom exceptions locally, e.g.\
+  to organize internal failures robustly without overlapping with existing
+  exceptions. Exceptions that are exposed in module signatures require extra
+  care, though, and should \<^emph>\<open>not\<close> be introduced by default. Surprise by users
+  of a module can be often minimized by using plain user errors instead.
 \<close>
 
 paragraph \<open>Interrupts.\<close>
-text \<open>These indicate arbitrary system events: both the ML runtime system and
-  the Isabelle/ML infrastructure signal various exceptional situations by
-  raising the special @{ML Exn.Interrupt} exception in user code.
+text \<open>
+  These indicate arbitrary system events: both the ML runtime system and the
+  Isabelle/ML infrastructure signal various exceptional situations by raising
+  the special @{ML Exn.Interrupt} exception in user code.
 
   This is the one and only way that physical events can intrude an Isabelle/ML
   program. Such an interrupt can mean out-of-memory, stack overflow, timeout,
@@ -1147,20 +1132,18 @@
   interrupts unintentionally and thus render the program semantics
   ill-defined.
 
-  Note that the Interrupt exception dates back to the original SML90
-  language definition.  It was excluded from the SML97 version to
-  avoid its malign impact on ML program semantics, but without
-  providing a viable alternative.  Isabelle/ML recovers physical
-  interruptibility (which is an indispensable tool to implement
-  managed evaluation of command transactions), but requires user code
-  to be strictly transparent wrt.\ interrupts.
+  Note that the Interrupt exception dates back to the original SML90 language
+  definition. It was excluded from the SML97 version to avoid its malign
+  impact on ML program semantics, but without providing a viable alternative.
+  Isabelle/ML recovers physical interruptibility (which is an indispensable
+  tool to implement managed evaluation of command transactions), but requires
+  user code to be strictly transparent wrt.\ interrupts.
 
   \begin{warn}
-  Isabelle/ML user code needs to terminate promptly on interruption,
-  without guessing at its meaning to the system infrastructure.
-  Temporary handling of interrupts for cleanup of global resources
-  etc.\ needs to be followed immediately by re-raising of the original
-  exception.
+  Isabelle/ML user code needs to terminate promptly on interruption, without
+  guessing at its meaning to the system infrastructure. Temporary handling of
+  interrupts for cleanup of global resources etc.\ needs to be followed
+  immediately by re-raising of the original exception.
   \end{warn}
 \<close>
 
@@ -1175,35 +1158,33 @@
   @{index_ML Runtime.exn_trace: "(unit -> 'a) -> 'a"} \\
   \end{mldecls}
 
-  \<^descr> @{ML try}~\<open>f x\<close> makes the partiality of evaluating
-  \<open>f x\<close> explicit via the option datatype.  Interrupts are
-  \<^emph>\<open>not\<close> handled here, i.e.\ this form serves as safe replacement
-  for the \<^emph>\<open>unsafe\<close> version @{ML_text "(SOME"}~\<open>f
-  x\<close>~@{ML_text "handle _ => NONE)"} that is occasionally seen in
-  books about SML97, but not in Isabelle/ML.
+  \<^descr> @{ML try}~\<open>f x\<close> makes the partiality of evaluating \<open>f x\<close> explicit via the
+  option datatype. Interrupts are \<^emph>\<open>not\<close> handled here, i.e.\ this form serves
+  as safe replacement for the \<^emph>\<open>unsafe\<close> version @{ML_text "(SOME"}~\<open>f
+  x\<close>~@{ML_text "handle _ => NONE)"} that is occasionally seen in books about
+  SML97, but not in Isabelle/ML.
 
   \<^descr> @{ML can} is similar to @{ML try} with more abstract result.
 
-  \<^descr> @{ML ERROR}~\<open>msg\<close> represents user errors; this
-  exception is normally raised indirectly via the @{ML error} function
-  (see \secref{sec:message-channels}).
+  \<^descr> @{ML ERROR}~\<open>msg\<close> represents user errors; this exception is normally
+  raised indirectly via the @{ML error} function (see
+  \secref{sec:message-channels}).
 
   \<^descr> @{ML Fail}~\<open>msg\<close> represents general program failures.
 
-  \<^descr> @{ML Exn.is_interrupt} identifies interrupts robustly, without
-  mentioning concrete exception constructors in user code.  Handled
-  interrupts need to be re-raised promptly!
+  \<^descr> @{ML Exn.is_interrupt} identifies interrupts robustly, without mentioning
+  concrete exception constructors in user code. Handled interrupts need to be
+  re-raised promptly!
+
+  \<^descr> @{ML reraise}~\<open>exn\<close> raises exception \<open>exn\<close> while preserving its implicit
+  position information (if possible, depending on the ML platform).
 
-  \<^descr> @{ML reraise}~\<open>exn\<close> raises exception \<open>exn\<close>
-  while preserving its implicit position information (if possible,
-  depending on the ML platform).
+  \<^descr> @{ML Runtime.exn_trace}~@{ML_text "(fn () =>"}~\<open>e\<close>@{ML_text ")"} evaluates
+  expression \<open>e\<close> while printing a full trace of its stack of nested exceptions
+  (if possible, depending on the ML platform).
 
-  \<^descr> @{ML Runtime.exn_trace}~@{ML_text "(fn () =>"}~\<open>e\<close>@{ML_text ")"} evaluates expression \<open>e\<close> while printing
-  a full trace of its stack of nested exceptions (if possible,
-  depending on the ML platform).
-
-  Inserting @{ML Runtime.exn_trace} into ML code temporarily is
-  useful for debugging, but not suitable for production code.
+  Inserting @{ML Runtime.exn_trace} into ML code temporarily is useful for
+  debugging, but not suitable for production code.
 \<close>
 
 text %mlantiq \<open>
@@ -1211,19 +1192,19 @@
   @{ML_antiquotation_def "assert"} & : & \<open>ML_antiquotation\<close> \\
   \end{matharray}
 
-  \<^descr> \<open>@{assert}\<close> inlines a function
-  @{ML_type "bool -> unit"} that raises @{ML Fail} if the argument is
-  @{ML false}.  Due to inlining the source position of failed
-  assertions is included in the error output.
+  \<^descr> \<open>@{assert}\<close> inlines a function @{ML_type "bool -> unit"} that raises @{ML
+  Fail} if the argument is @{ML false}. Due to inlining the source position of
+  failed assertions is included in the error output.
 \<close>
 
 
 section \<open>Strings of symbols \label{sec:symbols}\<close>
 
-text \<open>A \<^emph>\<open>symbol\<close> constitutes the smallest textual unit in Isabelle/ML --- raw
-  ML characters are normally not encountered at all. Isabelle strings consist
-  of a sequence of symbols, represented as a packed string or an exploded list
-  of strings. Each symbol is in itself a small string, which has either one of
+text \<open>
+  A \<^emph>\<open>symbol\<close> constitutes the smallest textual unit in Isabelle/ML --- raw ML
+  characters are normally not encountered at all. Isabelle strings consist of
+  a sequence of symbols, represented as a packed string or an exploded list of
+  strings. Each symbol is in itself a small string, which has either one of
   the following forms:
 
     \<^enum> a single ASCII character ``\<open>c\<close>'', for example ``\<^verbatim>\<open>a\<close>'',
@@ -1276,54 +1257,54 @@
   @{index_ML Symbol.decode: "Symbol.symbol -> Symbol.sym"} \\
   \end{mldecls}
 
-  \<^descr> Type @{ML_type "Symbol.symbol"} represents individual Isabelle
-  symbols.
+  \<^descr> Type @{ML_type "Symbol.symbol"} represents individual Isabelle symbols.
 
-  \<^descr> @{ML "Symbol.explode"}~\<open>str\<close> produces a symbol list
-  from the packed form.  This function supersedes @{ML
-  "String.explode"} for virtually all purposes of manipulating text in
-  Isabelle!\<^footnote>\<open>The runtime overhead for exploded strings is
-  mainly that of the list structure: individual symbols that happen to
-  be a singleton string do not require extra memory in Poly/ML.\<close>
+  \<^descr> @{ML "Symbol.explode"}~\<open>str\<close> produces a symbol list from the packed form.
+  This function supersedes @{ML "String.explode"} for virtually all purposes
+  of manipulating text in Isabelle!\<^footnote>\<open>The runtime overhead for exploded strings
+  is mainly that of the list structure: individual symbols that happen to be a
+  singleton string do not require extra memory in Poly/ML.\<close>
 
   \<^descr> @{ML "Symbol.is_letter"}, @{ML "Symbol.is_digit"}, @{ML
-  "Symbol.is_quasi"}, @{ML "Symbol.is_blank"} classify standard
-  symbols according to fixed syntactic conventions of Isabelle, cf.\
-  @{cite "isabelle-isar-ref"}.
+  "Symbol.is_quasi"}, @{ML "Symbol.is_blank"} classify standard symbols
+  according to fixed syntactic conventions of Isabelle, cf.\ @{cite
+  "isabelle-isar-ref"}.
 
-  \<^descr> Type @{ML_type "Symbol.sym"} is a concrete datatype that
-  represents the different kinds of symbols explicitly, with
-  constructors @{ML "Symbol.Char"}, @{ML "Symbol.UTF8"},
-  @{ML "Symbol.Sym"}, @{ML "Symbol.Control"}, @{ML "Symbol.Raw"},
-  @{ML "Symbol.Malformed"}.
+  \<^descr> Type @{ML_type "Symbol.sym"} is a concrete datatype that represents the
+  different kinds of symbols explicitly, with constructors @{ML
+  "Symbol.Char"}, @{ML "Symbol.UTF8"}, @{ML "Symbol.Sym"}, @{ML
+  "Symbol.Control"}, @{ML "Symbol.Raw"}, @{ML "Symbol.Malformed"}.
 
-  \<^descr> @{ML "Symbol.decode"} converts the string representation of a
-  symbol into the datatype version.
+  \<^descr> @{ML "Symbol.decode"} converts the string representation of a symbol into
+  the datatype version.
 \<close>
 
 paragraph \<open>Historical note.\<close>
-text \<open>In the original SML90 standard the primitive ML type @{ML_type char} did
-  not exists, and @{ML_text "explode: string -> string list"} produced a list
-  of singleton strings like @{ML "raw_explode: string -> string list"} in
+text \<open>
+  In the original SML90 standard the primitive ML type @{ML_type char} did not
+  exists, and @{ML_text "explode: string -> string list"} produced a list of
+  singleton strings like @{ML "raw_explode: string -> string list"} in
   Isabelle/ML today. When SML97 came out, Isabelle did not adopt its somewhat
   anachronistic 8-bit or 16-bit characters, but the idea of exploding a string
   into a list of small strings was extended to ``symbols'' as explained above.
   Thus Isabelle sources can refer to an infinite store of user-defined
   symbols, without having to worry about the multitude of Unicode encodings
-  that have emerged over the years.\<close>
+  that have emerged over the years.
+\<close>
 
 
 section \<open>Basic data types\<close>
 
-text \<open>The basis library proposal of SML97 needs to be treated with
-  caution.  Many of its operations simply do not fit with important
-  Isabelle/ML conventions (like ``canonical argument order'', see
-  \secref{sec:canonical-argument-order}), others cause problems with
-  the parallel evaluation model of Isabelle/ML (such as @{ML
-  TextIO.print} or @{ML OS.Process.system}).
+text \<open>
+  The basis library proposal of SML97 needs to be treated with caution. Many
+  of its operations simply do not fit with important Isabelle/ML conventions
+  (like ``canonical argument order'', see
+  \secref{sec:canonical-argument-order}), others cause problems with the
+  parallel evaluation model of Isabelle/ML (such as @{ML TextIO.print} or @{ML
+  OS.Process.system}).
 
-  Subsequently we give a brief overview of important operations on
-  basic ML data types.
+  Subsequently we give a brief overview of important operations on basic ML
+  data types.
 \<close>
 
 
@@ -1334,9 +1315,8 @@
   @{index_ML_type char} \\
   \end{mldecls}
 
-  \<^descr> Type @{ML_type char} is \<^emph>\<open>not\<close> used.  The smallest textual
-  unit in Isabelle is represented as a ``symbol'' (see
-  \secref{sec:symbols}).
+  \<^descr> Type @{ML_type char} is \<^emph>\<open>not\<close> used. The smallest textual unit in Isabelle
+  is represented as a ``symbol'' (see \secref{sec:symbols}).
 \<close>
 
 
@@ -1347,29 +1327,29 @@
   @{index_ML_type string} \\
   \end{mldecls}
 
-  \<^descr> Type @{ML_type string} represents immutable vectors of 8-bit
-  characters.  There are operations in SML to convert back and forth
-  to actual byte vectors, which are seldom used.
+  \<^descr> Type @{ML_type string} represents immutable vectors of 8-bit characters.
+  There are operations in SML to convert back and forth to actual byte
+  vectors, which are seldom used.
 
   This historically important raw text representation is used for
-  Isabelle-specific purposes with the following implicit substructures
-  packed into the string content:
+  Isabelle-specific purposes with the following implicit substructures packed
+  into the string content:
 
-    \<^enum> sequence of Isabelle symbols (see also \secref{sec:symbols}),
-    with @{ML Symbol.explode} as key operation;
+    \<^enum> sequence of Isabelle symbols (see also \secref{sec:symbols}), with @{ML
+    Symbol.explode} as key operation;
   
-    \<^enum> XML tree structure via YXML (see also @{cite "isabelle-system"}),
-    with @{ML YXML.parse_body} as key operation.
+    \<^enum> XML tree structure via YXML (see also @{cite "isabelle-system"}), with
+    @{ML YXML.parse_body} as key operation.
 
   Note that Isabelle/ML string literals may refer Isabelle symbols like
-  ``\<^verbatim>\<open>\<alpha>\<close>'' natively, \<^emph>\<open>without\<close> escaping the backslash. This is a
-  consequence of Isabelle treating all source text as strings of symbols,
-  instead of raw characters.
+  ``\<^verbatim>\<open>\<alpha>\<close>'' natively, \<^emph>\<open>without\<close> escaping the backslash. This is a consequence
+  of Isabelle treating all source text as strings of symbols, instead of raw
+  characters.
 \<close>
 
-text %mlex \<open>The subsequent example illustrates the difference of
-  physical addressing of bytes versus logical addressing of symbols in
-  Isabelle strings.
+text %mlex \<open>
+  The subsequent example illustrates the difference of physical addressing of
+  bytes versus logical addressing of symbols in Isabelle strings.
 \<close>
 
 ML_val \<open>
@@ -1379,12 +1359,14 @@
   @{assert} (size s = 4);
 \<close>
 
-text \<open>Note that in Unicode renderings of the symbol \<open>\<A>\<close>,
-  variations of encodings like UTF-8 or UTF-16 pose delicate questions
-  about the multi-byte representations of its codepoint, which is outside
-  of the 16-bit address space of the original Unicode standard from
-  the 1990-ies.  In Isabelle/ML it is just ``\<^verbatim>\<open>\<A>\<close>''
-  literally, using plain ASCII characters beyond any doubts.\<close>
+text \<open>
+  Note that in Unicode renderings of the symbol \<open>\<A>\<close>, variations of encodings
+  like UTF-8 or UTF-16 pose delicate questions about the multi-byte
+  representations of its codepoint, which is outside of the 16-bit address
+  space of the original Unicode standard from the 1990-ies. In Isabelle/ML it
+  is just ``\<^verbatim>\<open>\<A>\<close>'' literally, using plain ASCII characters beyond any
+  doubts.
+\<close>
 
 
 subsection \<open>Integers\<close>
@@ -1394,19 +1376,18 @@
   @{index_ML_type int} \\
   \end{mldecls}
 
-  \<^descr> Type @{ML_type int} represents regular mathematical integers, which
-  are \<^emph>\<open>unbounded\<close>. Overflow is treated properly, but should never happen
-  in practice.\<^footnote>\<open>The size limit for integer bit patterns in memory is
-  64\,MB for 32-bit Poly/ML, and much higher for 64-bit systems.\<close> This works
-  uniformly for all supported ML platforms (Poly/ML and SML/NJ).
+  \<^descr> Type @{ML_type int} represents regular mathematical integers, which are
+  \<^emph>\<open>unbounded\<close>. Overflow is treated properly, but should never happen in
+  practice.\<^footnote>\<open>The size limit for integer bit patterns in memory is 64\,MB for
+  32-bit Poly/ML, and much higher for 64-bit systems.\<close> This works uniformly
+  for all supported ML platforms (Poly/ML and SML/NJ).
 
-  Literal integers in ML text are forced to be of this one true
-  integer type --- adhoc overloading of SML97 is disabled.
+  Literal integers in ML text are forced to be of this one true integer type
+  --- adhoc overloading of SML97 is disabled.
 
   Structure @{ML_structure IntInf} of SML97 is obsolete and superseded by
-  @{ML_structure Int}.  Structure @{ML_structure Integer} in @{file
-  "~~/src/Pure/General/integer.ML"} provides some additional
-  operations.
+  @{ML_structure Int}. Structure @{ML_structure Integer} in @{file
+  "~~/src/Pure/General/integer.ML"} provides some additional operations.
 \<close>
 
 
@@ -1418,14 +1399,14 @@
   @{index_ML seconds: "real -> Time.time"} \\
   \end{mldecls}
 
-  \<^descr> Type @{ML_type Time.time} represents time abstractly according
-  to the SML97 basis library definition.  This is adequate for
-  internal ML operations, but awkward in concrete time specifications.
+  \<^descr> Type @{ML_type Time.time} represents time abstractly according to the
+  SML97 basis library definition. This is adequate for internal ML operations,
+  but awkward in concrete time specifications.
 
-  \<^descr> @{ML seconds}~\<open>s\<close> turns the concrete scalar \<open>s\<close> (measured in seconds) into an abstract time value.  Floating
-  point numbers are easy to use as configuration options in the
-  context (see \secref{sec:config-options}) or system options that
-  are maintained externally.
+  \<^descr> @{ML seconds}~\<open>s\<close> turns the concrete scalar \<open>s\<close> (measured in seconds) into
+  an abstract time value. Floating point numbers are easy to use as
+  configuration options in the context (see \secref{sec:config-options}) or
+  system options that are maintained externally.
 \<close>
 
 
@@ -1443,18 +1424,21 @@
   \end{mldecls}
 \<close>
 
-text \<open>Apart from @{ML Option.map} most other operations defined in
-  structure @{ML_structure Option} are alien to Isabelle/ML and never
-  used.  The operations shown above are defined in @{file
-  "~~/src/Pure/General/basics.ML"}.\<close>
+text \<open>
+  Apart from @{ML Option.map} most other operations defined in structure
+  @{ML_structure Option} are alien to Isabelle/ML and never used. The
+  operations shown above are defined in @{file
+  "~~/src/Pure/General/basics.ML"}.
+\<close>
 
 
 subsection \<open>Lists\<close>
 
-text \<open>Lists are ubiquitous in ML as simple and light-weight
-  ``collections'' for many everyday programming tasks.  Isabelle/ML
-  provides important additions and improvements over operations that
-  are predefined in the SML97 library.\<close>
+text \<open>
+  Lists are ubiquitous in ML as simple and light-weight ``collections'' for
+  many everyday programming tasks. Isabelle/ML provides important additions
+  and improvements over operations that are predefined in the SML97 library.
+\<close>
 
 text %mlref \<open>
   \begin{mldecls}
@@ -1467,30 +1451,28 @@
 
   \<^descr> @{ML cons}~\<open>x xs\<close> evaluates to \<open>x :: xs\<close>.
 
-  Tupled infix operators are a historical accident in Standard ML.
-  The curried @{ML cons} amends this, but it should be only used when
-  partial application is required.
+  Tupled infix operators are a historical accident in Standard ML. The curried
+  @{ML cons} amends this, but it should be only used when partial application
+  is required.
 
-  \<^descr> @{ML member}, @{ML insert}, @{ML remove}, @{ML update} treat
-  lists as a set-like container that maintains the order of elements.
-  See @{file "~~/src/Pure/library.ML"} for the full specifications
-  (written in ML).  There are some further derived operations like
-  @{ML union} or @{ML inter}.
+  \<^descr> @{ML member}, @{ML insert}, @{ML remove}, @{ML update} treat lists as a
+  set-like container that maintains the order of elements. See @{file
+  "~~/src/Pure/library.ML"} for the full specifications (written in ML). There
+  are some further derived operations like @{ML union} or @{ML inter}.
 
-  Note that @{ML insert} is conservative about elements that are
-  already a @{ML member} of the list, while @{ML update} ensures that
-  the latest entry is always put in front.  The latter discipline is
-  often more appropriate in declarations of context data
-  (\secref{sec:context-data}) that are issued by the user in Isar
-  source: later declarations take precedence over earlier ones.
-\<close>
+  Note that @{ML insert} is conservative about elements that are already a
+  @{ML member} of the list, while @{ML update} ensures that the latest entry
+  is always put in front. The latter discipline is often more appropriate in
+  declarations of context data (\secref{sec:context-data}) that are issued by
+  the user in Isar source: later declarations take precedence over earlier
+  ones. \<close>
 
-text %mlex \<open>Using canonical @{ML fold} together with @{ML cons} (or
-  similar standard operations) alternates the orientation of data.
-  The is quite natural and should not be altered forcible by inserting
-  extra applications of @{ML rev}.  The alternative @{ML fold_rev} can
-  be used in the few situations, where alternation should be
-  prevented.
+text %mlex \<open>
+  Using canonical @{ML fold} together with @{ML cons} (or similar standard
+  operations) alternates the orientation of data. The is quite natural and
+  should not be altered forcible by inserting extra applications of @{ML rev}.
+  The alternative @{ML fold_rev} can be used in the few situations, where
+  alternation should be prevented.
 \<close>
 
 ML_val \<open>
@@ -1503,31 +1485,33 @@
   @{assert} (list2 = items);
 \<close>
 
-text \<open>The subsequent example demonstrates how to \<^emph>\<open>merge\<close> two
-  lists in a natural way.\<close>
+text \<open>
+  The subsequent example demonstrates how to \<^emph>\<open>merge\<close> two lists in a natural
+  way.
+\<close>
 
 ML_val \<open>
   fun merge_lists eq (xs, ys) = fold_rev (insert eq) ys xs;
 \<close>
 
-text \<open>Here the first list is treated conservatively: only the new
-  elements from the second list are inserted.  The inside-out order of
-  insertion via @{ML fold_rev} attempts to preserve the order of
-  elements in the result.
+text \<open>
+  Here the first list is treated conservatively: only the new elements from
+  the second list are inserted. The inside-out order of insertion via @{ML
+  fold_rev} attempts to preserve the order of elements in the result.
 
   This way of merging lists is typical for context data
-  (\secref{sec:context-data}).  See also @{ML merge} as defined in
-  @{file "~~/src/Pure/library.ML"}.
+  (\secref{sec:context-data}). See also @{ML merge} as defined in @{file
+  "~~/src/Pure/library.ML"}.
 \<close>
 
 
 subsection \<open>Association lists\<close>
 
-text \<open>The operations for association lists interpret a concrete list
-  of pairs as a finite function from keys to values.  Redundant
-  representations with multiple occurrences of the same key are
-  implicitly normalized: lookup and update only take the first
-  occurrence into account.
+text \<open>
+  The operations for association lists interpret a concrete list of pairs as a
+  finite function from keys to values. Redundant representations with multiple
+  occurrences of the same key are implicitly normalized: lookup and update
+  only take the first occurrence into account.
 \<close>
 
 text \<open>
@@ -1537,20 +1521,18 @@
   @{index_ML AList.update: "('a * 'a -> bool) -> 'a * 'b -> ('a * 'b) list -> ('a * 'b) list"} \\
   \end{mldecls}
 
-  \<^descr> @{ML AList.lookup}, @{ML AList.defined}, @{ML AList.update}
-  implement the main ``framework operations'' for mappings in
-  Isabelle/ML, following standard conventions for their names and
-  types.
+  \<^descr> @{ML AList.lookup}, @{ML AList.defined}, @{ML AList.update} implement the
+  main ``framework operations'' for mappings in Isabelle/ML, following
+  standard conventions for their names and types.
 
-  Note that a function called \<^verbatim>\<open>lookup\<close> is obliged to express its
-  partiality via an explicit option element.  There is no choice to
-  raise an exception, without changing the name to something like
-  \<open>the_element\<close> or \<open>get\<close>.
+  Note that a function called \<^verbatim>\<open>lookup\<close> is obliged to express its partiality
+  via an explicit option element. There is no choice to raise an exception,
+  without changing the name to something like \<open>the_element\<close> or \<open>get\<close>.
 
-  The \<open>defined\<close> operation is essentially a contraction of @{ML
-  is_some} and \<^verbatim>\<open>lookup\<close>, but this is sufficiently frequent to
-  justify its independent existence.  This also gives the
-  implementation some opportunity for peep-hole optimization.
+  The \<open>defined\<close> operation is essentially a contraction of @{ML is_some} and
+  \<^verbatim>\<open>lookup\<close>, but this is sufficiently frequent to justify its independent
+  existence. This also gives the implementation some opportunity for peep-hole
+  optimization.
 
 
   Association lists are adequate as simple implementation of finite mappings
@@ -1571,18 +1553,18 @@
   \end{mldecls}
 \<close>
 
-text \<open>Due to ubiquitous parallelism in Isabelle/ML (see also
-  \secref{sec:multi-threading}), the mutable reference cells of
-  Standard ML are notorious for causing problems.  In a highly
-  parallel system, both correctness \<^emph>\<open>and\<close> performance are easily
-  degraded when using mutable data.
+text \<open>
+  Due to ubiquitous parallelism in Isabelle/ML (see also
+  \secref{sec:multi-threading}), the mutable reference cells of Standard ML
+  are notorious for causing problems. In a highly parallel system, both
+  correctness \<^emph>\<open>and\<close> performance are easily degraded when using mutable data.
 
-  The unwieldy name of @{ML Unsynchronized.ref} for the constructor
-  for references in Isabelle/ML emphasizes the inconveniences caused by
-  mutability.  Existing operations @{ML "!"}  and @{ML_op ":="} are
-  unchanged, but should be used with special precautions, say in a
-  strictly local situation that is guaranteed to be restricted to
-  sequential evaluation --- now and in the future.
+  The unwieldy name of @{ML Unsynchronized.ref} for the constructor for
+  references in Isabelle/ML emphasizes the inconveniences caused by
+  mutability. Existing operations @{ML "!"} and @{ML_op ":="} are unchanged,
+  but should be used with special precautions, say in a strictly local
+  situation that is guaranteed to be restricted to sequential evaluation ---
+  now and in the future.
 
   \begin{warn}
   Never @{ML_text "open Unsynchronized"}, not even in a local scope!
@@ -1593,62 +1575,58 @@
 
 section \<open>Thread-safe programming \label{sec:multi-threading}\<close>
 
-text \<open>Multi-threaded execution has become an everyday reality in
-  Isabelle since Poly/ML 5.2.1 and Isabelle2008.  Isabelle/ML provides
-  implicit and explicit parallelism by default, and there is no way
-  for user-space tools to ``opt out''.  ML programs that are purely
-  functional, output messages only via the official channels
-  (\secref{sec:message-channels}), and do not intercept interrupts
-  (\secref{sec:exceptions}) can participate in the multi-threaded
+text \<open>
+  Multi-threaded execution has become an everyday reality in Isabelle since
+  Poly/ML 5.2.1 and Isabelle2008. Isabelle/ML provides implicit and explicit
+  parallelism by default, and there is no way for user-space tools to ``opt
+  out''. ML programs that are purely functional, output messages only via the
+  official channels (\secref{sec:message-channels}), and do not intercept
+  interrupts (\secref{sec:exceptions}) can participate in the multi-threaded
   environment immediately without further ado.
 
-  More ambitious tools with more fine-grained interaction with the
-  environment need to observe the principles explained below.
+  More ambitious tools with more fine-grained interaction with the environment
+  need to observe the principles explained below.
 \<close>
 
 
 subsection \<open>Multi-threading with shared memory\<close>
 
-text \<open>Multiple threads help to organize advanced operations of the
-  system, such as real-time conditions on command transactions,
-  sub-components with explicit communication, general asynchronous
-  interaction etc.  Moreover, parallel evaluation is a prerequisite to
-  make adequate use of the CPU resources that are available on
-  multi-core systems.\<^footnote>\<open>Multi-core computing does not mean that
-  there are ``spare cycles'' to be wasted.  It means that the
-  continued exponential speedup of CPU performance due to ``Moore's
-  Law'' follows different rules: clock frequency has reached its peak
-  around 2005, and applications need to be parallelized in order to
-  avoid a perceived loss of performance.  See also
-  @{cite "Sutter:2005"}.\<close>
+text \<open>
+  Multiple threads help to organize advanced operations of the system, such as
+  real-time conditions on command transactions, sub-components with explicit
+  communication, general asynchronous interaction etc. Moreover, parallel
+  evaluation is a prerequisite to make adequate use of the CPU resources that
+  are available on multi-core systems.\<^footnote>\<open>Multi-core computing does not mean
+  that there are ``spare cycles'' to be wasted. It means that the continued
+  exponential speedup of CPU performance due to ``Moore's Law'' follows
+  different rules: clock frequency has reached its peak around 2005, and
+  applications need to be parallelized in order to avoid a perceived loss of
+  performance. See also @{cite "Sutter:2005"}.\<close>
 
   Isabelle/Isar exploits the inherent structure of theories and proofs to
-  support \<^emph>\<open>implicit parallelism\<close> to a large extent. LCF-style theorem
-  proving provides almost ideal conditions for that, see also
-  @{cite "Wenzel:2009"}. This means, significant parts of theory and proof
-  checking is parallelized by default. In Isabelle2013, a maximum
-  speedup-factor of 3.5 on 4 cores and 6.5 on 8 cores can be expected
-  @{cite "Wenzel:2013:ITP"}.
+  support \<^emph>\<open>implicit parallelism\<close> to a large extent. LCF-style theorem proving
+  provides almost ideal conditions for that, see also @{cite "Wenzel:2009"}.
+  This means, significant parts of theory and proof checking is parallelized
+  by default. In Isabelle2013, a maximum speedup-factor of 3.5 on 4 cores and
+  6.5 on 8 cores can be expected @{cite "Wenzel:2013:ITP"}.
 
   \<^medskip>
-  ML threads lack the memory protection of separate
-  processes, and operate concurrently on shared heap memory.  This has
-  the advantage that results of independent computations are directly
-  available to other threads: abstract values can be passed without
-  copying or awkward serialization that is typically required for
-  separate processes.
+  ML threads lack the memory protection of separate processes, and operate
+  concurrently on shared heap memory. This has the advantage that results of
+  independent computations are directly available to other threads: abstract
+  values can be passed without copying or awkward serialization that is
+  typically required for separate processes.
 
-  To make shared-memory multi-threading work robustly and efficiently,
-  some programming guidelines need to be observed.  While the ML
-  system is responsible to maintain basic integrity of the
-  representation of ML values in memory, the application programmer
-  needs to ensure that multi-threaded execution does not break the
-  intended semantics.
+  To make shared-memory multi-threading work robustly and efficiently, some
+  programming guidelines need to be observed. While the ML system is
+  responsible to maintain basic integrity of the representation of ML values
+  in memory, the application programmer needs to ensure that multi-threaded
+  execution does not break the intended semantics.
 
   \begin{warn}
-  To participate in implicit parallelism, tools need to be
-  thread-safe.  A single ill-behaved tool can affect the stability and
-  performance of the whole system.
+  To participate in implicit parallelism, tools need to be thread-safe. A
+  single ill-behaved tool can affect the stability and performance of the
+  whole system.
   \end{warn}
 
   Apart from observing the principles of thread-safeness passively, advanced
@@ -1656,76 +1634,69 @@
   functions for parallel list operations (\secref{sec:parlist}).
 
   \begin{warn}
-  Parallel computing resources are managed centrally by the
-  Isabelle/ML infrastructure.  User programs should not fork their own
-  ML threads to perform heavy computations.
+  Parallel computing resources are managed centrally by the Isabelle/ML
+  infrastructure. User programs should not fork their own ML threads to
+  perform heavy computations.
   \end{warn}
 \<close>
 
 
 subsection \<open>Critical shared resources\<close>
 
-text \<open>Thread-safeness is mainly concerned about concurrent
-  read/write access to shared resources, which are outside the purely
-  functional world of ML.  This covers the following in particular.
-
-  \<^item> Global references (or arrays), i.e.\ mutable memory cells that
-  persist over several invocations of associated
-  operations.\<^footnote>\<open>This is independent of the visibility of such
-  mutable values in the toplevel scope.\<close>
+text \<open>
+  Thread-safeness is mainly concerned about concurrent read/write access to
+  shared resources, which are outside the purely functional world of ML. This
+  covers the following in particular.
 
-  \<^item> Global state of the running Isabelle/ML process, i.e.\ raw I/O
-  channels, environment variables, current working directory.
-
-  \<^item> Writable resources in the file-system that are shared among
-  different threads or external processes.
+    \<^item> Global references (or arrays), i.e.\ mutable memory cells that persist
+    over several invocations of associated operations.\<^footnote>\<open>This is independent of
+    the visibility of such mutable values in the toplevel scope.\<close>
 
+    \<^item> Global state of the running Isabelle/ML process, i.e.\ raw I/O channels,
+    environment variables, current working directory.
 
-  Isabelle/ML provides various mechanisms to avoid critical shared
-  resources in most situations.  As last resort there are some
-  mechanisms for explicit synchronization.  The following guidelines
-  help to make Isabelle/ML programs work smoothly in a concurrent
-  environment.
+    \<^item> Writable resources in the file-system that are shared among different
+    threads or external processes.
 
-  \<^item> Avoid global references altogether.  Isabelle/Isar maintains a
-  uniform context that incorporates arbitrary data declared by user
-  programs (\secref{sec:context-data}).  This context is passed as
-  plain value and user tools can get/map their own data in a purely
-  functional manner.  Configuration options within the context
-  (\secref{sec:config-options}) provide simple drop-in replacements
-  for historic reference variables.
+  Isabelle/ML provides various mechanisms to avoid critical shared resources
+  in most situations. As last resort there are some mechanisms for explicit
+  synchronization. The following guidelines help to make Isabelle/ML programs
+  work smoothly in a concurrent environment.
+
+  \<^item> Avoid global references altogether. Isabelle/Isar maintains a uniform
+  context that incorporates arbitrary data declared by user programs
+  (\secref{sec:context-data}). This context is passed as plain value and user
+  tools can get/map their own data in a purely functional manner.
+  Configuration options within the context (\secref{sec:config-options})
+  provide simple drop-in replacements for historic reference variables.
 
-  \<^item> Keep components with local state information re-entrant.
-  Instead of poking initial values into (private) global references, a
-  new state record can be created on each invocation, and passed
-  through any auxiliary functions of the component.  The state record
-  contain mutable references in special situations, without requiring any
-  synchronization, as long as each invocation gets its own copy and the
-  tool itself is single-threaded.
+  \<^item> Keep components with local state information re-entrant. Instead of poking
+  initial values into (private) global references, a new state record can be
+  created on each invocation, and passed through any auxiliary functions of
+  the component. The state record contain mutable references in special
+  situations, without requiring any synchronization, as long as each
+  invocation gets its own copy and the tool itself is single-threaded.
 
-  \<^item> Avoid raw output on \<open>stdout\<close> or \<open>stderr\<close>.  The
-  Poly/ML library is thread-safe for each individual output operation,
-  but the ordering of parallel invocations is arbitrary.  This means
-  raw output will appear on some system console with unpredictable
-  interleaving of atomic chunks.
+  \<^item> Avoid raw output on \<open>stdout\<close> or \<open>stderr\<close>. The Poly/ML library is
+  thread-safe for each individual output operation, but the ordering of
+  parallel invocations is arbitrary. This means raw output will appear on some
+  system console with unpredictable interleaving of atomic chunks.
 
   Note that this does not affect regular message output channels
-  (\secref{sec:message-channels}).  An official message id is associated
-  with the command transaction from where it originates, independently
-  of other transactions.  This means each running Isar command has
-  effectively its own set of message channels, and interleaving can
-  only happen when commands use parallelism internally (and only at
-  message boundaries).
+  (\secref{sec:message-channels}). An official message id is associated with
+  the command transaction from where it originates, independently of other
+  transactions. This means each running Isar command has effectively its own
+  set of message channels, and interleaving can only happen when commands use
+  parallelism internally (and only at message boundaries).
 
-  \<^item> Treat environment variables and the current working directory
-  of the running process as read-only.
+  \<^item> Treat environment variables and the current working directory of the
+  running process as read-only.
 
-  \<^item> Restrict writing to the file-system to unique temporary files.
-  Isabelle already provides a temporary directory that is unique for
-  the running process, and there is a centralized source of unique
-  serial numbers in Isabelle/ML.  Thus temporary files that are passed
-  to to some external process will be always disjoint, and thus
-  thread-safe.
+  \<^item> Restrict writing to the file-system to unique temporary files. Isabelle
+  already provides a temporary directory that is unique for the running
+  process, and there is a centralized source of unique serial numbers in
+  Isabelle/ML. Thus temporary files that are passed to to some external
+  process will be always disjoint, and thus thread-safe.
 \<close>
 
 text %mlref \<open>
@@ -1734,16 +1705,15 @@
   @{index_ML serial_string: "unit -> string"} \\
   \end{mldecls}
 
-  \<^descr> @{ML File.tmp_path}~\<open>path\<close> relocates the base
-  component of \<open>path\<close> into the unique temporary directory of
-  the running Isabelle/ML process.
+  \<^descr> @{ML File.tmp_path}~\<open>path\<close> relocates the base component of \<open>path\<close> into the
+  unique temporary directory of the running Isabelle/ML process.
 
-  \<^descr> @{ML serial_string}~\<open>()\<close> creates a new serial number
-  that is unique over the runtime of the Isabelle/ML process.
+  \<^descr> @{ML serial_string}~\<open>()\<close> creates a new serial number that is unique over
+  the runtime of the Isabelle/ML process.
 \<close>
 
-text %mlex \<open>The following example shows how to create unique
-  temporary file names.
+text %mlex \<open>
+  The following example shows how to create unique temporary file names.
 \<close>
 
 ML_val \<open>
@@ -1755,7 +1725,8 @@
 
 subsection \<open>Explicit synchronization\<close>
 
-text \<open>Isabelle/ML provides explicit synchronization for mutable variables over
+text \<open>
+  Isabelle/ML provides explicit synchronization for mutable variables over
   immutable data, which may be updated atomically and exclusively. This
   addresses the rare situations where mutable shared resources are really
   required. Synchronization in Isabelle/ML is based on primitives of Poly/ML,
@@ -1763,11 +1734,12 @@
   Isabelle environment. User code should not break this abstraction, but stay
   within the confines of concurrent Isabelle/ML.
 
-  A \<^emph>\<open>synchronized variable\<close> is an explicit state component associated
-  with mechanisms for locking and signaling. There are operations to await a
+  A \<^emph>\<open>synchronized variable\<close> is an explicit state component associated with
+  mechanisms for locking and signaling. There are operations to await a
   condition, change the state, and signal the change to all other waiting
   threads. Synchronized access to the state variable is \<^emph>\<open>not\<close> re-entrant:
-  direct or indirect nesting within the same thread will cause a deadlock!\<close>
+  direct or indirect nesting within the same thread will cause a deadlock!
+\<close>
 
 text %mlref \<open>
   \begin{mldecls}
@@ -1777,31 +1749,29 @@
   ('a -> ('b * 'a) option) -> 'b"} \\
   \end{mldecls}
 
-  \<^descr> Type @{ML_type "'a Synchronized.var"} represents synchronized
-  variables with state of type @{ML_type 'a}.
+    \<^descr> Type @{ML_type "'a Synchronized.var"} represents synchronized variables
+    with state of type @{ML_type 'a}.
 
-  \<^descr> @{ML Synchronized.var}~\<open>name x\<close> creates a synchronized
-  variable that is initialized with value \<open>x\<close>.  The \<open>name\<close> is used for tracing.
+    \<^descr> @{ML Synchronized.var}~\<open>name x\<close> creates a synchronized variable that is
+    initialized with value \<open>x\<close>. The \<open>name\<close> is used for tracing.
 
-  \<^descr> @{ML Synchronized.guarded_access}~\<open>var f\<close> lets the
-  function \<open>f\<close> operate within a critical section on the state
-  \<open>x\<close> as follows: if \<open>f x\<close> produces @{ML NONE}, it
-  continues to wait on the internal condition variable, expecting that
-  some other thread will eventually change the content in a suitable
-  manner; if \<open>f x\<close> produces @{ML SOME}~\<open>(y, x')\<close> it is
-  satisfied and assigns the new state value \<open>x'\<close>, broadcasts a
-  signal to all waiting threads on the associated condition variable,
-  and returns the result \<open>y\<close>.
+    \<^descr> @{ML Synchronized.guarded_access}~\<open>var f\<close> lets the function \<open>f\<close> operate
+    within a critical section on the state \<open>x\<close> as follows: if \<open>f x\<close> produces
+    @{ML NONE}, it continues to wait on the internal condition variable,
+    expecting that some other thread will eventually change the content in a
+    suitable manner; if \<open>f x\<close> produces @{ML SOME}~\<open>(y, x')\<close> it is satisfied and
+    assigns the new state value \<open>x'\<close>, broadcasts a signal to all waiting threads
+    on the associated condition variable, and returns the result \<open>y\<close>.
 
-
-  There are some further variants of the @{ML
-  Synchronized.guarded_access} combinator, see @{file
-  "~~/src/Pure/Concurrent/synchronized.ML"} for details.
+  There are some further variants of the @{ML Synchronized.guarded_access}
+  combinator, see @{file "~~/src/Pure/Concurrent/synchronized.ML"} for
+  details.
 \<close>
 
-text %mlex \<open>The following example implements a counter that produces
-  positive integers that are unique over the runtime of the Isabelle
-  process:\<close>
+text %mlex \<open>
+  The following example implements a counter that produces positive integers
+  that are unique over the runtime of the Isabelle process:
+\<close>
 
 ML_val \<open>
   local
@@ -1821,60 +1791,56 @@
 
 text \<open>
   \<^medskip>
-  See @{file "~~/src/Pure/Concurrent/mailbox.ML"} how
-  to implement a mailbox as synchronized variable over a purely
-  functional list.\<close>
+  See @{file "~~/src/Pure/Concurrent/mailbox.ML"} how to implement a mailbox
+  as synchronized variable over a purely functional list.
+\<close>
 
 
 section \<open>Managed evaluation\<close>
 
-text \<open>Execution of Standard ML follows the model of strict
-  functional evaluation with optional exceptions.  Evaluation happens
-  whenever some function is applied to (sufficiently many)
-  arguments. The result is either an explicit value or an implicit
-  exception.
+text \<open>
+  Execution of Standard ML follows the model of strict functional evaluation
+  with optional exceptions. Evaluation happens whenever some function is
+  applied to (sufficiently many) arguments. The result is either an explicit
+  value or an implicit exception.
 
-  \<^emph>\<open>Managed evaluation\<close> in Isabelle/ML organizes expressions and
-  results to control certain physical side-conditions, to say more
-  specifically when and how evaluation happens.  For example, the
-  Isabelle/ML library supports lazy evaluation with memoing, parallel
-  evaluation via futures, asynchronous evaluation via promises,
-  evaluation with time limit etc.
+  \<^emph>\<open>Managed evaluation\<close> in Isabelle/ML organizes expressions and results to
+  control certain physical side-conditions, to say more specifically when and
+  how evaluation happens. For example, the Isabelle/ML library supports lazy
+  evaluation with memoing, parallel evaluation via futures, asynchronous
+  evaluation via promises, evaluation with time limit etc.
 
   \<^medskip>
-  An \<^emph>\<open>unevaluated expression\<close> is represented either as
-  unit abstraction \<^verbatim>\<open>fn () => a\<close> of type
-  \<^verbatim>\<open>unit -> 'a\<close> or as regular function
-  \<^verbatim>\<open>fn a => b\<close> of type \<^verbatim>\<open>'a -> 'b\<close>.  Both forms
-  occur routinely, and special care is required to tell them apart ---
-  the static type-system of SML is only of limited help here.
+  An \<^emph>\<open>unevaluated expression\<close> is represented either as unit abstraction \<^verbatim>\<open>fn
+  () => a\<close> of type \<^verbatim>\<open>unit -> 'a\<close> or as regular function \<^verbatim>\<open>fn a => b\<close> of type
+  \<^verbatim>\<open>'a -> 'b\<close>. Both forms occur routinely, and special care is required to
+  tell them apart --- the static type-system of SML is only of limited help
+  here.
 
-  The first form is more intuitive: some combinator \<open>(unit ->
-  'a) -> 'a\<close> applies the given function to \<open>()\<close> to initiate
-  the postponed evaluation process.  The second form is more flexible:
-  some combinator \<open>('a -> 'b) -> 'a -> 'b\<close> acts like a
-  modified form of function application; several such combinators may
-  be cascaded to modify a given function, before it is ultimately
-  applied to some argument.
+  The first form is more intuitive: some combinator \<open>(unit -> 'a) -> 'a\<close>
+  applies the given function to \<open>()\<close> to initiate the postponed evaluation
+  process. The second form is more flexible: some combinator \<open>('a -> 'b) -> 'a
+  -> 'b\<close> acts like a modified form of function application; several such
+  combinators may be cascaded to modify a given function, before it is
+  ultimately applied to some argument.
 
   \<^medskip>
-  \<^emph>\<open>Reified results\<close> make the disjoint sum of regular
-  values versions exceptional situations explicit as ML datatype:
-  \<open>'a result = Res of 'a | Exn of exn\<close>.  This is typically
-  used for administrative purposes, to store the overall outcome of an
-  evaluation process.
+  \<^emph>\<open>Reified results\<close> make the disjoint sum of regular values versions
+  exceptional situations explicit as ML datatype: \<open>'a result = Res of 'a | Exn
+  of exn\<close>. This is typically used for administrative purposes, to store the
+  overall outcome of an evaluation process.
 
-  \<^emph>\<open>Parallel exceptions\<close> aggregate reified results, such that
-  multiple exceptions are digested as a collection in canonical form
-  that identifies exceptions according to their original occurrence.
-  This is particular important for parallel evaluation via futures
-  \secref{sec:futures}, which are organized as acyclic graph of
-  evaluations that depend on other evaluations: exceptions stemming
-  from shared sub-graphs are exposed exactly once and in the order of
-  their original occurrence (e.g.\ when printed at the toplevel).
-  Interrupt counts as neutral element here: it is treated as minimal
-  information about some canceled evaluation process, and is absorbed
-  by the presence of regular program exceptions.\<close>
+  \<^emph>\<open>Parallel exceptions\<close> aggregate reified results, such that multiple
+  exceptions are digested as a collection in canonical form that identifies
+  exceptions according to their original occurrence. This is particular
+  important for parallel evaluation via futures \secref{sec:futures}, which
+  are organized as acyclic graph of evaluations that depend on other
+  evaluations: exceptions stemming from shared sub-graphs are exposed exactly
+  once and in the order of their original occurrence (e.g.\ when printed at
+  the toplevel). Interrupt counts as neutral element here: it is treated as
+  minimal information about some canceled evaluation process, and is absorbed
+  by the presence of regular program exceptions.
+\<close>
 
 text %mlref \<open>
   \begin{mldecls}
@@ -1886,54 +1852,49 @@
   @{index_ML Par_Exn.release_first: "'a Exn.result list -> 'a list"} \\
   \end{mldecls}
 
-  \<^descr> Type @{ML_type "'a Exn.result"} represents the disjoint sum of
-  ML results explicitly, with constructor @{ML Exn.Res} for regular
-  values and @{ML "Exn.Exn"} for exceptions.
+  \<^descr> Type @{ML_type "'a Exn.result"} represents the disjoint sum of ML results
+  explicitly, with constructor @{ML Exn.Res} for regular values and @{ML
+  "Exn.Exn"} for exceptions.
 
-  \<^descr> @{ML Exn.capture}~\<open>f x\<close> manages the evaluation of
-  \<open>f x\<close> such that exceptions are made explicit as @{ML
-  "Exn.Exn"}.  Note that this includes physical interrupts (see also
-  \secref{sec:exceptions}), so the same precautions apply to user
-  code: interrupts must not be absorbed accidentally!
+  \<^descr> @{ML Exn.capture}~\<open>f x\<close> manages the evaluation of \<open>f x\<close> such that
+  exceptions are made explicit as @{ML "Exn.Exn"}. Note that this includes
+  physical interrupts (see also \secref{sec:exceptions}), so the same
+  precautions apply to user code: interrupts must not be absorbed
+  accidentally!
 
-  \<^descr> @{ML Exn.interruptible_capture} is similar to @{ML
-  Exn.capture}, but interrupts are immediately re-raised as required
-  for user code.
+  \<^descr> @{ML Exn.interruptible_capture} is similar to @{ML Exn.capture}, but
+  interrupts are immediately re-raised as required for user code.
 
-  \<^descr> @{ML Exn.release}~\<open>result\<close> releases the original
-  runtime result, exposing its regular value or raising the reified
-  exception.
+  \<^descr> @{ML Exn.release}~\<open>result\<close> releases the original runtime result, exposing
+  its regular value or raising the reified exception.
 
-  \<^descr> @{ML Par_Exn.release_all}~\<open>results\<close> combines results
-  that were produced independently (e.g.\ by parallel evaluation).  If
-  all results are regular values, that list is returned.  Otherwise,
-  the collection of all exceptions is raised, wrapped-up as collective
-  parallel exception.  Note that the latter prevents access to
-  individual exceptions by conventional \<^verbatim>\<open>handle\<close> of ML.
+  \<^descr> @{ML Par_Exn.release_all}~\<open>results\<close> combines results that were produced
+  independently (e.g.\ by parallel evaluation). If all results are regular
+  values, that list is returned. Otherwise, the collection of all exceptions
+  is raised, wrapped-up as collective parallel exception. Note that the latter
+  prevents access to individual exceptions by conventional \<^verbatim>\<open>handle\<close> of ML.
 
-  \<^descr> @{ML Par_Exn.release_first} is similar to @{ML
-  Par_Exn.release_all}, but only the first (meaningful) exception that has
-  occurred in the original evaluation process is raised again, the others are
-  ignored.  That single exception may get handled by conventional
-  means in ML.
+  \<^descr> @{ML Par_Exn.release_first} is similar to @{ML Par_Exn.release_all}, but
+  only the first (meaningful) exception that has occurred in the original
+  evaluation process is raised again, the others are ignored. That single
+  exception may get handled by conventional means in ML.
 \<close>
 
 
 subsection \<open>Parallel skeletons \label{sec:parlist}\<close>
 
 text \<open>
-  Algorithmic skeletons are combinators that operate on lists in
-  parallel, in the manner of well-known \<open>map\<close>, \<open>exists\<close>,
-  \<open>forall\<close> etc.  Management of futures (\secref{sec:futures})
-  and their results as reified exceptions is wrapped up into simple
-  programming interfaces that resemble the sequential versions.
+  Algorithmic skeletons are combinators that operate on lists in parallel, in
+  the manner of well-known \<open>map\<close>, \<open>exists\<close>, \<open>forall\<close> etc. Management of
+  futures (\secref{sec:futures}) and their results as reified exceptions is
+  wrapped up into simple programming interfaces that resemble the sequential
+  versions.
 
-  What remains is the application-specific problem to present
-  expressions with suitable \<^emph>\<open>granularity\<close>: each list element
-  corresponds to one evaluation task.  If the granularity is too
-  coarse, the available CPUs are not saturated.  If it is too
-  fine-grained, CPU cycles are wasted due to the overhead of
-  organizing parallel processing.  In the worst case, parallel
+  What remains is the application-specific problem to present expressions with
+  suitable \<^emph>\<open>granularity\<close>: each list element corresponds to one evaluation
+  task. If the granularity is too coarse, the available CPUs are not
+  saturated. If it is too fine-grained, CPU cycles are wasted due to the
+  overhead of organizing parallel processing. In the worst case, parallel
   performance will be less than the sequential counterpart!
 \<close>
 
@@ -1943,31 +1904,29 @@
   @{index_ML Par_List.get_some: "('a -> 'b option) -> 'a list -> 'b option"} \\
   \end{mldecls}
 
-  \<^descr> @{ML Par_List.map}~\<open>f [x\<^sub>1, \<dots>, x\<^sub>n]\<close> is like @{ML
-  "map"}~\<open>f [x\<^sub>1, \<dots>, x\<^sub>n]\<close>, but the evaluation of \<open>f x\<^sub>i\<close>
-  for \<open>i = 1, \<dots>, n\<close> is performed in parallel.
+  \<^descr> @{ML Par_List.map}~\<open>f [x\<^sub>1, \<dots>, x\<^sub>n]\<close> is like @{ML "map"}~\<open>f [x\<^sub>1, \<dots>,
+  x\<^sub>n]\<close>, but the evaluation of \<open>f x\<^sub>i\<close> for \<open>i = 1, \<dots>, n\<close> is performed in
+  parallel.
 
-  An exception in any \<open>f x\<^sub>i\<close> cancels the overall evaluation
-  process.  The final result is produced via @{ML
-  Par_Exn.release_first} as explained above, which means the first
-  program exception that happened to occur in the parallel evaluation
-  is propagated, and all other failures are ignored.
+  An exception in any \<open>f x\<^sub>i\<close> cancels the overall evaluation process. The
+  final result is produced via @{ML Par_Exn.release_first} as explained above,
+  which means the first program exception that happened to occur in the
+  parallel evaluation is propagated, and all other failures are ignored.
 
-  \<^descr> @{ML Par_List.get_some}~\<open>f [x\<^sub>1, \<dots>, x\<^sub>n]\<close> produces some
-  \<open>f x\<^sub>i\<close> that is of the form \<open>SOME y\<^sub>i\<close>, if that
-  exists, otherwise \<open>NONE\<close>.  Thus it is similar to @{ML
-  Library.get_first}, but subject to a non-deterministic parallel
-  choice process.  The first successful result cancels the overall
-  evaluation process; other exceptions are propagated as for @{ML
-  Par_List.map}.
+  \<^descr> @{ML Par_List.get_some}~\<open>f [x\<^sub>1, \<dots>, x\<^sub>n]\<close> produces some \<open>f x\<^sub>i\<close> that is of
+  the form \<open>SOME y\<^sub>i\<close>, if that exists, otherwise \<open>NONE\<close>. Thus it is similar to
+  @{ML Library.get_first}, but subject to a non-deterministic parallel choice
+  process. The first successful result cancels the overall evaluation process;
+  other exceptions are propagated as for @{ML Par_List.map}.
 
-  This generic parallel choice combinator is the basis for derived
-  forms, such as @{ML Par_List.find_some}, @{ML Par_List.exists}, @{ML
-  Par_List.forall}.
+  This generic parallel choice combinator is the basis for derived forms, such
+  as @{ML Par_List.find_some}, @{ML Par_List.exists}, @{ML Par_List.forall}.
 \<close>
 
-text %mlex \<open>Subsequently, the Ackermann function is evaluated in
-  parallel for some ranges of arguments.\<close>
+text %mlex \<open>
+  Subsequently, the Ackermann function is evaluated in parallel for some
+  ranges of arguments.
+\<close>
 
 ML_val \<open>
   fun ackermann 0 n = n + 1
@@ -1982,19 +1941,19 @@
 subsection \<open>Lazy evaluation\<close>
 
 text \<open>
-  Classic lazy evaluation works via the \<open>lazy\<close>~/ \<open>force\<close> pair of
-  operations: \<open>lazy\<close> to wrap an unevaluated expression, and \<open>force\<close> to evaluate it once and store its result persistently. Later
-  invocations of \<open>force\<close> retrieve the stored result without another
-  evaluation. Isabelle/ML refines this idea to accommodate the aspects of
-  multi-threading, synchronous program exceptions and asynchronous interrupts.
+  Classic lazy evaluation works via the \<open>lazy\<close>~/ \<open>force\<close> pair of operations:
+  \<open>lazy\<close> to wrap an unevaluated expression, and \<open>force\<close> to evaluate it once
+  and store its result persistently. Later invocations of \<open>force\<close> retrieve the
+  stored result without another evaluation. Isabelle/ML refines this idea to
+  accommodate the aspects of multi-threading, synchronous program exceptions
+  and asynchronous interrupts.
 
-  The first thread that invokes \<open>force\<close> on an unfinished lazy value
-  changes its state into a \<^emph>\<open>promise\<close> of the eventual result and starts
-  evaluating it. Any other threads that \<open>force\<close> the same lazy value in
-  the meantime need to wait for it to finish, by producing a regular result or
-  program exception. If the evaluation attempt is interrupted, this event is
-  propagated to all waiting threads and the lazy value is reset to its
-  original state.
+  The first thread that invokes \<open>force\<close> on an unfinished lazy value changes
+  its state into a \<^emph>\<open>promise\<close> of the eventual result and starts evaluating it.
+  Any other threads that \<open>force\<close> the same lazy value in the meantime need to
+  wait for it to finish, by producing a regular result or program exception.
+  If the evaluation attempt is interrupted, this event is propagated to all
+  waiting threads and the lazy value is reset to its original state.
 
   This means a lazy value is completely evaluated at most once, in a
   thread-safe manner. There might be multiple interrupted evaluation attempts,
@@ -2013,11 +1972,11 @@
 
   \<^descr> Type @{ML_type "'a lazy"} represents lazy values over type \<^verbatim>\<open>'a\<close>.
 
-  \<^descr> @{ML Lazy.lazy}~\<open>(fn () => e)\<close> wraps the unevaluated
-  expression \<open>e\<close> as unfinished lazy value.
+  \<^descr> @{ML Lazy.lazy}~\<open>(fn () => e)\<close> wraps the unevaluated expression \<open>e\<close> as
+  unfinished lazy value.
 
-  \<^descr> @{ML Lazy.value}~\<open>a\<close> wraps the value \<open>a\<close> as finished lazy
-  value.  When forced, it returns \<open>a\<close> without any further evaluation.
+  \<^descr> @{ML Lazy.value}~\<open>a\<close> wraps the value \<open>a\<close> as finished lazy value. When
+  forced, it returns \<open>a\<close> without any further evaluation.
 
   There is very low overhead for this proforma wrapping of strict values as
   lazy values.
@@ -2032,19 +1991,19 @@
 
 text \<open>
   Futures help to organize parallel execution in a value-oriented manner, with
-  \<open>fork\<close>~/ \<open>join\<close> as the main pair of operations, and some further
-  variants; see also @{cite "Wenzel:2009" and "Wenzel:2013:ITP"}. Unlike lazy
-  values, futures are evaluated strictly and spontaneously on separate worker
-  threads. Futures may be canceled, which leads to interrupts on running
-  evaluation attempts, and forces structurally related futures to fail for all
-  time; already finished futures remain unchanged. Exceptions between related
+  \<open>fork\<close>~/ \<open>join\<close> as the main pair of operations, and some further variants;
+  see also @{cite "Wenzel:2009" and "Wenzel:2013:ITP"}. Unlike lazy values,
+  futures are evaluated strictly and spontaneously on separate worker threads.
+  Futures may be canceled, which leads to interrupts on running evaluation
+  attempts, and forces structurally related futures to fail for all time;
+  already finished futures remain unchanged. Exceptions between related
   futures are propagated as well, and turned into parallel exceptions (see
   above).
 
   Technically, a future is a single-assignment variable together with a
-  \<^emph>\<open>task\<close> that serves administrative purposes, notably within the
-  \<^emph>\<open>task queue\<close> where new futures are registered for eventual evaluation
-  and the worker threads retrieve their work.
+  \<^emph>\<open>task\<close> that serves administrative purposes, notably within the \<^emph>\<open>task
+  queue\<close> where new futures are registered for eventual evaluation and the
+  worker threads retrieve their work.
 
   The pool of worker threads is limited, in correlation with the number of
   physical cores on the machine. Note that allocation of runtime resources may
@@ -2055,27 +2014,27 @@
   timeout.
 
   \<^medskip>
-  Each future task belongs to some \<^emph>\<open>task group\<close>, which
-  represents the hierarchic structure of related tasks, together with the
-  exception status a that point. By default, the task group of a newly created
-  future is a new sub-group of the presently running one, but it is also
-  possible to indicate different group layouts under program control.
+  Each future task belongs to some \<^emph>\<open>task group\<close>, which represents the
+  hierarchic structure of related tasks, together with the exception status a
+  that point. By default, the task group of a newly created future is a new
+  sub-group of the presently running one, but it is also possible to indicate
+  different group layouts under program control.
 
   Cancellation of futures actually refers to the corresponding task group and
   all its sub-groups. Thus interrupts are propagated down the group hierarchy.
   Regular program exceptions are treated likewise: failure of the evaluation
   of some future task affects its own group and all sub-groups. Given a
-  particular task group, its \<^emph>\<open>group status\<close> cumulates all relevant
-  exceptions according to its position within the group hierarchy. Interrupted
-  tasks that lack regular result information, will pick up parallel exceptions
-  from the cumulative group status.
+  particular task group, its \<^emph>\<open>group status\<close> cumulates all relevant exceptions
+  according to its position within the group hierarchy. Interrupted tasks that
+  lack regular result information, will pick up parallel exceptions from the
+  cumulative group status.
 
   \<^medskip>
-  A \<^emph>\<open>passive future\<close> or \<^emph>\<open>promise\<close> is a future with slightly
-  different evaluation policies: there is only a single-assignment variable
-  and some expression to evaluate for the \<^emph>\<open>failed\<close> case (e.g.\ to clean
-  up resources when canceled). A regular result is produced by external means,
-  using a separate \<^emph>\<open>fulfill\<close> operation.
+  A \<^emph>\<open>passive future\<close> or \<^emph>\<open>promise\<close> is a future with slightly different
+  evaluation policies: there is only a single-assignment variable and some
+  expression to evaluate for the \<^emph>\<open>failed\<close> case (e.g.\ to clean up resources
+  when canceled). A regular result is produced by external means, using a
+  separate \<^emph>\<open>fulfill\<close> operation.
 
   Promises are managed in the same task queue, so regular futures may depend
   on them. This allows a form of reactive programming, where some promises are
@@ -2101,52 +2060,51 @@
 
   \<^descr> Type @{ML_type "'a future"} represents future values over type \<^verbatim>\<open>'a\<close>.
 
-  \<^descr> @{ML Future.fork}~\<open>(fn () => e)\<close> registers the unevaluated
-  expression \<open>e\<close> as unfinished future value, to be evaluated eventually
-  on the parallel worker-thread farm. This is a shorthand for @{ML
-  Future.forks} below, with default parameters and a single expression.
+  \<^descr> @{ML Future.fork}~\<open>(fn () => e)\<close> registers the unevaluated expression \<open>e\<close>
+  as unfinished future value, to be evaluated eventually on the parallel
+  worker-thread farm. This is a shorthand for @{ML Future.forks} below, with
+  default parameters and a single expression.
 
-  \<^descr> @{ML Future.forks}~\<open>params exprs\<close> is the general interface to
-  fork several futures simultaneously. The \<open>params\<close> consist of the
-  following fields:
+  \<^descr> @{ML Future.forks}~\<open>params exprs\<close> is the general interface to fork several
+  futures simultaneously. The \<open>params\<close> consist of the following fields:
 
-    \<^item> \<open>name : string\<close> (default @{ML "\"\""}) specifies a common name
-    for the tasks of the forked futures, which serves diagnostic purposes.
+    \<^item> \<open>name : string\<close> (default @{ML "\"\""}) specifies a common name for the
+    tasks of the forked futures, which serves diagnostic purposes.
 
-    \<^item> \<open>group : Future.group option\<close> (default @{ML NONE}) specifies
-    an optional task group for the forked futures. @{ML NONE} means that a new
-    sub-group of the current worker-thread task context is created. If this is
-    not a worker thread, the group will be a new root in the group hierarchy.
+    \<^item> \<open>group : Future.group option\<close> (default @{ML NONE}) specifies an optional
+    task group for the forked futures. @{ML NONE} means that a new sub-group
+    of the current worker-thread task context is created. If this is not a
+    worker thread, the group will be a new root in the group hierarchy.
 
-    \<^item> \<open>deps : Future.task list\<close> (default @{ML "[]"}) specifies
-    dependencies on other future tasks, i.e.\ the adjacency relation in the
-    global task queue. Dependencies on already finished tasks are ignored.
+    \<^item> \<open>deps : Future.task list\<close> (default @{ML "[]"}) specifies dependencies on
+    other future tasks, i.e.\ the adjacency relation in the global task queue.
+    Dependencies on already finished tasks are ignored.
 
-    \<^item> \<open>pri : int\<close> (default @{ML 0}) specifies a priority within the
-    task queue.
+    \<^item> \<open>pri : int\<close> (default @{ML 0}) specifies a priority within the task
+    queue.
 
-    Typically there is only little deviation from the default priority @{ML 0}.
-    As a rule of thumb, @{ML "~1"} means ``low priority" and @{ML 1} means
+    Typically there is only little deviation from the default priority @{ML
+    0}. As a rule of thumb, @{ML "~1"} means ``low priority" and @{ML 1} means
     ``high priority''.
 
-    Note that the task priority only affects the position in the queue, not the
-    thread priority. When a worker thread picks up a task for processing, it
-    runs with the normal thread priority to the end (or until canceled). Higher
-    priority tasks that are queued later need to wait until this (or another)
-    worker thread becomes free again.
+    Note that the task priority only affects the position in the queue, not
+    the thread priority. When a worker thread picks up a task for processing,
+    it runs with the normal thread priority to the end (or until canceled).
+    Higher priority tasks that are queued later need to wait until this (or
+    another) worker thread becomes free again.
 
-    \<^item> \<open>interrupts : bool\<close> (default @{ML true}) tells whether the
-    worker thread that processes the corresponding task is initially put into
-    interruptible state. This state may change again while running, by modifying
-    the thread attributes.
+    \<^item> \<open>interrupts : bool\<close> (default @{ML true}) tells whether the worker thread
+    that processes the corresponding task is initially put into interruptible
+    state. This state may change again while running, by modifying the thread
+    attributes.
 
-    With interrupts disabled, a running future task cannot be canceled.  It is
+    With interrupts disabled, a running future task cannot be canceled. It is
     the responsibility of the programmer that this special state is retained
     only briefly.
 
-  \<^descr> @{ML Future.join}~\<open>x\<close> retrieves the value of an already finished
-  future, which may lead to an exception, according to the result of its
-  previous evaluation.
+  \<^descr> @{ML Future.join}~\<open>x\<close> retrieves the value of an already finished future,
+  which may lead to an exception, according to the result of its previous
+  evaluation.
 
   For an unfinished future there are several cases depending on the role of
   the current thread and the status of the future. A non-worker thread waits
@@ -2163,11 +2121,11 @@
   some timeout.
 
   Whenever possible, static dependencies of futures should be specified
-  explicitly when forked (see \<open>deps\<close> above). Thus the evaluation can
-  work from the bottom up, without join conflicts and wait states.
+  explicitly when forked (see \<open>deps\<close> above). Thus the evaluation can work from
+  the bottom up, without join conflicts and wait states.
 
-  \<^descr> @{ML Future.joins}~\<open>xs\<close> joins the given list of futures
-  simultaneously, which is more efficient than @{ML "map Future.join"}~\<open>xs\<close>.
+  \<^descr> @{ML Future.joins}~\<open>xs\<close> joins the given list of futures simultaneously,
+  which is more efficient than @{ML "map Future.join"}~\<open>xs\<close>.
 
   Based on the dependency graph of tasks, the current thread takes over the
   responsibility to evaluate future expressions that are required for the main
@@ -2175,36 +2133,36 @@
   presently evaluated on other threads only happens as last resort, when no
   other unfinished futures are left over.
 
-  \<^descr> @{ML Future.value}~\<open>a\<close> wraps the value \<open>a\<close> as finished
-  future value, bypassing the worker-thread farm. When joined, it returns
-  \<open>a\<close> without any further evaluation.
+  \<^descr> @{ML Future.value}~\<open>a\<close> wraps the value \<open>a\<close> as finished future value,
+  bypassing the worker-thread farm. When joined, it returns \<open>a\<close> without any
+  further evaluation.
 
   There is very low overhead for this proforma wrapping of strict values as
   futures.
 
   \<^descr> @{ML Future.map}~\<open>f x\<close> is a fast-path implementation of @{ML
-  Future.fork}~\<open>(fn () => f (\<close>@{ML Future.join}~\<open>x))\<close>, which
-  avoids the full overhead of the task queue and worker-thread farm as far as
-  possible. The function \<open>f\<close> is supposed to be some trivial
-  post-processing or projection of the future result.
+  Future.fork}~\<open>(fn () => f (\<close>@{ML Future.join}~\<open>x))\<close>, which avoids the full
+  overhead of the task queue and worker-thread farm as far as possible. The
+  function \<open>f\<close> is supposed to be some trivial post-processing or projection of
+  the future result.
 
-  \<^descr> @{ML Future.cancel}~\<open>x\<close> cancels the task group of the given
-  future, using @{ML Future.cancel_group} below.
+  \<^descr> @{ML Future.cancel}~\<open>x\<close> cancels the task group of the given future, using
+  @{ML Future.cancel_group} below.
 
-  \<^descr> @{ML Future.cancel_group}~\<open>group\<close> cancels all tasks of the
-  given task group for all time. Threads that are presently processing a task
-  of the given group are interrupted: it may take some time until they are
-  actually terminated. Tasks that are queued but not yet processed are
-  dequeued and forced into interrupted state. Since the task group is itself
-  invalidated, any further attempt to fork a future that belongs to it will
-  yield a canceled result as well.
+  \<^descr> @{ML Future.cancel_group}~\<open>group\<close> cancels all tasks of the given task
+  group for all time. Threads that are presently processing a task of the
+  given group are interrupted: it may take some time until they are actually
+  terminated. Tasks that are queued but not yet processed are dequeued and
+  forced into interrupted state. Since the task group is itself invalidated,
+  any further attempt to fork a future that belongs to it will yield a
+  canceled result as well.
 
-  \<^descr> @{ML Future.promise}~\<open>abort\<close> registers a passive future with the
-  given \<open>abort\<close> operation: it is invoked when the future task group is
-  canceled.
+  \<^descr> @{ML Future.promise}~\<open>abort\<close> registers a passive future with the given
+  \<open>abort\<close> operation: it is invoked when the future task group is canceled.
 
-  \<^descr> @{ML Future.fulfill}~\<open>x a\<close> finishes the passive future \<open>x\<close> by the given value \<open>a\<close>. If the promise has already been canceled,
-  the attempt to fulfill it causes an exception.
+  \<^descr> @{ML Future.fulfill}~\<open>x a\<close> finishes the passive future \<open>x\<close> by the given
+  value \<open>a\<close>. If the promise has already been canceled, the attempt to fulfill
+  it causes an exception.
 \<close>
 
 end
--- a/src/Doc/Implementation/Prelim.thy	Tue Dec 15 14:41:47 2015 +0000
+++ b/src/Doc/Implementation/Prelim.thy	Wed Dec 16 17:30:30 2015 +0100
@@ -9,88 +9,78 @@
 section \<open>Contexts \label{sec:context}\<close>
 
 text \<open>
-  A logical context represents the background that is required for
-  formulating statements and composing proofs.  It acts as a medium to
-  produce formal content, depending on earlier material (declarations,
-  results etc.).
+  A logical context represents the background that is required for formulating
+  statements and composing proofs. It acts as a medium to produce formal
+  content, depending on earlier material (declarations, results etc.).
 
-  For example, derivations within the Isabelle/Pure logic can be
-  described as a judgment \<open>\<Gamma> \<turnstile>\<^sub>\<Theta> \<phi>\<close>, which means that a
-  proposition \<open>\<phi>\<close> is derivable from hypotheses \<open>\<Gamma>\<close>
-  within the theory \<open>\<Theta>\<close>.  There are logical reasons for
-  keeping \<open>\<Theta>\<close> and \<open>\<Gamma>\<close> separate: theories can be
-  liberal about supporting type constructors and schematic
-  polymorphism of constants and axioms, while the inner calculus of
-  \<open>\<Gamma> \<turnstile> \<phi>\<close> is strictly limited to Simple Type Theory (with
+  For example, derivations within the Isabelle/Pure logic can be described as
+  a judgment \<open>\<Gamma> \<turnstile>\<^sub>\<Theta> \<phi>\<close>, which means that a proposition \<open>\<phi>\<close> is derivable from
+  hypotheses \<open>\<Gamma>\<close> within the theory \<open>\<Theta>\<close>. There are logical reasons for keeping
+  \<open>\<Theta>\<close> and \<open>\<Gamma>\<close> separate: theories can be liberal about supporting type
+  constructors and schematic polymorphism of constants and axioms, while the
+  inner calculus of \<open>\<Gamma> \<turnstile> \<phi>\<close> is strictly limited to Simple Type Theory (with
   fixed type variables in the assumptions).
 
   \<^medskip>
-  Contexts and derivations are linked by the following key
-  principles:
+  Contexts and derivations are linked by the following key principles:
 
-  \<^item> Transfer: monotonicity of derivations admits results to be
-  transferred into a \<^emph>\<open>larger\<close> context, i.e.\ \<open>\<Gamma> \<turnstile>\<^sub>\<Theta>
-  \<phi>\<close> implies \<open>\<Gamma>' \<turnstile>\<^sub>\<Theta>\<^sub>' \<phi>\<close> for contexts \<open>\<Theta>'
-  \<supseteq> \<Theta>\<close> and \<open>\<Gamma>' \<supseteq> \<Gamma>\<close>.
+  \<^item> Transfer: monotonicity of derivations admits results to be transferred
+  into a \<^emph>\<open>larger\<close> context, i.e.\ \<open>\<Gamma> \<turnstile>\<^sub>\<Theta> \<phi>\<close> implies \<open>\<Gamma>' \<turnstile>\<^sub>\<Theta>\<^sub>' \<phi>\<close> for contexts
+  \<open>\<Theta>' \<supseteq> \<Theta>\<close> and \<open>\<Gamma>' \<supseteq> \<Gamma>\<close>.
 
-  \<^item> Export: discharge of hypotheses admits results to be exported
-  into a \<^emph>\<open>smaller\<close> context, i.e.\ \<open>\<Gamma>' \<turnstile>\<^sub>\<Theta> \<phi>\<close>
-  implies \<open>\<Gamma> \<turnstile>\<^sub>\<Theta> \<Delta> \<Longrightarrow> \<phi>\<close> where \<open>\<Gamma>' \<supseteq> \<Gamma>\<close> and
-  \<open>\<Delta> = \<Gamma>' - \<Gamma>\<close>.  Note that \<open>\<Theta>\<close> remains unchanged here,
-  only the \<open>\<Gamma>\<close> part is affected.
+  \<^item> Export: discharge of hypotheses admits results to be exported into a
+  \<^emph>\<open>smaller\<close> context, i.e.\ \<open>\<Gamma>' \<turnstile>\<^sub>\<Theta> \<phi>\<close> implies \<open>\<Gamma> \<turnstile>\<^sub>\<Theta> \<Delta> \<Longrightarrow> \<phi>\<close> where \<open>\<Gamma>' \<supseteq> \<Gamma>\<close>
+  and \<open>\<Delta> = \<Gamma>' - \<Gamma>\<close>. Note that \<open>\<Theta>\<close> remains unchanged here, only the \<open>\<Gamma>\<close> part is
+  affected.
 
 
   \<^medskip>
-  By modeling the main characteristics of the primitive
-  \<open>\<Theta>\<close> and \<open>\<Gamma>\<close> above, and abstracting over any
-  particular logical content, we arrive at the fundamental notions of
-  \<^emph>\<open>theory context\<close> and \<^emph>\<open>proof context\<close> in Isabelle/Isar.
-  These implement a certain policy to manage arbitrary \<^emph>\<open>context
-  data\<close>.  There is a strongly-typed mechanism to declare new kinds of
+  By modeling the main characteristics of the primitive \<open>\<Theta>\<close> and \<open>\<Gamma>\<close> above, and
+  abstracting over any particular logical content, we arrive at the
+  fundamental notions of \<^emph>\<open>theory context\<close> and \<^emph>\<open>proof context\<close> in
+  Isabelle/Isar. These implement a certain policy to manage arbitrary
+  \<^emph>\<open>context data\<close>. There is a strongly-typed mechanism to declare new kinds of
   data at compile time.
 
-  The internal bootstrap process of Isabelle/Pure eventually reaches a
-  stage where certain data slots provide the logical content of \<open>\<Theta>\<close> and \<open>\<Gamma>\<close> sketched above, but this does not stop there!
-  Various additional data slots support all kinds of mechanisms that
-  are not necessarily part of the core logic.
+  The internal bootstrap process of Isabelle/Pure eventually reaches a stage
+  where certain data slots provide the logical content of \<open>\<Theta>\<close> and \<open>\<Gamma>\<close> sketched
+  above, but this does not stop there! Various additional data slots support
+  all kinds of mechanisms that are not necessarily part of the core logic.
 
-  For example, there would be data for canonical introduction and
-  elimination rules for arbitrary operators (depending on the
-  object-logic and application), which enables users to perform
-  standard proof steps implicitly (cf.\ the \<open>rule\<close> method
-  @{cite "isabelle-isar-ref"}).
+  For example, there would be data for canonical introduction and elimination
+  rules for arbitrary operators (depending on the object-logic and
+  application), which enables users to perform standard proof steps implicitly
+  (cf.\ the \<open>rule\<close> method @{cite "isabelle-isar-ref"}).
 
   \<^medskip>
-  Thus Isabelle/Isar is able to bring forth more and more
-  concepts successively.  In particular, an object-logic like
-  Isabelle/HOL continues the Isabelle/Pure setup by adding specific
-  components for automated reasoning (classical reasoner, tableau
-  prover, structured induction etc.) and derived specification
-  mechanisms (inductive predicates, recursive functions etc.).  All of
-  this is ultimately based on the generic data management by theory
-  and proof contexts introduced here.
+  Thus Isabelle/Isar is able to bring forth more and more concepts
+  successively. In particular, an object-logic like Isabelle/HOL continues the
+  Isabelle/Pure setup by adding specific components for automated reasoning
+  (classical reasoner, tableau prover, structured induction etc.) and derived
+  specification mechanisms (inductive predicates, recursive functions etc.).
+  All of this is ultimately based on the generic data management by theory and
+  proof contexts introduced here.
 \<close>
 
 
 subsection \<open>Theory context \label{sec:context-theory}\<close>
 
-text \<open>A \<^emph>\<open>theory\<close> is a data container with explicit name and
-  unique identifier.  Theories are related by a (nominal) sub-theory
-  relation, which corresponds to the dependency graph of the original
-  construction; each theory is derived from a certain sub-graph of
-  ancestor theories.  To this end, the system maintains a set of
-  symbolic ``identification stamps'' within each theory.
+text \<open>
+  A \<^emph>\<open>theory\<close> is a data container with explicit name and unique identifier.
+  Theories are related by a (nominal) sub-theory relation, which corresponds
+  to the dependency graph of the original construction; each theory is derived
+  from a certain sub-graph of ancestor theories. To this end, the system
+  maintains a set of symbolic ``identification stamps'' within each theory.
 
-  The \<open>begin\<close> operation starts a new theory by importing several
-  parent theories (with merged contents) and entering a special mode of
-  nameless incremental updates, until the final \<open>end\<close> operation is
-  performed.
+  The \<open>begin\<close> operation starts a new theory by importing several parent
+  theories (with merged contents) and entering a special mode of nameless
+  incremental updates, until the final \<open>end\<close> operation is performed.
 
   \<^medskip>
-  The example in \figref{fig:ex-theory} below shows a theory
-  graph derived from \<open>Pure\<close>, with theory \<open>Length\<close>
-  importing \<open>Nat\<close> and \<open>List\<close>.  The body of \<open>Length\<close> consists of a sequence of updates, resulting in locally a
-  linear sub-theory relation for each intermediate step.
+  The example in \figref{fig:ex-theory} below shows a theory graph derived
+  from \<open>Pure\<close>, with theory \<open>Length\<close> importing \<open>Nat\<close> and \<open>List\<close>. The body of
+  \<open>Length\<close> consists of a sequence of updates, resulting in locally a linear
+  sub-theory relation for each intermediate step.
 
   \begin{figure}[htb]
   \begin{center}
@@ -111,10 +101,10 @@
   \end{figure}
 
   \<^medskip>
-  Derived formal entities may retain a reference to the
-  background theory in order to indicate the formal context from which
-  they were produced.  This provides an immutable certificate of the
-  background theory.\<close>
+  Derived formal entities may retain a reference to the background theory in
+  order to indicate the formal context from which they were produced. This
+  provides an immutable certificate of the background theory.
+\<close>
 
 text %mlref \<open>
   \begin{mldecls}
@@ -128,24 +118,21 @@
 
   \<^descr> Type @{ML_type theory} represents theory contexts.
 
-  \<^descr> @{ML "Context.eq_thy"}~\<open>(thy\<^sub>1, thy\<^sub>2)\<close> check strict
-  identity of two theories.
+  \<^descr> @{ML "Context.eq_thy"}~\<open>(thy\<^sub>1, thy\<^sub>2)\<close> check strict identity of two
+  theories.
 
-  \<^descr> @{ML "Context.subthy"}~\<open>(thy\<^sub>1, thy\<^sub>2)\<close> compares theories
-  according to the intrinsic graph structure of the construction.
-  This sub-theory relation is a nominal approximation of inclusion
-  (\<open>\<subseteq>\<close>) of the corresponding content (according to the
-  semantics of the ML modules that implement the data).
+  \<^descr> @{ML "Context.subthy"}~\<open>(thy\<^sub>1, thy\<^sub>2)\<close> compares theories according to the
+  intrinsic graph structure of the construction. This sub-theory relation is a
+  nominal approximation of inclusion (\<open>\<subseteq>\<close>) of the corresponding content
+  (according to the semantics of the ML modules that implement the data).
 
-  \<^descr> @{ML "Theory.begin_theory"}~\<open>name parents\<close> constructs
-  a new theory based on the given parents.  This ML function is
-  normally not invoked directly.
+  \<^descr> @{ML "Theory.begin_theory"}~\<open>name parents\<close> constructs a new theory based
+  on the given parents. This ML function is normally not invoked directly.
 
-  \<^descr> @{ML "Theory.parents_of"}~\<open>thy\<close> returns the direct
-  ancestors of \<open>thy\<close>.
+  \<^descr> @{ML "Theory.parents_of"}~\<open>thy\<close> returns the direct ancestors of \<open>thy\<close>.
 
-  \<^descr> @{ML "Theory.ancestors_of"}~\<open>thy\<close> returns all
-  ancestors of \<open>thy\<close> (not including \<open>thy\<close> itself).
+  \<^descr> @{ML "Theory.ancestors_of"}~\<open>thy\<close> returns all ancestors of \<open>thy\<close> (not
+  including \<open>thy\<close> itself).
 \<close>
 
 text %mlantiq \<open>
@@ -160,45 +147,43 @@
   @@{ML_antiquotation theory_context} nameref
   \<close>}
 
-  \<^descr> \<open>@{theory}\<close> refers to the background theory of the
-  current context --- as abstract value.
+  \<^descr> \<open>@{theory}\<close> refers to the background theory of the current context --- as
+  abstract value.
 
-  \<^descr> \<open>@{theory A}\<close> refers to an explicitly named ancestor
-  theory \<open>A\<close> of the background theory of the current context
-  --- as abstract value.
+  \<^descr> \<open>@{theory A}\<close> refers to an explicitly named ancestor theory \<open>A\<close> of the
+  background theory of the current context --- as abstract value.
 
-  \<^descr> \<open>@{theory_context A}\<close> is similar to \<open>@{theory
-  A}\<close>, but presents the result as initial @{ML_type Proof.context}
-  (see also @{ML Proof_Context.init_global}).
+  \<^descr> \<open>@{theory_context A}\<close> is similar to \<open>@{theory A}\<close>, but presents the result
+  as initial @{ML_type Proof.context} (see also @{ML
+  Proof_Context.init_global}).
 \<close>
 
 
 subsection \<open>Proof context \label{sec:context-proof}\<close>
 
-text \<open>A proof context is a container for pure data that refers to
-  the theory from which it is derived. The \<open>init\<close> operation
-  creates a proof context from a given theory. There is an explicit
-  \<open>transfer\<close> operation to force resynchronization with updates
-  to the background theory -- this is rarely required in practice.
+text \<open>
+  A proof context is a container for pure data that refers to the theory from
+  which it is derived. The \<open>init\<close> operation creates a proof context from a
+  given theory. There is an explicit \<open>transfer\<close> operation to force
+  resynchronization with updates to the background theory -- this is rarely
+  required in practice.
 
-  Entities derived in a proof context need to record logical
-  requirements explicitly, since there is no separate context
-  identification or symbolic inclusion as for theories.  For example,
-  hypotheses used in primitive derivations (cf.\ \secref{sec:thms})
-  are recorded separately within the sequent \<open>\<Gamma> \<turnstile> \<phi>\<close>, just to
-  make double sure.  Results could still leak into an alien proof
-  context due to programming errors, but Isabelle/Isar includes some
-  extra validity checks in critical positions, notably at the end of a
+  Entities derived in a proof context need to record logical requirements
+  explicitly, since there is no separate context identification or symbolic
+  inclusion as for theories. For example, hypotheses used in primitive
+  derivations (cf.\ \secref{sec:thms}) are recorded separately within the
+  sequent \<open>\<Gamma> \<turnstile> \<phi>\<close>, just to make double sure. Results could still leak into an
+  alien proof context due to programming errors, but Isabelle/Isar includes
+  some extra validity checks in critical positions, notably at the end of a
   sub-proof.
 
   Proof contexts may be manipulated arbitrarily, although the common
-  discipline is to follow block structure as a mental model: a given
-  context is extended consecutively, and results are exported back
-  into the original context.  Note that an Isar proof state models
-  block-structured reasoning explicitly, using a stack of proof
-  contexts internally.  For various technical reasons, the background
-  theory of an Isar proof state must not be changed while the proof is
-  still under construction!
+  discipline is to follow block structure as a mental model: a given context
+  is extended consecutively, and results are exported back into the original
+  context. Note that an Isar proof state models block-structured reasoning
+  explicitly, using a stack of proof contexts internally. For various
+  technical reasons, the background theory of an Isar proof state must not be
+  changed while the proof is still under construction!
 \<close>
 
 text %mlref \<open>
@@ -211,14 +196,14 @@
 
   \<^descr> Type @{ML_type Proof.context} represents proof contexts.
 
-  \<^descr> @{ML Proof_Context.init_global}~\<open>thy\<close> produces a proof
-  context derived from \<open>thy\<close>, initializing all data.
+  \<^descr> @{ML Proof_Context.init_global}~\<open>thy\<close> produces a proof context derived
+  from \<open>thy\<close>, initializing all data.
 
-  \<^descr> @{ML Proof_Context.theory_of}~\<open>ctxt\<close> selects the
-  background theory from \<open>ctxt\<close>.
+  \<^descr> @{ML Proof_Context.theory_of}~\<open>ctxt\<close> selects the background theory from
+  \<open>ctxt\<close>.
 
-  \<^descr> @{ML Proof_Context.transfer}~\<open>thy ctxt\<close> promotes the
-  background theory of \<open>ctxt\<close> to the super theory \<open>thy\<close>.
+  \<^descr> @{ML Proof_Context.transfer}~\<open>thy ctxt\<close> promotes the background theory of
+  \<open>ctxt\<close> to the super theory \<open>thy\<close>.
 \<close>
 
 text %mlantiq \<open>
@@ -226,9 +211,9 @@
   @{ML_antiquotation_def "context"} & : & \<open>ML_antiquotation\<close> \\
   \end{matharray}
 
-  \<^descr> \<open>@{context}\<close> refers to \<^emph>\<open>the\<close> context at
-  compile-time --- as abstract value.  Independently of (local) theory
-  or proof mode, this always produces a meaningful result.
+  \<^descr> \<open>@{context}\<close> refers to \<^emph>\<open>the\<close> context at compile-time --- as abstract
+  value. Independently of (local) theory or proof mode, this always produces a
+  meaningful result.
 
   This is probably the most common antiquotation in interactive
   experimentation with ML inside Isar.
@@ -238,17 +223,16 @@
 subsection \<open>Generic contexts \label{sec:generic-context}\<close>
 
 text \<open>
-  A generic context is the disjoint sum of either a theory or proof
-  context.  Occasionally, this enables uniform treatment of generic
-  context data, typically extra-logical information.  Operations on
-  generic contexts include the usual injections, partial selections,
-  and combinators for lifting operations on either component of the
-  disjoint sum.
+  A generic context is the disjoint sum of either a theory or proof context.
+  Occasionally, this enables uniform treatment of generic context data,
+  typically extra-logical information. Operations on generic contexts include
+  the usual injections, partial selections, and combinators for lifting
+  operations on either component of the disjoint sum.
 
-  Moreover, there are total operations \<open>theory_of\<close> and \<open>proof_of\<close> to convert a generic context into either kind: a theory
-  can always be selected from the sum, while a proof context might
-  have to be constructed by an ad-hoc \<open>init\<close> operation, which
-  incurs a small runtime overhead.
+  Moreover, there are total operations \<open>theory_of\<close> and \<open>proof_of\<close> to convert a
+  generic context into either kind: a theory can always be selected from the
+  sum, while a proof context might have to be constructed by an ad-hoc \<open>init\<close>
+  operation, which incurs a small runtime overhead.
 \<close>
 
 text %mlref \<open>
@@ -258,27 +242,27 @@
   @{index_ML Context.proof_of: "Context.generic -> Proof.context"} \\
   \end{mldecls}
 
-  \<^descr> Type @{ML_type Context.generic} is the direct sum of @{ML_type
-  "theory"} and @{ML_type "Proof.context"}, with the datatype
-  constructors @{ML "Context.Theory"} and @{ML "Context.Proof"}.
+  \<^descr> Type @{ML_type Context.generic} is the direct sum of @{ML_type "theory"}
+  and @{ML_type "Proof.context"}, with the datatype constructors @{ML
+  "Context.Theory"} and @{ML "Context.Proof"}.
 
-  \<^descr> @{ML Context.theory_of}~\<open>context\<close> always produces a
-  theory from the generic \<open>context\<close>, using @{ML
-  "Proof_Context.theory_of"} as required.
+  \<^descr> @{ML Context.theory_of}~\<open>context\<close> always produces a theory from the
+  generic \<open>context\<close>, using @{ML "Proof_Context.theory_of"} as required.
 
-  \<^descr> @{ML Context.proof_of}~\<open>context\<close> always produces a
-  proof context from the generic \<open>context\<close>, using @{ML
-  "Proof_Context.init_global"} as required (note that this re-initializes the
-  context data with each invocation).
+  \<^descr> @{ML Context.proof_of}~\<open>context\<close> always produces a proof context from the
+  generic \<open>context\<close>, using @{ML "Proof_Context.init_global"} as required (note
+  that this re-initializes the context data with each invocation).
 \<close>
 
 
 subsection \<open>Context data \label{sec:context-data}\<close>
 
-text \<open>The main purpose of theory and proof contexts is to manage
-  arbitrary (pure) data.  New data types can be declared incrementally
-  at compile time.  There are separate declaration mechanisms for any
-  of the three kinds of contexts: theory, proof, generic.\<close>
+text \<open>
+  The main purpose of theory and proof contexts is to manage arbitrary (pure)
+  data. New data types can be declared incrementally at compile time. There
+  are separate declaration mechanisms for any of the three kinds of contexts:
+  theory, proof, generic.
+\<close>
 
 paragraph \<open>Theory data\<close>
 text \<open>declarations need to implement the following ML signature:
@@ -292,18 +276,18 @@
   \end{tabular}
   \<^medskip>
 
-  The \<open>empty\<close> value acts as initial default for \<^emph>\<open>any\<close>
-  theory that does not declare actual data content; \<open>extend\<close>
-  is acts like a unitary version of \<open>merge\<close>.
+  The \<open>empty\<close> value acts as initial default for \<^emph>\<open>any\<close> theory that does not
+  declare actual data content; \<open>extend\<close> is acts like a unitary version of
+  \<open>merge\<close>.
 
-  Implementing \<open>merge\<close> can be tricky.  The general idea is
-  that \<open>merge (data\<^sub>1, data\<^sub>2)\<close> inserts those parts of \<open>data\<^sub>2\<close> into \<open>data\<^sub>1\<close> that are not yet present, while
-  keeping the general order of things.  The @{ML Library.merge}
+  Implementing \<open>merge\<close> can be tricky. The general idea is that \<open>merge (data\<^sub>1,
+  data\<^sub>2)\<close> inserts those parts of \<open>data\<^sub>2\<close> into \<open>data\<^sub>1\<close> that are not yet
+  present, while keeping the general order of things. The @{ML Library.merge}
   function on plain lists may serve as canonical template.
 
-  Particularly note that shared parts of the data must not be
-  duplicated by naive concatenation, or a theory graph that is like a
-  chain of diamonds would cause an exponential blowup!
+  Particularly note that shared parts of the data must not be duplicated by
+  naive concatenation, or a theory graph that is like a chain of diamonds
+  would cause an exponential blowup!
 \<close>
 
 paragraph \<open>Proof context data\<close>
@@ -316,26 +300,25 @@
   \end{tabular}
   \<^medskip>
 
-  The \<open>init\<close> operation is supposed to produce a pure value
-  from the given background theory and should be somehow
-  ``immediate''.  Whenever a proof context is initialized, which
-  happens frequently, the the system invokes the \<open>init\<close>
-  operation of \<^emph>\<open>all\<close> theory data slots ever declared.  This also
-  means that one needs to be economic about the total number of proof
-  data declarations in the system, i.e.\ each ML module should declare
-  at most one, sometimes two data slots for its internal use.
-  Repeated data declarations to simulate a record type should be
-  avoided!
+  The \<open>init\<close> operation is supposed to produce a pure value from the given
+  background theory and should be somehow ``immediate''. Whenever a proof
+  context is initialized, which happens frequently, the the system invokes the
+  \<open>init\<close> operation of \<^emph>\<open>all\<close> theory data slots ever declared. This also means
+  that one needs to be economic about the total number of proof data
+  declarations in the system, i.e.\ each ML module should declare at most one,
+  sometimes two data slots for its internal use. Repeated data declarations to
+  simulate a record type should be avoided!
 \<close>
 
 paragraph \<open>Generic data\<close>
-text \<open>provides a hybrid interface for both theory and proof data. The \<open>init\<close>
-  operation for proof contexts is predefined to select the current data
-  value from the background theory.
+text \<open>
+  provides a hybrid interface for both theory and proof data. The \<open>init\<close>
+  operation for proof contexts is predefined to select the current data value
+  from the background theory.
 
   \<^bigskip>
-  Any of the above data declarations over type \<open>T\<close>
-  result in an ML structure with the following signature:
+  Any of the above data declarations over type \<open>T\<close> result in an ML structure
+  with the following signature:
 
   \<^medskip>
   \begin{tabular}{ll}
@@ -345,12 +328,12 @@
   \end{tabular}
   \<^medskip>
 
-  These other operations provide exclusive access for the particular
-  kind of context (theory, proof, or generic context).  This interface
-  observes the ML discipline for types and scopes: there is no other
-  way to access the corresponding data slot of a context.  By keeping
-  these operations private, an Isabelle/ML module may maintain
-  abstract values authentically.\<close>
+  These other operations provide exclusive access for the particular kind of
+  context (theory, proof, or generic context). This interface observes the ML
+  discipline for types and scopes: there is no other way to access the
+  corresponding data slot of a context. By keeping these operations private,
+  an Isabelle/ML module may maintain abstract values authentically.
+\<close>
 
 text %mlref \<open>
   \begin{mldecls}
@@ -359,22 +342,20 @@
   @{index_ML_functor Generic_Data} \\
   \end{mldecls}
 
-  \<^descr> @{ML_functor Theory_Data}\<open>(spec)\<close> declares data for
-  type @{ML_type theory} according to the specification provided as
-  argument structure.  The resulting structure provides data init and
-  access operations as described above.
+  \<^descr> @{ML_functor Theory_Data}\<open>(spec)\<close> declares data for type @{ML_type theory}
+  according to the specification provided as argument structure. The resulting
+  structure provides data init and access operations as described above.
 
-  \<^descr> @{ML_functor Proof_Data}\<open>(spec)\<close> is analogous to
-  @{ML_functor Theory_Data} for type @{ML_type Proof.context}.
+  \<^descr> @{ML_functor Proof_Data}\<open>(spec)\<close> is analogous to @{ML_functor Theory_Data}
+  for type @{ML_type Proof.context}.
 
-  \<^descr> @{ML_functor Generic_Data}\<open>(spec)\<close> is analogous to
-  @{ML_functor Theory_Data} for type @{ML_type Context.generic}.
-\<close>
+  \<^descr> @{ML_functor Generic_Data}\<open>(spec)\<close> is analogous to @{ML_functor
+  Theory_Data} for type @{ML_type Context.generic}. \<close>
 
 text %mlex \<open>
-  The following artificial example demonstrates theory
-  data: we maintain a set of terms that are supposed to be wellformed
-  wrt.\ the enclosing theory.  The public interface is as follows:
+  The following artificial example demonstrates theory data: we maintain a set
+  of terms that are supposed to be wellformed wrt.\ the enclosing theory. The
+  public interface is as follows:
 \<close>
 
 ML \<open>
@@ -385,9 +366,10 @@
   end;
 \<close>
 
-text \<open>The implementation uses private theory data internally, and
-  only exposes an operation that involves explicit argument checking
-  wrt.\ the given theory.\<close>
+text \<open>
+  The implementation uses private theory data internally, and only exposes an
+  operation that involves explicit argument checking wrt.\ the given theory.
+\<close>
 
 ML \<open>
   structure Wellformed_Terms: WELLFORMED_TERMS =
@@ -414,59 +396,57 @@
   end;
 \<close>
 
-text \<open>Type @{ML_type "term Ord_List.T"} is used for reasonably
-  efficient representation of a set of terms: all operations are
-  linear in the number of stored elements.  Here we assume that users
-  of this module do not care about the declaration order, since that
-  data structure forces its own arrangement of elements.
+text \<open>
+  Type @{ML_type "term Ord_List.T"} is used for reasonably efficient
+  representation of a set of terms: all operations are linear in the number of
+  stored elements. Here we assume that users of this module do not care about
+  the declaration order, since that data structure forces its own arrangement
+  of elements.
 
-  Observe how the @{ML_text merge} operation joins the data slots of
-  the two constituents: @{ML Ord_List.union} prevents duplication of
-  common data from different branches, thus avoiding the danger of
-  exponential blowup.  Plain list append etc.\ must never be used for
-  theory data merges!
+  Observe how the @{ML_text merge} operation joins the data slots of the two
+  constituents: @{ML Ord_List.union} prevents duplication of common data from
+  different branches, thus avoiding the danger of exponential blowup. Plain
+  list append etc.\ must never be used for theory data merges!
 
   \<^medskip>
   Our intended invariant is achieved as follows:
 
-  \<^enum> @{ML Wellformed_Terms.add} only admits terms that have passed
-  the @{ML Sign.cert_term} check of the given theory at that point.
-
-  \<^enum> Wellformedness in the sense of @{ML Sign.cert_term} is
-  monotonic wrt.\ the sub-theory relation.  So our data can move
-  upwards in the hierarchy (via extension or merges), and maintain
-  wellformedness without further checks.
-
+    \<^enum> @{ML Wellformed_Terms.add} only admits terms that have passed the @{ML
+    Sign.cert_term} check of the given theory at that point.
+  
+    \<^enum> Wellformedness in the sense of @{ML Sign.cert_term} is monotonic wrt.\
+    the sub-theory relation. So our data can move upwards in the hierarchy
+    (via extension or merges), and maintain wellformedness without further
+    checks.
 
-  Note that all basic operations of the inference kernel (which
-  includes @{ML Sign.cert_term}) observe this monotonicity principle,
-  but other user-space tools don't.  For example, fully-featured
-  type-inference via @{ML Syntax.check_term} (cf.\
-  \secref{sec:term-check}) is not necessarily monotonic wrt.\ the
-  background theory, since constraints of term constants can be
-  modified by later declarations, for example.
+  Note that all basic operations of the inference kernel (which includes @{ML
+  Sign.cert_term}) observe this monotonicity principle, but other user-space
+  tools don't. For example, fully-featured type-inference via @{ML
+  Syntax.check_term} (cf.\ \secref{sec:term-check}) is not necessarily
+  monotonic wrt.\ the background theory, since constraints of term constants
+  can be modified by later declarations, for example.
 
-  In most cases, user-space context data does not have to take such
-  invariants too seriously.  The situation is different in the
-  implementation of the inference kernel itself, which uses the very
-  same data mechanisms for types, constants, axioms etc.
+  In most cases, user-space context data does not have to take such invariants
+  too seriously. The situation is different in the implementation of the
+  inference kernel itself, which uses the very same data mechanisms for types,
+  constants, axioms etc.
 \<close>
 
 
 subsection \<open>Configuration options \label{sec:config-options}\<close>
 
-text \<open>A \<^emph>\<open>configuration option\<close> is a named optional value of
-  some basic type (Boolean, integer, string) that is stored in the
-  context.  It is a simple application of general context data
-  (\secref{sec:context-data}) that is sufficiently common to justify
-  customized setup, which includes some concrete declarations for
-  end-users using existing notation for attributes (cf.\
-  \secref{sec:attributes}).
+text \<open>
+  A \<^emph>\<open>configuration option\<close> is a named optional value of some basic type
+  (Boolean, integer, string) that is stored in the context. It is a simple
+  application of general context data (\secref{sec:context-data}) that is
+  sufficiently common to justify customized setup, which includes some
+  concrete declarations for end-users using existing notation for attributes
+  (cf.\ \secref{sec:attributes}).
 
-  For example, the predefined configuration option @{attribute
-  show_types} controls output of explicit type constraints for
-  variables in printed terms (cf.\ \secref{sec:read-print}).  Its
-  value can be modified within Isar text like this:
+  For example, the predefined configuration option @{attribute show_types}
+  controls output of explicit type constraints for variables in printed terms
+  (cf.\ \secref{sec:read-print}). Its value can be modified within Isar text
+  like this:
 \<close>
 
 experiment
@@ -489,16 +469,18 @@
 
 end
 
-text \<open>Configuration options that are not set explicitly hold a
-  default value that can depend on the application context.  This
-  allows to retrieve the value from another slot within the context,
-  or fall back on a global preference mechanism, for example.
+text \<open>
+  Configuration options that are not set explicitly hold a default value that
+  can depend on the application context. This allows to retrieve the value
+  from another slot within the context, or fall back on a global preference
+  mechanism, for example.
 
-  The operations to declare configuration options and get/map their
-  values are modeled as direct replacements for historic global
-  references, only that the context is made explicit.  This allows
-  easy configuration of tools, without relying on the execution order
-  as required for old-style mutable references.\<close>
+  The operations to declare configuration options and get/map their values are
+  modeled as direct replacements for historic global references, only that the
+  context is made explicit. This allows easy configuration of tools, without
+  relying on the execution order as required for old-style mutable
+  references.
+\<close>
 
 text %mlref \<open>
   \begin{mldecls}
@@ -514,37 +496,38 @@
   string Config.T"} \\
   \end{mldecls}
 
-  \<^descr> @{ML Config.get}~\<open>ctxt config\<close> gets the value of
-  \<open>config\<close> in the given context.
+  \<^descr> @{ML Config.get}~\<open>ctxt config\<close> gets the value of \<open>config\<close> in the given
+  context.
 
-  \<^descr> @{ML Config.map}~\<open>config f ctxt\<close> updates the context
-  by updating the value of \<open>config\<close>.
+  \<^descr> @{ML Config.map}~\<open>config f ctxt\<close> updates the context by updating the value
+  of \<open>config\<close>.
 
-  \<^descr> \<open>config =\<close>~@{ML Attrib.setup_config_bool}~\<open>name
-  default\<close> creates a named configuration option of type @{ML_type
-  bool}, with the given \<open>default\<close> depending on the application
-  context.  The resulting \<open>config\<close> can be used to get/map its
-  value in a given context.  There is an implicit update of the
-  background theory that registers the option as attribute with some
-  concrete syntax.
+  \<^descr> \<open>config =\<close>~@{ML Attrib.setup_config_bool}~\<open>name default\<close> creates a named
+  configuration option of type @{ML_type bool}, with the given \<open>default\<close>
+  depending on the application context. The resulting \<open>config\<close> can be used to
+  get/map its value in a given context. There is an implicit update of the
+  background theory that registers the option as attribute with some concrete
+  syntax.
 
   \<^descr> @{ML Attrib.config_int}, @{ML Attrib.config_real}, and @{ML
-  Attrib.config_string} work like @{ML Attrib.config_bool}, but for
-  types @{ML_type int} and @{ML_type string}, respectively.
+  Attrib.config_string} work like @{ML Attrib.config_bool}, but for types
+  @{ML_type int} and @{ML_type string}, respectively.
 \<close>
 
-text %mlex \<open>The following example shows how to declare and use a
-  Boolean configuration option called \<open>my_flag\<close> with constant
-  default value @{ML false}.\<close>
+text %mlex \<open>
+  The following example shows how to declare and use a Boolean configuration
+  option called \<open>my_flag\<close> with constant default value @{ML false}.
+\<close>
 
 ML \<open>
   val my_flag =
     Attrib.setup_config_bool @{binding my_flag} (K false)
 \<close>
 
-text \<open>Now the user can refer to @{attribute my_flag} in
-  declarations, while ML tools can retrieve the current value from the
-  context via @{ML Config.get}.\<close>
+text \<open>
+  Now the user can refer to @{attribute my_flag} in declarations, while ML
+  tools can retrieve the current value from the context via @{ML Config.get}.
+\<close>
 
 ML_val \<open>@{assert} (Config.get @{context} my_flag = false)\<close>
 
@@ -561,8 +544,10 @@
   ML_val \<open>@{assert} (Config.get @{context} my_flag = true)\<close>
 end
 
-text \<open>Here is another example involving ML type @{ML_type real}
-  (floating-point numbers).\<close>
+text \<open>
+  Here is another example involving ML type @{ML_type real} (floating-point
+  numbers).
+\<close>
 
 ML \<open>
   val airspeed_velocity =
@@ -575,62 +560,58 @@
 
 section \<open>Names \label{sec:names}\<close>
 
-text \<open>In principle, a name is just a string, but there are various
-  conventions for representing additional structure.  For example,
-  ``\<open>Foo.bar.baz\<close>'' is considered as a long name consisting of
-  qualifier \<open>Foo.bar\<close> and base name \<open>baz\<close>.  The
-  individual constituents of a name may have further substructure,
-  e.g.\ the string ``\<^verbatim>\<open>\<alpha>\<close>'' encodes as a single
-  symbol (\secref{sec:symbols}).
+text \<open>
+  In principle, a name is just a string, but there are various conventions for
+  representing additional structure. For example, ``\<open>Foo.bar.baz\<close>'' is
+  considered as a long name consisting of qualifier \<open>Foo.bar\<close> and base name
+  \<open>baz\<close>. The individual constituents of a name may have further substructure,
+  e.g.\ the string ``\<^verbatim>\<open>\<alpha>\<close>'' encodes as a single symbol (\secref{sec:symbols}).
 
   \<^medskip>
-  Subsequently, we shall introduce specific categories of
-  names.  Roughly speaking these correspond to logical entities as
-  follows:
+  Subsequently, we shall introduce specific categories of names. Roughly
+  speaking these correspond to logical entities as follows:
 
-  \<^item> Basic names (\secref{sec:basic-name}): free and bound
-  variables.
+  \<^item> Basic names (\secref{sec:basic-name}): free and bound variables.
 
   \<^item> Indexed names (\secref{sec:indexname}): schematic variables.
 
-  \<^item> Long names (\secref{sec:long-name}): constants of any kind
-  (type constructors, term constants, other concepts defined in user
-  space).  Such entities are typically managed via name spaces
-  (\secref{sec:name-space}).
+  \<^item> Long names (\secref{sec:long-name}): constants of any kind (type
+  constructors, term constants, other concepts defined in user space). Such
+  entities are typically managed via name spaces (\secref{sec:name-space}).
 \<close>
 
 
 subsection \<open>Basic names \label{sec:basic-name}\<close>
 
 text \<open>
-  A \<^emph>\<open>basic name\<close> essentially consists of a single Isabelle
-  identifier.  There are conventions to mark separate classes of basic
-  names, by attaching a suffix of underscores: one underscore means
-  \<^emph>\<open>internal name\<close>, two underscores means \<^emph>\<open>Skolem name\<close>,
-  three underscores means \<^emph>\<open>internal Skolem name\<close>.
+  A \<^emph>\<open>basic name\<close> essentially consists of a single Isabelle identifier. There
+  are conventions to mark separate classes of basic names, by attaching a
+  suffix of underscores: one underscore means \<^emph>\<open>internal name\<close>, two
+  underscores means \<^emph>\<open>Skolem name\<close>, three underscores means \<^emph>\<open>internal Skolem
+  name\<close>.
 
-  For example, the basic name \<open>foo\<close> has the internal version
-  \<open>foo_\<close>, with Skolem versions \<open>foo__\<close> and \<open>foo___\<close>, respectively.
+  For example, the basic name \<open>foo\<close> has the internal version \<open>foo_\<close>, with
+  Skolem versions \<open>foo__\<close> and \<open>foo___\<close>, respectively.
 
-  These special versions provide copies of the basic name space, apart
-  from anything that normally appears in the user text.  For example,
-  system generated variables in Isar proof contexts are usually marked
-  as internal, which prevents mysterious names like \<open>xaa\<close> to
-  appear in human-readable text.
+  These special versions provide copies of the basic name space, apart from
+  anything that normally appears in the user text. For example, system
+  generated variables in Isar proof contexts are usually marked as internal,
+  which prevents mysterious names like \<open>xaa\<close> to appear in human-readable text.
 
   \<^medskip>
-  Manipulating binding scopes often requires on-the-fly
-  renamings.  A \<^emph>\<open>name context\<close> contains a collection of already
-  used names.  The \<open>declare\<close> operation adds names to the
-  context.
+  Manipulating binding scopes often requires on-the-fly renamings. A \<^emph>\<open>name
+  context\<close> contains a collection of already used names. The \<open>declare\<close>
+  operation adds names to the context.
 
-  The \<open>invents\<close> operation derives a number of fresh names from
-  a given starting point.  For example, the first three names derived
-  from \<open>a\<close> are \<open>a\<close>, \<open>b\<close>, \<open>c\<close>.
+  The \<open>invents\<close> operation derives a number of fresh names from a given
+  starting point. For example, the first three names derived from \<open>a\<close> are \<open>a\<close>,
+  \<open>b\<close>, \<open>c\<close>.
 
-  The \<open>variants\<close> operation produces fresh names by
-  incrementing tentative names as base-26 numbers (with digits \<open>a..z\<close>) until all clashes are resolved.  For example, name \<open>foo\<close> results in variants \<open>fooa\<close>, \<open>foob\<close>, \<open>fooc\<close>, \dots, \<open>fooaa\<close>, \<open>fooab\<close> etc.; each renaming
-  step picks the next unused variant from this sequence.
+  The \<open>variants\<close> operation produces fresh names by incrementing tentative
+  names as base-26 numbers (with digits \<open>a..z\<close>) until all clashes are
+  resolved. For example, name \<open>foo\<close> results in variants \<open>fooa\<close>, \<open>foob\<close>,
+  \<open>fooc\<close>, \dots, \<open>fooaa\<close>, \<open>fooab\<close> etc.; each renaming step picks the next
+  unused variant from this sequence.
 \<close>
 
 text %mlref \<open>
@@ -649,35 +630,35 @@
   @{index_ML Variable.names_of: "Proof.context -> Name.context"} \\
   \end{mldecls}
 
-  \<^descr> @{ML Name.internal}~\<open>name\<close> produces an internal name
-  by adding one underscore.
+  \<^descr> @{ML Name.internal}~\<open>name\<close> produces an internal name by adding one
+  underscore.
 
-  \<^descr> @{ML Name.skolem}~\<open>name\<close> produces a Skolem name by
-  adding two underscores.
+  \<^descr> @{ML Name.skolem}~\<open>name\<close> produces a Skolem name by adding two underscores.
 
-  \<^descr> Type @{ML_type Name.context} represents the context of already
-  used names; the initial value is @{ML "Name.context"}.
+  \<^descr> Type @{ML_type Name.context} represents the context of already used names;
+  the initial value is @{ML "Name.context"}.
 
-  \<^descr> @{ML Name.declare}~\<open>name\<close> enters a used name into the
-  context.
+  \<^descr> @{ML Name.declare}~\<open>name\<close> enters a used name into the context.
 
-  \<^descr> @{ML Name.invent}~\<open>context name n\<close> produces \<open>n\<close> fresh names derived from \<open>name\<close>.
+  \<^descr> @{ML Name.invent}~\<open>context name n\<close> produces \<open>n\<close> fresh names derived from
+  \<open>name\<close>.
 
-  \<^descr> @{ML Name.variant}~\<open>name context\<close> produces a fresh
-  variant of \<open>name\<close>; the result is declared to the context.
+  \<^descr> @{ML Name.variant}~\<open>name context\<close> produces a fresh variant of \<open>name\<close>; the
+  result is declared to the context.
 
-  \<^descr> @{ML Variable.names_of}~\<open>ctxt\<close> retrieves the context
-  of declared type and term variable names.  Projecting a proof
-  context down to a primitive name context is occasionally useful when
-  invoking lower-level operations.  Regular management of ``fresh
-  variables'' is done by suitable operations of structure @{ML_structure
-  Variable}, which is also able to provide an official status of
-  ``locally fixed variable'' within the logical environment (cf.\
-  \secref{sec:variables}).
+  \<^descr> @{ML Variable.names_of}~\<open>ctxt\<close> retrieves the context of declared type and
+  term variable names. Projecting a proof context down to a primitive name
+  context is occasionally useful when invoking lower-level operations. Regular
+  management of ``fresh variables'' is done by suitable operations of
+  structure @{ML_structure Variable}, which is also able to provide an
+  official status of ``locally fixed variable'' within the logical environment
+  (cf.\ \secref{sec:variables}).
 \<close>
 
-text %mlex \<open>The following simple examples demonstrate how to produce
-  fresh names from the initial @{ML Name.context}.\<close>
+text %mlex \<open>
+  The following simple examples demonstrate how to produce fresh names from
+  the initial @{ML Name.context}.
+\<close>
 
 ML_val \<open>
   val list1 = Name.invent Name.context "a" 5;
@@ -712,36 +693,32 @@
 subsection \<open>Indexed names \label{sec:indexname}\<close>
 
 text \<open>
-  An \<^emph>\<open>indexed name\<close> (or \<open>indexname\<close>) is a pair of a basic
-  name and a natural number.  This representation allows efficient
-  renaming by incrementing the second component only.  The canonical
-  way to rename two collections of indexnames apart from each other is
-  this: determine the maximum index \<open>maxidx\<close> of the first
-  collection, then increment all indexes of the second collection by
-  \<open>maxidx + 1\<close>; the maximum index of an empty collection is
+  An \<^emph>\<open>indexed name\<close> (or \<open>indexname\<close>) is a pair of a basic name and a natural
+  number. This representation allows efficient renaming by incrementing the
+  second component only. The canonical way to rename two collections of
+  indexnames apart from each other is this: determine the maximum index
+  \<open>maxidx\<close> of the first collection, then increment all indexes of the second
+  collection by \<open>maxidx + 1\<close>; the maximum index of an empty collection is
   \<open>-1\<close>.
 
-  Occasionally, basic names are injected into the same pair type of
-  indexed names: then \<open>(x, -1)\<close> is used to encode the basic
-  name \<open>x\<close>.
+  Occasionally, basic names are injected into the same pair type of indexed
+  names: then \<open>(x, -1)\<close> is used to encode the basic name \<open>x\<close>.
 
   \<^medskip>
-  Isabelle syntax observes the following rules for
-  representing an indexname \<open>(x, i)\<close> as a packed string:
+  Isabelle syntax observes the following rules for representing an indexname
+  \<open>(x, i)\<close> as a packed string:
 
-  \<^item> \<open>?x\<close> if \<open>x\<close> does not end with a digit and \<open>i = 0\<close>,
+    \<^item> \<open>?x\<close> if \<open>x\<close> does not end with a digit and \<open>i = 0\<close>,
 
-  \<^item> \<open>?xi\<close> if \<open>x\<close> does not end with a digit,
+    \<^item> \<open>?xi\<close> if \<open>x\<close> does not end with a digit,
 
-  \<^item> \<open>?x.i\<close> otherwise.
-
+    \<^item> \<open>?x.i\<close> otherwise.
 
-  Indexnames may acquire large index numbers after several maxidx
-  shifts have been applied.  Results are usually normalized towards
-  \<open>0\<close> at certain checkpoints, notably at the end of a proof.
-  This works by producing variants of the corresponding basic name
-  components.  For example, the collection \<open>?x1, ?x7, ?x42\<close>
-  becomes \<open>?x, ?xa, ?xb\<close>.
+  Indexnames may acquire large index numbers after several maxidx shifts have
+  been applied. Results are usually normalized towards \<open>0\<close> at certain
+  checkpoints, notably at the end of a proof. This works by producing variants
+  of the corresponding basic name components. For example, the collection
+  \<open>?x1, ?x7, ?x42\<close> becomes \<open>?x, ?xa, ?xb\<close>.
 \<close>
 
 text %mlref \<open>
@@ -749,36 +726,34 @@
   @{index_ML_type indexname: "string * int"} \\
   \end{mldecls}
 
-  \<^descr> Type @{ML_type indexname} represents indexed names.  This is
-  an abbreviation for @{ML_type "string * int"}.  The second component
-  is usually non-negative, except for situations where \<open>(x,
-  -1)\<close> is used to inject basic names into this type.  Other negative
-  indexes should not be used.
+  \<^descr> Type @{ML_type indexname} represents indexed names. This is an
+  abbreviation for @{ML_type "string * int"}. The second component is usually
+  non-negative, except for situations where \<open>(x, -1)\<close> is used to inject basic
+  names into this type. Other negative indexes should not be used.
 \<close>
 
 
 subsection \<open>Long names \label{sec:long-name}\<close>
 
-text \<open>A \<^emph>\<open>long name\<close> consists of a sequence of non-empty name
-  components.  The packed representation uses a dot as separator, as
-  in ``\<open>A.b.c\<close>''.  The last component is called \<^emph>\<open>base
-  name\<close>, the remaining prefix is called \<^emph>\<open>qualifier\<close> (which may be
-  empty).  The qualifier can be understood as the access path to the
-  named entity while passing through some nested block-structure,
-  although our free-form long names do not really enforce any strict
-  discipline.
+text \<open>
+  A \<^emph>\<open>long name\<close> consists of a sequence of non-empty name components. The
+  packed representation uses a dot as separator, as in ``\<open>A.b.c\<close>''. The last
+  component is called \<^emph>\<open>base name\<close>, the remaining prefix is called
+  \<^emph>\<open>qualifier\<close> (which may be empty). The qualifier can be understood as the
+  access path to the named entity while passing through some nested
+  block-structure, although our free-form long names do not really enforce any
+  strict discipline.
 
-  For example, an item named ``\<open>A.b.c\<close>'' may be understood as
-  a local entity \<open>c\<close>, within a local structure \<open>b\<close>,
-  within a global structure \<open>A\<close>.  In practice, long names
-  usually represent 1--3 levels of qualification.  User ML code should
-  not make any assumptions about the particular structure of long
+  For example, an item named ``\<open>A.b.c\<close>'' may be understood as a local entity
+  \<open>c\<close>, within a local structure \<open>b\<close>, within a global structure \<open>A\<close>. In
+  practice, long names usually represent 1--3 levels of qualification. User ML
+  code should not make any assumptions about the particular structure of long
   names!
 
-  The empty name is commonly used as an indication of unnamed
-  entities, or entities that are not entered into the corresponding
-  name space, whenever this makes any sense.  The basic operations on
-  long names map empty names again to empty names.
+  The empty name is commonly used as an indication of unnamed entities, or
+  entities that are not entered into the corresponding name space, whenever
+  this makes any sense. The basic operations on long names map empty names
+  again to empty names.
 \<close>
 
 text %mlref \<open>
@@ -790,45 +765,42 @@
   @{index_ML Long_Name.explode: "string -> string list"} \\
   \end{mldecls}
 
-  \<^descr> @{ML Long_Name.base_name}~\<open>name\<close> returns the base name
-  of a long name.
+  \<^descr> @{ML Long_Name.base_name}~\<open>name\<close> returns the base name of a long name.
 
-  \<^descr> @{ML Long_Name.qualifier}~\<open>name\<close> returns the qualifier
-  of a long name.
+  \<^descr> @{ML Long_Name.qualifier}~\<open>name\<close> returns the qualifier of a long name.
 
-  \<^descr> @{ML Long_Name.append}~\<open>name\<^sub>1 name\<^sub>2\<close> appends two long
-  names.
+  \<^descr> @{ML Long_Name.append}~\<open>name\<^sub>1 name\<^sub>2\<close> appends two long names.
 
-  \<^descr> @{ML Long_Name.implode}~\<open>names\<close> and @{ML
-  Long_Name.explode}~\<open>name\<close> convert between the packed string
-  representation and the explicit list form of long names.
+  \<^descr> @{ML Long_Name.implode}~\<open>names\<close> and @{ML Long_Name.explode}~\<open>name\<close> convert
+  between the packed string representation and the explicit list form of long
+  names.
 \<close>
 
 
 subsection \<open>Name spaces \label{sec:name-space}\<close>
 
-text \<open>A \<open>name space\<close> manages a collection of long names,
-  together with a mapping between partially qualified external names
-  and fully qualified internal names (in both directions).  Note that
-  the corresponding \<open>intern\<close> and \<open>extern\<close> operations
-  are mostly used for parsing and printing only!  The \<open>declare\<close> operation augments a name space according to the accesses
+text \<open>
+  A \<open>name space\<close> manages a collection of long names, together with a mapping
+  between partially qualified external names and fully qualified internal
+  names (in both directions). Note that the corresponding \<open>intern\<close> and
+  \<open>extern\<close> operations are mostly used for parsing and printing only! The
+  \<open>declare\<close> operation augments a name space according to the accesses
   determined by a given binding, and a naming policy from the context.
 
   \<^medskip>
-  A \<open>binding\<close> specifies details about the prospective
-  long name of a newly introduced formal entity.  It consists of a
-  base name, prefixes for qualification (separate ones for system
-  infrastructure and user-space mechanisms), a slot for the original
-  source position, and some additional flags.
+  A \<open>binding\<close> specifies details about the prospective long name of a newly
+  introduced formal entity. It consists of a base name, prefixes for
+  qualification (separate ones for system infrastructure and user-space
+  mechanisms), a slot for the original source position, and some additional
+  flags.
 
   \<^medskip>
-  A \<open>naming\<close> provides some additional details for
-  producing a long name from a binding.  Normally, the naming is
-  implicit in the theory or proof context.  The \<open>full\<close>
-  operation (and its variants for different context types) produces a
-  fully qualified internal name to be entered into a name space.  The
-  main equation of this ``chemical reaction'' when binding new
-  entities in a context is as follows:
+  A \<open>naming\<close> provides some additional details for producing a long name from a
+  binding. Normally, the naming is implicit in the theory or proof context.
+  The \<open>full\<close> operation (and its variants for different context types) produces
+  a fully qualified internal name to be entered into a name space. The main
+  equation of this ``chemical reaction'' when binding new entities in a
+  context is as follows:
 
   \<^medskip>
   \begin{tabular}{l}
@@ -836,21 +808,20 @@
   \end{tabular}
 
   \<^bigskip>
-  As a general principle, there is a separate name space for
-  each kind of formal entity, e.g.\ fact, logical constant, type
-  constructor, type class.  It is usually clear from the occurrence in
-  concrete syntax (or from the scope) which kind of entity a name
-  refers to.  For example, the very same name \<open>c\<close> may be used
-  uniformly for a constant, type constructor, and type class.
+  As a general principle, there is a separate name space for each kind of
+  formal entity, e.g.\ fact, logical constant, type constructor, type class.
+  It is usually clear from the occurrence in concrete syntax (or from the
+  scope) which kind of entity a name refers to. For example, the very same
+  name \<open>c\<close> may be used uniformly for a constant, type constructor, and type
+  class.
 
-  There are common schemes to name derived entities systematically
-  according to the name of the main logical entity involved, e.g.\
-  fact \<open>c.intro\<close> for a canonical introduction rule related to
-  constant \<open>c\<close>.  This technique of mapping names from one
-  space into another requires some care in order to avoid conflicts.
-  In particular, theorem names derived from a type constructor or type
-  class should get an additional suffix in addition to the usual
-  qualification.  This leads to the following conventions for derived
+  There are common schemes to name derived entities systematically according
+  to the name of the main logical entity involved, e.g.\ fact \<open>c.intro\<close> for a
+  canonical introduction rule related to constant \<open>c\<close>. This technique of
+  mapping names from one space into another requires some care in order to
+  avoid conflicts. In particular, theorem names derived from a type
+  constructor or type class should get an additional suffix in addition to the
+  usual qualification. This leads to the following conventions for derived
   names:
 
   \<^medskip>
@@ -889,81 +860,76 @@
   @{index_ML Name_Space.is_concealed: "Name_Space.T -> string -> bool"}
   \end{mldecls}
 
-  \<^descr> Type @{ML_type binding} represents the abstract concept of
-  name bindings.
+  \<^descr> Type @{ML_type binding} represents the abstract concept of name bindings.
 
   \<^descr> @{ML Binding.empty} is the empty binding.
 
-  \<^descr> @{ML Binding.name}~\<open>name\<close> produces a binding with base
-  name \<open>name\<close>.  Note that this lacks proper source position
-  information; see also the ML antiquotation @{ML_antiquotation
-  binding}.
+  \<^descr> @{ML Binding.name}~\<open>name\<close> produces a binding with base name \<open>name\<close>. Note
+  that this lacks proper source position information; see also the ML
+  antiquotation @{ML_antiquotation binding}.
 
-  \<^descr> @{ML Binding.qualify}~\<open>mandatory name binding\<close>
-  prefixes qualifier \<open>name\<close> to \<open>binding\<close>.  The \<open>mandatory\<close> flag tells if this name component always needs to be
-  given in name space accesses --- this is mostly \<open>false\<close> in
-  practice.  Note that this part of qualification is typically used in
-  derived specification mechanisms.
+  \<^descr> @{ML Binding.qualify}~\<open>mandatory name binding\<close> prefixes qualifier \<open>name\<close>
+  to \<open>binding\<close>. The \<open>mandatory\<close> flag tells if this name component always needs
+  to be given in name space accesses --- this is mostly \<open>false\<close> in practice.
+  Note that this part of qualification is typically used in derived
+  specification mechanisms.
 
-  \<^descr> @{ML Binding.prefix} is similar to @{ML Binding.qualify}, but
-  affects the system prefix.  This part of extra qualification is
-  typically used in the infrastructure for modular specifications,
-  notably ``local theory targets'' (see also \chref{ch:local-theory}).
+  \<^descr> @{ML Binding.prefix} is similar to @{ML Binding.qualify}, but affects the
+  system prefix. This part of extra qualification is typically used in the
+  infrastructure for modular specifications, notably ``local theory targets''
+  (see also \chref{ch:local-theory}).
 
-  \<^descr> @{ML Binding.concealed}~\<open>binding\<close> indicates that the
-  binding shall refer to an entity that serves foundational purposes
-  only.  This flag helps to mark implementation details of
-  specification mechanism etc.  Other tools should not depend on the
-  particulars of concealed entities (cf.\ @{ML
+  \<^descr> @{ML Binding.concealed}~\<open>binding\<close> indicates that the binding shall refer
+  to an entity that serves foundational purposes only. This flag helps to mark
+  implementation details of specification mechanism etc. Other tools should
+  not depend on the particulars of concealed entities (cf.\ @{ML
   Name_Space.is_concealed}).
 
-  \<^descr> @{ML Binding.print}~\<open>binding\<close> produces a string
-  representation for human-readable output, together with some formal
-  markup that might get used in GUI front-ends, for example.
+  \<^descr> @{ML Binding.print}~\<open>binding\<close> produces a string representation for
+  human-readable output, together with some formal markup that might get used
+  in GUI front-ends, for example.
 
-  \<^descr> Type @{ML_type Name_Space.naming} represents the abstract
-  concept of a naming policy.
+  \<^descr> Type @{ML_type Name_Space.naming} represents the abstract concept of a
+  naming policy.
 
-  \<^descr> @{ML Name_Space.global_naming} is the default naming policy: it is
-  global and lacks any path prefix.  In a regular theory context this is
-  augmented by a path prefix consisting of the theory name.
+  \<^descr> @{ML Name_Space.global_naming} is the default naming policy: it is global
+  and lacks any path prefix. In a regular theory context this is augmented by
+  a path prefix consisting of the theory name.
 
-  \<^descr> @{ML Name_Space.add_path}~\<open>path naming\<close> augments the
-  naming policy by extending its path component.
+  \<^descr> @{ML Name_Space.add_path}~\<open>path naming\<close> augments the naming policy by
+  extending its path component.
 
-  \<^descr> @{ML Name_Space.full_name}~\<open>naming binding\<close> turns a
-  name binding (usually a basic name) into the fully qualified
-  internal name, according to the given naming policy.
+  \<^descr> @{ML Name_Space.full_name}~\<open>naming binding\<close> turns a name binding (usually
+  a basic name) into the fully qualified internal name, according to the given
+  naming policy.
 
   \<^descr> Type @{ML_type Name_Space.T} represents name spaces.
 
-  \<^descr> @{ML Name_Space.empty}~\<open>kind\<close> and @{ML Name_Space.merge}~\<open>(space\<^sub>1, space\<^sub>2)\<close> are the canonical operations for
-  maintaining name spaces according to theory data management
-  (\secref{sec:context-data}); \<open>kind\<close> is a formal comment
-  to characterize the purpose of a name space.
+  \<^descr> @{ML Name_Space.empty}~\<open>kind\<close> and @{ML Name_Space.merge}~\<open>(space\<^sub>1,
+  space\<^sub>2)\<close> are the canonical operations for maintaining name spaces according
+  to theory data management (\secref{sec:context-data}); \<open>kind\<close> is a formal
+  comment to characterize the purpose of a name space.
 
-  \<^descr> @{ML Name_Space.declare}~\<open>context strict binding
-  space\<close> enters a name binding as fully qualified internal name into
-  the name space, using the naming of the context.
+  \<^descr> @{ML Name_Space.declare}~\<open>context strict binding space\<close> enters a name
+  binding as fully qualified internal name into the name space, using the
+  naming of the context.
 
-  \<^descr> @{ML Name_Space.intern}~\<open>space name\<close> internalizes a
-  (partially qualified) external name.
+  \<^descr> @{ML Name_Space.intern}~\<open>space name\<close> internalizes a (partially qualified)
+  external name.
 
-  This operation is mostly for parsing!  Note that fully qualified
-  names stemming from declarations are produced via @{ML
-  "Name_Space.full_name"} and @{ML "Name_Space.declare"}
-  (or their derivatives for @{ML_type theory} and
+  This operation is mostly for parsing! Note that fully qualified names
+  stemming from declarations are produced via @{ML "Name_Space.full_name"} and
+  @{ML "Name_Space.declare"} (or their derivatives for @{ML_type theory} and
   @{ML_type Proof.context}).
 
-  \<^descr> @{ML Name_Space.extern}~\<open>ctxt space name\<close> externalizes a
-  (fully qualified) internal name.
+  \<^descr> @{ML Name_Space.extern}~\<open>ctxt space name\<close> externalizes a (fully qualified)
+  internal name.
 
-  This operation is mostly for printing!  User code should not rely on
-  the precise result too much.
+  This operation is mostly for printing! User code should not rely on the
+  precise result too much.
 
-  \<^descr> @{ML Name_Space.is_concealed}~\<open>space name\<close> indicates
-  whether \<open>name\<close> refers to a strictly private entity that
-  other tools are supposed to ignore!
+  \<^descr> @{ML Name_Space.is_concealed}~\<open>space name\<close> indicates whether \<open>name\<close> refers
+  to a strictly private entity that other tools are supposed to ignore!
 \<close>
 
 text %mlantiq \<open>
@@ -975,35 +941,37 @@
   @@{ML_antiquotation binding} name
   \<close>}
 
-  \<^descr> \<open>@{binding name}\<close> produces a binding with base name
-  \<open>name\<close> and the source position taken from the concrete
-  syntax of this antiquotation.  In many situations this is more
-  appropriate than the more basic @{ML Binding.name} function.
+  \<^descr> \<open>@{binding name}\<close> produces a binding with base name \<open>name\<close> and the source
+  position taken from the concrete syntax of this antiquotation. In many
+  situations this is more appropriate than the more basic @{ML Binding.name}
+  function.
 \<close>
 
-text %mlex \<open>The following example yields the source position of some
-  concrete binding inlined into the text:
+text %mlex \<open>
+  The following example yields the source position of some concrete binding
+  inlined into the text:
 \<close>
 
 ML_val \<open>Binding.pos_of @{binding here}\<close>
 
 text \<open>
   \<^medskip>
-  That position can be also printed in a message as follows:\<close>
+  That position can be also printed in a message as follows:
+\<close>
 
 ML_command
   \<open>writeln
     ("Look here" ^ Position.here (Binding.pos_of @{binding here}))\<close>
 
-text \<open>This illustrates a key virtue of formalized bindings as
-  opposed to raw specifications of base names: the system can use this
-  additional information for feedback given to the user (error
-  messages etc.).
+text \<open>
+  This illustrates a key virtue of formalized bindings as opposed to raw
+  specifications of base names: the system can use this additional information
+  for feedback given to the user (error messages etc.).
 
   \<^medskip>
-  The following example refers to its source position
-  directly, which is occasionally useful for experimentation and
-  diagnostic purposes:\<close>
+  The following example refers to its source position directly, which is
+  occasionally useful for experimentation and diagnostic purposes:
+\<close>
 
 ML_command \<open>warning ("Look here" ^ Position.here @{here})\<close>
 
--- a/src/Doc/Implementation/Proof.thy	Tue Dec 15 14:41:47 2015 +0000
+++ b/src/Doc/Implementation/Proof.thy	Wed Dec 16 17:30:30 2015 +0100
@@ -9,52 +9,48 @@
 section \<open>Variables \label{sec:variables}\<close>
 
 text \<open>
-  Any variable that is not explicitly bound by \<open>\<lambda>\<close>-abstraction
-  is considered as ``free''.  Logically, free variables act like
-  outermost universal quantification at the sequent level: \<open>A\<^sub>1(x), \<dots>, A\<^sub>n(x) \<turnstile> B(x)\<close> means that the result
-  holds \<^emph>\<open>for all\<close> values of \<open>x\<close>.  Free variables for
-  terms (not types) can be fully internalized into the logic: \<open>\<turnstile> B(x)\<close> and \<open>\<turnstile> \<And>x. B(x)\<close> are interchangeable, provided
-  that \<open>x\<close> does not occur elsewhere in the context.
-  Inspecting \<open>\<turnstile> \<And>x. B(x)\<close> more closely, we see that inside the
-  quantifier, \<open>x\<close> is essentially ``arbitrary, but fixed'',
-  while from outside it appears as a place-holder for instantiation
-  (thanks to \<open>\<And>\<close> elimination).
+  Any variable that is not explicitly bound by \<open>\<lambda>\<close>-abstraction is considered
+  as ``free''. Logically, free variables act like outermost universal
+  quantification at the sequent level: \<open>A\<^sub>1(x), \<dots>, A\<^sub>n(x) \<turnstile> B(x)\<close> means that
+  the result holds \<^emph>\<open>for all\<close> values of \<open>x\<close>. Free variables for terms (not
+  types) can be fully internalized into the logic: \<open>\<turnstile> B(x)\<close> and \<open>\<turnstile> \<And>x. B(x)\<close>
+  are interchangeable, provided that \<open>x\<close> does not occur elsewhere in the
+  context. Inspecting \<open>\<turnstile> \<And>x. B(x)\<close> more closely, we see that inside the
+  quantifier, \<open>x\<close> is essentially ``arbitrary, but fixed'', while from outside
+  it appears as a place-holder for instantiation (thanks to \<open>\<And>\<close> elimination).
 
-  The Pure logic represents the idea of variables being either inside
-  or outside the current scope by providing separate syntactic
-  categories for \<^emph>\<open>fixed variables\<close> (e.g.\ \<open>x\<close>) vs.\
-  \<^emph>\<open>schematic variables\<close> (e.g.\ \<open>?x\<close>).  Incidently, a
-  universal result \<open>\<turnstile> \<And>x. B(x)\<close> has the HHF normal form \<open>\<turnstile> B(?x)\<close>, which represents its generality without requiring an
-  explicit quantifier.  The same principle works for type variables:
-  \<open>\<turnstile> B(?\<alpha>)\<close> represents the idea of ``\<open>\<turnstile> \<forall>\<alpha>. B(\<alpha>)\<close>''
-  without demanding a truly polymorphic framework.
+  The Pure logic represents the idea of variables being either inside or
+  outside the current scope by providing separate syntactic categories for
+  \<^emph>\<open>fixed variables\<close> (e.g.\ \<open>x\<close>) vs.\ \<^emph>\<open>schematic variables\<close> (e.g.\ \<open>?x\<close>).
+  Incidently, a universal result \<open>\<turnstile> \<And>x. B(x)\<close> has the HHF normal form \<open>\<turnstile>
+  B(?x)\<close>, which represents its generality without requiring an explicit
+  quantifier. The same principle works for type variables: \<open>\<turnstile> B(?\<alpha>)\<close>
+  represents the idea of ``\<open>\<turnstile> \<forall>\<alpha>. B(\<alpha>)\<close>'' without demanding a truly
+  polymorphic framework.
 
   \<^medskip>
-  Additional care is required to treat type variables in a
-  way that facilitates type-inference.  In principle, term variables
-  depend on type variables, which means that type variables would have
-  to be declared first.  For example, a raw type-theoretic framework
-  would demand the context to be constructed in stages as follows:
-  \<open>\<Gamma> = \<alpha>: type, x: \<alpha>, a: A(x\<^sub>\<alpha>)\<close>.
+  Additional care is required to treat type variables in a way that
+  facilitates type-inference. In principle, term variables depend on type
+  variables, which means that type variables would have to be declared first.
+  For example, a raw type-theoretic framework would demand the context to be
+  constructed in stages as follows: \<open>\<Gamma> = \<alpha>: type, x: \<alpha>, a: A(x\<^sub>\<alpha>)\<close>.
 
-  We allow a slightly less formalistic mode of operation: term
-  variables \<open>x\<close> are fixed without specifying a type yet
-  (essentially \<^emph>\<open>all\<close> potential occurrences of some instance
-  \<open>x\<^sub>\<tau>\<close> are fixed); the first occurrence of \<open>x\<close>
+  We allow a slightly less formalistic mode of operation: term variables \<open>x\<close>
+  are fixed without specifying a type yet (essentially \<^emph>\<open>all\<close> potential
+  occurrences of some instance \<open>x\<^sub>\<tau>\<close> are fixed); the first occurrence of \<open>x\<close>
   within a specific term assigns its most general type, which is then
-  maintained consistently in the context.  The above example becomes
-  \<open>\<Gamma> = x: term, \<alpha>: type, A(x\<^sub>\<alpha>)\<close>, where type \<open>\<alpha>\<close> is fixed \<^emph>\<open>after\<close> term \<open>x\<close>, and the constraint
-  \<open>x :: \<alpha>\<close> is an implicit consequence of the occurrence of
-  \<open>x\<^sub>\<alpha>\<close> in the subsequent proposition.
+  maintained consistently in the context. The above example becomes \<open>\<Gamma> = x:
+  term, \<alpha>: type, A(x\<^sub>\<alpha>)\<close>, where type \<open>\<alpha>\<close> is fixed \<^emph>\<open>after\<close> term \<open>x\<close>, and the
+  constraint \<open>x :: \<alpha>\<close> is an implicit consequence of the occurrence of \<open>x\<^sub>\<alpha>\<close> in
+  the subsequent proposition.
 
-  This twist of dependencies is also accommodated by the reverse
-  operation of exporting results from a context: a type variable
-  \<open>\<alpha>\<close> is considered fixed as long as it occurs in some fixed
-  term variable of the context.  For example, exporting \<open>x:
-  term, \<alpha>: type \<turnstile> x\<^sub>\<alpha> \<equiv> x\<^sub>\<alpha>\<close> produces in the first step \<open>x: term
-  \<turnstile> x\<^sub>\<alpha> \<equiv> x\<^sub>\<alpha>\<close> for fixed \<open>\<alpha>\<close>, and only in the second step
-  \<open>\<turnstile> ?x\<^sub>?\<^sub>\<alpha> \<equiv> ?x\<^sub>?\<^sub>\<alpha>\<close> for schematic \<open>?x\<close> and \<open>?\<alpha>\<close>.
-  The following Isar source text illustrates this scenario.
+  This twist of dependencies is also accommodated by the reverse operation of
+  exporting results from a context: a type variable \<open>\<alpha>\<close> is considered fixed as
+  long as it occurs in some fixed term variable of the context. For example,
+  exporting \<open>x: term, \<alpha>: type \<turnstile> x\<^sub>\<alpha> \<equiv> x\<^sub>\<alpha>\<close> produces in the first step \<open>x: term
+  \<turnstile> x\<^sub>\<alpha> \<equiv> x\<^sub>\<alpha>\<close> for fixed \<open>\<alpha>\<close>, and only in the second step \<open>\<turnstile> ?x\<^sub>?\<^sub>\<alpha> \<equiv> ?x\<^sub>?\<^sub>\<alpha>\<close>
+  for schematic \<open>?x\<close> and \<open>?\<alpha>\<close>. The following Isar source text illustrates this
+  scenario.
 \<close>
 
 notepad
@@ -70,29 +66,28 @@
   thm this  \<comment> \<open>fully general result for arbitrary \<open>?x::?'a\<close>\<close>
 end
 
-text \<open>The Isabelle/Isar proof context manages the details of term
-  vs.\ type variables, with high-level principles for moving the
-  frontier between fixed and schematic variables.
+text \<open>
+  The Isabelle/Isar proof context manages the details of term vs.\ type
+  variables, with high-level principles for moving the frontier between fixed
+  and schematic variables.
 
-  The \<open>add_fixes\<close> operation explicitly declares fixed
-  variables; the \<open>declare_term\<close> operation absorbs a term into
-  a context by fixing new type variables and adding syntactic
-  constraints.
+  The \<open>add_fixes\<close> operation explicitly declares fixed variables; the
+  \<open>declare_term\<close> operation absorbs a term into a context by fixing new type
+  variables and adding syntactic constraints.
 
-  The \<open>export\<close> operation is able to perform the main work of
-  generalizing term and type variables as sketched above, assuming
-  that fixing variables and terms have been declared properly.
+  The \<open>export\<close> operation is able to perform the main work of generalizing term
+  and type variables as sketched above, assuming that fixing variables and
+  terms have been declared properly.
 
-  There \<open>import\<close> operation makes a generalized fact a genuine
-  part of the context, by inventing fixed variables for the schematic
-  ones.  The effect can be reversed by using \<open>export\<close> later,
-  potentially with an extended context; the result is equivalent to
-  the original modulo renaming of schematic variables.
+  There \<open>import\<close> operation makes a generalized fact a genuine part of the
+  context, by inventing fixed variables for the schematic ones. The effect can
+  be reversed by using \<open>export\<close> later, potentially with an extended context;
+  the result is equivalent to the original modulo renaming of schematic
+  variables.
 
-  The \<open>focus\<close> operation provides a variant of \<open>import\<close>
-  for nested propositions (with explicit quantification): \<open>\<And>x\<^sub>1 \<dots> x\<^sub>n. B(x\<^sub>1, \<dots>, x\<^sub>n)\<close> is
-  decomposed by inventing fixed variables \<open>x\<^sub>1, \<dots>,
-  x\<^sub>n\<close> for the body.
+  The \<open>focus\<close> operation provides a variant of \<open>import\<close> for nested propositions
+  (with explicit quantification): \<open>\<And>x\<^sub>1 \<dots> x\<^sub>n. B(x\<^sub>1, \<dots>, x\<^sub>n)\<close> is decomposed
+  by inventing fixed variables \<open>x\<^sub>1, \<dots>, x\<^sub>n\<close> for the body.
 \<close>
 
 text %mlref \<open>
@@ -111,48 +106,47 @@
   ((string * (string * typ)) list * term) * Proof.context"} \\
   \end{mldecls}
 
-  \<^descr> @{ML Variable.add_fixes}~\<open>xs ctxt\<close> fixes term
-  variables \<open>xs\<close>, returning the resulting internal names.  By
-  default, the internal representation coincides with the external
-  one, which also means that the given variables must not be fixed
-  already.  There is a different policy within a local proof body: the
-  given names are just hints for newly invented Skolem variables.
+  \<^descr> @{ML Variable.add_fixes}~\<open>xs ctxt\<close> fixes term variables \<open>xs\<close>, returning
+  the resulting internal names. By default, the internal representation
+  coincides with the external one, which also means that the given variables
+  must not be fixed already. There is a different policy within a local proof
+  body: the given names are just hints for newly invented Skolem variables.
 
-  \<^descr> @{ML Variable.variant_fixes} is similar to @{ML
-  Variable.add_fixes}, but always produces fresh variants of the given
-  names.
+  \<^descr> @{ML Variable.variant_fixes} is similar to @{ML Variable.add_fixes}, but
+  always produces fresh variants of the given names.
 
-  \<^descr> @{ML Variable.declare_term}~\<open>t ctxt\<close> declares term
-  \<open>t\<close> to belong to the context.  This automatically fixes new
-  type variables, but not term variables.  Syntactic constraints for
-  type and term variables are declared uniformly, though.
+  \<^descr> @{ML Variable.declare_term}~\<open>t ctxt\<close> declares term \<open>t\<close> to belong to the
+  context. This automatically fixes new type variables, but not term
+  variables. Syntactic constraints for type and term variables are declared
+  uniformly, though.
+
+  \<^descr> @{ML Variable.declare_constraints}~\<open>t ctxt\<close> declares syntactic constraints
+  from term \<open>t\<close>, without making it part of the context yet.
 
-  \<^descr> @{ML Variable.declare_constraints}~\<open>t ctxt\<close> declares
-  syntactic constraints from term \<open>t\<close>, without making it part
-  of the context yet.
-
-  \<^descr> @{ML Variable.export}~\<open>inner outer thms\<close> generalizes
-  fixed type and term variables in \<open>thms\<close> according to the
-  difference of the \<open>inner\<close> and \<open>outer\<close> context,
-  following the principles sketched above.
+  \<^descr> @{ML Variable.export}~\<open>inner outer thms\<close> generalizes fixed type and term
+  variables in \<open>thms\<close> according to the difference of the \<open>inner\<close> and \<open>outer\<close>
+  context, following the principles sketched above.
 
-  \<^descr> @{ML Variable.polymorphic}~\<open>ctxt ts\<close> generalizes type
-  variables in \<open>ts\<close> as far as possible, even those occurring
-  in fixed term variables.  The default policy of type-inference is to
-  fix newly introduced type variables, which is essentially reversed
-  with @{ML Variable.polymorphic}: here the given terms are detached
-  from the context as far as possible.
+  \<^descr> @{ML Variable.polymorphic}~\<open>ctxt ts\<close> generalizes type variables in \<open>ts\<close> as
+  far as possible, even those occurring in fixed term variables. The default
+  policy of type-inference is to fix newly introduced type variables, which is
+  essentially reversed with @{ML Variable.polymorphic}: here the given terms
+  are detached from the context as far as possible.
 
-  \<^descr> @{ML Variable.import}~\<open>open thms ctxt\<close> invents fixed
-  type and term variables for the schematic ones occurring in \<open>thms\<close>.  The \<open>open\<close> flag indicates whether the fixed names
-  should be accessible to the user, otherwise newly introduced names
-  are marked as ``internal'' (\secref{sec:names}).
+  \<^descr> @{ML Variable.import}~\<open>open thms ctxt\<close> invents fixed type and term
+  variables for the schematic ones occurring in \<open>thms\<close>. The \<open>open\<close> flag
+  indicates whether the fixed names should be accessible to the user,
+  otherwise newly introduced names are marked as ``internal''
+  (\secref{sec:names}).
 
-  \<^descr> @{ML Variable.focus}~\<open>bindings B\<close> decomposes the outermost \<open>\<And>\<close> prefix of proposition \<open>B\<close>, using the given name bindings.
+  \<^descr> @{ML Variable.focus}~\<open>bindings B\<close> decomposes the outermost \<open>\<And>\<close> prefix of
+  proposition \<open>B\<close>, using the given name bindings.
 \<close>
 
-text %mlex \<open>The following example shows how to work with fixed term
-  and type parameters and with type-inference.\<close>
+text %mlex \<open>
+  The following example shows how to work with fixed term and type parameters
+  and with type-inference.
+\<close>
 
 ML_val \<open>
   (*static compile-time context -- for testing only*)
@@ -173,12 +167,13 @@
   val t2 = Syntax.read_term ctxt2 "x";  (*x::nat is enforced*)
 \<close>
 
-text \<open>In the above example, the starting context is derived from the
-  toplevel theory, which means that fixed variables are internalized
-  literally: \<open>x\<close> is mapped again to \<open>x\<close>, and
-  attempting to fix it again in the subsequent context is an error.
-  Alternatively, fixed parameters can be renamed explicitly as
-  follows:\<close>
+text \<open>
+  In the above example, the starting context is derived from the toplevel
+  theory, which means that fixed variables are internalized literally: \<open>x\<close> is
+  mapped again to \<open>x\<close>, and attempting to fix it again in the subsequent
+  context is an error. Alternatively, fixed parameters can be renamed
+  explicitly as follows:
+\<close>
 
 ML_val \<open>
   val ctxt0 = @{context};
@@ -186,11 +181,12 @@
     ctxt0 |> Variable.variant_fixes ["x", "x", "x"];
 \<close>
 
-text \<open>The following ML code can now work with the invented names of
-  \<open>x1\<close>, \<open>x2\<close>, \<open>x3\<close>, without depending on
-  the details on the system policy for introducing these variants.
-  Recall that within a proof body the system always invents fresh
-  ``Skolem constants'', e.g.\ as follows:\<close>
+text \<open>
+  The following ML code can now work with the invented names of \<open>x1\<close>, \<open>x2\<close>,
+  \<open>x3\<close>, without depending on the details on the system policy for introducing
+  these variants. Recall that within a proof body the system always invents
+  fresh ``Skolem constants'', e.g.\ as follows:
+\<close>
 
 notepad
 begin
@@ -205,68 +201,63 @@
       ctxt3 |> Variable.variant_fixes ["y", "y"];\<close>
 end
 
-text \<open>In this situation @{ML Variable.add_fixes} and @{ML
-  Variable.variant_fixes} are very similar, but identical name
-  proposals given in a row are only accepted by the second version.
+text \<open>
+  In this situation @{ML Variable.add_fixes} and @{ML Variable.variant_fixes}
+  are very similar, but identical name proposals given in a row are only
+  accepted by the second version.
 \<close>
 
 
 section \<open>Assumptions \label{sec:assumptions}\<close>
 
 text \<open>
-  An \<^emph>\<open>assumption\<close> is a proposition that it is postulated in the
-  current context.  Local conclusions may use assumptions as
-  additional facts, but this imposes implicit hypotheses that weaken
-  the overall statement.
+  An \<^emph>\<open>assumption\<close> is a proposition that it is postulated in the current
+  context. Local conclusions may use assumptions as additional facts, but this
+  imposes implicit hypotheses that weaken the overall statement.
 
-  Assumptions are restricted to fixed non-schematic statements, i.e.\
-  all generality needs to be expressed by explicit quantifiers.
-  Nevertheless, the result will be in HHF normal form with outermost
-  quantifiers stripped.  For example, by assuming \<open>\<And>x :: \<alpha>. P
-  x\<close> we get \<open>\<And>x :: \<alpha>. P x \<turnstile> P ?x\<close> for schematic \<open>?x\<close>
-  of fixed type \<open>\<alpha>\<close>.  Local derivations accumulate more and
-  more explicit references to hypotheses: \<open>A\<^sub>1, \<dots>,
-  A\<^sub>n \<turnstile> B\<close> where \<open>A\<^sub>1, \<dots>, A\<^sub>n\<close> needs to
-  be covered by the assumptions of the current context.
+  Assumptions are restricted to fixed non-schematic statements, i.e.\ all
+  generality needs to be expressed by explicit quantifiers. Nevertheless, the
+  result will be in HHF normal form with outermost quantifiers stripped. For
+  example, by assuming \<open>\<And>x :: \<alpha>. P x\<close> we get \<open>\<And>x :: \<alpha>. P x \<turnstile> P ?x\<close> for
+  schematic \<open>?x\<close> of fixed type \<open>\<alpha>\<close>. Local derivations accumulate more and more
+  explicit references to hypotheses: \<open>A\<^sub>1, \<dots>, A\<^sub>n \<turnstile> B\<close> where \<open>A\<^sub>1, \<dots>, A\<^sub>n\<close>
+  needs to be covered by the assumptions of the current context.
 
   \<^medskip>
-  The \<open>add_assms\<close> operation augments the context by
-  local assumptions, which are parameterized by an arbitrary \<open>export\<close> rule (see below).
+  The \<open>add_assms\<close> operation augments the context by local assumptions, which
+  are parameterized by an arbitrary \<open>export\<close> rule (see below).
 
-  The \<open>export\<close> operation moves facts from a (larger) inner
-  context into a (smaller) outer context, by discharging the
-  difference of the assumptions as specified by the associated export
-  rules.  Note that the discharged portion is determined by the
-  difference of contexts, not the facts being exported!  There is a
-  separate flag to indicate a goal context, where the result is meant
-  to refine an enclosing sub-goal of a structured proof state.
+  The \<open>export\<close> operation moves facts from a (larger) inner context into a
+  (smaller) outer context, by discharging the difference of the assumptions as
+  specified by the associated export rules. Note that the discharged portion
+  is determined by the difference of contexts, not the facts being exported!
+  There is a separate flag to indicate a goal context, where the result is
+  meant to refine an enclosing sub-goal of a structured proof state.
 
   \<^medskip>
-  The most basic export rule discharges assumptions directly
-  by means of the \<open>\<Longrightarrow>\<close> introduction rule:
+  The most basic export rule discharges assumptions directly by means of the
+  \<open>\<Longrightarrow>\<close> introduction rule:
   \[
   \infer[(\<open>\<Longrightarrow>\<hyphen>intro\<close>)]{\<open>\<Gamma> - A \<turnstile> A \<Longrightarrow> B\<close>}{\<open>\<Gamma> \<turnstile> B\<close>}
   \]
 
-  The variant for goal refinements marks the newly introduced
-  premises, which causes the canonical Isar goal refinement scheme to
-  enforce unification with local premises within the goal:
+  The variant for goal refinements marks the newly introduced premises, which
+  causes the canonical Isar goal refinement scheme to enforce unification with
+  local premises within the goal:
   \[
   \infer[(\<open>#\<Longrightarrow>\<hyphen>intro\<close>)]{\<open>\<Gamma> - A \<turnstile> #A \<Longrightarrow> B\<close>}{\<open>\<Gamma> \<turnstile> B\<close>}
   \]
 
   \<^medskip>
-  Alternative versions of assumptions may perform arbitrary
-  transformations on export, as long as the corresponding portion of
-  hypotheses is removed from the given facts.  For example, a local
-  definition works by fixing \<open>x\<close> and assuming \<open>x \<equiv> t\<close>,
-  with the following export rule to reverse the effect:
+  Alternative versions of assumptions may perform arbitrary transformations on
+  export, as long as the corresponding portion of hypotheses is removed from
+  the given facts. For example, a local definition works by fixing \<open>x\<close> and
+  assuming \<open>x \<equiv> t\<close>, with the following export rule to reverse the effect:
   \[
   \infer[(\<open>\<equiv>\<hyphen>expand\<close>)]{\<open>\<Gamma> - (x \<equiv> t) \<turnstile> B t\<close>}{\<open>\<Gamma> \<turnstile> B x\<close>}
   \]
-  This works, because the assumption \<open>x \<equiv> t\<close> was introduced in
-  a context with \<open>x\<close> being fresh, so \<open>x\<close> does not
-  occur in \<open>\<Gamma>\<close> here.
+  This works, because the assumption \<open>x \<equiv> t\<close> was introduced in a context with
+  \<open>x\<close> being fresh, so \<open>x\<close> does not occur in \<open>\<Gamma>\<close> here.
 \<close>
 
 text %mlref \<open>
@@ -281,36 +272,33 @@
   @{index_ML Assumption.export: "bool -> Proof.context -> Proof.context -> thm -> thm"} \\
   \end{mldecls}
 
-  \<^descr> Type @{ML_type Assumption.export} represents arbitrary export
-  rules, which is any function of type @{ML_type "bool -> cterm list
-  -> thm -> thm"}, where the @{ML_type "bool"} indicates goal mode,
-  and the @{ML_type "cterm list"} the collection of assumptions to be
-  discharged simultaneously.
+  \<^descr> Type @{ML_type Assumption.export} represents arbitrary export rules, which
+  is any function of type @{ML_type "bool -> cterm list -> thm -> thm"}, where
+  the @{ML_type "bool"} indicates goal mode, and the @{ML_type "cterm list"}
+  the collection of assumptions to be discharged simultaneously.
 
-  \<^descr> @{ML Assumption.assume}~\<open>ctxt A\<close> turns proposition \<open>A\<close> into a primitive assumption \<open>A \<turnstile> A'\<close>, where the
-  conclusion \<open>A'\<close> is in HHF normal form.
+  \<^descr> @{ML Assumption.assume}~\<open>ctxt A\<close> turns proposition \<open>A\<close> into a primitive
+  assumption \<open>A \<turnstile> A'\<close>, where the conclusion \<open>A'\<close> is in HHF normal form.
 
-  \<^descr> @{ML Assumption.add_assms}~\<open>r As\<close> augments the context
-  by assumptions \<open>As\<close> with export rule \<open>r\<close>.  The
-  resulting facts are hypothetical theorems as produced by the raw
-  @{ML Assumption.assume}.
+  \<^descr> @{ML Assumption.add_assms}~\<open>r As\<close> augments the context by assumptions \<open>As\<close>
+  with export rule \<open>r\<close>. The resulting facts are hypothetical theorems as
+  produced by the raw @{ML Assumption.assume}.
 
-  \<^descr> @{ML Assumption.add_assumes}~\<open>As\<close> is a special case of
-  @{ML Assumption.add_assms} where the export rule performs \<open>\<Longrightarrow>\<hyphen>intro\<close> or \<open>#\<Longrightarrow>\<hyphen>intro\<close>, depending on goal
-  mode.
+  \<^descr> @{ML Assumption.add_assumes}~\<open>As\<close> is a special case of @{ML
+  Assumption.add_assms} where the export rule performs \<open>\<Longrightarrow>\<hyphen>intro\<close> or
+  \<open>#\<Longrightarrow>\<hyphen>intro\<close>, depending on goal mode.
 
-  \<^descr> @{ML Assumption.export}~\<open>is_goal inner outer thm\<close>
-  exports result \<open>thm\<close> from the the \<open>inner\<close> context
-  back into the \<open>outer\<close> one; \<open>is_goal = true\<close> means
-  this is a goal context.  The result is in HHF normal form.  Note
-  that @{ML "Proof_Context.export"} combines @{ML "Variable.export"}
-  and @{ML "Assumption.export"} in the canonical way.
+  \<^descr> @{ML Assumption.export}~\<open>is_goal inner outer thm\<close> exports result \<open>thm\<close>
+  from the the \<open>inner\<close> context back into the \<open>outer\<close> one; \<open>is_goal = true\<close>
+  means this is a goal context. The result is in HHF normal form. Note that
+  @{ML "Proof_Context.export"} combines @{ML "Variable.export"} and @{ML
+  "Assumption.export"} in the canonical way.
 \<close>
 
-text %mlex \<open>The following example demonstrates how rules can be
-  derived by building up a context of assumptions first, and exporting
-  some local fact afterwards.  We refer to @{theory Pure} equality
-  here for testing purposes.
+text %mlex \<open>
+  The following example demonstrates how rules can be derived by building up a
+  context of assumptions first, and exporting some local fact afterwards. We
+  refer to @{theory Pure} equality here for testing purposes.
 \<close>
 
 ML_val \<open>
@@ -325,52 +313,50 @@
   val r = Assumption.export false ctxt1 ctxt0 eq';
 \<close>
 
-text \<open>Note that the variables of the resulting rule are not
-  generalized.  This would have required to fix them properly in the
-  context beforehand, and export wrt.\ variables afterwards (cf.\ @{ML
-  Variable.export} or the combined @{ML "Proof_Context.export"}).\<close>
+text \<open>
+  Note that the variables of the resulting rule are not generalized. This
+  would have required to fix them properly in the context beforehand, and
+  export wrt.\ variables afterwards (cf.\ @{ML Variable.export} or the
+  combined @{ML "Proof_Context.export"}).
+\<close>
 
 
 section \<open>Structured goals and results \label{sec:struct-goals}\<close>
 
 text \<open>
-  Local results are established by monotonic reasoning from facts
-  within a context.  This allows common combinations of theorems,
-  e.g.\ via \<open>\<And>/\<Longrightarrow>\<close> elimination, resolution rules, or equational
-  reasoning, see \secref{sec:thms}.  Unaccounted context manipulations
-  should be avoided, notably raw \<open>\<And>/\<Longrightarrow>\<close> introduction or ad-hoc
-  references to free variables or assumptions not present in the proof
-  context.
+  Local results are established by monotonic reasoning from facts within a
+  context. This allows common combinations of theorems, e.g.\ via \<open>\<And>/\<Longrightarrow>\<close>
+  elimination, resolution rules, or equational reasoning, see
+  \secref{sec:thms}. Unaccounted context manipulations should be avoided,
+  notably raw \<open>\<And>/\<Longrightarrow>\<close> introduction or ad-hoc references to free variables or
+  assumptions not present in the proof context.
 
   \<^medskip>
-  The \<open>SUBPROOF\<close> combinator allows to structure a
-  tactical proof recursively by decomposing a selected sub-goal:
-  \<open>(\<And>x. A(x) \<Longrightarrow> B(x)) \<Longrightarrow> \<dots>\<close> is turned into \<open>B(x) \<Longrightarrow> \<dots>\<close>
-  after fixing \<open>x\<close> and assuming \<open>A(x)\<close>.  This means
-  the tactic needs to solve the conclusion, but may use the premise as
-  a local fact, for locally fixed variables.
+  The \<open>SUBPROOF\<close> combinator allows to structure a tactical proof recursively
+  by decomposing a selected sub-goal: \<open>(\<And>x. A(x) \<Longrightarrow> B(x)) \<Longrightarrow> \<dots>\<close> is turned into
+  \<open>B(x) \<Longrightarrow> \<dots>\<close> after fixing \<open>x\<close> and assuming \<open>A(x)\<close>. This means the tactic needs
+  to solve the conclusion, but may use the premise as a local fact, for
+  locally fixed variables.
 
-  The family of \<open>FOCUS\<close> combinators is similar to \<open>SUBPROOF\<close>, but allows to retain schematic variables and pending
-  subgoals in the resulting goal state.
+  The family of \<open>FOCUS\<close> combinators is similar to \<open>SUBPROOF\<close>, but allows to
+  retain schematic variables and pending subgoals in the resulting goal state.
 
-  The \<open>prove\<close> operation provides an interface for structured
-  backwards reasoning under program control, with some explicit sanity
-  checks of the result.  The goal context can be augmented by
-  additional fixed variables (cf.\ \secref{sec:variables}) and
-  assumptions (cf.\ \secref{sec:assumptions}), which will be available
-  as local facts during the proof and discharged into implications in
-  the result.  Type and term variables are generalized as usual,
-  according to the context.
+  The \<open>prove\<close> operation provides an interface for structured backwards
+  reasoning under program control, with some explicit sanity checks of the
+  result. The goal context can be augmented by additional fixed variables
+  (cf.\ \secref{sec:variables}) and assumptions (cf.\
+  \secref{sec:assumptions}), which will be available as local facts during the
+  proof and discharged into implications in the result. Type and term
+  variables are generalized as usual, according to the context.
 
-  The \<open>obtain\<close> operation produces results by eliminating
-  existing facts by means of a given tactic.  This acts like a dual
-  conclusion: the proof demonstrates that the context may be augmented
-  by parameters and assumptions, without affecting any conclusions
-  that do not mention these parameters.  See also
-  @{cite "isabelle-isar-ref"} for the user-level @{command obtain} and
-  @{command guess} elements.  Final results, which may not refer to
-  the parameters in the conclusion, need to exported explicitly into
-  the original context.\<close>
+  The \<open>obtain\<close> operation produces results by eliminating existing facts by
+  means of a given tactic. This acts like a dual conclusion: the proof
+  demonstrates that the context may be augmented by parameters and
+  assumptions, without affecting any conclusions that do not mention these
+  parameters. See also @{cite "isabelle-isar-ref"} for the user-level
+  @{command obtain} and @{command guess} elements. Final results, which may
+  not refer to the parameters in the conclusion, need to exported explicitly
+  into the original context.\<close>
 
 text %mlref \<open>
   \begin{mldecls}
@@ -402,54 +388,54 @@
   Proof.context -> ((string * cterm) list * thm list) * Proof.context"} \\
   \end{mldecls}
 
-  \<^descr> @{ML SUBPROOF}~\<open>tac ctxt i\<close> decomposes the structure
-  of the specified sub-goal, producing an extended context and a
-  reduced goal, which needs to be solved by the given tactic.  All
-  schematic parameters of the goal are imported into the context as
-  fixed ones, which may not be instantiated in the sub-proof.
+  \<^descr> @{ML SUBPROOF}~\<open>tac ctxt i\<close> decomposes the structure of the specified
+  sub-goal, producing an extended context and a reduced goal, which needs to
+  be solved by the given tactic. All schematic parameters of the goal are
+  imported into the context as fixed ones, which may not be instantiated in
+  the sub-proof.
 
   \<^descr> @{ML Subgoal.FOCUS}, @{ML Subgoal.FOCUS_PREMS}, and @{ML
-  Subgoal.FOCUS_PARAMS} are similar to @{ML SUBPROOF}, but are
-  slightly more flexible: only the specified parts of the subgoal are
-  imported into the context, and the body tactic may introduce new
-  subgoals and schematic variables.
+  Subgoal.FOCUS_PARAMS} are similar to @{ML SUBPROOF}, but are slightly more
+  flexible: only the specified parts of the subgoal are imported into the
+  context, and the body tactic may introduce new subgoals and schematic
+  variables.
 
-  \<^descr> @{ML Subgoal.focus}, @{ML Subgoal.focus_prems}, @{ML
-  Subgoal.focus_params} extract the focus information from a goal
-  state in the same way as the corresponding tacticals above.  This is
-  occasionally useful to experiment without writing actual tactics
-  yet.
+  \<^descr> @{ML Subgoal.focus}, @{ML Subgoal.focus_prems}, @{ML Subgoal.focus_params}
+  extract the focus information from a goal state in the same way as the
+  corresponding tacticals above. This is occasionally useful to experiment
+  without writing actual tactics yet.
 
-  \<^descr> @{ML Goal.prove}~\<open>ctxt xs As C tac\<close> states goal \<open>C\<close> in the context augmented by fixed variables \<open>xs\<close> and
-  assumptions \<open>As\<close>, and applies tactic \<open>tac\<close> to solve
-  it.  The latter may depend on the local assumptions being presented
-  as facts.  The result is in HHF normal form.
+  \<^descr> @{ML Goal.prove}~\<open>ctxt xs As C tac\<close> states goal \<open>C\<close> in the context
+  augmented by fixed variables \<open>xs\<close> and assumptions \<open>As\<close>, and applies tactic
+  \<open>tac\<close> to solve it. The latter may depend on the local assumptions being
+  presented as facts. The result is in HHF normal form.
 
-  \<^descr> @{ML Goal.prove_common}~\<open>ctxt fork_pri\<close> is the common form
-  to state and prove a simultaneous goal statement, where @{ML Goal.prove}
-  is a convenient shorthand that is most frequently used in applications.
+  \<^descr> @{ML Goal.prove_common}~\<open>ctxt fork_pri\<close> is the common form to state and
+  prove a simultaneous goal statement, where @{ML Goal.prove} is a convenient
+  shorthand that is most frequently used in applications.
 
   The given list of simultaneous conclusions is encoded in the goal state by
-  means of Pure conjunction: @{ML Goal.conjunction_tac} will turn this into
-  a collection of individual subgoals, but note that the original multi-goal
+  means of Pure conjunction: @{ML Goal.conjunction_tac} will turn this into a
+  collection of individual subgoals, but note that the original multi-goal
   state is usually required for advanced induction.
 
-  It is possible to provide an optional priority for a forked proof,
-  typically @{ML "SOME ~1"}, while @{ML NONE} means the proof is immediate
-  (sequential) as for @{ML Goal.prove}. Note that a forked proof does not
-  exhibit any failures in the usual way via exceptions in ML, but
-  accumulates error situations under the execution id of the running
-  transaction. Thus the system is able to expose error messages ultimately
-  to the end-user, even though the subsequent ML code misses them.
+  It is possible to provide an optional priority for a forked proof, typically
+  @{ML "SOME ~1"}, while @{ML NONE} means the proof is immediate (sequential)
+  as for @{ML Goal.prove}. Note that a forked proof does not exhibit any
+  failures in the usual way via exceptions in ML, but accumulates error
+  situations under the execution id of the running transaction. Thus the
+  system is able to expose error messages ultimately to the end-user, even
+  though the subsequent ML code misses them.
 
-  \<^descr> @{ML Obtain.result}~\<open>tac thms ctxt\<close> eliminates the
-  given facts using a tactic, which results in additional fixed
-  variables and assumptions in the context.  Final results need to be
-  exported explicitly.
+  \<^descr> @{ML Obtain.result}~\<open>tac thms ctxt\<close> eliminates the given facts using a
+  tactic, which results in additional fixed variables and assumptions in the
+  context. Final results need to be exported explicitly.
 \<close>
 
-text %mlex \<open>The following minimal example illustrates how to access
-  the focus information of a structured goal state.\<close>
+text %mlex \<open>
+  The following minimal example illustrates how to access the focus
+  information of a structured goal state.
+\<close>
 
 notepad
 begin
@@ -467,8 +453,9 @@
 
 text \<open>
   \<^medskip>
-  The next example demonstrates forward-elimination in
-  a local context, using @{ML Obtain.result}.\<close>
+  The next example demonstrates forward-elimination in a local context, using
+  @{ML Obtain.result}.
+\<close>
 
 notepad
 begin
--- a/src/Doc/Implementation/Syntax.thy	Tue Dec 15 14:41:47 2015 +0000
+++ b/src/Doc/Implementation/Syntax.thy	Wed Dec 16 17:30:30 2015 +0100
@@ -6,51 +6,49 @@
 
 chapter \<open>Concrete syntax and type-checking\<close>
 
-text \<open>Pure \<open>\<lambda>\<close>-calculus as introduced in \chref{ch:logic} is
-  an adequate foundation for logical languages --- in the tradition of
-  \<^emph>\<open>higher-order abstract syntax\<close> --- but end-users require
-  additional means for reading and printing of terms and types.  This
-  important add-on outside the logical core is called \<^emph>\<open>inner
-  syntax\<close> in Isabelle jargon, as opposed to the \<^emph>\<open>outer syntax\<close> of
-  the theory and proof language @{cite "isabelle-isar-ref"}.
+text \<open>
+  Pure \<open>\<lambda>\<close>-calculus as introduced in \chref{ch:logic} is an adequate
+  foundation for logical languages --- in the tradition of \<^emph>\<open>higher-order
+  abstract syntax\<close> --- but end-users require additional means for reading and
+  printing of terms and types. This important add-on outside the logical core
+  is called \<^emph>\<open>inner syntax\<close> in Isabelle jargon, as opposed to the \<^emph>\<open>outer
+  syntax\<close> of the theory and proof language @{cite "isabelle-isar-ref"}.
 
   For example, according to @{cite church40} quantifiers are represented as
-  higher-order constants \<open>All :: ('a \<Rightarrow> bool) \<Rightarrow> bool\<close> such that \<open>All (\<lambda>x::'a. B x)\<close> faithfully represents the idea that is displayed in
-  Isabelle as \<open>\<forall>x::'a. B x\<close> via @{keyword "binder"} notation.
-  Moreover, type-inference in the style of Hindley-Milner @{cite hindleymilner}
-  (and extensions) enables users to write \<open>\<forall>x. B x\<close> concisely, when
-  the type \<open>'a\<close> is already clear from the
-  context.\<^footnote>\<open>Type-inference taken to the extreme can easily confuse
-  users. Beginners often stumble over unexpectedly general types inferred by
-  the system.\<close>
+  higher-order constants \<open>All :: ('a \<Rightarrow> bool) \<Rightarrow> bool\<close> such that \<open>All (\<lambda>x::'a. B
+  x)\<close> faithfully represents the idea that is displayed in Isabelle as \<open>\<forall>x::'a.
+  B x\<close> via @{keyword "binder"} notation. Moreover, type-inference in the style
+  of Hindley-Milner @{cite hindleymilner} (and extensions) enables users to
+  write \<open>\<forall>x. B x\<close> concisely, when the type \<open>'a\<close> is already clear from the
+  context.\<^footnote>\<open>Type-inference taken to the extreme can easily confuse users.
+  Beginners often stumble over unexpectedly general types inferred by the
+  system.\<close>
 
   \<^medskip>
-  The main inner syntax operations are \<^emph>\<open>read\<close> for
-  parsing together with type-checking, and \<^emph>\<open>pretty\<close> for formatted
-  output.  See also \secref{sec:read-print}.
+  The main inner syntax operations are \<^emph>\<open>read\<close> for parsing together with
+  type-checking, and \<^emph>\<open>pretty\<close> for formatted output. See also
+  \secref{sec:read-print}.
 
   Furthermore, the input and output syntax layers are sub-divided into
-  separate phases for \<^emph>\<open>concrete syntax\<close> versus \<^emph>\<open>abstract
-  syntax\<close>, see also \secref{sec:parse-unparse} and
-  \secref{sec:term-check}, respectively.  This results in the
-  following decomposition of the main operations:
+  separate phases for \<^emph>\<open>concrete syntax\<close> versus \<^emph>\<open>abstract syntax\<close>, see also
+  \secref{sec:parse-unparse} and \secref{sec:term-check}, respectively. This
+  results in the following decomposition of the main operations:
 
-  \<^item> \<open>read = parse; check\<close>
+    \<^item> \<open>read = parse; check\<close>
 
-  \<^item> \<open>pretty = uncheck; unparse\<close>
-
+    \<^item> \<open>pretty = uncheck; unparse\<close>
 
   For example, some specification package might thus intercept syntax
-  processing at a well-defined stage after \<open>parse\<close>, to a augment the
-  resulting pre-term before full type-reconstruction is performed by \<open>check\<close>. Note that the formal status of bound variables, versus free
-  variables, versus constants must not be changed between these phases.
+  processing at a well-defined stage after \<open>parse\<close>, to a augment the resulting
+  pre-term before full type-reconstruction is performed by \<open>check\<close>. Note that
+  the formal status of bound variables, versus free variables, versus
+  constants must not be changed between these phases.
 
   \<^medskip>
-  In general, \<open>check\<close> and \<open>uncheck\<close> operate
-  simultaneously on a list of terms. This is particular important for
-  type-checking, to reconstruct types for several terms of the same context
-  and scope. In contrast, \<open>parse\<close> and \<open>unparse\<close> operate separately
-  on single terms.
+  In general, \<open>check\<close> and \<open>uncheck\<close> operate simultaneously on a list of terms.
+  This is particular important for type-checking, to reconstruct types for
+  several terms of the same context and scope. In contrast, \<open>parse\<close> and
+  \<open>unparse\<close> operate separately on single terms.
 
   There are analogous operations to read and print types, with the same
   sub-division into phases.
@@ -61,11 +59,12 @@
 
 text \<open>
   Read and print operations are roughly dual to each other, such that for the
-  user \<open>s' = pretty (read s)\<close> looks similar to the original source
-  text \<open>s\<close>, but the details depend on many side-conditions. There are
-  also explicit options to control the removal of type information in the
-  output. The default configuration routinely looses information, so \<open>t' = read (pretty t)\<close> might fail, or produce a differently typed term, or
-  a completely different term in the face of syntactic overloading.
+  user \<open>s' = pretty (read s)\<close> looks similar to the original source text \<open>s\<close>,
+  but the details depend on many side-conditions. There are also explicit
+  options to control the removal of type information in the output. The
+  default configuration routinely looses information, so \<open>t' = read (pretty
+  t)\<close> might fail, or produce a differently typed term, or a completely
+  different term in the face of syntactic overloading.
 \<close>
 
 text %mlref \<open>
@@ -82,54 +81,53 @@
   @{index_ML Syntax.string_of_term: "Proof.context -> term -> string"} \\
   \end{mldecls}
 
-  \<^descr> @{ML Syntax.read_typs}~\<open>ctxt strs\<close> parses and checks a
-  simultaneous list of source strings as types of the logic.
+  \<^descr> @{ML Syntax.read_typs}~\<open>ctxt strs\<close> parses and checks a simultaneous list
+  of source strings as types of the logic.
 
-  \<^descr> @{ML Syntax.read_terms}~\<open>ctxt strs\<close> parses and checks a
-  simultaneous list of source strings as terms of the logic.
-  Type-reconstruction puts all parsed terms into the same scope: types of
-  free variables ultimately need to coincide.
+  \<^descr> @{ML Syntax.read_terms}~\<open>ctxt strs\<close> parses and checks a simultaneous list
+  of source strings as terms of the logic. Type-reconstruction puts all parsed
+  terms into the same scope: types of free variables ultimately need to
+  coincide.
 
   If particular type-constraints are required for some of the arguments, the
   read operations needs to be split into its parse and check phases. Then it
   is possible to use @{ML Type.constraint} on the intermediate pre-terms
   (\secref{sec:term-check}).
 
-  \<^descr> @{ML Syntax.read_props}~\<open>ctxt strs\<close> parses and checks a
-  simultaneous list of source strings as terms of the logic, with an implicit
-  type-constraint for each argument to enforce type @{typ prop}; this also
-  affects the inner syntax for parsing. The remaining type-reconstruction
-  works as for @{ML Syntax.read_terms}.
+  \<^descr> @{ML Syntax.read_props}~\<open>ctxt strs\<close> parses and checks a simultaneous list
+  of source strings as terms of the logic, with an implicit type-constraint
+  for each argument to enforce type @{typ prop}; this also affects the inner
+  syntax for parsing. The remaining type-reconstruction works as for @{ML
+  Syntax.read_terms}.
 
-  \<^descr> @{ML Syntax.read_typ}, @{ML Syntax.read_term}, @{ML Syntax.read_prop}
-  are like the simultaneous versions, but operate on a single argument only.
-  This convenient shorthand is adequate in situations where a single item in
-  its own scope is processed. Do not use @{ML "map o Syntax.read_term"} where
-  @{ML Syntax.read_terms} is actually intended!
+  \<^descr> @{ML Syntax.read_typ}, @{ML Syntax.read_term}, @{ML Syntax.read_prop} are
+  like the simultaneous versions, but operate on a single argument only. This
+  convenient shorthand is adequate in situations where a single item in its
+  own scope is processed. Do not use @{ML "map o Syntax.read_term"} where @{ML
+  Syntax.read_terms} is actually intended!
 
-  \<^descr> @{ML Syntax.pretty_typ}~\<open>ctxt T\<close> and @{ML
-  Syntax.pretty_term}~\<open>ctxt t\<close> uncheck and pretty-print the given type
-  or term, respectively. Although the uncheck phase acts on a simultaneous
-  list as well, this is rarely used in practice, so only the singleton case is
-  provided as combined pretty operation. There is no distinction of term vs.\
-  proposition.
+  \<^descr> @{ML Syntax.pretty_typ}~\<open>ctxt T\<close> and @{ML Syntax.pretty_term}~\<open>ctxt t\<close>
+  uncheck and pretty-print the given type or term, respectively. Although the
+  uncheck phase acts on a simultaneous list as well, this is rarely used in
+  practice, so only the singleton case is provided as combined pretty
+  operation. There is no distinction of term vs.\ proposition.
 
-  \<^descr> @{ML Syntax.string_of_typ} and @{ML Syntax.string_of_term} are
-  convenient compositions of @{ML Syntax.pretty_typ} and @{ML
-  Syntax.pretty_term} with @{ML Pretty.string_of} for output. The result may
-  be concatenated with other strings, as long as there is no further
-  formatting and line-breaking involved.
+  \<^descr> @{ML Syntax.string_of_typ} and @{ML Syntax.string_of_term} are convenient
+  compositions of @{ML Syntax.pretty_typ} and @{ML Syntax.pretty_term} with
+  @{ML Pretty.string_of} for output. The result may be concatenated with other
+  strings, as long as there is no further formatting and line-breaking
+  involved.
 
 
   @{ML Syntax.read_term}, @{ML Syntax.read_prop}, and @{ML
   Syntax.string_of_term} are the most important operations in practice.
 
   \<^medskip>
-  Note that the string values that are passed in and out are
-  annotated by the system, to carry further markup that is relevant for the
-  Prover IDE @{cite "isabelle-jedit"}. User code should neither compose its
-  own input strings, nor try to analyze the output strings. Conceptually this
-  is an abstract datatype, encoded as concrete string for historical reasons.
+  Note that the string values that are passed in and out are annotated by the
+  system, to carry further markup that is relevant for the Prover IDE @{cite
+  "isabelle-jedit"}. User code should neither compose its own input strings,
+  nor try to analyze the output strings. Conceptually this is an abstract
+  datatype, encoded as concrete string for historical reasons.
 
   The standard way to provide the required position markup for input works via
   the outer syntax parser wrapper @{ML Parse.inner_syntax}, which is already
@@ -147,8 +145,8 @@
   \<^emph>\<open>pre-term\<close> format, where all bindings and scopes are already resolved
   faithfully. Thus the names of free variables or constants are determined in
   the sense of the logical context, but type information might be still
-  missing. Pre-terms support an explicit language of \<^emph>\<open>type constraints\<close>
-  that may be augmented by user code to guide the later \<^emph>\<open>check\<close> phase.
+  missing. Pre-terms support an explicit language of \<^emph>\<open>type constraints\<close> that
+  may be augmented by user code to guide the later \<^emph>\<open>check\<close> phase.
 
   Actual parsing is based on traditional lexical analysis and Earley parsing
   for arbitrary context-free grammars. The user can specify the grammar
@@ -170,24 +168,23 @@
   @{index_ML Syntax.unparse_term: "Proof.context -> term -> Pretty.T"} \\
   \end{mldecls}
 
-  \<^descr> @{ML Syntax.parse_typ}~\<open>ctxt str\<close> parses a source string as
-  pre-type that is ready to be used with subsequent check operations.
+  \<^descr> @{ML Syntax.parse_typ}~\<open>ctxt str\<close> parses a source string as pre-type that
+  is ready to be used with subsequent check operations.
 
-  \<^descr> @{ML Syntax.parse_term}~\<open>ctxt str\<close> parses a source string as
-  pre-term that is ready to be used with subsequent check operations.
+  \<^descr> @{ML Syntax.parse_term}~\<open>ctxt str\<close> parses a source string as pre-term that
+  is ready to be used with subsequent check operations.
 
-  \<^descr> @{ML Syntax.parse_prop}~\<open>ctxt str\<close> parses a source string as
-  pre-term that is ready to be used with subsequent check operations. The
-  inner syntax category is @{typ prop} and a suitable type-constraint is
-  included to ensure that this information is observed in subsequent type
-  reconstruction.
+  \<^descr> @{ML Syntax.parse_prop}~\<open>ctxt str\<close> parses a source string as pre-term that
+  is ready to be used with subsequent check operations. The inner syntax
+  category is @{typ prop} and a suitable type-constraint is included to ensure
+  that this information is observed in subsequent type reconstruction.
 
-  \<^descr> @{ML Syntax.unparse_typ}~\<open>ctxt T\<close> unparses a type after
-  uncheck operations, to turn it into a pretty tree.
+  \<^descr> @{ML Syntax.unparse_typ}~\<open>ctxt T\<close> unparses a type after uncheck
+  operations, to turn it into a pretty tree.
 
-  \<^descr> @{ML Syntax.unparse_term}~\<open>ctxt T\<close> unparses a term after
-  uncheck operations, to turn it into a pretty tree. There is no distinction
-  for propositions here.
+  \<^descr> @{ML Syntax.unparse_term}~\<open>ctxt T\<close> unparses a term after uncheck
+  operations, to turn it into a pretty tree. There is no distinction for
+  propositions here.
 
 
   These operations always operate on a single item; use the combinator @{ML
@@ -197,32 +194,31 @@
 
 section \<open>Checking and unchecking \label{sec:term-check}\<close>
 
-text \<open>These operations define the transition from pre-terms and
-  fully-annotated terms in the sense of the logical core
-  (\chref{ch:logic}).
+text \<open>
+  These operations define the transition from pre-terms and fully-annotated
+  terms in the sense of the logical core (\chref{ch:logic}).
 
-  The \<^emph>\<open>check\<close> phase is meant to subsume a variety of mechanisms
-  in the manner of ``type-inference'' or ``type-reconstruction'' or
-  ``type-improvement'', not just type-checking in the narrow sense.
-  The \<^emph>\<open>uncheck\<close> phase is roughly dual, it prunes type-information
-  before pretty printing.
+  The \<^emph>\<open>check\<close> phase is meant to subsume a variety of mechanisms in the manner
+  of ``type-inference'' or ``type-reconstruction'' or ``type-improvement'',
+  not just type-checking in the narrow sense. The \<^emph>\<open>uncheck\<close> phase is roughly
+  dual, it prunes type-information before pretty printing.
 
   A typical add-on for the check/uncheck syntax layer is the @{command
   abbreviation} mechanism @{cite "isabelle-isar-ref"}. Here the user specifies
-  syntactic definitions that are managed by the system as polymorphic \<open>let\<close> bindings. These are expanded during the \<open>check\<close> phase, and
-  contracted during the \<open>uncheck\<close> phase, without affecting the
-  type-assignment of the given terms.
+  syntactic definitions that are managed by the system as polymorphic \<open>let\<close>
+  bindings. These are expanded during the \<open>check\<close> phase, and contracted during
+  the \<open>uncheck\<close> phase, without affecting the type-assignment of the given
+  terms.
 
   \<^medskip>
-  The precise meaning of type checking depends on the context ---
-  additional check/uncheck modules might be defined in user space.
+  The precise meaning of type checking depends on the context --- additional
+  check/uncheck modules might be defined in user space.
 
-  For example, the @{command class} command defines a context where
-  \<open>check\<close> treats certain type instances of overloaded
-  constants according to the ``dictionary construction'' of its
-  logical foundation.  This involves ``type improvement''
-  (specialization of slightly too general types) and replacement by
-  certain locale parameters.  See also @{cite "Haftmann-Wenzel:2009"}.
+  For example, the @{command class} command defines a context where \<open>check\<close>
+  treats certain type instances of overloaded constants according to the
+  ``dictionary construction'' of its logical foundation. This involves ``type
+  improvement'' (specialization of slightly too general types) and replacement
+  by certain locale parameters. See also @{cite "Haftmann-Wenzel:2009"}.
 \<close>
 
 text %mlref \<open>
@@ -234,13 +230,13 @@
   @{index_ML Syntax.uncheck_terms: "Proof.context -> term list -> term list"} \\
   \end{mldecls}
 
-  \<^descr> @{ML Syntax.check_typs}~\<open>ctxt Ts\<close> checks a simultaneous list
-  of pre-types as types of the logic.  Typically, this involves normalization
-  of type synonyms.
+  \<^descr> @{ML Syntax.check_typs}~\<open>ctxt Ts\<close> checks a simultaneous list of pre-types
+  as types of the logic. Typically, this involves normalization of type
+  synonyms.
 
-  \<^descr> @{ML Syntax.check_terms}~\<open>ctxt ts\<close> checks a simultaneous list
-  of pre-terms as terms of the logic. Typically, this involves type-inference
-  and normalization term abbreviations. The types within the given terms are
+  \<^descr> @{ML Syntax.check_terms}~\<open>ctxt ts\<close> checks a simultaneous list of pre-terms
+  as terms of the logic. Typically, this involves type-inference and
+  normalization term abbreviations. The types within the given terms are
   treated in the same way as for @{ML Syntax.check_typs}.
 
   Applications sometimes need to check several types and terms together. The
@@ -249,17 +245,17 @@
   is checked; afterwards the type arguments are recovered with @{ML
   Logic.dest_type}.
 
-  \<^descr> @{ML Syntax.check_props}~\<open>ctxt ts\<close> checks a simultaneous list
-  of pre-terms as terms of the logic, such that all terms are constrained by
-  type @{typ prop}. The remaining check operation works as @{ML
-  Syntax.check_terms} above.
+  \<^descr> @{ML Syntax.check_props}~\<open>ctxt ts\<close> checks a simultaneous list of pre-terms
+  as terms of the logic, such that all terms are constrained by type @{typ
+  prop}. The remaining check operation works as @{ML Syntax.check_terms}
+  above.
 
-  \<^descr> @{ML Syntax.uncheck_typs}~\<open>ctxt Ts\<close> unchecks a simultaneous
-  list of types of the logic, in preparation of pretty printing.
+  \<^descr> @{ML Syntax.uncheck_typs}~\<open>ctxt Ts\<close> unchecks a simultaneous list of types
+  of the logic, in preparation of pretty printing.
 
-  \<^descr> @{ML Syntax.uncheck_terms}~\<open>ctxt ts\<close> unchecks a simultaneous
-  list of terms of the logic, in preparation of pretty printing. There is no
-  distinction for propositions here.
+  \<^descr> @{ML Syntax.uncheck_terms}~\<open>ctxt ts\<close> unchecks a simultaneous list of terms
+  of the logic, in preparation of pretty printing. There is no distinction for
+  propositions here.
 
 
   These operations always operate simultaneously on a list; use the combinator
--- a/src/Doc/Implementation/Tactic.thy	Tue Dec 15 14:41:47 2015 +0000
+++ b/src/Doc/Implementation/Tactic.thy	Wed Dec 16 17:30:30 2015 +0100
@@ -6,42 +6,41 @@
 
 chapter \<open>Tactical reasoning\<close>
 
-text \<open>Tactical reasoning works by refining an initial claim in a
-  backwards fashion, until a solved form is reached.  A \<open>goal\<close>
-  consists of several subgoals that need to be solved in order to
-  achieve the main statement; zero subgoals means that the proof may
-  be finished.  A \<open>tactic\<close> is a refinement operation that maps
-  a goal to a lazy sequence of potential successors.  A \<open>tactical\<close> is a combinator for composing tactics.\<close>
+text \<open>
+  Tactical reasoning works by refining an initial claim in a backwards
+  fashion, until a solved form is reached. A \<open>goal\<close> consists of several
+  subgoals that need to be solved in order to achieve the main statement; zero
+  subgoals means that the proof may be finished. A \<open>tactic\<close> is a refinement
+  operation that maps a goal to a lazy sequence of potential successors. A
+  \<open>tactical\<close> is a combinator for composing tactics.
+\<close>
 
 
 section \<open>Goals \label{sec:tactical-goals}\<close>
 
 text \<open>
-  Isabelle/Pure represents a goal as a theorem stating that the
-  subgoals imply the main goal: \<open>A\<^sub>1 \<Longrightarrow> \<dots> \<Longrightarrow> A\<^sub>n \<Longrightarrow>
-  C\<close>.  The outermost goal structure is that of a Horn Clause: i.e.\
-  an iterated implication without any quantifiers\<^footnote>\<open>Recall that
-  outermost \<open>\<And>x. \<phi>[x]\<close> is always represented via schematic
-  variables in the body: \<open>\<phi>[?x]\<close>.  These variables may get
-  instantiated during the course of reasoning.\<close>.  For \<open>n = 0\<close>
-  a goal is called ``solved''.
+  Isabelle/Pure represents a goal as a theorem stating that the subgoals imply
+  the main goal: \<open>A\<^sub>1 \<Longrightarrow> \<dots> \<Longrightarrow> A\<^sub>n \<Longrightarrow> C\<close>. The outermost goal structure is that of
+  a Horn Clause: i.e.\ an iterated implication without any quantifiers\<^footnote>\<open>Recall
+  that outermost \<open>\<And>x. \<phi>[x]\<close> is always represented via schematic variables in
+  the body: \<open>\<phi>[?x]\<close>. These variables may get instantiated during the course of
+  reasoning.\<close>. For \<open>n = 0\<close> a goal is called ``solved''.
 
-  The structure of each subgoal \<open>A\<^sub>i\<close> is that of a
-  general Hereditary Harrop Formula \<open>\<And>x\<^sub>1 \<dots>
-  \<And>x\<^sub>k. H\<^sub>1 \<Longrightarrow> \<dots> \<Longrightarrow> H\<^sub>m \<Longrightarrow> B\<close>.  Here \<open>x\<^sub>1, \<dots>, x\<^sub>k\<close> are goal parameters, i.e.\
-  arbitrary-but-fixed entities of certain types, and \<open>H\<^sub>1, \<dots>, H\<^sub>m\<close> are goal hypotheses, i.e.\ facts that may
-  be assumed locally.  Together, this forms the goal context of the
-  conclusion \<open>B\<close> to be established.  The goal hypotheses may be
-  again arbitrary Hereditary Harrop Formulas, although the level of
-  nesting rarely exceeds 1--2 in practice.
+  The structure of each subgoal \<open>A\<^sub>i\<close> is that of a general Hereditary Harrop
+  Formula \<open>\<And>x\<^sub>1 \<dots> \<And>x\<^sub>k. H\<^sub>1 \<Longrightarrow> \<dots> \<Longrightarrow> H\<^sub>m \<Longrightarrow> B\<close>. Here \<open>x\<^sub>1, \<dots>, x\<^sub>k\<close> are goal
+  parameters, i.e.\ arbitrary-but-fixed entities of certain types, and \<open>H\<^sub>1,
+  \<dots>, H\<^sub>m\<close> are goal hypotheses, i.e.\ facts that may be assumed locally.
+  Together, this forms the goal context of the conclusion \<open>B\<close> to be
+  established. The goal hypotheses may be again arbitrary Hereditary Harrop
+  Formulas, although the level of nesting rarely exceeds 1--2 in practice.
 
-  The main conclusion \<open>C\<close> is internally marked as a protected
-  proposition, which is represented explicitly by the notation \<open>#C\<close> here.  This ensures that the decomposition into subgoals and
-  main conclusion is well-defined for arbitrarily structured claims.
+  The main conclusion \<open>C\<close> is internally marked as a protected proposition,
+  which is represented explicitly by the notation \<open>#C\<close> here. This ensures that
+  the decomposition into subgoals and main conclusion is well-defined for
+  arbitrarily structured claims.
 
   \<^medskip>
-  Basic goal management is performed via the following
-  Isabelle/Pure rules:
+  Basic goal management is performed via the following Isabelle/Pure rules:
 
   \[
   \infer[\<open>(init)\<close>]{\<open>C \<Longrightarrow> #C\<close>}{} \qquad
@@ -49,8 +48,8 @@
   \]
 
   \<^medskip>
-  The following low-level variants admit general reasoning
-  with protected propositions:
+  The following low-level variants admit general reasoning with protected
+  propositions:
 
   \[
   \infer[\<open>(protect n)\<close>]{\<open>A\<^sub>1 \<Longrightarrow> \<dots> \<Longrightarrow> A\<^sub>n \<Longrightarrow> #C\<close>}{\<open>A\<^sub>1 \<Longrightarrow> \<dots> \<Longrightarrow> A\<^sub>n \<Longrightarrow> C\<close>}
@@ -68,101 +67,91 @@
   @{index_ML Goal.conclude: "thm -> thm"} \\
   \end{mldecls}
 
-  \<^descr> @{ML "Goal.init"}~\<open>C\<close> initializes a tactical goal from
-  the well-formed proposition \<open>C\<close>.
+  \<^descr> @{ML "Goal.init"}~\<open>C\<close> initializes a tactical goal from the well-formed
+  proposition \<open>C\<close>.
 
-  \<^descr> @{ML "Goal.finish"}~\<open>ctxt thm\<close> checks whether theorem
-  \<open>thm\<close> is a solved goal (no subgoals), and concludes the
-  result by removing the goal protection.  The context is only
-  required for printing error messages.
+  \<^descr> @{ML "Goal.finish"}~\<open>ctxt thm\<close> checks whether theorem \<open>thm\<close> is a solved
+  goal (no subgoals), and concludes the result by removing the goal
+  protection. The context is only required for printing error messages.
 
-  \<^descr> @{ML "Goal.protect"}~\<open>n thm\<close> protects the statement
-  of theorem \<open>thm\<close>.  The parameter \<open>n\<close> indicates the
-  number of premises to be retained.
+  \<^descr> @{ML "Goal.protect"}~\<open>n thm\<close> protects the statement of theorem \<open>thm\<close>. The
+  parameter \<open>n\<close> indicates the number of premises to be retained.
 
-  \<^descr> @{ML "Goal.conclude"}~\<open>thm\<close> removes the goal
-  protection, even if there are pending subgoals.
+  \<^descr> @{ML "Goal.conclude"}~\<open>thm\<close> removes the goal protection, even if there are
+  pending subgoals.
 \<close>
 
 
 section \<open>Tactics\label{sec:tactics}\<close>
 
-text \<open>A \<open>tactic\<close> is a function \<open>goal \<rightarrow> goal\<^sup>*\<^sup>*\<close> that
-  maps a given goal state (represented as a theorem, cf.\
-  \secref{sec:tactical-goals}) to a lazy sequence of potential
-  successor states.  The underlying sequence implementation is lazy
-  both in head and tail, and is purely functional in \<^emph>\<open>not\<close>
-  supporting memoing.\<^footnote>\<open>The lack of memoing and the strict
-  nature of ML requires some care when working with low-level
-  sequence operations, to avoid duplicate or premature evaluation of
-  results.  It also means that modified runtime behavior, such as
-  timeout, is very hard to achieve for general tactics.\<close>
+text \<open>
+  A \<open>tactic\<close> is a function \<open>goal \<rightarrow> goal\<^sup>*\<^sup>*\<close> that maps a given goal state
+  (represented as a theorem, cf.\ \secref{sec:tactical-goals}) to a lazy
+  sequence of potential successor states. The underlying sequence
+  implementation is lazy both in head and tail, and is purely functional in
+  \<^emph>\<open>not\<close> supporting memoing.\<^footnote>\<open>The lack of memoing and the strict nature of ML
+  requires some care when working with low-level sequence operations, to avoid
+  duplicate or premature evaluation of results. It also means that modified
+  runtime behavior, such as timeout, is very hard to achieve for general
+  tactics.\<close>
 
-  An \<^emph>\<open>empty result sequence\<close> means that the tactic has failed: in
-  a compound tactic expression other tactics might be tried instead,
-  or the whole refinement step might fail outright, producing a
-  toplevel error message in the end.  When implementing tactics from
-  scratch, one should take care to observe the basic protocol of
-  mapping regular error conditions to an empty result; only serious
-  faults should emerge as exceptions.
+  An \<^emph>\<open>empty result sequence\<close> means that the tactic has failed: in a compound
+  tactic expression other tactics might be tried instead, or the whole
+  refinement step might fail outright, producing a toplevel error message in
+  the end. When implementing tactics from scratch, one should take care to
+  observe the basic protocol of mapping regular error conditions to an empty
+  result; only serious faults should emerge as exceptions.
 
-  By enumerating \<^emph>\<open>multiple results\<close>, a tactic can easily express
-  the potential outcome of an internal search process.  There are also
-  combinators for building proof tools that involve search
-  systematically, see also \secref{sec:tacticals}.
+  By enumerating \<^emph>\<open>multiple results\<close>, a tactic can easily express the
+  potential outcome of an internal search process. There are also combinators
+  for building proof tools that involve search systematically, see also
+  \secref{sec:tacticals}.
 
   \<^medskip>
-  As explained before, a goal state essentially consists of a
-  list of subgoals that imply the main goal (conclusion).  Tactics may
-  operate on all subgoals or on a particularly specified subgoal, but
-  must not change the main conclusion (apart from instantiating
-  schematic goal variables).
+  As explained before, a goal state essentially consists of a list of subgoals
+  that imply the main goal (conclusion). Tactics may operate on all subgoals
+  or on a particularly specified subgoal, but must not change the main
+  conclusion (apart from instantiating schematic goal variables).
 
-  Tactics with explicit \<^emph>\<open>subgoal addressing\<close> are of the form
-  \<open>int \<rightarrow> tactic\<close> and may be applied to a particular subgoal
-  (counting from 1).  If the subgoal number is out of range, the
-  tactic should fail with an empty result sequence, but must not raise
-  an exception!
+  Tactics with explicit \<^emph>\<open>subgoal addressing\<close> are of the form \<open>int \<rightarrow> tactic\<close>
+  and may be applied to a particular subgoal (counting from 1). If the subgoal
+  number is out of range, the tactic should fail with an empty result
+  sequence, but must not raise an exception!
 
-  Operating on a particular subgoal means to replace it by an interval
-  of zero or more subgoals in the same place; other subgoals must not
-  be affected, apart from instantiating schematic variables ranging
-  over the whole goal state.
+  Operating on a particular subgoal means to replace it by an interval of zero
+  or more subgoals in the same place; other subgoals must not be affected,
+  apart from instantiating schematic variables ranging over the whole goal
+  state.
 
-  A common pattern of composing tactics with subgoal addressing is to
-  try the first one, and then the second one only if the subgoal has
-  not been solved yet.  Special care is required here to avoid bumping
-  into unrelated subgoals that happen to come after the original
-  subgoal.  Assuming that there is only a single initial subgoal is a
-  very common error when implementing tactics!
+  A common pattern of composing tactics with subgoal addressing is to try the
+  first one, and then the second one only if the subgoal has not been solved
+  yet. Special care is required here to avoid bumping into unrelated subgoals
+  that happen to come after the original subgoal. Assuming that there is only
+  a single initial subgoal is a very common error when implementing tactics!
 
-  Tactics with internal subgoal addressing should expose the subgoal
-  index as \<open>int\<close> argument in full generality; a hardwired
-  subgoal 1 is not acceptable.
+  Tactics with internal subgoal addressing should expose the subgoal index as
+  \<open>int\<close> argument in full generality; a hardwired subgoal 1 is not acceptable.
   
   \<^medskip>
-  The main well-formedness conditions for proper tactics are
-  summarized as follows.
+  The main well-formedness conditions for proper tactics are summarized as
+  follows.
 
-  \<^item> General tactic failure is indicated by an empty result, only
-  serious faults may produce an exception.
+    \<^item> General tactic failure is indicated by an empty result, only serious
+    faults may produce an exception.
 
-  \<^item> The main conclusion must not be changed, apart from
-  instantiating schematic variables.
+    \<^item> The main conclusion must not be changed, apart from instantiating
+    schematic variables.
 
-  \<^item> A tactic operates either uniformly on all subgoals, or
-  specifically on a selected subgoal (without bumping into unrelated
-  subgoals).
+    \<^item> A tactic operates either uniformly on all subgoals, or specifically on a
+    selected subgoal (without bumping into unrelated subgoals).
 
-  \<^item> Range errors in subgoal addressing produce an empty result.
-
+    \<^item> Range errors in subgoal addressing produce an empty result.
 
-  Some of these conditions are checked by higher-level goal
-  infrastructure (\secref{sec:struct-goals}); others are not checked
-  explicitly, and violating them merely results in ill-behaved tactics
-  experienced by the user (e.g.\ tactics that insist in being
-  applicable only to singleton goals, or prevent composition via
-  standard tacticals such as @{ML REPEAT}).
+  Some of these conditions are checked by higher-level goal infrastructure
+  (\secref{sec:struct-goals}); others are not checked explicitly, and
+  violating them merely results in ill-behaved tactics experienced by the user
+  (e.g.\ tactics that insist in being applicable only to singleton goals, or
+  prevent composition via standard tacticals such as @{ML REPEAT}).
 \<close>
 
 text %mlref \<open>
@@ -178,90 +167,79 @@
   @{index_ML PREFER_GOAL: "tactic -> int -> tactic"} \\
   \end{mldecls}
 
-  \<^descr> Type @{ML_type tactic} represents tactics.  The
-  well-formedness conditions described above need to be observed.  See
-  also @{file "~~/src/Pure/General/seq.ML"} for the underlying
-  implementation of lazy sequences.
+  \<^descr> Type @{ML_type tactic} represents tactics. The well-formedness conditions
+  described above need to be observed. See also @{file
+  "~~/src/Pure/General/seq.ML"} for the underlying implementation of lazy
+  sequences.
 
-  \<^descr> Type @{ML_type "int -> tactic"} represents tactics with
-  explicit subgoal addressing, with well-formedness conditions as
-  described above.
+  \<^descr> Type @{ML_type "int -> tactic"} represents tactics with explicit subgoal
+  addressing, with well-formedness conditions as described above.
 
-  \<^descr> @{ML no_tac} is a tactic that always fails, returning the
-  empty sequence.
+  \<^descr> @{ML no_tac} is a tactic that always fails, returning the empty sequence.
 
-  \<^descr> @{ML all_tac} is a tactic that always succeeds, returning a
-  singleton sequence with unchanged goal state.
+  \<^descr> @{ML all_tac} is a tactic that always succeeds, returning a singleton
+  sequence with unchanged goal state.
 
-  \<^descr> @{ML print_tac}~\<open>ctxt message\<close> is like @{ML all_tac}, but
-  prints a message together with the goal state on the tracing
-  channel.
+  \<^descr> @{ML print_tac}~\<open>ctxt message\<close> is like @{ML all_tac}, but prints a message
+  together with the goal state on the tracing channel.
 
-  \<^descr> @{ML PRIMITIVE}~\<open>rule\<close> turns a primitive inference rule
-  into a tactic with unique result.  Exception @{ML THM} is considered
-  a regular tactic failure and produces an empty result; other
-  exceptions are passed through.
+  \<^descr> @{ML PRIMITIVE}~\<open>rule\<close> turns a primitive inference rule into a tactic with
+  unique result. Exception @{ML THM} is considered a regular tactic failure
+  and produces an empty result; other exceptions are passed through.
 
-  \<^descr> @{ML SUBGOAL}~\<open>(fn (subgoal, i) => tactic)\<close> is the
-  most basic form to produce a tactic with subgoal addressing.  The
-  given abstraction over the subgoal term and subgoal number allows to
-  peek at the relevant information of the full goal state.  The
-  subgoal range is checked as required above.
+  \<^descr> @{ML SUBGOAL}~\<open>(fn (subgoal, i) => tactic)\<close> is the most basic form to
+  produce a tactic with subgoal addressing. The given abstraction over the
+  subgoal term and subgoal number allows to peek at the relevant information
+  of the full goal state. The subgoal range is checked as required above.
 
-  \<^descr> @{ML CSUBGOAL} is similar to @{ML SUBGOAL}, but passes the
-  subgoal as @{ML_type cterm} instead of raw @{ML_type term}.  This
-  avoids expensive re-certification in situations where the subgoal is
-  used directly for primitive inferences.
+  \<^descr> @{ML CSUBGOAL} is similar to @{ML SUBGOAL}, but passes the subgoal as
+  @{ML_type cterm} instead of raw @{ML_type term}. This avoids expensive
+  re-certification in situations where the subgoal is used directly for
+  primitive inferences.
 
-  \<^descr> @{ML SELECT_GOAL}~\<open>tac i\<close> confines a tactic to the
-  specified subgoal \<open>i\<close>.  This rearranges subgoals and the
-  main goal protection (\secref{sec:tactical-goals}), while retaining
-  the syntactic context of the overall goal state (concerning
-  schematic variables etc.).
+  \<^descr> @{ML SELECT_GOAL}~\<open>tac i\<close> confines a tactic to the specified subgoal \<open>i\<close>.
+  This rearranges subgoals and the main goal protection
+  (\secref{sec:tactical-goals}), while retaining the syntactic context of the
+  overall goal state (concerning schematic variables etc.).
 
-  \<^descr> @{ML PREFER_GOAL}~\<open>tac i\<close> rearranges subgoals to put
-  \<open>i\<close> in front.  This is similar to @{ML SELECT_GOAL}, but
-  without changing the main goal protection.
+  \<^descr> @{ML PREFER_GOAL}~\<open>tac i\<close> rearranges subgoals to put \<open>i\<close> in front. This is
+  similar to @{ML SELECT_GOAL}, but without changing the main goal protection.
 \<close>
 
 
 subsection \<open>Resolution and assumption tactics \label{sec:resolve-assume-tac}\<close>
 
-text \<open>\<^emph>\<open>Resolution\<close> is the most basic mechanism for refining a
-  subgoal using a theorem as object-level rule.
-  \<^emph>\<open>Elim-resolution\<close> is particularly suited for elimination rules:
-  it resolves with a rule, proves its first premise by assumption, and
-  finally deletes that assumption from any new subgoals.
-  \<^emph>\<open>Destruct-resolution\<close> is like elim-resolution, but the given
-  destruction rules are first turned into canonical elimination
-  format.  \<^emph>\<open>Forward-resolution\<close> is like destruct-resolution, but
-  without deleting the selected assumption.  The \<open>r/e/d/f\<close>
-  naming convention is maintained for several different kinds of
-  resolution rules and tactics.
+text \<open>
+  \<^emph>\<open>Resolution\<close> is the most basic mechanism for refining a subgoal using a
+  theorem as object-level rule. \<^emph>\<open>Elim-resolution\<close> is particularly suited for
+  elimination rules: it resolves with a rule, proves its first premise by
+  assumption, and finally deletes that assumption from any new subgoals.
+  \<^emph>\<open>Destruct-resolution\<close> is like elim-resolution, but the given destruction
+  rules are first turned into canonical elimination format.
+  \<^emph>\<open>Forward-resolution\<close> is like destruct-resolution, but without deleting the
+  selected assumption. The \<open>r/e/d/f\<close> naming convention is maintained for
+  several different kinds of resolution rules and tactics.
 
-  Assumption tactics close a subgoal by unifying some of its premises
-  against its conclusion.
+  Assumption tactics close a subgoal by unifying some of its premises against
+  its conclusion.
 
   \<^medskip>
-  All the tactics in this section operate on a subgoal
-  designated by a positive integer.  Other subgoals might be affected
-  indirectly, due to instantiation of schematic variables.
+  All the tactics in this section operate on a subgoal designated by a
+  positive integer. Other subgoals might be affected indirectly, due to
+  instantiation of schematic variables.
 
-  There are various sources of non-determinism, the tactic result
-  sequence enumerates all possibilities of the following choices (if
-  applicable):
-
-  \<^enum> selecting one of the rules given as argument to the tactic;
+  There are various sources of non-determinism, the tactic result sequence
+  enumerates all possibilities of the following choices (if applicable):
 
-  \<^enum> selecting a subgoal premise to eliminate, unifying it against
-  the first premise of the rule;
+    \<^enum> selecting one of the rules given as argument to the tactic;
+
+    \<^enum> selecting a subgoal premise to eliminate, unifying it against the first
+    premise of the rule;
 
-  \<^enum> unifying the conclusion of the subgoal to the conclusion of
-  the rule.
+    \<^enum> unifying the conclusion of the subgoal to the conclusion of the rule.
 
-
-  Recall that higher-order unification may produce multiple results
-  that are enumerated here.
+  Recall that higher-order unification may produce multiple results that are
+  enumerated here.
 \<close>
 
 text %mlref \<open>
@@ -279,55 +257,53 @@
   @{index_ML bimatch_tac: "Proof.context -> (bool * thm) list -> int -> tactic"} \\
   \end{mldecls}
 
-  \<^descr> @{ML resolve_tac}~\<open>ctxt thms i\<close> refines the goal state
-  using the given theorems, which should normally be introduction
-  rules.  The tactic resolves a rule's conclusion with subgoal \<open>i\<close>, replacing it by the corresponding versions of the rule's
-  premises.
+  \<^descr> @{ML resolve_tac}~\<open>ctxt thms i\<close> refines the goal state using the given
+  theorems, which should normally be introduction rules. The tactic resolves a
+  rule's conclusion with subgoal \<open>i\<close>, replacing it by the corresponding
+  versions of the rule's premises.
 
-  \<^descr> @{ML eresolve_tac}~\<open>ctxt thms i\<close> performs elim-resolution
-  with the given theorems, which are normally be elimination rules.
+  \<^descr> @{ML eresolve_tac}~\<open>ctxt thms i\<close> performs elim-resolution with the given
+  theorems, which are normally be elimination rules.
 
   Note that @{ML_text "eresolve_tac ctxt [asm_rl]"} is equivalent to @{ML_text
   "assume_tac ctxt"}, which facilitates mixing of assumption steps with
   genuine eliminations.
 
-  \<^descr> @{ML dresolve_tac}~\<open>ctxt thms i\<close> performs
-  destruct-resolution with the given theorems, which should normally
-  be destruction rules.  This replaces an assumption by the result of
-  applying one of the rules.
+  \<^descr> @{ML dresolve_tac}~\<open>ctxt thms i\<close> performs destruct-resolution with the
+  given theorems, which should normally be destruction rules. This replaces an
+  assumption by the result of applying one of the rules.
 
-  \<^descr> @{ML forward_tac} is like @{ML dresolve_tac} except that the
-  selected assumption is not deleted.  It applies a rule to an
-  assumption, adding the result as a new assumption.
+  \<^descr> @{ML forward_tac} is like @{ML dresolve_tac} except that the selected
+  assumption is not deleted. It applies a rule to an assumption, adding the
+  result as a new assumption.
 
-  \<^descr> @{ML biresolve_tac}~\<open>ctxt brls i\<close> refines the proof state
-  by resolution or elim-resolution on each rule, as indicated by its
-  flag.  It affects subgoal \<open>i\<close> of the proof state.
+  \<^descr> @{ML biresolve_tac}~\<open>ctxt brls i\<close> refines the proof state by resolution or
+  elim-resolution on each rule, as indicated by its flag. It affects subgoal
+  \<open>i\<close> of the proof state.
 
-  For each pair \<open>(flag, rule)\<close>, it applies resolution if the
-  flag is \<open>false\<close> and elim-resolution if the flag is \<open>true\<close>.  A single tactic call handles a mixture of introduction and
-  elimination rules, which is useful to organize the search process
-  systematically in proof tools.
+  For each pair \<open>(flag, rule)\<close>, it applies resolution if the flag is \<open>false\<close>
+  and elim-resolution if the flag is \<open>true\<close>. A single tactic call handles a
+  mixture of introduction and elimination rules, which is useful to organize
+  the search process systematically in proof tools.
 
-  \<^descr> @{ML assume_tac}~\<open>ctxt i\<close> attempts to solve subgoal \<open>i\<close>
-  by assumption (modulo higher-order unification).
+  \<^descr> @{ML assume_tac}~\<open>ctxt i\<close> attempts to solve subgoal \<open>i\<close> by assumption
+  (modulo higher-order unification).
 
-  \<^descr> @{ML eq_assume_tac} is similar to @{ML assume_tac}, but checks
-  only for immediate \<open>\<alpha>\<close>-convertibility instead of using
-  unification.  It succeeds (with a unique next state) if one of the
-  assumptions is equal to the subgoal's conclusion.  Since it does not
-  instantiate variables, it cannot make other subgoals unprovable.
+  \<^descr> @{ML eq_assume_tac} is similar to @{ML assume_tac}, but checks only for
+  immediate \<open>\<alpha>\<close>-convertibility instead of using unification. It succeeds (with
+  a unique next state) if one of the assumptions is equal to the subgoal's
+  conclusion. Since it does not instantiate variables, it cannot make other
+  subgoals unprovable.
 
-  \<^descr> @{ML match_tac}, @{ML ematch_tac}, @{ML dmatch_tac}, and @{ML
-  bimatch_tac} are similar to @{ML resolve_tac}, @{ML eresolve_tac},
-  @{ML dresolve_tac}, and @{ML biresolve_tac}, respectively, but do
-  not instantiate schematic variables in the goal state.\<^footnote>\<open>Strictly speaking,
-  matching means to treat the unknowns in the goal
-  state as constants, but these tactics merely discard unifiers that would
-  update the goal state. In rare situations (where the conclusion and 
-  goal state have flexible terms at the same position), the tactic
-  will fail even though an acceptable unifier exists.\<close>
-  These tactics were written for a specific application within the classical reasoner.
+  \<^descr> @{ML match_tac}, @{ML ematch_tac}, @{ML dmatch_tac}, and @{ML bimatch_tac}
+  are similar to @{ML resolve_tac}, @{ML eresolve_tac}, @{ML dresolve_tac},
+  and @{ML biresolve_tac}, respectively, but do not instantiate schematic
+  variables in the goal state.\<^footnote>\<open>Strictly speaking, matching means to treat the
+  unknowns in the goal state as constants, but these tactics merely discard
+  unifiers that would update the goal state. In rare situations (where the
+  conclusion and goal state have flexible terms at the same position), the
+  tactic will fail even though an acceptable unifier exists.\<close> These tactics
+  were written for a specific application within the classical reasoner.
 
   Flexible subgoals are not updated at will, but are left alone.
 \<close>
@@ -335,42 +311,44 @@
 
 subsection \<open>Explicit instantiation within a subgoal context\<close>
 
-text \<open>The main resolution tactics (\secref{sec:resolve-assume-tac})
-  use higher-order unification, which works well in many practical
-  situations despite its daunting theoretical properties.
-  Nonetheless, there are important problem classes where unguided
-  higher-order unification is not so useful.  This typically involves
-  rules like universal elimination, existential introduction, or
-  equational substitution.  Here the unification problem involves
-  fully flexible \<open>?P ?x\<close> schemes, which are hard to manage
+text \<open>
+  The main resolution tactics (\secref{sec:resolve-assume-tac}) use
+  higher-order unification, which works well in many practical situations
+  despite its daunting theoretical properties. Nonetheless, there are
+  important problem classes where unguided higher-order unification is not so
+  useful. This typically involves rules like universal elimination,
+  existential introduction, or equational substitution. Here the unification
+  problem involves fully flexible \<open>?P ?x\<close> schemes, which are hard to manage
   without further hints.
 
-  By providing a (small) rigid term for \<open>?x\<close> explicitly, the
-  remaining unification problem is to assign a (large) term to \<open>?P\<close>, according to the shape of the given subgoal.  This is
-  sufficiently well-behaved in most practical situations.
+  By providing a (small) rigid term for \<open>?x\<close> explicitly, the remaining
+  unification problem is to assign a (large) term to \<open>?P\<close>, according to the
+  shape of the given subgoal. This is sufficiently well-behaved in most
+  practical situations.
 
   \<^medskip>
-  Isabelle provides separate versions of the standard \<open>r/e/d/f\<close> resolution tactics that allow to provide explicit
-  instantiations of unknowns of the given rule, wrt.\ terms that refer
-  to the implicit context of the selected subgoal.
+  Isabelle provides separate versions of the standard \<open>r/e/d/f\<close> resolution
+  tactics that allow to provide explicit instantiations of unknowns of the
+  given rule, wrt.\ terms that refer to the implicit context of the selected
+  subgoal.
 
-  An instantiation consists of a list of pairs of the form \<open>(?x, t)\<close>, where \<open>?x\<close> is a schematic variable occurring in
-  the given rule, and \<open>t\<close> is a term from the current proof
-  context, augmented by the local goal parameters of the selected
-  subgoal; cf.\ the \<open>focus\<close> operation described in
+  An instantiation consists of a list of pairs of the form \<open>(?x, t)\<close>, where
+  \<open>?x\<close> is a schematic variable occurring in the given rule, and \<open>t\<close> is a term
+  from the current proof context, augmented by the local goal parameters of
+  the selected subgoal; cf.\ the \<open>focus\<close> operation described in
   \secref{sec:variables}.
 
-  Entering the syntactic context of a subgoal is a brittle operation,
-  because its exact form is somewhat accidental, and the choice of
-  bound variable names depends on the presence of other local and
-  global names.  Explicit renaming of subgoal parameters prior to
-  explicit instantiation might help to achieve a bit more robustness.
+  Entering the syntactic context of a subgoal is a brittle operation, because
+  its exact form is somewhat accidental, and the choice of bound variable
+  names depends on the presence of other local and global names. Explicit
+  renaming of subgoal parameters prior to explicit instantiation might help to
+  achieve a bit more robustness.
 
-  Type instantiations may be given as well, via pairs like \<open>(?'a, \<tau>)\<close>.  Type instantiations are distinguished from term
-  instantiations by the syntactic form of the schematic variable.
-  Types are instantiated before terms are.  Since term instantiation
-  already performs simple type-inference, so explicit type
-  instantiations are seldom necessary.
+  Type instantiations may be given as well, via pairs like \<open>(?'a, \<tau>)\<close>. Type
+  instantiations are distinguished from term instantiations by the syntactic
+  form of the schematic variable. Types are instantiated before terms are.
+  Since term instantiation already performs simple type-inference, so explicit
+  type instantiations are seldom necessary.
 \<close>
 
 text %mlref \<open>
@@ -394,48 +372,48 @@
   @{index_ML rename_tac: "string list -> int -> tactic"} \\
   \end{mldecls}
 
-  \<^descr> @{ML Rule_Insts.res_inst_tac}~\<open>ctxt insts thm i\<close> instantiates the
-  rule \<open>thm\<close> with the instantiations \<open>insts\<close>, as described
-  above, and then performs resolution on subgoal \<open>i\<close>.
+  \<^descr> @{ML Rule_Insts.res_inst_tac}~\<open>ctxt insts thm i\<close> instantiates the rule
+  \<open>thm\<close> with the instantiations \<open>insts\<close>, as described above, and then performs
+  resolution on subgoal \<open>i\<close>.
   
-  \<^descr> @{ML Rule_Insts.eres_inst_tac} is like @{ML Rule_Insts.res_inst_tac},
-  but performs elim-resolution.
+  \<^descr> @{ML Rule_Insts.eres_inst_tac} is like @{ML Rule_Insts.res_inst_tac}, but
+  performs elim-resolution.
 
-  \<^descr> @{ML Rule_Insts.dres_inst_tac} is like @{ML Rule_Insts.res_inst_tac},
-  but performs destruct-resolution.
+  \<^descr> @{ML Rule_Insts.dres_inst_tac} is like @{ML Rule_Insts.res_inst_tac}, but
+  performs destruct-resolution.
 
   \<^descr> @{ML Rule_Insts.forw_inst_tac} is like @{ML Rule_Insts.dres_inst_tac}
   except that the selected assumption is not deleted.
 
-  \<^descr> @{ML Rule_Insts.subgoal_tac}~\<open>ctxt \<phi> i\<close> adds the proposition
-  \<open>\<phi>\<close> as local premise to subgoal \<open>i\<close>, and poses the
-  same as a new subgoal \<open>i + 1\<close> (in the original context).
+  \<^descr> @{ML Rule_Insts.subgoal_tac}~\<open>ctxt \<phi> i\<close> adds the proposition \<open>\<phi>\<close> as local
+  premise to subgoal \<open>i\<close>, and poses the same as a new subgoal \<open>i + 1\<close> (in the
+  original context).
 
-  \<^descr> @{ML Rule_Insts.thin_tac}~\<open>ctxt \<phi> i\<close> deletes the specified
-  premise from subgoal \<open>i\<close>.  Note that \<open>\<phi>\<close> may contain
-  schematic variables, to abbreviate the intended proposition; the
-  first matching subgoal premise will be deleted.  Removing useless
-  premises from a subgoal increases its readability and can make
-  search tactics run faster.
+  \<^descr> @{ML Rule_Insts.thin_tac}~\<open>ctxt \<phi> i\<close> deletes the specified premise from
+  subgoal \<open>i\<close>. Note that \<open>\<phi>\<close> may contain schematic variables, to abbreviate
+  the intended proposition; the first matching subgoal premise will be
+  deleted. Removing useless premises from a subgoal increases its readability
+  and can make search tactics run faster.
 
-  \<^descr> @{ML rename_tac}~\<open>names i\<close> renames the innermost
-  parameters of subgoal \<open>i\<close> according to the provided \<open>names\<close> (which need to be distinct identifiers).
+  \<^descr> @{ML rename_tac}~\<open>names i\<close> renames the innermost parameters of subgoal \<open>i\<close>
+  according to the provided \<open>names\<close> (which need to be distinct identifiers).
 
 
-  For historical reasons, the above instantiation tactics take
-  unparsed string arguments, which makes them hard to use in general
-  ML code.  The slightly more advanced @{ML Subgoal.FOCUS} combinator
-  of \secref{sec:struct-goals} allows to refer to internal goal
-  structure with explicit context management.
+  For historical reasons, the above instantiation tactics take unparsed string
+  arguments, which makes them hard to use in general ML code. The slightly
+  more advanced @{ML Subgoal.FOCUS} combinator of \secref{sec:struct-goals}
+  allows to refer to internal goal structure with explicit context management.
 \<close>
 
 
 subsection \<open>Rearranging goal states\<close>
 
-text \<open>In rare situations there is a need to rearrange goal states:
-  either the overall collection of subgoals, or the local structure of
-  a subgoal.  Various administrative tactics allow to operate on the
-  concrete presentation these conceptual sets of formulae.\<close>
+text \<open>
+  In rare situations there is a need to rearrange goal states: either the
+  overall collection of subgoals, or the local structure of a subgoal. Various
+  administrative tactics allow to operate on the concrete presentation these
+  conceptual sets of formulae.
+\<close>
 
 text %mlref \<open>
   \begin{mldecls}
@@ -444,33 +422,32 @@
   @{index_ML flexflex_tac: "Proof.context -> tactic"} \\
   \end{mldecls}
 
-  \<^descr> @{ML rotate_tac}~\<open>n i\<close> rotates the premises of subgoal
-  \<open>i\<close> by \<open>n\<close> positions: from right to left if \<open>n\<close> is
-  positive, and from left to right if \<open>n\<close> is negative.
+  \<^descr> @{ML rotate_tac}~\<open>n i\<close> rotates the premises of subgoal \<open>i\<close> by \<open>n\<close>
+  positions: from right to left if \<open>n\<close> is positive, and from left to right if
+  \<open>n\<close> is negative.
 
-  \<^descr> @{ML distinct_subgoals_tac} removes duplicate subgoals from a
-  proof state.  This is potentially inefficient.
+  \<^descr> @{ML distinct_subgoals_tac} removes duplicate subgoals from a proof state.
+  This is potentially inefficient.
 
-  \<^descr> @{ML flexflex_tac} removes all flex-flex pairs from the proof
-  state by applying the trivial unifier.  This drastic step loses
-  information.  It is already part of the Isar infrastructure for
-  facts resulting from goals, and rarely needs to be invoked manually.
+  \<^descr> @{ML flexflex_tac} removes all flex-flex pairs from the proof state by
+  applying the trivial unifier. This drastic step loses information. It is
+  already part of the Isar infrastructure for facts resulting from goals, and
+  rarely needs to be invoked manually.
 
   Flex-flex constraints arise from difficult cases of higher-order
-  unification.  To prevent this, use @{ML Rule_Insts.res_inst_tac} to
-  instantiate some variables in a rule.  Normally flex-flex constraints
-  can be ignored; they often disappear as unknowns get instantiated.
+  unification. To prevent this, use @{ML Rule_Insts.res_inst_tac} to
+  instantiate some variables in a rule. Normally flex-flex constraints can be
+  ignored; they often disappear as unknowns get instantiated.
 \<close>
 
 
 subsection \<open>Raw composition: resolution without lifting\<close>
 
 text \<open>
-  Raw composition of two rules means resolving them without prior
-  lifting or renaming of unknowns.  This low-level operation, which
-  underlies the resolution tactics, may occasionally be useful for
-  special effects.  Schematic variables are not renamed by default, so
-  beware of clashes!
+  Raw composition of two rules means resolving them without prior lifting or
+  renaming of unknowns. This low-level operation, which underlies the
+  resolution tactics, may occasionally be useful for special effects.
+  Schematic variables are not renamed by default, so beware of clashes!
 \<close>
 
 text %mlref \<open>
@@ -480,50 +457,48 @@
   @{index_ML_op COMP: "thm * thm -> thm"} \\
   \end{mldecls}
 
-  \<^descr> @{ML compose_tac}~\<open>ctxt (flag, rule, m) i\<close> refines subgoal
-  \<open>i\<close> using \<open>rule\<close>, without lifting.  The \<open>rule\<close> is taken to have the form \<open>\<psi>\<^sub>1 \<Longrightarrow> \<dots> \<psi>\<^sub>m \<Longrightarrow> \<psi>\<close>, where
-  \<open>\<psi>\<close> need not be atomic; thus \<open>m\<close> determines the
-  number of new subgoals.  If \<open>flag\<close> is \<open>true\<close> then it
-  performs elim-resolution --- it solves the first premise of \<open>rule\<close> by assumption and deletes that assumption.
+  \<^descr> @{ML compose_tac}~\<open>ctxt (flag, rule, m) i\<close> refines subgoal \<open>i\<close> using
+  \<open>rule\<close>, without lifting. The \<open>rule\<close> is taken to have the form \<open>\<psi>\<^sub>1 \<Longrightarrow> \<dots> \<psi>\<^sub>m \<Longrightarrow>
+  \<psi>\<close>, where \<open>\<psi>\<close> need not be atomic; thus \<open>m\<close> determines the number of new
+  subgoals. If \<open>flag\<close> is \<open>true\<close> then it performs elim-resolution --- it solves
+  the first premise of \<open>rule\<close> by assumption and deletes that assumption.
 
-  \<^descr> @{ML Drule.compose}~\<open>(thm\<^sub>1, i, thm\<^sub>2)\<close> uses \<open>thm\<^sub>1\<close>,
-  regarded as an atomic formula, to solve premise \<open>i\<close> of
-  \<open>thm\<^sub>2\<close>.  Let \<open>thm\<^sub>1\<close> and \<open>thm\<^sub>2\<close> be \<open>\<psi>\<close> and \<open>\<phi>\<^sub>1 \<Longrightarrow> \<dots> \<phi>\<^sub>n \<Longrightarrow> \<phi>\<close>.  The unique \<open>s\<close> that
-  unifies \<open>\<psi>\<close> and \<open>\<phi>\<^sub>i\<close> yields the theorem \<open>(\<phi>\<^sub>1 \<Longrightarrow>
-  \<dots> \<phi>\<^sub>i\<^sub>-\<^sub>1 \<Longrightarrow> \<phi>\<^sub>i\<^sub>+\<^sub>1 \<Longrightarrow> \<dots> \<phi>\<^sub>n \<Longrightarrow> \<phi>)s\<close>.  Multiple results are considered as
-  error (exception @{ML THM}).
+  \<^descr> @{ML Drule.compose}~\<open>(thm\<^sub>1, i, thm\<^sub>2)\<close> uses \<open>thm\<^sub>1\<close>, regarded as an
+  atomic formula, to solve premise \<open>i\<close> of \<open>thm\<^sub>2\<close>. Let \<open>thm\<^sub>1\<close> and \<open>thm\<^sub>2\<close> be
+  \<open>\<psi>\<close> and \<open>\<phi>\<^sub>1 \<Longrightarrow> \<dots> \<phi>\<^sub>n \<Longrightarrow> \<phi>\<close>. The unique \<open>s\<close> that unifies \<open>\<psi>\<close> and \<open>\<phi>\<^sub>i\<close> yields
+  the theorem \<open>(\<phi>\<^sub>1 \<Longrightarrow> \<dots> \<phi>\<^sub>i\<^sub>-\<^sub>1 \<Longrightarrow> \<phi>\<^sub>i\<^sub>+\<^sub>1 \<Longrightarrow> \<dots> \<phi>\<^sub>n \<Longrightarrow> \<phi>)s\<close>. Multiple results are
+  considered as error (exception @{ML THM}).
 
-  \<^descr> \<open>thm\<^sub>1 COMP thm\<^sub>2\<close> is the same as \<open>Drule.compose
-  (thm\<^sub>1, 1, thm\<^sub>2)\<close>.
+  \<^descr> \<open>thm\<^sub>1 COMP thm\<^sub>2\<close> is the same as \<open>Drule.compose (thm\<^sub>1, 1, thm\<^sub>2)\<close>.
 
 
   \begin{warn}
-  These low-level operations are stepping outside the structure
-  imposed by regular rule resolution.  Used without understanding of
-  the consequences, they may produce results that cause problems with
-  standard rules and tactics later on.
+  These low-level operations are stepping outside the structure imposed by
+  regular rule resolution. Used without understanding of the consequences,
+  they may produce results that cause problems with standard rules and tactics
+  later on.
   \end{warn}
 \<close>
 
 
 section \<open>Tacticals \label{sec:tacticals}\<close>
 
-text \<open>A \<^emph>\<open>tactical\<close> is a functional combinator for building up
-  complex tactics from simpler ones.  Common tacticals perform
-  sequential composition, disjunctive choice, iteration, or goal
-  addressing.  Various search strategies may be expressed via
-  tacticals.
+text \<open>
+  A \<^emph>\<open>tactical\<close> is a functional combinator for building up complex tactics
+  from simpler ones. Common tacticals perform sequential composition,
+  disjunctive choice, iteration, or goal addressing. Various search strategies
+  may be expressed via tacticals.
 \<close>
 
 
 subsection \<open>Combining tactics\<close>
 
-text \<open>Sequential composition and alternative choices are the most
-  basic ways to combine tactics, similarly to ``\<^verbatim>\<open>,\<close>'' and
-  ``\<^verbatim>\<open>|\<close>'' in Isar method notation.  This corresponds to
-  @{ML_op "THEN"} and @{ML_op "ORELSE"} in ML, but there are further
-  possibilities for fine-tuning alternation of tactics such as @{ML_op
-  "APPEND"}.  Further details become visible in ML due to explicit
+text \<open>
+  Sequential composition and alternative choices are the most basic ways to
+  combine tactics, similarly to ``\<^verbatim>\<open>,\<close>'' and ``\<^verbatim>\<open>|\<close>'' in Isar method notation.
+  This corresponds to @{ML_op "THEN"} and @{ML_op "ORELSE"} in ML, but there
+  are further possibilities for fine-tuning alternation of tactics such as
+  @{ML_op "APPEND"}. Further details become visible in ML due to explicit
   subgoal addressing.
 \<close>
 
@@ -542,35 +517,34 @@
   @{index_ML "FIRST'": "('a -> tactic) list -> 'a -> tactic"} \\
   \end{mldecls}
 
-  \<^descr> \<open>tac\<^sub>1\<close>~@{ML_op THEN}~\<open>tac\<^sub>2\<close> is the sequential
-  composition of \<open>tac\<^sub>1\<close> and \<open>tac\<^sub>2\<close>.  Applied to a goal
-  state, it returns all states reachable in two steps by applying
-  \<open>tac\<^sub>1\<close> followed by \<open>tac\<^sub>2\<close>.  First, it applies \<open>tac\<^sub>1\<close> to the goal state, getting a sequence of possible next
-  states; then, it applies \<open>tac\<^sub>2\<close> to each of these and
-  concatenates the results to produce again one flat sequence of
-  states.
+  \<^descr> \<open>tac\<^sub>1\<close>~@{ML_op THEN}~\<open>tac\<^sub>2\<close> is the sequential composition of \<open>tac\<^sub>1\<close> and
+  \<open>tac\<^sub>2\<close>. Applied to a goal state, it returns all states reachable in two
+  steps by applying \<open>tac\<^sub>1\<close> followed by \<open>tac\<^sub>2\<close>. First, it applies \<open>tac\<^sub>1\<close> to
+  the goal state, getting a sequence of possible next states; then, it applies
+  \<open>tac\<^sub>2\<close> to each of these and concatenates the results to produce again one
+  flat sequence of states.
 
-  \<^descr> \<open>tac\<^sub>1\<close>~@{ML_op ORELSE}~\<open>tac\<^sub>2\<close> makes a choice
-  between \<open>tac\<^sub>1\<close> and \<open>tac\<^sub>2\<close>.  Applied to a state, it
-  tries \<open>tac\<^sub>1\<close> and returns the result if non-empty; if \<open>tac\<^sub>1\<close> fails then it uses \<open>tac\<^sub>2\<close>.  This is a deterministic
-  choice: if \<open>tac\<^sub>1\<close> succeeds then \<open>tac\<^sub>2\<close> is excluded
-  from the result.
+  \<^descr> \<open>tac\<^sub>1\<close>~@{ML_op ORELSE}~\<open>tac\<^sub>2\<close> makes a choice between \<open>tac\<^sub>1\<close> and
+  \<open>tac\<^sub>2\<close>. Applied to a state, it tries \<open>tac\<^sub>1\<close> and returns the result if
+  non-empty; if \<open>tac\<^sub>1\<close> fails then it uses \<open>tac\<^sub>2\<close>. This is a deterministic
+  choice: if \<open>tac\<^sub>1\<close> succeeds then \<open>tac\<^sub>2\<close> is excluded from the result.
 
-  \<^descr> \<open>tac\<^sub>1\<close>~@{ML_op APPEND}~\<open>tac\<^sub>2\<close> concatenates the
-  possible results of \<open>tac\<^sub>1\<close> and \<open>tac\<^sub>2\<close>.  Unlike
-  @{ML_op "ORELSE"} there is \<^emph>\<open>no commitment\<close> to either tactic, so
-  @{ML_op "APPEND"} helps to avoid incompleteness during search, at
-  the cost of potential inefficiencies.
+  \<^descr> \<open>tac\<^sub>1\<close>~@{ML_op APPEND}~\<open>tac\<^sub>2\<close> concatenates the possible results of
+  \<open>tac\<^sub>1\<close> and \<open>tac\<^sub>2\<close>. Unlike @{ML_op "ORELSE"} there is \<^emph>\<open>no commitment\<close> to
+  either tactic, so @{ML_op "APPEND"} helps to avoid incompleteness during
+  search, at the cost of potential inefficiencies.
 
-  \<^descr> @{ML EVERY}~\<open>[tac\<^sub>1, \<dots>, tac\<^sub>n]\<close> abbreviates \<open>tac\<^sub>1\<close>~@{ML_op THEN}~\<open>\<dots>\<close>~@{ML_op THEN}~\<open>tac\<^sub>n\<close>.
-  Note that @{ML "EVERY []"} is the same as @{ML all_tac}: it always
-  succeeds.
+  \<^descr> @{ML EVERY}~\<open>[tac\<^sub>1, \<dots>, tac\<^sub>n]\<close> abbreviates \<open>tac\<^sub>1\<close>~@{ML_op
+  THEN}~\<open>\<dots>\<close>~@{ML_op THEN}~\<open>tac\<^sub>n\<close>. Note that @{ML "EVERY []"} is the same as
+  @{ML all_tac}: it always succeeds.
 
-  \<^descr> @{ML FIRST}~\<open>[tac\<^sub>1, \<dots>, tac\<^sub>n]\<close> abbreviates \<open>tac\<^sub>1\<close>~@{ML_op ORELSE}~\<open>\<dots>\<close>~@{ML_op "ORELSE"}~\<open>tac\<^sub>n\<close>.  Note that @{ML "FIRST []"} is the same as @{ML no_tac}: it
-  always fails.
+  \<^descr> @{ML FIRST}~\<open>[tac\<^sub>1, \<dots>, tac\<^sub>n]\<close> abbreviates \<open>tac\<^sub>1\<close>~@{ML_op
+  ORELSE}~\<open>\<dots>\<close>~@{ML_op "ORELSE"}~\<open>tac\<^sub>n\<close>. Note that @{ML "FIRST []"} is the
+  same as @{ML no_tac}: it always fails.
 
-  \<^descr> @{ML_op "THEN'"} is the lifted version of @{ML_op "THEN"}, for
-  tactics with explicit subgoal addressing.  So \<open>(tac\<^sub>1\<close>~@{ML_op THEN'}~\<open>tac\<^sub>2) i\<close> is the same as \<open>(tac\<^sub>1 i\<close>~@{ML_op THEN}~\<open>tac\<^sub>2 i)\<close>.
+  \<^descr> @{ML_op "THEN'"} is the lifted version of @{ML_op "THEN"}, for tactics
+  with explicit subgoal addressing. So \<open>(tac\<^sub>1\<close>~@{ML_op THEN'}~\<open>tac\<^sub>2) i\<close> is
+  the same as \<open>(tac\<^sub>1 i\<close>~@{ML_op THEN}~\<open>tac\<^sub>2 i)\<close>.
 
   The other primed tacticals work analogously.
 \<close>
@@ -578,9 +552,10 @@
 
 subsection \<open>Repetition tacticals\<close>
 
-text \<open>These tacticals provide further control over repetition of
-  tactics, beyond the stylized forms of ``\<^verbatim>\<open>?\<close>'' and
-  ``\<^verbatim>\<open>+\<close>'' in Isar method expressions.\<close>
+text \<open>
+  These tacticals provide further control over repetition of tactics, beyond
+  the stylized forms of ``\<^verbatim>\<open>?\<close>'' and ``\<^verbatim>\<open>+\<close>'' in Isar method expressions.
+\<close>
 
 text %mlref \<open>
   \begin{mldecls}
@@ -591,50 +566,44 @@
   @{index_ML "REPEAT_DETERM_N": "int -> tactic -> tactic"} \\
   \end{mldecls}
 
-  \<^descr> @{ML TRY}~\<open>tac\<close> applies \<open>tac\<close> to the goal
-  state and returns the resulting sequence, if non-empty; otherwise it
-  returns the original state.  Thus, it applies \<open>tac\<close> at most
-  once.
+  \<^descr> @{ML TRY}~\<open>tac\<close> applies \<open>tac\<close> to the goal state and returns the resulting
+  sequence, if non-empty; otherwise it returns the original state. Thus, it
+  applies \<open>tac\<close> at most once.
 
-  Note that for tactics with subgoal addressing, the combinator can be
-  applied via functional composition: @{ML "TRY"}~@{ML_op o}~\<open>tac\<close>.
-  There is no need for \<^verbatim>\<open>TRY'\<close>.
+  Note that for tactics with subgoal addressing, the combinator can be applied
+  via functional composition: @{ML "TRY"}~@{ML_op o}~\<open>tac\<close>. There is no need
+  for \<^verbatim>\<open>TRY'\<close>.
 
-  \<^descr> @{ML REPEAT}~\<open>tac\<close> applies \<open>tac\<close> to the goal
-  state and, recursively, to each element of the resulting sequence.
-  The resulting sequence consists of those states that make \<open>tac\<close> fail.  Thus, it applies \<open>tac\<close> as many times as
+  \<^descr> @{ML REPEAT}~\<open>tac\<close> applies \<open>tac\<close> to the goal state and, recursively, to
+  each element of the resulting sequence. The resulting sequence consists of
+  those states that make \<open>tac\<close> fail. Thus, it applies \<open>tac\<close> as many times as
   possible (including zero times), and allows backtracking over each
-  invocation of \<open>tac\<close>.  @{ML REPEAT} is more general than @{ML
-  REPEAT_DETERM}, but requires more space.
+  invocation of \<open>tac\<close>. @{ML REPEAT} is more general than @{ML REPEAT_DETERM},
+  but requires more space.
 
-  \<^descr> @{ML REPEAT1}~\<open>tac\<close> is like @{ML REPEAT}~\<open>tac\<close>
-  but it always applies \<open>tac\<close> at least once, failing if this
-  is impossible.
+  \<^descr> @{ML REPEAT1}~\<open>tac\<close> is like @{ML REPEAT}~\<open>tac\<close> but it always applies \<open>tac\<close>
+  at least once, failing if this is impossible.
 
-  \<^descr> @{ML REPEAT_DETERM}~\<open>tac\<close> applies \<open>tac\<close> to the
-  goal state and, recursively, to the head of the resulting sequence.
-  It returns the first state to make \<open>tac\<close> fail.  It is
-  deterministic, discarding alternative outcomes.
+  \<^descr> @{ML REPEAT_DETERM}~\<open>tac\<close> applies \<open>tac\<close> to the goal state and,
+  recursively, to the head of the resulting sequence. It returns the first
+  state to make \<open>tac\<close> fail. It is deterministic, discarding alternative
+  outcomes.
 
-  \<^descr> @{ML REPEAT_DETERM_N}~\<open>n tac\<close> is like @{ML
-  REPEAT_DETERM}~\<open>tac\<close> but the number of repetitions is bound
-  by \<open>n\<close> (where @{ML "~1"} means \<open>\<infinity>\<close>).
+  \<^descr> @{ML REPEAT_DETERM_N}~\<open>n tac\<close> is like @{ML REPEAT_DETERM}~\<open>tac\<close> but the
+  number of repetitions is bound by \<open>n\<close> (where @{ML "~1"} means \<open>\<infinity>\<close>).
 \<close>
 
-text %mlex \<open>The basic tactics and tacticals considered above follow
-  some algebraic laws:
+text %mlex \<open>
+  The basic tactics and tacticals considered above follow some algebraic laws:
 
-  \<^item> @{ML all_tac} is the identity element of the tactical @{ML_op
-  "THEN"}.
+  \<^item> @{ML all_tac} is the identity element of the tactical @{ML_op "THEN"}.
 
-  \<^item> @{ML no_tac} is the identity element of @{ML_op "ORELSE"} and
-  @{ML_op "APPEND"}.  Also, it is a zero element for @{ML_op "THEN"},
-  which means that \<open>tac\<close>~@{ML_op THEN}~@{ML no_tac} is
-  equivalent to @{ML no_tac}.
+  \<^item> @{ML no_tac} is the identity element of @{ML_op "ORELSE"} and @{ML_op
+  "APPEND"}. Also, it is a zero element for @{ML_op "THEN"}, which means that
+  \<open>tac\<close>~@{ML_op THEN}~@{ML no_tac} is equivalent to @{ML no_tac}.
 
-  \<^item> @{ML TRY} and @{ML REPEAT} can be expressed as (recursive)
-  functions over more basic combinators (ignoring some internal
-  implementation tricks):
+  \<^item> @{ML TRY} and @{ML REPEAT} can be expressed as (recursive) functions over
+  more basic combinators (ignoring some internal implementation tricks):
 \<close>
 
 ML \<open>
@@ -642,17 +611,16 @@
   fun REPEAT tac st = ((tac THEN REPEAT tac) ORELSE all_tac) st;
 \<close>
 
-text \<open>If \<open>tac\<close> can return multiple outcomes then so can @{ML
-  REPEAT}~\<open>tac\<close>.  @{ML REPEAT} uses @{ML_op "ORELSE"} and not
-  @{ML_op "APPEND"}, it applies \<open>tac\<close> as many times as
-  possible in each outcome.
+text \<open>
+  If \<open>tac\<close> can return multiple outcomes then so can @{ML REPEAT}~\<open>tac\<close>. @{ML
+  REPEAT} uses @{ML_op "ORELSE"} and not @{ML_op "APPEND"}, it applies \<open>tac\<close>
+  as many times as possible in each outcome.
 
   \begin{warn}
-  Note the explicit abstraction over the goal state in the ML
-  definition of @{ML REPEAT}.  Recursive tacticals must be coded in
-  this awkward fashion to avoid infinite recursion of eager functional
-  evaluation in Standard ML.  The following attempt would make @{ML
-  REPEAT}~\<open>tac\<close> loop:
+  Note the explicit abstraction over the goal state in the ML definition of
+  @{ML REPEAT}. Recursive tacticals must be coded in this awkward fashion to
+  avoid infinite recursion of eager functional evaluation in Standard ML. The
+  following attempt would make @{ML REPEAT}~\<open>tac\<close> loop:
   \end{warn}
 \<close>
 
@@ -664,18 +632,18 @@
 
 subsection \<open>Applying tactics to subgoal ranges\<close>
 
-text \<open>Tactics with explicit subgoal addressing
-  @{ML_type "int -> tactic"} can be used together with tacticals that
-  act like ``subgoal quantifiers'': guided by success of the body
-  tactic a certain range of subgoals is covered.  Thus the body tactic
-  is applied to \<^emph>\<open>all\<close> subgoals, \<^emph>\<open>some\<close> subgoal etc.
+text \<open>
+  Tactics with explicit subgoal addressing @{ML_type "int -> tactic"} can be
+  used together with tacticals that act like ``subgoal quantifiers'': guided
+  by success of the body tactic a certain range of subgoals is covered. Thus
+  the body tactic is applied to \<^emph>\<open>all\<close> subgoals, \<^emph>\<open>some\<close> subgoal etc.
 
-  Suppose that the goal state has \<open>n \<ge> 0\<close> subgoals.  Many of
-  these tacticals address subgoal ranges counting downwards from
-  \<open>n\<close> towards \<open>1\<close>.  This has the fortunate effect that
-  newly emerging subgoals are concatenated in the result, without
-  interfering each other.  Nonetheless, there might be situations
-  where a different order is desired.\<close>
+  Suppose that the goal state has \<open>n \<ge> 0\<close> subgoals. Many of these tacticals
+  address subgoal ranges counting downwards from \<open>n\<close> towards \<open>1\<close>. This has the
+  fortunate effect that newly emerging subgoals are concatenated in the
+  result, without interfering each other. Nonetheless, there might be
+  situations where a different order is desired.
+\<close>
 
 text %mlref \<open>
   \begin{mldecls}
@@ -688,43 +656,40 @@
   @{index_ML RANGE: "(int -> tactic) list -> int -> tactic"} \\
   \end{mldecls}
 
-  \<^descr> @{ML ALLGOALS}~\<open>tac\<close> is equivalent to \<open>tac
-  n\<close>~@{ML_op THEN}~\<open>\<dots>\<close>~@{ML_op THEN}~\<open>tac 1\<close>.  It
-  applies the \<open>tac\<close> to all the subgoals, counting downwards.
+  \<^descr> @{ML ALLGOALS}~\<open>tac\<close> is equivalent to \<open>tac n\<close>~@{ML_op THEN}~\<open>\<dots>\<close>~@{ML_op
+  THEN}~\<open>tac 1\<close>. It applies the \<open>tac\<close> to all the subgoals, counting downwards.
 
-  \<^descr> @{ML SOMEGOAL}~\<open>tac\<close> is equivalent to \<open>tac
-  n\<close>~@{ML_op ORELSE}~\<open>\<dots>\<close>~@{ML_op ORELSE}~\<open>tac 1\<close>.  It
-  applies \<open>tac\<close> to one subgoal, counting downwards.
+  \<^descr> @{ML SOMEGOAL}~\<open>tac\<close> is equivalent to \<open>tac n\<close>~@{ML_op ORELSE}~\<open>\<dots>\<close>~@{ML_op
+  ORELSE}~\<open>tac 1\<close>. It applies \<open>tac\<close> to one subgoal, counting downwards.
 
-  \<^descr> @{ML FIRSTGOAL}~\<open>tac\<close> is equivalent to \<open>tac
-  1\<close>~@{ML_op ORELSE}~\<open>\<dots>\<close>~@{ML_op ORELSE}~\<open>tac n\<close>.  It
-  applies \<open>tac\<close> to one subgoal, counting upwards.
+  \<^descr> @{ML FIRSTGOAL}~\<open>tac\<close> is equivalent to \<open>tac 1\<close>~@{ML_op ORELSE}~\<open>\<dots>\<close>~@{ML_op
+  ORELSE}~\<open>tac n\<close>. It applies \<open>tac\<close> to one subgoal, counting upwards.
 
-  \<^descr> @{ML HEADGOAL}~\<open>tac\<close> is equivalent to \<open>tac 1\<close>.
-  It applies \<open>tac\<close> unconditionally to the first subgoal.
+  \<^descr> @{ML HEADGOAL}~\<open>tac\<close> is equivalent to \<open>tac 1\<close>. It applies \<open>tac\<close>
+  unconditionally to the first subgoal.
 
-  \<^descr> @{ML REPEAT_SOME}~\<open>tac\<close> applies \<open>tac\<close> once or
-  more to a subgoal, counting downwards.
+  \<^descr> @{ML REPEAT_SOME}~\<open>tac\<close> applies \<open>tac\<close> once or more to a subgoal, counting
+  downwards.
 
-  \<^descr> @{ML REPEAT_FIRST}~\<open>tac\<close> applies \<open>tac\<close> once or
-  more to a subgoal, counting upwards.
+  \<^descr> @{ML REPEAT_FIRST}~\<open>tac\<close> applies \<open>tac\<close> once or more to a subgoal, counting
+  upwards.
 
-  \<^descr> @{ML RANGE}~\<open>[tac\<^sub>1, \<dots>, tac\<^sub>k] i\<close> is equivalent to
-  \<open>tac\<^sub>k (i + k - 1)\<close>~@{ML_op THEN}~\<open>\<dots>\<close>~@{ML_op
-  THEN}~\<open>tac\<^sub>1 i\<close>.  It applies the given list of tactics to the
-  corresponding range of subgoals, counting downwards.
+  \<^descr> @{ML RANGE}~\<open>[tac\<^sub>1, \<dots>, tac\<^sub>k] i\<close> is equivalent to \<open>tac\<^sub>k (i + k -
+  1)\<close>~@{ML_op THEN}~\<open>\<dots>\<close>~@{ML_op THEN}~\<open>tac\<^sub>1 i\<close>. It applies the given list of
+  tactics to the corresponding range of subgoals, counting downwards.
 \<close>
 
 
 subsection \<open>Control and search tacticals\<close>
 
-text \<open>A predicate on theorems @{ML_type "thm -> bool"} can test
-  whether a goal state enjoys some desirable property --- such as
-  having no subgoals.  Tactics that search for satisfactory goal
-  states are easy to express.  The main search procedures,
-  depth-first, breadth-first and best-first, are provided as
-  tacticals.  They generate the search tree by repeatedly applying a
-  given tactic.\<close>
+text \<open>
+  A predicate on theorems @{ML_type "thm -> bool"} can test whether a goal
+  state enjoys some desirable property --- such as having no subgoals. Tactics
+  that search for satisfactory goal states are easy to express. The main
+  search procedures, depth-first, breadth-first and best-first, are provided
+  as tacticals. They generate the search tree by repeatedly applying a given
+  tactic.
+\<close>
 
 
 text %mlref ""
@@ -737,14 +702,13 @@
   @{index_ML CHANGED: "tactic -> tactic"} \\
   \end{mldecls}
 
-  \<^descr> @{ML FILTER}~\<open>sat tac\<close> applies \<open>tac\<close> to the
-  goal state and returns a sequence consisting of those result goal
-  states that are satisfactory in the sense of \<open>sat\<close>.
+  \<^descr> @{ML FILTER}~\<open>sat tac\<close> applies \<open>tac\<close> to the goal state and returns a
+  sequence consisting of those result goal states that are satisfactory in the
+  sense of \<open>sat\<close>.
 
-  \<^descr> @{ML CHANGED}~\<open>tac\<close> applies \<open>tac\<close> to the goal
-  state and returns precisely those states that differ from the
-  original state (according to @{ML Thm.eq_thm}).  Thus @{ML
-  CHANGED}~\<open>tac\<close> always has some effect on the state.
+  \<^descr> @{ML CHANGED}~\<open>tac\<close> applies \<open>tac\<close> to the goal state and returns precisely
+  those states that differ from the original state (according to @{ML
+  Thm.eq_thm}). Thus @{ML CHANGED}~\<open>tac\<close> always has some effect on the state.
 \<close>
 
 
@@ -757,19 +721,17 @@
   @{index_ML DEPTH_SOLVE_1: "tactic -> tactic"} \\
   \end{mldecls}
 
-  \<^descr> @{ML DEPTH_FIRST}~\<open>sat tac\<close> returns the goal state if
-  \<open>sat\<close> returns true.  Otherwise it applies \<open>tac\<close>,
-  then recursively searches from each element of the resulting
-  sequence.  The code uses a stack for efficiency, in effect applying
-  \<open>tac\<close>~@{ML_op THEN}~@{ML DEPTH_FIRST}~\<open>sat tac\<close> to
-  the state.
+  \<^descr> @{ML DEPTH_FIRST}~\<open>sat tac\<close> returns the goal state if \<open>sat\<close> returns true.
+  Otherwise it applies \<open>tac\<close>, then recursively searches from each element of
+  the resulting sequence. The code uses a stack for efficiency, in effect
+  applying \<open>tac\<close>~@{ML_op THEN}~@{ML DEPTH_FIRST}~\<open>sat tac\<close> to the state.
 
-  \<^descr> @{ML DEPTH_SOLVE}\<open>tac\<close> uses @{ML DEPTH_FIRST} to
-  search for states having no subgoals.
+  \<^descr> @{ML DEPTH_SOLVE}\<open>tac\<close> uses @{ML DEPTH_FIRST} to search for states having
+  no subgoals.
 
-  \<^descr> @{ML DEPTH_SOLVE_1}~\<open>tac\<close> uses @{ML DEPTH_FIRST} to
-  search for states having fewer subgoals than the given state.  Thus,
-  it insists upon solving at least one subgoal.
+  \<^descr> @{ML DEPTH_SOLVE_1}~\<open>tac\<close> uses @{ML DEPTH_FIRST} to search for states
+  having fewer subgoals than the given state. Thus, it insists upon solving at
+  least one subgoal.
 \<close>
 
 
@@ -782,31 +744,28 @@
   @{index_ML THEN_BEST_FIRST: "tactic -> (thm -> bool) * (thm -> int) -> tactic -> tactic"} \\
   \end{mldecls}
 
-  These search strategies will find a solution if one exists.
-  However, they do not enumerate all solutions; they terminate after
-  the first satisfactory result from \<open>tac\<close>.
+  These search strategies will find a solution if one exists. However, they do
+  not enumerate all solutions; they terminate after the first satisfactory
+  result from \<open>tac\<close>.
 
-  \<^descr> @{ML BREADTH_FIRST}~\<open>sat tac\<close> uses breadth-first
-  search to find states for which \<open>sat\<close> is true.  For most
-  applications, it is too slow.
+  \<^descr> @{ML BREADTH_FIRST}~\<open>sat tac\<close> uses breadth-first search to find states for
+  which \<open>sat\<close> is true. For most applications, it is too slow.
 
-  \<^descr> @{ML BEST_FIRST}~\<open>(sat, dist) tac\<close> does a heuristic
-  search, using \<open>dist\<close> to estimate the distance from a
-  satisfactory state (in the sense of \<open>sat\<close>).  It maintains a
-  list of states ordered by distance.  It applies \<open>tac\<close> to the
-  head of this list; if the result contains any satisfactory states,
-  then it returns them.  Otherwise, @{ML BEST_FIRST} adds the new
-  states to the list, and continues.
+  \<^descr> @{ML BEST_FIRST}~\<open>(sat, dist) tac\<close> does a heuristic search, using \<open>dist\<close>
+  to estimate the distance from a satisfactory state (in the sense of \<open>sat\<close>).
+  It maintains a list of states ordered by distance. It applies \<open>tac\<close> to the
+  head of this list; if the result contains any satisfactory states, then it
+  returns them. Otherwise, @{ML BEST_FIRST} adds the new states to the list,
+  and continues.
 
-  The distance function is typically @{ML size_of_thm}, which computes
-  the size of the state.  The smaller the state, the fewer and simpler
-  subgoals it has.
+  The distance function is typically @{ML size_of_thm}, which computes the
+  size of the state. The smaller the state, the fewer and simpler subgoals it
+  has.
 
-  \<^descr> @{ML THEN_BEST_FIRST}~\<open>tac\<^sub>0 (sat, dist) tac\<close> is like
-  @{ML BEST_FIRST}, except that the priority queue initially contains
-  the result of applying \<open>tac\<^sub>0\<close> to the goal state.  This
-  tactical permits separate tactics for starting the search and
-  continuing the search.
+  \<^descr> @{ML THEN_BEST_FIRST}~\<open>tac\<^sub>0 (sat, dist) tac\<close> is like @{ML BEST_FIRST},
+  except that the priority queue initially contains the result of applying
+  \<open>tac\<^sub>0\<close> to the goal state. This tactical permits separate tactics for
+  starting the search and continuing the search.
 \<close>
 
 
@@ -820,24 +779,22 @@
   @{index_ML DETERM: "tactic -> tactic"} \\
   \end{mldecls}
 
-  \<^descr> @{ML COND}~\<open>sat tac\<^sub>1 tac\<^sub>2\<close> applies \<open>tac\<^sub>1\<close> to
-  the goal state if it satisfies predicate \<open>sat\<close>, and applies
-  \<open>tac\<^sub>2\<close>.  It is a conditional tactical in that only one of
-  \<open>tac\<^sub>1\<close> and \<open>tac\<^sub>2\<close> is applied to a goal state.
-  However, both \<open>tac\<^sub>1\<close> and \<open>tac\<^sub>2\<close> are evaluated
-  because ML uses eager evaluation.
+  \<^descr> @{ML COND}~\<open>sat tac\<^sub>1 tac\<^sub>2\<close> applies \<open>tac\<^sub>1\<close> to the goal state if it
+  satisfies predicate \<open>sat\<close>, and applies \<open>tac\<^sub>2\<close>. It is a conditional tactical
+  in that only one of \<open>tac\<^sub>1\<close> and \<open>tac\<^sub>2\<close> is applied to a goal state. However,
+  both \<open>tac\<^sub>1\<close> and \<open>tac\<^sub>2\<close> are evaluated because ML uses eager evaluation.
 
-  \<^descr> @{ML IF_UNSOLVED}~\<open>tac\<close> applies \<open>tac\<close> to the
-  goal state if it has any subgoals, and simply returns the goal state
-  otherwise.  Many common tactics, such as @{ML resolve_tac}, fail if
-  applied to a goal state that has no subgoals.
+  \<^descr> @{ML IF_UNSOLVED}~\<open>tac\<close> applies \<open>tac\<close> to the goal state if it has any
+  subgoals, and simply returns the goal state otherwise. Many common tactics,
+  such as @{ML resolve_tac}, fail if applied to a goal state that has no
+  subgoals.
 
-  \<^descr> @{ML SOLVE}~\<open>tac\<close> applies \<open>tac\<close> to the goal
-  state and then fails iff there are subgoals left.
+  \<^descr> @{ML SOLVE}~\<open>tac\<close> applies \<open>tac\<close> to the goal state and then fails iff there
+  are subgoals left.
 
-  \<^descr> @{ML DETERM}~\<open>tac\<close> applies \<open>tac\<close> to the goal
-  state and returns the head of the resulting sequence.  @{ML DETERM}
-  limits the search space by making its argument deterministic.
+  \<^descr> @{ML DETERM}~\<open>tac\<close> applies \<open>tac\<close> to the goal state and returns the head of
+  the resulting sequence. @{ML DETERM} limits the search space by making its
+  argument deterministic.
 \<close>
 
 
@@ -851,19 +808,20 @@
   @{index_ML size_of_thm: "thm -> int"} \\
   \end{mldecls}
 
-  \<^descr> @{ML has_fewer_prems}~\<open>n thm\<close> reports whether \<open>thm\<close> has fewer than \<open>n\<close> premises.
-
-  \<^descr> @{ML Thm.eq_thm}~\<open>(thm\<^sub>1, thm\<^sub>2)\<close> reports whether \<open>thm\<^sub>1\<close> and \<open>thm\<^sub>2\<close> are equal.  Both theorems must have the
-  same conclusions, the same set of hypotheses, and the same set of sort
-  hypotheses.  Names of bound variables are ignored as usual.
+  \<^descr> @{ML has_fewer_prems}~\<open>n thm\<close> reports whether \<open>thm\<close> has fewer than \<open>n\<close>
+  premises.
 
-  \<^descr> @{ML Thm.eq_thm_prop}~\<open>(thm\<^sub>1, thm\<^sub>2)\<close> reports whether
-  the propositions of \<open>thm\<^sub>1\<close> and \<open>thm\<^sub>2\<close> are equal.
-  Names of bound variables are ignored.
+  \<^descr> @{ML Thm.eq_thm}~\<open>(thm\<^sub>1, thm\<^sub>2)\<close> reports whether \<open>thm\<^sub>1\<close> and \<open>thm\<^sub>2\<close> are
+  equal. Both theorems must have the same conclusions, the same set of
+  hypotheses, and the same set of sort hypotheses. Names of bound variables
+  are ignored as usual.
 
-  \<^descr> @{ML size_of_thm}~\<open>thm\<close> computes the size of \<open>thm\<close>, namely the number of variables, constants and abstractions
-  in its conclusion.  It may serve as a distance function for
-  @{ML BEST_FIRST}.
+  \<^descr> @{ML Thm.eq_thm_prop}~\<open>(thm\<^sub>1, thm\<^sub>2)\<close> reports whether the propositions of
+  \<open>thm\<^sub>1\<close> and \<open>thm\<^sub>2\<close> are equal. Names of bound variables are ignored.
+
+  \<^descr> @{ML size_of_thm}~\<open>thm\<close> computes the size of \<open>thm\<close>, namely the number of
+  variables, constants and abstractions in its conclusion. It may serve as a
+  distance function for @{ML BEST_FIRST}.
 \<close>
 
 end
--- a/src/Doc/Isar_Ref/Spec.thy	Tue Dec 15 14:41:47 2015 +0000
+++ b/src/Doc/Isar_Ref/Spec.thy	Wed Dec 16 17:30:30 2015 +0100
@@ -1066,7 +1066,7 @@
 (*<*)experiment begin(*>*)
   attribute_setup my_rule =
     \<open>Attrib.thms >> (fn ths =>
-      Thm.rule_attribute
+      Thm.rule_attribute ths
         (fn context: Context.generic => fn th: thm =>
           let val th' = th OF ths
           in th' end))\<close>
--- a/src/HOL/Eisbach/Eisbach.thy	Tue Dec 15 14:41:47 2015 +0000
+++ b/src/HOL/Eisbach/Eisbach.thy	Wed Dec 16 17:30:30 2015 +0100
@@ -21,14 +21,6 @@
 ML_file "match_method.ML"
 ML_file "eisbach_antiquotations.ML"
 
-(* FIXME reform Isabelle/Pure attributes to make this work by default *)
-setup \<open>
-  fold (Method_Closure.wrap_attribute {handle_all_errs = true, declaration = true})
-    [@{binding intro}, @{binding elim}, @{binding dest}, @{binding simp}] #>
-  fold (Method_Closure.wrap_attribute {handle_all_errs = false, declaration = false})
-    [@{binding THEN}, @{binding OF}, @{binding rotated}, @{binding simplified}]
-\<close>
-
 method solves methods m = (m; fail)
 
 end
--- a/src/HOL/Eisbach/Eisbach_Tools.thy	Tue Dec 15 14:41:47 2015 +0000
+++ b/src/HOL/Eisbach/Eisbach_Tools.thy	Wed Dec 16 17:30:30 2015 +0100
@@ -76,7 +76,7 @@
       in Conjunction.curry_balanced (length conjs) thm end;
 \<close>
 
-attribute_setup uncurry = \<open>Scan.succeed (Thm.rule_attribute (K uncurry_rule))\<close>
-attribute_setup curry = \<open>Scan.succeed (Thm.rule_attribute (K curry_rule))\<close>
+attribute_setup uncurry = \<open>Scan.succeed (Thm.rule_attribute [] (K uncurry_rule))\<close>
+attribute_setup curry = \<open>Scan.succeed (Thm.rule_attribute [] (K curry_rule))\<close>
 
 end
--- a/src/HOL/Eisbach/eisbach_rule_insts.ML	Tue Dec 15 14:41:47 2015 +0000
+++ b/src/HOL/Eisbach/eisbach_rule_insts.ML	Wed Dec 16 17:30:30 2015 +0100
@@ -228,8 +228,11 @@
     (Attrib.setup @{binding "where"}
       (Scan.lift
         (Parse.and_list1 (Args.var -- (Args.$$$ "=" |-- Parse_Tools.name_term)) -- Parse.for_fixes)
-        >> (fn args => let val args' = read_where_insts args in Thm.rule_attribute (fn context =>
-            read_instantiate_closed (Context.proof_of context) args') end))
+        >> (fn args =>
+            let val args' = read_where_insts args in
+              Thm.mixed_attribute (fn (context, thm) =>
+                (context, read_instantiate_closed (Context.proof_of context) args' thm))
+            end))
       "named instantiation of theorem");
 
 val _ =
@@ -244,10 +247,12 @@
         let
           val read_insts = read_of_insts (not unchecked) context args;
         in
-          Thm.rule_attribute (fn context => fn thm =>
-            if Method_Closure.is_free_thm thm andalso unchecked
-            then Method_Closure.dummy_free_thm
-            else read_instantiate_closed (Context.proof_of context) (read_insts thm) thm)
+          Thm.mixed_attribute (fn (context, thm) =>
+            let val thm' =
+              if Thm.is_free_dummy thm andalso unchecked
+              then Drule.free_dummy_thm
+              else read_instantiate_closed (Context.proof_of context) (read_insts thm) thm
+            in (context, thm') end)
         end))
       "positional instantiation of theorem");
 
--- a/src/HOL/Eisbach/match_method.ML	Tue Dec 15 14:41:47 2015 +0000
+++ b/src/HOL/Eisbach/match_method.ML	Wed Dec 16 17:30:30 2015 +0100
@@ -171,7 +171,7 @@
                   val param_thm = map (Drule.mk_term o Thm.cterm_of ctxt' o Free) abs_nms
                     |> Conjunction.intr_balanced
                     |> Drule.generalize ([], map fst abs_nms)
-                    |> Method_Closure.tag_free_thm;
+                    |> Thm.tag_free_dummy;
 
                   val atts = map (Attrib.attribute ctxt') att;
                   val (param_thm', ctxt'') = Thm.proof_attributes atts param_thm ctxt';
@@ -183,7 +183,7 @@
 
                   val [head_thm, body_thm] =
                     Drule.zero_var_indexes_list (map label_thm [param_thm, param_thm'])
-                    |> map Method_Closure.tag_free_thm;
+                    |> map Thm.tag_free_dummy;
 
                   val ctxt''' =
                     Attrib.local_notes "" [((b, []), [([body_thm], [])])] ctxt''
--- a/src/HOL/Eisbach/method_closure.ML	Tue Dec 15 14:41:47 2015 +0000
+++ b/src/HOL/Eisbach/method_closure.ML	Wed Dec 16 17:30:30 2015 +0100
@@ -10,12 +10,6 @@
 
 signature METHOD_CLOSURE =
 sig
-  val tag_free_thm: thm -> thm
-  val is_free_thm: thm -> bool
-  val dummy_free_thm: thm
-  val free_aware_rule_attribute: thm list -> (Context.generic -> thm -> thm) -> Thm.attribute
-  val wrap_attribute: {handle_all_errs : bool, declaration : bool} ->
-    Binding.binding -> theory -> theory
   val read: Proof.context -> Token.src -> Method.text
   val read_closure: Proof.context -> Token.src -> Method.text * Token.src
   val read_closure_input: Proof.context -> Input.source -> Method.text * Token.src
@@ -68,52 +62,6 @@
 val put_recursive_method = Local_Data.map o apsnd o K;
 
 
-(* free thm *)
-
-val free_thmN = "Method_Closure.free_thm";
-fun tag_free_thm thm = Thm.tag_rule (free_thmN, "") thm;
-
-val dummy_free_thm = tag_free_thm Drule.dummy_thm;
-
-fun is_free_thm thm = Properties.defined (Thm.get_tags thm) free_thmN;
-
-fun free_aware_rule_attribute args f =
-  Thm.rule_attribute (fn context => fn thm =>
-    if exists is_free_thm (thm :: args) then dummy_free_thm
-    else f context thm);
-
-fun free_aware_attribute thy {handle_all_errs, declaration} src (context, thm) =
-  let
-    val src' = map Token.init_assignable src;
-    fun apply_att thm = (Attrib.attribute_global thy src') (context, thm);
-    val _ =
-      if handle_all_errs then (try apply_att Drule.dummy_thm; ())
-      else (apply_att Drule.dummy_thm; ()) handle THM _ => () | TERM _ => () | TYPE _ => ();
-
-    val src'' = map Token.closure src';
-    val thms =
-      map_filter Token.get_value (Token.args_of_src src'')
-      |> map_filter (fn (Token.Fact (_, f)) => SOME f | _ => NONE)
-      |> flat;
-  in
-    if exists is_free_thm (thm :: thms) then
-      if declaration then (NONE, NONE)
-      else (NONE, SOME dummy_free_thm)
-    else apply_att thm
-  end;
-
-fun wrap_attribute args binding thy =
-  let
-    val name = Binding.name_of binding;
-    val name' = Attrib.check_name_generic (Context.Theory thy) (name, Position.none);
-    fun get_src src =
-      Token.make_src (name', Position.set_range (Token.range_of src)) (Token.args_of_src src);
-  in
-    Attrib.define_global binding (free_aware_attribute thy args o get_src) "" thy
-    |> snd
-  end;
-
-
 (* read method text *)
 
 fun read ctxt src =
@@ -247,7 +195,7 @@
 fun dummy_named_thm named_thm ctxt =
   let
     val ctxt' = empty_named_thm named_thm ctxt;
-    val (_, ctxt'') = Thm.proof_attributes [Named_Theorems.add named_thm] dummy_free_thm ctxt';
+    val (_, ctxt'') = Thm.proof_attributes [Named_Theorems.add named_thm] Drule.free_dummy_thm ctxt';
   in ctxt'' end;
 
 fun parse_method_args method_names =
--- a/src/HOL/TLA/Action.thy	Tue Dec 15 14:41:47 2015 +0000
+++ b/src/HOL/TLA/Action.thy	Wed Dec 16 17:30:30 2015 +0100
@@ -123,11 +123,11 @@
 \<close>
 
 attribute_setup action_unlift =
-  \<open>Scan.succeed (Thm.rule_attribute (action_unlift o Context.proof_of))\<close>
+  \<open>Scan.succeed (Thm.rule_attribute [] (action_unlift o Context.proof_of))\<close>
 attribute_setup action_rewrite =
-  \<open>Scan.succeed (Thm.rule_attribute (action_rewrite o Context.proof_of))\<close>
+  \<open>Scan.succeed (Thm.rule_attribute [] (action_rewrite o Context.proof_of))\<close>
 attribute_setup action_use =
-  \<open>Scan.succeed (Thm.rule_attribute (action_use o Context.proof_of))\<close>
+  \<open>Scan.succeed (Thm.rule_attribute [] (action_use o Context.proof_of))\<close>
 
 
 (* =========================== square / angle brackets =========================== *)
--- a/src/HOL/TLA/Intensional.thy	Tue Dec 15 14:41:47 2015 +0000
+++ b/src/HOL/TLA/Intensional.thy	Wed Dec 16 17:30:30 2015 +0100
@@ -256,12 +256,13 @@
 \<close>
 
 attribute_setup int_unlift =
-  \<open>Scan.succeed (Thm.rule_attribute (int_unlift o Context.proof_of))\<close>
+  \<open>Scan.succeed (Thm.rule_attribute [] (int_unlift o Context.proof_of))\<close>
 attribute_setup int_rewrite =
-  \<open>Scan.succeed (Thm.rule_attribute (int_rewrite o Context.proof_of))\<close>
-attribute_setup flatten = \<open>Scan.succeed (Thm.rule_attribute (K flatten))\<close>
+  \<open>Scan.succeed (Thm.rule_attribute [] (int_rewrite o Context.proof_of))\<close>
+attribute_setup flatten =
+  \<open>Scan.succeed (Thm.rule_attribute [] (K flatten))\<close>
 attribute_setup int_use =
-  \<open>Scan.succeed (Thm.rule_attribute (int_use o Context.proof_of))\<close>
+  \<open>Scan.succeed (Thm.rule_attribute [] (int_use o Context.proof_of))\<close>
 
 lemma Not_Rall: "\<turnstile> (\<not>(\<forall>x. F x)) = (\<exists>x. \<not>F x)"
   by (simp add: Valid_def)
--- a/src/HOL/TLA/TLA.thy	Tue Dec 15 14:41:47 2015 +0000
+++ b/src/HOL/TLA/TLA.thy	Wed Dec 16 17:30:30 2015 +0100
@@ -124,13 +124,13 @@
 \<close>
 
 attribute_setup temp_unlift =
-  \<open>Scan.succeed (Thm.rule_attribute (temp_unlift o Context.proof_of))\<close>
+  \<open>Scan.succeed (Thm.rule_attribute [] (temp_unlift o Context.proof_of))\<close>
 attribute_setup temp_rewrite =
-  \<open>Scan.succeed (Thm.rule_attribute (temp_rewrite o Context.proof_of))\<close>
+  \<open>Scan.succeed (Thm.rule_attribute [] (temp_rewrite o Context.proof_of))\<close>
 attribute_setup temp_use =
-  \<open>Scan.succeed (Thm.rule_attribute (temp_use o Context.proof_of))\<close>
+  \<open>Scan.succeed (Thm.rule_attribute [] (temp_use o Context.proof_of))\<close>
 attribute_setup try_rewrite =
-  \<open>Scan.succeed (Thm.rule_attribute (try_rewrite o Context.proof_of))\<close>
+  \<open>Scan.succeed (Thm.rule_attribute [] (try_rewrite o Context.proof_of))\<close>
 
 
 (* ------------------------------------------------------------------------- *)
--- a/src/HOL/Tools/Quotient/quotient_tacs.ML	Tue Dec 15 14:41:47 2015 +0000
+++ b/src/HOL/Tools/Quotient/quotient_tacs.ML	Wed Dec 16 17:30:30 2015 +0100
@@ -755,7 +755,7 @@
 
 (* lifting as an attribute *)
 
-val lifted_attrib = Thm.rule_attribute (fn context =>
+val lifted_attrib = Thm.rule_attribute [] (fn context =>
   let
     val ctxt = Context.proof_of context
     val qtys = map #qtyp (Quotient_Info.dest_quotients ctxt)
--- a/src/HOL/Tools/Transfer/transfer.ML	Tue Dec 15 14:41:47 2015 +0000
+++ b/src/HOL/Tools/Transfer/transfer.ML	Wed Dec 16 17:30:30 2015 +0100
@@ -847,11 +847,11 @@
 
 (* Attributes for transferred rules *)
 
-fun transferred_attribute thms = Thm.rule_attribute
-  (fn context => transferred (Context.proof_of context) thms)
+fun transferred_attribute thms =
+  Thm.rule_attribute thms (fn context => transferred (Context.proof_of context) thms)
 
-fun untransferred_attribute thms = Thm.rule_attribute
-  (fn context => untransferred (Context.proof_of context) thms)
+fun untransferred_attribute thms =
+  Thm.rule_attribute thms (fn context => untransferred (Context.proof_of context) thms)
 
 val transferred_attribute_parser =
   Attrib.thms >> transferred_attribute
--- a/src/HOL/Tools/inductive_set.ML	Tue Dec 15 14:41:47 2015 +0000
+++ b/src/HOL/Tools/inductive_set.ML	Wed Dec 16 17:30:30 2015 +0100
@@ -349,7 +349,7 @@
     Rule_Cases.save thm
   end;
 
-val to_pred_att = Thm.rule_attribute o to_pred;
+val to_pred_att = Thm.rule_attribute [] o to_pred;
 
 
 (**** convert theorem in predicate notation to set notation ****)
@@ -385,7 +385,7 @@
     Rule_Cases.save thm
   end;
 
-val to_set_att = Thm.rule_attribute o to_set;
+val to_set_att = Thm.rule_attribute [] o to_set;
 
 
 (**** definition of inductive sets ****)
--- a/src/HOL/Tools/legacy_transfer.ML	Tue Dec 15 14:41:47 2015 +0000
+++ b/src/HOL/Tools/legacy_transfer.ML	Wed Dec 16 17:30:30 2015 +0100
@@ -253,7 +253,7 @@
     "install transfer data" #>
     Attrib.setup @{binding transferred}
       (selection -- these (keyword_colon leavingN |-- names)
-        >> (fn (selection, leave) => Thm.rule_attribute (fn context =>
+        >> (fn (selection, leave) => Thm.rule_attribute [] (fn context =>
               Conjunction.intr_balanced o transfer context selection leave)))
     "transfer theorems");
 
--- a/src/HOL/Tools/split_rule.ML	Tue Dec 15 14:41:47 2015 +0000
+++ b/src/HOL/Tools/split_rule.ML	Wed Dec 16 17:30:30 2015 +0100
@@ -113,10 +113,10 @@
   Theory.setup
    (Attrib.setup @{binding split_format}
       (Scan.lift (Args.parens (Args.$$$ "complete")
-        >> K (Thm.rule_attribute (complete_split_rule o Context.proof_of))))
+        >> K (Thm.rule_attribute [] (complete_split_rule o Context.proof_of))))
       "split pair-typed subterms in premises, or function arguments" #>
     Attrib.setup @{binding split_rule}
-      (Scan.succeed (Thm.rule_attribute (split_rule o Context.proof_of)))
+      (Scan.succeed (Thm.rule_attribute [] (split_rule o Context.proof_of)))
       "curries ALL function variables occurring in a rule's conclusion");
 
 end;
--- a/src/HOL/UNITY/Comp/Alloc.thy	Tue Dec 15 14:41:47 2015 +0000
+++ b/src/HOL/UNITY/Comp/Alloc.thy	Wed Dec 16 17:30:30 2015 +0100
@@ -321,7 +321,7 @@
       handle THM _ => th RS (@{thm guarantees_INT_right_iff} RS iffD1))
     handle THM _ => th;
 in
-  Scan.succeed (Thm.rule_attribute (K normalized))
+  Scan.succeed (Thm.rule_attribute [] (K normalized))
 end
 *}
 
--- a/src/Provers/classical.ML	Tue Dec 15 14:41:47 2015 +0000
+++ b/src/Provers/classical.ML	Wed Dec 16 17:30:30 2015 +0100
@@ -190,7 +190,7 @@
 
 (*Creates rules to eliminate ~A, from rules to introduce A*)
 fun swapify intrs = intrs RLN (2, [Data.swap]);
-val swapped = Thm.rule_attribute (fn _ => fn th => th RSN (2, Data.swap));
+val swapped = Thm.rule_attribute [] (fn _ => fn th => th RSN (2, Data.swap));
 
 (*Uses introduction rules in the normal way, or on negated assumptions,
   trying rules in order. *)
--- a/src/Pure/Isar/attrib.ML	Tue Dec 15 14:41:47 2015 +0000
+++ b/src/Pure/Isar/attrib.ML	Wed Dec 16 17:30:30 2015 +0100
@@ -167,9 +167,6 @@
 
 (* pretty printing *)
 
-fun markup_extern ctxt =
-  Name_Space.markup_extern ctxt (Name_Space.space_of_table (get_attributes ctxt));
-
 fun pretty_attribs _ [] = []
   | pretty_attribs ctxt srcs = [Pretty.enum "," "[" "]" (map (Token.pretty_src ctxt) srcs)];
 
--- a/src/Pure/Isar/calculation.ML	Tue Dec 15 14:41:47 2015 +0000
+++ b/src/Pure/Isar/calculation.ML	Wed Dec 16 17:30:30 2015 +0100
@@ -85,12 +85,13 @@
 
 (* symmetric *)
 
-val symmetric = Thm.rule_attribute (fn context => fn th =>
-  (case Seq.chop 2
-      (Drule.multi_resolves (SOME (Context.proof_of context)) [th] (#2 (#1 (Data.get context)))) of
-    ([th'], _) => Drule.zero_var_indexes th'
-  | ([], _) => raise THM ("symmetric: no unifiers", 1, [th])
-  | _ => raise THM ("symmetric: multiple unifiers", 1, [th])));
+val symmetric =
+  Thm.rule_attribute [] (fn context => fn th =>
+    (case Seq.chop 2
+        (Drule.multi_resolves (SOME (Context.proof_of context)) [th] (#2 (#1 (Data.get context)))) of
+      ([th'], _) => Drule.zero_var_indexes th'
+    | ([], _) => raise THM ("symmetric: no unifiers", 1, [th])
+    | _ => raise THM ("symmetric: multiple unifiers", 1, [th])));
 
 
 (* concrete syntax *)
--- a/src/Pure/Isar/object_logic.ML	Tue Dec 15 14:41:47 2015 +0000
+++ b/src/Pure/Isar/object_logic.ML	Wed Dec 16 17:30:30 2015 +0100
@@ -213,7 +213,7 @@
 val rulify = gen_rulify true;
 val rulify_no_asm = gen_rulify false;
 
-val rule_format = Thm.rule_attribute (rulify o Context.proof_of);
-val rule_format_no_asm = Thm.rule_attribute (rulify_no_asm o Context.proof_of);
+val rule_format = Thm.rule_attribute [] (rulify o Context.proof_of);
+val rule_format_no_asm = Thm.rule_attribute [] (rulify_no_asm o Context.proof_of);
 
 end;
--- a/src/Pure/Isar/rule_cases.ML	Tue Dec 15 14:41:47 2015 +0000
+++ b/src/Pure/Isar/rule_cases.ML	Wed Dec 16 17:30:30 2015 +0100
@@ -404,7 +404,7 @@
   |> fold_index (fn (i, xs) => Thm.rename_params_rule (xs, i + 1)) xss
   |> save th;
 
-fun params xss = Thm.rule_attribute (K (rename_params xss));
+fun params xss = Thm.rule_attribute [] (K (rename_params xss));
 
 
 
--- a/src/Pure/Pure.thy	Tue Dec 15 14:41:47 2015 +0000
+++ b/src/Pure/Pure.thy	Wed Dec 16 17:30:30 2015 +0100
@@ -128,26 +128,26 @@
 
 attribute_setup THEN =
   \<open>Scan.lift (Scan.optional (Args.bracks Parse.nat) 1) -- Attrib.thm
-    >> (fn (i, B) => Thm.rule_attribute (fn _ => fn A => A RSN (i, B)))\<close>
+    >> (fn (i, B) => Thm.rule_attribute [B] (fn _ => fn A => A RSN (i, B)))\<close>
   "resolution with rule"
 
 attribute_setup OF =
-  \<open>Attrib.thms >> (fn Bs => Thm.rule_attribute (fn _ => fn A => A OF Bs))\<close>
+  \<open>Attrib.thms >> (fn Bs => Thm.rule_attribute Bs (fn _ => fn A => A OF Bs))\<close>
   "rule resolved with facts"
 
 attribute_setup rename_abs =
   \<open>Scan.lift (Scan.repeat (Args.maybe Args.name)) >> (fn vs =>
-    Thm.rule_attribute (K (Drule.rename_bvars' vs)))\<close>
+    Thm.rule_attribute [] (K (Drule.rename_bvars' vs)))\<close>
   "rename bound variables in abstractions"
 
 attribute_setup unfolded =
   \<open>Attrib.thms >> (fn ths =>
-    Thm.rule_attribute (fn context => Local_Defs.unfold (Context.proof_of context) ths))\<close>
+    Thm.rule_attribute ths (fn context => Local_Defs.unfold (Context.proof_of context) ths))\<close>
   "unfolded definitions"
 
 attribute_setup folded =
   \<open>Attrib.thms >> (fn ths =>
-    Thm.rule_attribute (fn context => Local_Defs.fold (Context.proof_of context) ths))\<close>
+    Thm.rule_attribute ths (fn context => Local_Defs.fold (Context.proof_of context) ths))\<close>
   "folded definitions"
 
 attribute_setup consumes =
@@ -181,11 +181,11 @@
   "result put into canonical rule format"
 
 attribute_setup elim_format =
-  \<open>Scan.succeed (Thm.rule_attribute (K Tactic.make_elim))\<close>
+  \<open>Scan.succeed (Thm.rule_attribute [] (K Tactic.make_elim))\<close>
   "destruct rule turned into elimination rule format"
 
 attribute_setup no_vars =
-  \<open>Scan.succeed (Thm.rule_attribute (fn context => fn th =>
+  \<open>Scan.succeed (Thm.rule_attribute [] (fn context => fn th =>
     let
       val ctxt = Variable.set_body false (Context.proof_of context);
       val ((_, [th']), _) = Variable.import true [th] ctxt;
@@ -202,7 +202,7 @@
 
 attribute_setup rotated =
   \<open>Scan.lift (Scan.optional Parse.int 1
-    >> (fn n => Thm.rule_attribute (fn _ => rotate_prems n)))\<close>
+    >> (fn n => Thm.rule_attribute [] (fn _ => rotate_prems n)))\<close>
   "rotated theorem premises"
 
 attribute_setup defn =
@@ -210,7 +210,7 @@
   "declaration of definitional transformations"
 
 attribute_setup abs_def =
-  \<open>Scan.succeed (Thm.rule_attribute (fn context =>
+  \<open>Scan.succeed (Thm.rule_attribute [] (fn context =>
     Local_Defs.meta_rewrite_rule (Context.proof_of context) #> Drule.abs_def))\<close>
   "abstract over free variables of definitional theorem"
 
--- a/src/Pure/Tools/rule_insts.ML	Tue Dec 15 14:41:47 2015 +0000
+++ b/src/Pure/Tools/rule_insts.ML	Wed Dec 16 17:30:30 2015 +0100
@@ -182,7 +182,7 @@
 val _ = Theory.setup
   (Attrib.setup @{binding "where"}
     (Scan.lift named_insts >> (fn args =>
-      Thm.rule_attribute (fn context => uncurry (where_rule (Context.proof_of context)) args)))
+      Thm.rule_attribute [] (fn context => uncurry (where_rule (Context.proof_of context)) args)))
     "named instantiation of theorem");
 
 
@@ -202,7 +202,7 @@
 val _ = Theory.setup
   (Attrib.setup @{binding "of"}
     (Scan.lift (insts -- Parse.for_fixes) >> (fn args =>
-      Thm.rule_attribute (fn context => uncurry (of_rule (Context.proof_of context)) args)))
+      Thm.rule_attribute [] (fn context => uncurry (of_rule (Context.proof_of context)) args)))
     "positional instantiation of theorem");
 
 end;
--- a/src/Pure/drule.ML	Tue Dec 15 14:41:47 2015 +0000
+++ b/src/Pure/drule.ML	Wed Dec 16 17:30:30 2015 +0100
@@ -94,6 +94,7 @@
   val cterm_add_frees: cterm -> cterm list -> cterm list
   val cterm_add_vars: cterm -> cterm list -> cterm list
   val dummy_thm: thm
+  val free_dummy_thm: thm
   val is_sort_constraint: term -> bool
   val sort_constraintI: thm
   val sort_constraint_eq: thm
@@ -628,6 +629,7 @@
 val cterm_add_vars = Thm.add_vars o mk_term;
 
 val dummy_thm = mk_term (certify Term.dummy_prop);
+val free_dummy_thm = Thm.tag_free_dummy dummy_thm;
 
 
 (* sort_constraint *)
--- a/src/Pure/more_thm.ML	Tue Dec 15 14:41:47 2015 +0000
+++ b/src/Pure/more_thm.ML	Wed Dec 16 17:30:30 2015 +0100
@@ -84,19 +84,10 @@
   type attribute = Context.generic * thm -> Context.generic option * thm option
   type binding = binding * attribute list
   val empty_binding: binding
-  val rule_attribute: (Context.generic -> thm -> thm) -> attribute
-  val declaration_attribute: (thm -> Context.generic -> Context.generic) -> attribute
-  val mixed_attribute: (Context.generic * thm -> Context.generic * thm) -> attribute
-  val apply_attribute: attribute -> thm -> Context.generic -> thm * Context.generic
-  val attribute_declaration: attribute -> thm -> Context.generic -> Context.generic
-  val theory_attributes: attribute list -> thm -> theory -> thm * theory
-  val proof_attributes: attribute list -> thm -> Proof.context -> thm * Proof.context
-  val no_attributes: 'a -> 'a * 'b list
-  val simple_fact: 'a -> ('a * 'b list) list
   val tag_rule: string * string -> thm -> thm
   val untag_rule: string -> thm -> thm
-  val tag: string * string -> attribute
-  val untag: string -> attribute
+  val is_free_dummy: thm -> bool
+  val tag_free_dummy: thm -> thm
   val def_name: string -> string
   val def_name_optional: string -> string -> string
   val def_binding: Binding.binding -> Binding.binding
@@ -107,6 +98,17 @@
   val theoremK: string
   val legacy_get_kind: thm -> string
   val kind_rule: string -> thm -> thm
+  val rule_attribute: thm list -> (Context.generic -> thm -> thm) -> attribute
+  val declaration_attribute: (thm -> Context.generic -> Context.generic) -> attribute
+  val mixed_attribute: (Context.generic * thm -> Context.generic * thm) -> attribute
+  val apply_attribute: attribute -> thm -> Context.generic -> thm * Context.generic
+  val attribute_declaration: attribute -> thm -> Context.generic -> Context.generic
+  val theory_attributes: attribute list -> thm -> theory -> thm * theory
+  val proof_attributes: attribute list -> thm -> Proof.context -> thm * Proof.context
+  val no_attributes: 'a -> 'a * 'b list
+  val simple_fact: 'a -> ('a * 'b list) list
+  val tag: string * string -> attribute
+  val untag: string -> attribute
   val kind: string -> attribute
   val register_proofs: thm list -> theory -> theory
   val join_theory_proofs: theory -> unit
@@ -541,38 +543,6 @@
 
 
 
-(** attributes **)
-
-(*attributes subsume any kind of rules or context modifiers*)
-type attribute = Context.generic * thm -> Context.generic option * thm option;
-
-type binding = binding * attribute list;
-val empty_binding: binding = (Binding.empty, []);
-
-fun rule_attribute f (x, th) = (NONE, SOME (f x th));
-fun declaration_attribute f (x, th) = (SOME (f th x), NONE);
-fun mixed_attribute f (x, th) = let val (x', th') = f (x, th) in (SOME x', SOME th') end;
-
-fun apply_attribute (att: attribute) th x =
-  let val (x', th') = att (x, check_hyps x (Thm.transfer (Context.theory_of x) th))
-  in (the_default th th', the_default x x') end;
-
-fun attribute_declaration att th x = #2 (apply_attribute att th x);
-
-fun apply_attributes mk dest =
-  let
-    fun app [] th x = (th, x)
-      | app (att :: atts) th x = apply_attribute att th (mk x) ||> dest |-> app atts;
-  in app end;
-
-val theory_attributes = apply_attributes Context.Theory Context.the_theory;
-val proof_attributes = apply_attributes Context.Proof Context.the_proof;
-
-fun no_attributes x = (x, []);
-fun simple_fact x = [(x, [])];
-
-
-
 (*** theorem tags ***)
 
 (* add / delete tags *)
@@ -580,8 +550,12 @@
 fun tag_rule tg = Thm.map_tags (insert (op =) tg);
 fun untag_rule s = Thm.map_tags (filter_out (fn (s', _) => s = s'));
 
-fun tag tg = rule_attribute (K (tag_rule tg));
-fun untag s = rule_attribute (K (untag_rule s));
+
+(* free dummy thm -- for abstract closure *)
+
+val free_dummyN = "free_dummy";
+fun is_free_dummy thm = Properties.defined (Thm.get_tags thm) free_dummyN;
+val tag_free_dummy = tag_rule (free_dummyN, "");
 
 
 (* def_name *)
@@ -614,7 +588,50 @@
 fun legacy_get_kind thm = the_default "" (Properties.get (Thm.get_tags thm) Markup.kindN);
 
 fun kind_rule k = tag_rule (Markup.kindN, k) o untag_rule Markup.kindN;
-fun kind k = rule_attribute (K (k <> "" ? kind_rule k));
+
+
+
+(** attributes **)
+
+(*attributes subsume any kind of rules or context modifiers*)
+type attribute = Context.generic * thm -> Context.generic option * thm option;
+
+type binding = binding * attribute list;
+val empty_binding: binding = (Binding.empty, []);
+
+fun rule_attribute ths f (x, th) =
+  (NONE,
+    (case find_first is_free_dummy (th :: ths) of
+      SOME th' => SOME th'
+    | NONE => SOME (f x th)));
+
+fun declaration_attribute f (x, th) =
+  (if is_free_dummy th then NONE else SOME (f th x), NONE);
+
+fun mixed_attribute f (x, th) =
+  let val (x', th') = f (x, th) in (SOME x', SOME th') end;
+
+fun apply_attribute (att: attribute) th x =
+  let val (x', th') = att (x, check_hyps x (Thm.transfer (Context.theory_of x) th))
+  in (the_default th th', the_default x x') end;
+
+fun attribute_declaration att th x = #2 (apply_attribute att th x);
+
+fun apply_attributes mk dest =
+  let
+    fun app [] th x = (th, x)
+      | app (att :: atts) th x = apply_attribute att th (mk x) ||> dest |-> app atts;
+  in app end;
+
+val theory_attributes = apply_attributes Context.Theory Context.the_theory;
+val proof_attributes = apply_attributes Context.Proof Context.the_proof;
+
+fun no_attributes x = (x, []);
+fun simple_fact x = [(x, [])];
+
+fun tag tg = rule_attribute [] (K (tag_rule tg));
+fun untag s = rule_attribute [] (K (untag_rule s));
+fun kind k = rule_attribute [] (K (k <> "" ? kind_rule k));
 
 
 (* forked proofs *)
--- a/src/Pure/simplifier.ML	Tue Dec 15 14:41:47 2015 +0000
+++ b/src/Pure/simplifier.ML	Wed Dec 16 17:30:30 2015 +0100
@@ -309,7 +309,7 @@
 in
 
 val simplified = conv_mode -- Attrib.thms >>
-  (fn (f, ths) => Thm.rule_attribute (fn context =>
+  (fn (f, ths) => Thm.rule_attribute ths (fn context =>
     f ((if null ths then I else Raw_Simplifier.clear_simpset)
         (Context.proof_of context) addsimps ths)));
 
--- a/src/Tools/case_product.ML	Tue Dec 15 14:41:47 2015 +0000
+++ b/src/Tools/case_product.ML	Wed Dec 16 17:30:30 2015 +0100
@@ -96,7 +96,8 @@
     Rule_Cases.name names o Rule_Cases.put_consumes (SOME (cons1 + cons2))
   end
 
-fun annotation thm1 thm2 = Thm.rule_attribute (K (annotation_rule thm1 thm2))
+fun annotation thm1 thm2 =
+  Thm.rule_attribute [thm1, thm2] (K (annotation_rule thm1 thm2))
 
 fun combine_annotated ctxt thm1 thm2 =
   combine ctxt thm1 thm2
@@ -111,7 +112,7 @@
       let
         fun combine_list ctxt = fold (fn x => fn y => combine_annotated ctxt y x)
       in
-        Attrib.thms >> (fn thms => Thm.rule_attribute (fn ctxt => fn thm =>
+        Attrib.thms >> (fn thms => Thm.rule_attribute thms (fn ctxt => fn thm =>
           combine_list (Context.proof_of ctxt) thms thm))
       end
     "product with other case rules")
--- a/src/Tools/induct.ML	Tue Dec 15 14:41:47 2015 +0000
+++ b/src/Tools/induct.ML	Wed Dec 16 17:30:30 2015 +0100
@@ -312,7 +312,9 @@
   Thm.mixed_attribute (fn (context, thm) =>
     let
       val thm' = g thm;
-      val context' = Data.map (f (name, Thm.trim_context thm')) context;
+      val context' =
+        if Thm.is_free_dummy thm then context
+        else Data.map (f (name, Thm.trim_context thm')) context;
     in (context', thm') end);
 
 fun del_att which =