merged
authorhaftmann
Mon, 08 Feb 2010 14:08:32 +0100
changeset 35039 e682bb587071
parent 35031 2ddc7edce107 (diff)
parent 35038 a1d93ce94235 (current diff)
child 35040 e42e7f133d94
child 35041 6eb917794a5c
merged
NEWS
src/HOL/IsaMakefile
src/HOL/Library/Library.thy
--- a/Admin/isatest/settings/mac-poly64-M4	Mon Feb 08 14:06:58 2010 +0100
+++ b/Admin/isatest/settings/mac-poly64-M4	Mon Feb 08 14:08:32 2010 +0100
@@ -4,7 +4,7 @@
   ML_SYSTEM="polyml-5.3.0"
   ML_PLATFORM="x86_64-darwin"
   ML_HOME="$POLYML_HOME/$ML_PLATFORM"
-  ML_OPTIONS="--mutable 12000 --immutable 4000"
+  ML_OPTIONS="--mutable 4000 --immutable 4000"
 
 
 ISABELLE_HOME_USER=~/isabelle-mac-poly64-M4
--- a/Admin/isatest/settings/mac-poly64-M8	Mon Feb 08 14:06:58 2010 +0100
+++ b/Admin/isatest/settings/mac-poly64-M8	Mon Feb 08 14:08:32 2010 +0100
@@ -4,7 +4,7 @@
   ML_SYSTEM="polyml-5.3.0"
   ML_PLATFORM="x86_64-darwin"
   ML_HOME="$POLYML_HOME/$ML_PLATFORM"
-  ML_OPTIONS="--mutable 12000 --immutable 4000"
+  ML_OPTIONS="--mutable 4000 --immutable 4000"
 
 
 ISABELLE_HOME_USER=~/isabelle-mac-poly64-M8
--- a/NEWS	Mon Feb 08 14:06:58 2010 +0100
+++ b/NEWS	Mon Feb 08 14:08:32 2010 +0100
@@ -63,10 +63,13 @@
 
 INCOMPATIBILITY.
 
-* new theory Algebras.thy contains generic algebraic structures and
+* Index syntax for structures must be imported explicitly from library
+theory Structure_Syntax.  INCOMPATIBILITY.
+
+* New theory Algebras contains generic algebraic structures and
 generic algebraic operations.  INCOMPATIBILTY: constants zero, one,
-plus, minus, uminus, times, inverse, divide, abs, sgn, less_eq and less
-have been moved from HOL.thy to Algebras.thy.
+plus, minus, uminus, times, inverse, divide, abs, sgn, less_eq and
+less have been moved from HOL.thy to Algebras.thy.
 
 * HOLogic.strip_psplit: types are returned in syntactic order, similar
 to other strip and tuple operations.  INCOMPATIBILITY.
@@ -75,24 +78,24 @@
 replaced by new-style primrec, especially in theory List.  The corresponding
 constants now have authentic syntax.  INCOMPATIBILITY.
 
-* Reorganized theory Multiset.thy: less duplication, less historical
+* Reorganized theory Multiset: less duplication, less historical
 organization of sections, conversion from associations lists to
 multisets, rudimentary code generation.  Use insert_DiffM2 [symmetric]
 instead of elem_imp_eq_diff_union, if needed.  INCOMPATIBILITY.
 
-* Reorganized theory Sum_Type.thy; Inl and Inr now have
-authentic syntax.  INCOMPATIBILITY.
+* Reorganized theory Sum_Type; Inl and Inr now have authentic syntax.
+INCOMPATIBILITY.
 
 * Code generation: ML and OCaml code is decorated with signatures.
 
-* Complete_Lattice.thy: lemmas top_def and bot_def have been replaced
-by the more convenient lemmas Inf_empty and Sup_empty.  Dropped lemmas
-Inf_insert_simp and Sup_insert_simp, which are subsumed by Inf_insert
-and Sup_insert.  Lemmas Inf_UNIV and Sup_UNIV replace former Inf_Univ
-and Sup_Univ.  Lemmas inf_top_right and sup_bot_right subsume inf_top
-and sup_bot respectively.  INCOMPATIBILITY.
-
-* Finite_Set.thy and List.thy: some lemmas have been generalized from
+* Theory Complete_Lattice: lemmas top_def and bot_def have been
+replaced by the more convenient lemmas Inf_empty and Sup_empty.
+Dropped lemmas Inf_insert_simp and Sup_insert_simp, which are subsumed
+by Inf_insert and Sup_insert.  Lemmas Inf_UNIV and Sup_UNIV replace
+former Inf_Univ and Sup_Univ.  Lemmas inf_top_right and sup_bot_right
+subsume inf_top and sup_bot respectively.  INCOMPATIBILITY.
+
+* Theory Finite_Set and List: some lemmas have been generalized from
 sets to lattices:
 
   fun_left_comm_idem_inter      ~> fun_left_comm_idem_inf
@@ -106,10 +109,15 @@
   INTER_fold_inter              ~> INFI_fold_inf
   UNION_fold_union              ~> SUPR_fold_sup
 
-* Added transpose to List.thy.
+* Theory List: added transpose.
+
 
 *** ML ***
 
+* Renamed old-style Drule.standard to Drule.export_without_context, to
+emphasize that this is in no way a standard operation.
+INCOMPATIBILITY.
+
 * Curried take and drop in library.ML; negative length is interpreted
 as infinity (as in chop).  INCOMPATIBILITY.
 
--- a/doc-src/IsarImplementation/Thy/Integration.thy	Mon Feb 08 14:06:58 2010 +0100
+++ b/doc-src/IsarImplementation/Thy/Integration.thy	Mon Feb 08 14:08:32 2010 +0100
@@ -79,8 +79,9 @@
   \item @{ML Toplevel.is_toplevel}~@{text "state"} checks for an empty
   toplevel state.
 
-  \item @{ML Toplevel.theory_of}~@{text "state"} selects the theory of
-  a theory or proof (!), otherwise raises @{ML Toplevel.UNDEF}.
+  \item @{ML Toplevel.theory_of}~@{text "state"} selects the
+  background theory of @{text "state"}, raises @{ML Toplevel.UNDEF}
+  for an empty toplevel state.
 
   \item @{ML Toplevel.proof_of}~@{text "state"} selects the Isar proof
   state if available, otherwise raises @{ML Toplevel.UNDEF}.
@@ -114,16 +115,16 @@
   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} acts
+  alternative state transitions.  For example, \isakeyword{qed} works
   differently for a local proofs vs.\ the global ending of the main
   proof.
 
   Toplevel transitions are composed via transition transformers.
   Internally, Isar commands are put together from an empty transition
-  extended by name and source position (and optional source text).  It
-  is then left to the individual command parser to turn the given
-  concrete syntax into a suitable transition transformer that adjoins
-  actual operations on a theory or proof state etc.
+  extended by name and source position.  It is then left to the
+  individual command parser to turn the given concrete syntax into a
+  suitable transition transformer that adjoins actual operations on a
+  theory or proof state etc.
 *}
 
 text %mlref {*
@@ -188,7 +189,7 @@
   the toplevel itself, and only make sense in interactive mode.  Under
   normal circumstances, the user encounters these only implicitly as
   part of the protocol between the Isabelle/Isar system and a
-  user-interface such as ProofGeneral.
+  user-interface such as Proof~General.
 
   \begin{description}
 
@@ -323,9 +324,7 @@
   associated with each theory, by declaring these dependencies in the
   theory header as @{text \<USES>}, and loading them consecutively
   within the theory context.  The system keeps track of incoming {\ML}
-  sources and associates them with the current theory.  The file
-  @{text A}\verb,.ML, is loaded after a theory has been concluded, in
-  order to support legacy proof {\ML} proof scripts.
+  sources and associates them with the current theory.
 
   The basic internal actions of the theory database are @{text
   "update"}, @{text "outdate"}, and @{text "remove"}:
@@ -386,12 +385,15 @@
 
   \item @{ML use_thy}~@{text A} ensures that theory @{text A} is fully
   up-to-date wrt.\ the external file store, reloading outdated
-  ancestors as required.
+  ancestors as required.  In batch mode, the simultaneous @{ML
+  use_thys} should be used exclusively.
 
   \item @{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, though.
+  result.  By loading a whole sub-graph of theories like that, the
+  intrinsic parallelism can be exploited by the system, to speedup
+  loading.
 
   \item @{ML ThyInfo.touch_thy}~@{text A} performs and @{text outdate} action
   on theory @{text A} and all descendants.
@@ -400,7 +402,7 @@
   descendants from the theory database.
 
   \item @{ML ThyInfo.begin_theory} is the basic operation behind a
-  @{text \<THEORY>} header declaration.  This is {\ML} functions is
+  @{text \<THEORY>} header declaration.  This {\ML} function is
   normally not invoked directly.
 
   \item @{ML ThyInfo.end_theory} concludes the loading of a theory
--- a/doc-src/IsarImplementation/Thy/Local_Theory.thy	Mon Feb 08 14:06:58 2010 +0100
+++ b/doc-src/IsarImplementation/Thy/Local_Theory.thy	Mon Feb 08 14:08:32 2010 +0100
@@ -2,7 +2,7 @@
 imports Base
 begin
 
-chapter {* Local theory specifications *}
+chapter {* Local theory specifications \label{ch:local-theory} *}
 
 text {*
   A \emph{local theory} combines aspects of both theory and proof
--- a/doc-src/IsarImplementation/Thy/Logic.thy	Mon Feb 08 14:06:58 2010 +0100
+++ b/doc-src/IsarImplementation/Thy/Logic.thy	Mon Feb 08 14:08:32 2010 +0100
@@ -24,7 +24,9 @@
   schematic polymorphism, which is strictly speaking outside the
   logic.\footnote{This is the deeper logical reason, why the theory
   context @{text "\<Theta>"} is separate from the proof context @{text "\<Gamma>"}
-  of the core calculus.}
+  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!}
 *}
 
 
@@ -41,18 +43,17 @@
   internally.  The resulting relation is an ordering: reflexive,
   transitive, and antisymmetric.
 
-  A \emph{sort} is a list of type classes written as @{text "s =
-  {c\<^isub>1, \<dots>, c\<^isub>m}"}, which represents symbolic
-  intersection.  Notationally, the curly braces are omitted for
-  singleton intersections, i.e.\ any class @{text "c"} may be read as
-  a sort @{text "{c}"}.  The ordering on type classes is extended to
-  sorts according to the meaning of intersections: @{text
-  "{c\<^isub>1, \<dots> c\<^isub>m} \<subseteq> {d\<^isub>1, \<dots>, d\<^isub>n}"} iff
-  @{text "\<forall>j. \<exists>i. c\<^isub>i \<subseteq> d\<^isub>j"}.  The empty intersection
-  @{text "{}"} refers to the universal sort, which is the largest
-  element wrt.\ the sort order.  The intersections of all (finitely
-  many) classes declared in the current theory are the minimal
-  elements wrt.\ the sort order.
+  A \emph{sort} is a list of type classes written as @{text "s = {c\<^isub>1,
+  \<dots>, c\<^isub>m}"}, it represents symbolic intersection.  Notationally, the
+  curly braces are omitted for singleton intersections, i.e.\ any
+  class @{text "c"} may be read as a sort @{text "{c}"}.  The ordering
+  on type classes is extended to sorts according to the meaning of
+  intersections: @{text "{c\<^isub>1, \<dots> c\<^isub>m} \<subseteq> {d\<^isub>1, \<dots>, d\<^isub>n}"} iff @{text
+  "\<forall>j. \<exists>i. c\<^isub>i \<subseteq> d\<^isub>j"}.  The empty intersection @{text "{}"} refers to
+  the universal sort, which is the largest element wrt.\ the sort
+  order.  Thus @{text "{}"} 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{fixed type variable} is a pair of a basic name
   (starting with a @{text "'"} character) and a sort constraint, e.g.\
@@ -62,10 +63,10 @@
   printed as @{text "?\<alpha>\<^isub>s"}.
 
   Note that \emph{all} syntactic components contribute to the identity
-  of type variables, including the sort constraint.  The core logic
-  handles type variables with the same name but different sorts as
-  different, although some outer layers of the system make it hard to
-  produce anything like this.
+  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{type constructor} @{text "\<kappa>"} is a @{text "k"}-ary operator
   on types declared in the theory.  Type constructor application is
@@ -77,8 +78,8 @@
   right-associative infix @{text "\<alpha> \<Rightarrow> \<beta>"} instead of @{text "(\<alpha>,
   \<beta>)fun"}.
   
-  A \emph{type} is defined inductively over type variables and type
-  constructors as follows: @{text "\<tau> = \<alpha>\<^isub>s | ?\<alpha>\<^isub>s |
+  The logical category \emph{type} is defined inductively over type
+  variables and type constructors as follows: @{text "\<tau> = \<alpha>\<^isub>s | ?\<alpha>\<^isub>s |
   (\<tau>\<^sub>1, \<dots>, \<tau>\<^sub>k)\<kappa>"}.
 
   A \emph{type abbreviation} is a syntactic definition @{text
@@ -116,9 +117,9 @@
 
 text %mlref {*
   \begin{mldecls}
-  @{index_ML_type class} \\
-  @{index_ML_type sort} \\
-  @{index_ML_type arity} \\
+  @{index_ML_type class: string} \\
+  @{index_ML_type sort: "class list"} \\
+  @{index_ML_type arity: "string * sort list * sort"} \\
   @{index_ML_type typ} \\
   @{index_ML map_atyps: "(typ -> typ) -> typ -> typ"} \\
   @{index_ML fold_atyps: "(typ -> 'a -> 'a) -> typ -> 'a -> 'a"} \\
@@ -136,15 +137,15 @@
 
   \begin{description}
 
-  \item @{ML_type class} represents type classes; this is an alias for
-  @{ML_type string}.
+  \item @{ML_type class} represents type classes.
 
-  \item @{ML_type sort} represents sorts; this is an alias for
-  @{ML_type "class list"}.
+  \item @{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''.
 
-  \item @{ML_type arity} represents type arities; this is an alias for
-  triples of the form @{text "(\<kappa>, \<^vec>s, s)"} for @{text "\<kappa> ::
-  (\<^vec>s)s"} described above.
+  \item @{ML_type arity} represents type arities.  A triple @{text
+  "(\<kappa>, \<^vec>s, s) : arity"} represents @{text "\<kappa> :: (\<^vec>s)s"} as
+  described above.
 
   \item @{ML_type typ} represents types; this is a datatype with
   constructors @{ML TFree}, @{ML TVar}, @{ML Type}.
@@ -193,15 +194,13 @@
   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
-  are have an explicit name and type in each occurrence.
+  have an explicit name and type in each occurrence.
 
   \medskip A \emph{bound variable} is a natural number @{text "b"},
   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 @{text
-  "\<lambda>\<^bsub>nat\<^esub>. \<lambda>\<^bsub>nat\<^esub>. 1 + 0"} would
-  correspond to @{text
-  "\<lambda>x\<^bsub>nat\<^esub>. \<lambda>y\<^bsub>nat\<^esub>. x + y"} in a named
+  example, the de-Bruijn term @{text "\<lambda>\<^bsub>bool\<^esub>. \<lambda>\<^bsub>bool\<^esub>. 1 \<and> 0"} would
+  correspond to @{text "\<lambda>x\<^bsub>bool\<^esub>. \<lambda>y\<^bsub>bool\<^esub>. x \<and> y"} 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.
@@ -213,27 +212,27 @@
   without any loose variables.
 
   A \emph{fixed variable} is a pair of a basic name and a type, e.g.\
-  @{text "(x, \<tau>)"} which is usually printed @{text "x\<^isub>\<tau>"}.  A
+  @{text "(x, \<tau>)"} which is usually printed @{text "x\<^isub>\<tau>"} here.  A
   \emph{schematic variable} is a pair of an indexname and a type,
-  e.g.\ @{text "((x, 0), \<tau>)"} which is usually printed as @{text
+  e.g.\ @{text "((x, 0), \<tau>)"} which is likewise printed as @{text
   "?x\<^isub>\<tau>"}.
 
   \medskip A \emph{constant} is a pair of a basic name and a type,
-  e.g.\ @{text "(c, \<tau>)"} which is usually printed as @{text
-  "c\<^isub>\<tau>"}.  Constants are declared in the context as polymorphic
-  families @{text "c :: \<sigma>"}, meaning that all substitution instances
-  @{text "c\<^isub>\<tau>"} for @{text "\<tau> = \<sigma>\<vartheta>"} are valid.
+  e.g.\ @{text "(c, \<tau>)"} which is usually printed as @{text "c\<^isub>\<tau>"}
+  here.  Constants are declared in the context as polymorphic families
+  @{text "c :: \<sigma>"}, meaning that all substitution instances @{text
+  "c\<^isub>\<tau>"} for @{text "\<tau> = \<sigma>\<vartheta>"} are valid.
 
-  The vector of \emph{type arguments} of constant @{text "c\<^isub>\<tau>"}
-  wrt.\ the declaration @{text "c :: \<sigma>"} is defined as the codomain of
-  the matcher @{text "\<vartheta> = {?\<alpha>\<^isub>1 \<mapsto> \<tau>\<^isub>1, \<dots>,
-  ?\<alpha>\<^isub>n \<mapsto> \<tau>\<^isub>n}"} presented in canonical order @{text
-  "(\<tau>\<^isub>1, \<dots>, \<tau>\<^isub>n)"}.  Within a given theory context,
-  there is a one-to-one correspondence between any constant @{text
-  "c\<^isub>\<tau>"} and the application @{text "c(\<tau>\<^isub>1, \<dots>,
-  \<tau>\<^isub>n)"} of its type arguments.  For example, with @{text "plus
-  :: \<alpha> \<Rightarrow> \<alpha> \<Rightarrow> \<alpha>"}, the instance @{text "plus\<^bsub>nat \<Rightarrow> nat \<Rightarrow>
-  nat\<^esub>"} corresponds to @{text "plus(nat)"}.
+  The vector of \emph{type arguments} of constant @{text "c\<^isub>\<tau>"} wrt.\
+  the declaration @{text "c :: \<sigma>"} is defined as the codomain of the
+  matcher @{text "\<vartheta> = {?\<alpha>\<^isub>1 \<mapsto> \<tau>\<^isub>1, \<dots>, ?\<alpha>\<^isub>n \<mapsto> \<tau>\<^isub>n}"} presented in
+  canonical order @{text "(\<tau>\<^isub>1, \<dots>, \<tau>\<^isub>n)"}, corresponding to the
+  left-to-right occurrences of the @{text "\<alpha>\<^isub>i"} in @{text "\<sigma>"}.
+  Within a given theory context, there is a one-to-one correspondence
+  between any constant @{text "c\<^isub>\<tau>"} and the application @{text "c(\<tau>\<^isub>1,
+  \<dots>, \<tau>\<^isub>n)"} of its type arguments.  For example, with @{text "plus :: \<alpha>
+  \<Rightarrow> \<alpha> \<Rightarrow> \<alpha>"}, the instance @{text "plus\<^bsub>nat \<Rightarrow> nat \<Rightarrow> nat\<^esub>"} corresponds to
+  @{text "plus(nat)"}.
 
   Constant declarations @{text "c :: \<sigma>"} may contain sort constraints
   for type variables in @{text "\<sigma>"}.  These are observed by
@@ -242,14 +241,13 @@
   polymorphic constants that the user-level type-checker would reject
   due to violation of type class restrictions.
 
-  \medskip An \emph{atomic} term is either a variable or constant.  A
-  \emph{term} is defined inductively over atomic terms, with
-  abstraction and application as follows: @{text "t = b | x\<^isub>\<tau> |
-  ?x\<^isub>\<tau> | c\<^isub>\<tau> | \<lambda>\<^isub>\<tau>. t | t\<^isub>1 t\<^isub>2"}.
-  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.
+  \medskip An \emph{atomic} term is either a variable or constant.
+  The logical category \emph{term} is defined inductively over atomic
+  terms, with abstraction and application as follows: @{text "t = b |
+  x\<^isub>\<tau> | ?x\<^isub>\<tau> | c\<^isub>\<tau> | \<lambda>\<^isub>\<tau>. t | t\<^isub>1 t\<^isub>2"}.  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 @{text "t :: \<tau>"} assigns a (unique) type to a
   term according to the structure of atomic terms, abstractions, and
@@ -271,18 +269,17 @@
 
   The identity of atomic terms consists both of the name and the type
   component.  This means that different variables @{text
-  "x\<^bsub>\<tau>\<^isub>1\<^esub>"} and @{text
-  "x\<^bsub>\<tau>\<^isub>2\<^esub>"} may become the same after type
-  instantiation.  Some outer layers of the system make it hard to
-  produce variables of the same name, but different types.  In
-  contrast, mixed instances of polymorphic constants occur frequently.
+  "x\<^bsub>\<tau>\<^isub>1\<^esub>"} and @{text "x\<^bsub>\<tau>\<^isub>2\<^esub>"} 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{hidden polymorphism} of a term @{text "t :: \<sigma>"}
   is the set of type variables occurring in @{text "t"}, but not in
-  @{text "\<sigma>"}.  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 @{text "t\<vartheta> :: \<sigma>"} and @{text
-  "t\<vartheta>' :: \<sigma>"} with the same type.  This slightly
+  its type @{text "\<sigma>"}.  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 @{text "t\<vartheta> :: \<sigma>"} and
+  @{text "t\<vartheta>' :: \<sigma>"} with the same type.  This slightly
   pathological situation notoriously demands additional care.
 
   \medskip A \emph{term abbreviation} is a syntactic definition @{text
@@ -431,14 +428,14 @@
   \infer[@{text "(assume)"}]{@{text "A \<turnstile> A"}}{}
   \]
   \[
-  \infer[@{text "(\<And>_intro)"}]{@{text "\<Gamma> \<turnstile> \<And>x. b[x]"}}{@{text "\<Gamma> \<turnstile> b[x]"} & @{text "x \<notin> \<Gamma>"}}
+  \infer[@{text "(\<And>\<dash>intro)"}]{@{text "\<Gamma> \<turnstile> \<And>x. b[x]"}}{@{text "\<Gamma> \<turnstile> b[x]"} & @{text "x \<notin> \<Gamma>"}}
   \qquad
-  \infer[@{text "(\<And>_elim)"}]{@{text "\<Gamma> \<turnstile> b[a]"}}{@{text "\<Gamma> \<turnstile> \<And>x. b[x]"}}
+  \infer[@{text "(\<And>\<dash>elim)"}]{@{text "\<Gamma> \<turnstile> b[a]"}}{@{text "\<Gamma> \<turnstile> \<And>x. b[x]"}}
   \]
   \[
-  \infer[@{text "(\<Longrightarrow>_intro)"}]{@{text "\<Gamma> - A \<turnstile> A \<Longrightarrow> B"}}{@{text "\<Gamma> \<turnstile> B"}}
+  \infer[@{text "(\<Longrightarrow>\<dash>intro)"}]{@{text "\<Gamma> - A \<turnstile> A \<Longrightarrow> B"}}{@{text "\<Gamma> \<turnstile> B"}}
   \qquad
-  \infer[@{text "(\<Longrightarrow>_elim)"}]{@{text "\<Gamma>\<^sub>1 \<union> \<Gamma>\<^sub>2 \<turnstile> B"}}{@{text "\<Gamma>\<^sub>1 \<turnstile> A \<Longrightarrow> B"} & @{text "\<Gamma>\<^sub>2 \<turnstile> A"}}
+  \infer[@{text "(\<Longrightarrow>\<dash>elim)"}]{@{text "\<Gamma>\<^sub>1 \<union> \<Gamma>\<^sub>2 \<turnstile> B"}}{@{text "\<Gamma>\<^sub>1 \<turnstile> A \<Longrightarrow> B"} & @{text "\<Gamma>\<^sub>2 \<turnstile> A"}}
   \]
   \caption{Primitive inferences of Pure}\label{fig:prim-rules}
   \end{center}
@@ -467,21 +464,21 @@
   terms, and @{text "\<And>/\<Longrightarrow>"} for proofs (cf.\
   \cite{Berghofer-Nipkow:2000:TPHOL}).
 
-  Observe that locally fixed parameters (as in @{text "\<And>_intro"}) need
-  not be recorded in the hypotheses, because the simple syntactic
-  types of Pure are always inhabitable.  ``Assumptions'' @{text "x ::
-  \<tau>"} for type-membership are only present as long as some @{text
-  "x\<^isub>\<tau>"} occurs in the statement body.\footnote{This is the key
-  difference to ``@{text "\<lambda>HOL"}'' in the PTS framework
-  \cite{Barendregt-Geuvers:2001}, where hypotheses @{text "x : A"} are
-  treated uniformly for propositions and types.}
+  Observe that locally fixed parameters (as in @{text
+  "\<And>\<dash>intro"}) need not be recorded in the hypotheses, because
+  the simple syntactic types of Pure are always inhabitable.
+  ``Assumptions'' @{text "x :: \<tau>"} for type-membership are only
+  present as long as some @{text "x\<^isub>\<tau>"} occurs in the statement
+  body.\footnote{This is the key difference to ``@{text "\<lambda>HOL"}'' in
+  the PTS framework \cite{Barendregt-Geuvers:2001}, where hypotheses
+  @{text "x : A"} are treated uniformly for propositions and types.}
 
   \medskip The axiomatization of a theory is implicitly closed by
   forming all instances of type and term variables: @{text "\<turnstile>
   A\<vartheta>"} holds for any substitution instance of an axiom
   @{text "\<turnstile> A"}.  By pushing substitutions through derivations
   inductively, we also get admissible @{text "generalize"} and @{text
-  "instance"} rules as shown in \figref{fig:subst-rules}.
+  "instantiate"} rules as shown in \figref{fig:subst-rules}.
 
   \begin{figure}[htb]
   \begin{center}
@@ -588,11 +585,11 @@
   Every @{ML thm} value contains a sliding back-reference to the
   enclosing theory, cf.\ \secref{sec:context-theory}.
 
-  \item @{ML proofs} determines the detail of proof recording within
+  \item @{ML proofs} specifies the detail of proof recording within
   @{ML_type thm} values: @{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 always recorded.
+  to a result are recorded in any case.
 
   \item @{ML Thm.assume}, @{ML Thm.forall_intr}, @{ML
   Thm.forall_elim}, @{ML Thm.implies_intr}, and @{ML Thm.implies_elim}
@@ -644,8 +641,8 @@
   \begin{figure}[htb]
   \begin{center}
   \begin{tabular}{ll}
-  @{text "conjunction :: prop \<Rightarrow> prop \<Rightarrow> prop"} & (infix @{text "&"}) \\
-  @{text "\<turnstile> A & B \<equiv> (\<And>C. (A \<Longrightarrow> B \<Longrightarrow> C) \<Longrightarrow> C)"} \\[1ex]
+  @{text "conjunction :: prop \<Rightarrow> prop \<Rightarrow> prop"} & (infix @{text "&&&"}) \\
+  @{text "\<turnstile> A &&& B \<equiv> (\<And>C. (A \<Longrightarrow> B \<Longrightarrow> C) \<Longrightarrow> C)"} \\[1ex]
   @{text "prop :: prop \<Rightarrow> prop"} & (prefix @{text "#"}, suppressed) \\
   @{text "#A \<equiv> A"} \\[1ex]
   @{text "term :: \<alpha> \<Rightarrow> prop"} & (prefix @{text "TERM"}) \\
@@ -657,13 +654,15 @@
   \end{center}
   \end{figure}
 
-  Derived conjunction rules include introduction @{text "A \<Longrightarrow> B \<Longrightarrow> A &
-  B"}, and destructions @{text "A & B \<Longrightarrow> A"} and @{text "A & B \<Longrightarrow> B"}.
-  Conjunction allows to treat simultaneous assumptions and conclusions
-  uniformly.  For example, multiple claims are intermediately
-  represented as explicit conjunction, but this is refined into
-  separate sub-goals before the user continues the proof; the final
-  result is projected into a list of theorems (cf.\
+  The introduction @{text "A \<Longrightarrow> B \<Longrightarrow> A &&& B"}, and eliminations
+  (projections) @{text "A &&& B \<Longrightarrow> A"} and @{text "A &&& B \<Longrightarrow> B"} are
+  available as derived rules.  Conjunction allows to treat
+  simultaneous assumptions and conclusions uniformly, e.g.\ consider
+  @{text "A \<Longrightarrow> B \<Longrightarrow> C &&& D"}.  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 @{text "prop"} marker (@{text "#"}) makes arbitrarily complex
@@ -681,7 +680,7 @@
   language of types into that of terms.  There is specific notation
   @{text "TYPE(\<tau>)"} for @{text "TYPE\<^bsub>\<tau>
  itself\<^esub>"}.
-  Although being devoid of any particular meaning, the @{text
+  Although being devoid of any particular meaning, the term @{text
   "TYPE(\<tau>)"} accounts for the type @{text "\<tau>"} within the term
   language.  In particular, @{text "TYPE(\<alpha>)"} may be used as formal
   argument in primitive definitions, in order to circumvent hidden
@@ -703,11 +702,11 @@
 
   \begin{description}
 
-  \item @{ML Conjunction.intr} derives @{text "A & B"} from @{text
+  \item @{ML Conjunction.intr} derives @{text "A &&& B"} from @{text
   "A"} and @{text "B"}.
 
   \item @{ML Conjunction.elim} derives @{text "A"} and @{text "B"}
-  from @{text "A & B"}.
+  from @{text "A &&& B"}.
 
   \item @{ML Drule.mk_term} derives @{text "TERM t"}.
 
@@ -784,7 +783,8 @@
   "(A \<Longrightarrow> B) \<Longrightarrow> A \<longrightarrow> B"} or mathematical induction @{text "P 0 \<Longrightarrow> (\<And>n. P n
   \<Longrightarrow> P (Suc n)) \<Longrightarrow> P n"}.  Even deeper nesting occurs in well-founded
   induction @{text "(\<And>x. (\<And>y. y \<prec> x \<Longrightarrow> P y) \<Longrightarrow> P x) \<Longrightarrow> P x"}, but this
-  already marks the limit of rule complexity seen in practice.
+  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:
@@ -890,11 +890,10 @@
 
   \begin{description}
 
-  \item @{text "rule\<^sub>1 RS rule\<^sub>2"} resolves @{text
-  "rule\<^sub>1"} with @{text "rule\<^sub>2"} according to the
-  @{inference resolution} principle explained above.  Note that the
-  corresponding attribute in the Isar language is called @{attribute
-  THEN}.
+  \item @{text "rule\<^sub>1 RS rule\<^sub>2"} resolves @{text "rule\<^sub>1"} with @{text
+  "rule\<^sub>2"} according to the @{inference resolution} principle
+  explained above.  Note that the corresponding rule attribute in the
+  Isar language is called @{attribute THEN}.
 
   \item @{text "rule OF rules"} resolves a list of rules with the
   first rule, addressing its premises @{text "1, \<dots>, length rules"}
--- a/doc-src/IsarImplementation/Thy/Prelim.thy	Mon Feb 08 14:06:58 2010 +0100
+++ b/doc-src/IsarImplementation/Thy/Prelim.thy	Mon Feb 08 14:08:32 2010 +0100
@@ -73,41 +73,47 @@
 
 subsection {* Theory context \label{sec:context-theory} *}
 
-text {*
-  A \emph{theory} is a data container with explicit name and unique
-  identifier.  Theories are related by a (nominal) sub-theory
+text {* A \emph{theory} 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.
-
-  The @{text "merge"} operation produces the least upper bound of two
-  theories, which actually degenerates into absorption of one theory
-  into the other (due to the nominal sub-theory relation).
+  ancestor theories.  To this end, the system maintains a set of
+  symbolic ``identification stamps'' within each theory.
 
-  The @{text "begin"} operation starts a new theory by importing
-  several parent theories and entering a special @{text "draft"} mode,
-  which is sustained until the final @{text "end"} operation.  A draft
-  theory acts like a linear type, where updates invalidate earlier
-  versions.  An invalidated draft is called ``stale''.
+  In order to avoid the full-scale overhead of explicit sub-theory
+  identification of arbitrary intermediate stages, a theory is
+  switched into @{text "draft"} mode under certain circumstances.  A
+  draft theory acts like a linear type, where updates invalidate
+  earlier versions.  An invalidated draft is called \emph{stale}.
 
-  The @{text "checkpoint"} operation produces an intermediate stepping
-  stone that will survive the next update: both the original and the
-  changed theory remain valid and are related by the sub-theory
-  relation.  Checkpointing essentially recovers purely functional
-  theory values, at the expense of some extra internal bookkeeping.
+  The @{text "checkpoint"} operation produces a safe stepping stone
+  that will survive the next update without becoming stale: both the
+  old and the new theory remain valid and are related by the
+  sub-theory relation.  Checkpointing essentially recovers purely
+  functional theory values, at the expense of some extra internal
+  bookkeeping.
 
   The @{text "copy"} operation produces an auxiliary version that has
   the same data content, but is unrelated to the original: updates of
   the copy do not affect the original, neither does the sub-theory
   relation hold.
 
+  The @{text "merge"} operation produces the least upper bound of two
+  theories, which actually degenerates into absorption of one theory
+  into the other (according to the nominal sub-theory relation).
+
+  The @{text "begin"} operation starts a new theory by importing
+  several parent theories and entering a special mode of nameless
+  incremental updates, until the final @{text "end"} operation is
+  performed.
+
   \medskip The example in \figref{fig:ex-theory} below shows a theory
   graph derived from @{text "Pure"}, with theory @{text "Length"}
   importing @{text "Nat"} and @{text "List"}.  The body of @{text
   "Length"} consists of a sequence of updates, working mostly on
-  drafts.  Intermediate checkpoints may occur as well, due to the
-  history mechanism provided by the Isar top-level, cf.\
-  \secref{sec:isar-toplevel}.
+  drafts internally, while transaction boundaries of Isar top-level
+  commands (\secref{sec:isar-toplevel}) are guaranteed to be safe
+  checkpoints.
 
   \begin{figure}[htb]
   \begin{center}
@@ -147,9 +153,10 @@
   \begin{mldecls}
   @{index_ML_type theory} \\
   @{index_ML Theory.subthy: "theory * theory -> bool"} \\
-  @{index_ML Theory.merge: "theory * theory -> theory"} \\
   @{index_ML Theory.checkpoint: "theory -> theory"} \\
   @{index_ML Theory.copy: "theory -> theory"} \\
+  @{index_ML Theory.merge: "theory * theory -> theory"} \\
+  @{index_ML Theory.begin_theory: "string -> theory list -> theory"} \\
   \end{mldecls}
   \begin{mldecls}
   @{index_ML_type theory_ref} \\
@@ -160,25 +167,32 @@
   \begin{description}
 
   \item @{ML_type theory} represents theory contexts.  This is
-  essentially a linear type!  Most operations destroy the original
-  version, which then becomes ``stale''.
+  essentially a linear type, with explicit runtime checking!  Most
+  internal theory operations destroy the original version, which then
+  becomes ``stale''.
 
-  \item @{ML "Theory.subthy"}~@{text "(thy\<^sub>1, thy\<^sub>2)"}
-  compares theories according to the inherent graph structure of the
-  construction.  This sub-theory relation is a nominal approximation
-  of inclusion (@{text "\<subseteq>"}) of the corresponding content.
-
-  \item @{ML "Theory.merge"}~@{text "(thy\<^sub>1, thy\<^sub>2)"}
-  absorbs one theory into the other.  This fails for unrelated
-  theories!
+  \item @{ML "Theory.subthy"}~@{text "(thy\<^sub>1, thy\<^sub>2)"} compares theories
+  according to the intrinsic graph structure of the construction.
+  This sub-theory relation is a nominal approximation of inclusion
+  (@{text "\<subseteq>"}) of the corresponding content (according to the
+  semantics of the ML modules that implement the data).
 
   \item @{ML "Theory.checkpoint"}~@{text "thy"} produces a safe
-  stepping stone in the linear development of @{text "thy"}.  The next
-  update will result in two related, valid theories.
+  stepping stone in the linear development of @{text "thy"}.  This
+  changes the old theory, but the next update will result in two
+  related, valid theories.
 
   \item @{ML "Theory.copy"}~@{text "thy"} produces a variant of @{text
-  "thy"} with the same data.  The result is not related to the
-  original; the original is unchanged.
+  "thy"} with the same data.  The copy is not related to the original,
+  but the original is unchanged.
+
+  \item @{ML "Theory.merge"}~@{text "(thy\<^sub>1, thy\<^sub>2)"} absorbs one theory
+  into the other, without changing @{text "thy\<^sub>1"} or @{text "thy\<^sub>2"}.
+  This version of ad-hoc theory merge fails for unrelated theories!
+
+  \item @{ML "Theory.begin_theory"}~@{text "name parents"} constructs
+  a new theory based on the given parents.  This {\ML} function is
+  normally not invoked directly.
 
   \item @{ML_type theory_ref} represents a sliding reference to an
   always valid theory; updates on the original are propagated
@@ -198,31 +212,32 @@
 
 subsection {* Proof context \label{sec:context-proof} *}
 
-text {*
-  A proof context is a container for pure data with a back-reference
-  to the theory it belongs to.  The @{text "init"} operation creates a
-  proof context from a given theory.  Modifications to draft theories
-  are propagated to the proof context as usual, but there is also an
-  explicit @{text "transfer"} operation to force resynchronization
-  with more substantial updates to the underlying theory.  The actual
-  context data does not require any special bookkeeping, thanks to the
-  lack of destructive features.
+text {* A proof context is a container for pure data with a
+  back-reference to the theory it belongs to.  The @{text "init"}
+  operation creates a proof context from a given theory.
+  Modifications to draft theories are propagated to the proof context
+  as usual, but there is also an explicit @{text "transfer"} operation
+  to force resynchronization with more substantial updates to the
+  underlying theory.
 
-  Entities derived in a proof context need to record inherent logical
+  Entities derived in a proof context need to record logical
   requirements explicitly, since there is no separate context
-  identification as for theories.  For example, hypotheses used in
-  primitive derivations (cf.\ \secref{sec:thms}) are recorded
-  separately within the sequent @{text "\<Gamma> \<turnstile> \<phi>"}, 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.
+  identification or symbolic inclusion as for theories.  For example,
+  hypotheses used in primitive derivations (cf.\ \secref{sec:thms})
+  are recorded separately within the sequent @{text "\<Gamma> \<turnstile> \<phi>"}, 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 the Isar proof states model
+  into the original context.  Note that an Isar proof state models
   block-structured reasoning explicitly, using a stack of proof
-  contexts internally.
+  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!
 *}
 
 text %mlref {*
@@ -267,7 +282,8 @@
   Moreover, there are total operations @{text "theory_of"} and @{text
   "proof_of"} 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 @{text "init"} operation.
+  have to be constructed by an ad-hoc @{text "init"} operation, which
+  incurs a small runtime overhead.
 *}
 
 text %mlref {*
@@ -319,6 +335,16 @@
   \emph{any} theory that does not declare actual data content; @{text
   "extend"} is acts like a unitary version of @{text "merge"}.
 
+  Implementing @{text "merge"} can be tricky.  The general idea is
+  that @{text "merge (data\<^sub>1, data\<^sub>2)"} inserts those parts of @{text
+  "data\<^sub>2"} into @{text "data\<^sub>1"} 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!
+
   \paragraph{Proof context data} declarations need to implement the
   following SML signature:
 
@@ -330,15 +356,18 @@
   \medskip
 
   \noindent The @{text "init"} operation is supposed to produce a pure
-  value from the given background theory.
+  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 @{text "init"}
+  operation of \emph{all} theory data slots ever declared.
 
   \paragraph{Generic data} provides a hybrid interface for both theory
   and proof data.  The @{text "init"} operation for proof contexts is
   predefined to select the current data value from the background
   theory.
 
-  \bigskip A data declaration of type @{text "T"} results in the
-  following interface:
+  \bigskip Any of these data declaration over type @{text "T"} result
+  in an ML structure with the following signature:
 
   \medskip
   \begin{tabular}{ll}
@@ -348,12 +377,12 @@
   \end{tabular}
   \medskip
 
-  \noindent These other operations provide access for the particular
-  kind of context (theory, proof, or generic context).  Note that this
-  is a safe interface: there is no other way to access the
-  corresponding data slot of a context.  By keeping these operations
-  private, a component may maintain abstract values authentically,
-  without other components interfering.
+  \noindent These other operations provide exclusive access for the
+  particular kind of context (theory, proof, or generic context).
+  This interface fully 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.
 *}
 
 text %mlref {*
@@ -379,32 +408,127 @@
   \end{description}
 *}
 
+text %mlex {*
+  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:
+*}
+
+ML {*
+  signature WELLFORMED_TERMS =
+  sig
+    val get: theory -> term list
+    val add: term -> theory -> theory
+  end;
+*}
+
+text {* \noindent The implementation uses private theory data
+  internally, and only exposes an operation that involves explicit
+  argument checking wrt.\ the given theory. *}
+
+ML {*
+  structure Wellformed_Terms: WELLFORMED_TERMS =
+  struct
+
+  structure Terms = Theory_Data
+  (
+    type T = term OrdList.T;
+    val empty = [];
+    val extend = I;
+    fun merge (ts1, ts2) =
+      OrdList.union TermOrd.fast_term_ord ts1 ts2;
+  )
+
+  val get = Terms.get;
+
+  fun add raw_t thy =
+    let val t = Sign.cert_term thy raw_t
+    in Terms.map (OrdList.insert TermOrd.fast_term_ord t) thy end;
+
+  end;
+*}
+
+text {* We use @{ML_type "term OrdList.T"} for reasonably efficient
+  representation of a set of terms: all operations are linear in the
+  number of stored elements.  Here we assume that our users do not
+  care about the declaration order, since that data structure forces
+  its own arrangement of elements.
+
+  Observe how the @{verbatim merge} operation joins the data slots of
+  the two constituents: @{ML OrdList.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:
+  \begin{enumerate}
+
+  \item @{ML Wellformed_Terms.add} only admits terms that have passed
+  the @{ML Sign.cert_term} check of the given theory at that point.
+
+  \item 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.
+
+  \end{enumerate}
+
+  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
+  strengthened 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.
+*}
+
 
 section {* Names \label{sec:names} *}
 
-text {*
-  In principle, a name is just a string, but there are various
-  convention for encoding additional structure.  For example, ``@{text
-  "Foo.bar.baz"}'' is considered as a qualified name consisting of
-  three basic name components.  The individual constituents of a name
-  may have further substructure, e.g.\ the string
-  ``\verb,\,\verb,<alpha>,'' encodes as a single symbol.
+text {* In principle, a name is just a string, but there are various
+  conventions for representing additional structure.  For example,
+  ``@{text "Foo.bar.baz"}'' is considered as a long name consisting of
+  qualifier @{text "Foo.bar"} and base name @{text "baz"}.  The
+  individual constituents of a name may have further substructure,
+  e.g.\ the string ``\verb,\,\verb,<alpha>,'' encodes as a single
+  symbol.
+
+  \medskip Subsequently, we shall introduce specific categories of
+  names.  Roughly speaking these correspond to logical entities as
+  follows:
+  \begin{itemize}
+
+  \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}).
+
+  \end{itemize}
 *}
 
 
 subsection {* Strings of symbols *}
 
-text {*
-  A \emph{symbol} constitutes the smallest textual unit in Isabelle
-  --- raw characters are normally not encountered at all.  Isabelle
-  strings consist of a sequence of symbols, represented as a packed
-  string or a list of strings.  Each symbol is in itself a small
-  string, which has either one of the following forms:
+text {* A \emph{symbol} constitutes the smallest textual unit in
+  Isabelle --- 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:
 
   \begin{enumerate}
 
-  \item a single ASCII character ``@{text "c"}'', for example
-  ``\verb,a,'',
+  \item a single ASCII character ``@{text "c"}'' or raw byte in the
+  range of 128\dots 255, for example ``\verb,a,'',
 
   \item a regular symbol ``\verb,\,\verb,<,@{text "ident"}\verb,>,'',
   for example ``\verb,\,\verb,<alpha>,'',
@@ -413,7 +537,7 @@
   for example ``\verb,\,\verb,<^bold>,'',
 
   \item a raw symbol ``\verb,\,\verb,<^raw:,@{text text}\verb,>,''
-  where @{text text} constists of printable characters excluding
+  where @{text text} consists of printable characters excluding
   ``\verb,.,'' and ``\verb,>,'', for example
   ``\verb,\,\verb,<^raw:$\sum_{i = 1}^n$>,'',
 
@@ -432,22 +556,31 @@
   may occur within regular Isabelle identifiers.
 
   Since the character set underlying Isabelle symbols is 7-bit ASCII
-  and 8-bit characters are passed through transparently, Isabelle may
-  also process Unicode/UCS data in UTF-8 encoding.  Unicode provides
-  its own collection of mathematical symbols, but there is no built-in
-  link to the standard collection of Isabelle.
+  and 8-bit characters are passed through transparently, Isabelle can
+  also process Unicode/UCS data in UTF-8 encoding.\footnote{When
+  counting precise source positions internally, bytes in the range of
+  128\dots 191 are ignored.  In UTF-8 encoding, this interval covers
+  the additional trailer bytes, so Isabelle happens to count Unicode
+  characters here, not bytes in memory.  In ISO-Latin encoding, the
+  ignored range merely includes some extra punctuation characters that
+  even have replacements within the standard collection of Isabelle
+  symbols; the accented letters range is counted properly.} Unicode
+  provides its own collection of mathematical symbols, but within the
+  core Isabelle/ML world there is no link to the standard collection
+  of Isabelle regular symbols.
 
   \medskip Output of Isabelle symbols depends on the print mode
   (\secref{print-mode}).  For example, the standard {\LaTeX} setup of
   the Isabelle document preparation system would present
   ``\verb,\,\verb,<alpha>,'' as @{text "\<alpha>"}, and
   ``\verb,\,\verb,<^bold>,\verb,\,\verb,<alpha>,'' as @{text
-  "\<^bold>\<alpha>"}.
+  "\<^bold>\<alpha>"}.  On-screen rendering usually works by mapping a finite
+  subset of Isabelle symbols to suitable Unicode characters.
 *}
 
 text %mlref {*
   \begin{mldecls}
-  @{index_ML_type "Symbol.symbol"} \\
+  @{index_ML_type "Symbol.symbol": string} \\
   @{index_ML Symbol.explode: "string -> Symbol.symbol list"} \\
   @{index_ML Symbol.is_letter: "Symbol.symbol -> bool"} \\
   @{index_ML Symbol.is_digit: "Symbol.symbol -> bool"} \\
@@ -462,12 +595,15 @@
   \begin{description}
 
   \item @{ML_type "Symbol.symbol"} represents individual Isabelle
-  symbols; this is an alias for @{ML_type "string"}.
+  symbols.
 
   \item @{ML "Symbol.explode"}~@{text "str"} produces a symbol list
   from the packed form.  This function supercedes @{ML
   "String.explode"} for virtually all purposes of manipulating text in
-  Isabelle!
+  Isabelle!\footnote{The runtime overhead for exploded strings is
+  mainly that of the list structure: individual symbols that happen to
+  be a singleton string --- which is the most common case --- do not
+  require extra memory in Poly/ML.}
 
   \item @{ML "Symbol.is_letter"}, @{ML "Symbol.is_digit"}, @{ML
   "Symbol.is_quasi"}, @{ML "Symbol.is_blank"} classify standard
@@ -483,10 +619,20 @@
   symbol into the datatype version.
 
   \end{description}
+
+  \paragraph{Historical note.} In the original SML90 standard the
+  primitive ML type @{ML_type char} did not exists, and the basic @{ML
+  "explode: string -> string list"} operation would produce a list of
+  singleton strings as in Isabelle/ML today.  When SML97 came out,
+  Isabelle did not adopt its slightly anachronistic 8-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.
 *}
 
 
-subsection {* Basic names \label{sec:basic-names} *}
+subsection {* Basic names \label{sec:basic-name} *}
 
 text {*
   A \emph{basic name} essentially consists of a single Isabelle
@@ -502,8 +648,8 @@
   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 name references like @{text
-  "xaa"} to appear in the text.
+  as internal, which prevents mysterious names like @{text "xaa"} to
+  appear in human-readable text.
 
   \medskip Manipulating binding scopes often requires on-the-fly
   renamings.  A \emph{name context} contains a collection of already
@@ -534,6 +680,9 @@
   @{index_ML Name.invents: "Name.context -> string -> int -> string list"} \\
   @{index_ML Name.variants: "string list -> Name.context -> string list * Name.context"} \\
   \end{mldecls}
+  \begin{mldecls}
+  @{index_ML Variable.names_of: "Proof.context -> Name.context"} \\
+  \end{mldecls}
 
   \begin{description}
 
@@ -555,11 +704,20 @@
   \item @{ML Name.variants}~@{text "names context"} produces fresh
   variants of @{text "names"}; the result is entered into the context.
 
+  \item @{ML Variable.names_of}~@{text "ctxt"} 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_struct
+  Variable}, which is also able to provide an official status of
+  ``locally fixed variable'' within the logical environment (cf.\
+  \secref{sec:variables}).
+
   \end{description}
 *}
 
 
-subsection {* Indexed names *}
+subsection {* Indexed names \label{sec:indexname} *}
 
 text {*
   An \emph{indexed name} (or @{text "indexname"}) is a pair of a basic
@@ -571,9 +729,9 @@
   @{text "maxidx + 1"}; the maximum index of an empty collection is
   @{text "-1"}.
 
-  Occasionally, basic names and indexed names are injected into the
-  same pair type: the (improper) indexname @{text "(x, -1)"} is used
-  to encode basic names.
+  Occasionally, basic names are injected into the same pair type of
+  indexed names: then @{text "(x, -1)"} is used to encode the basic
+  name @{text "x"}.
 
   \medskip Isabelle syntax observes the following rules for
   representing an indexname @{text "(x, i)"} as a packed string:
@@ -588,11 +746,12 @@
 
   \end{itemize}
 
-  Indexnames may acquire large index numbers over time.  Results are
-  normalized towards @{text "0"} 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
-  @{text "?x1, ?x7, ?x42"} becomes @{text "?x, ?xa, ?xb"}.
+  Indexnames may acquire large index numbers after several maxidx
+  shifts have been applied.  Results are usually normalized towards
+  @{text "0"} 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 @{text "?x1, ?x7, ?x42"}
+  becomes @{text "?x, ?xa, ?xb"}.
 *}
 
 text %mlref {*
@@ -605,65 +764,35 @@
   \item @{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 @{text "(x, -1)"}
-  is used to embed basic names into this type.
+  is used to inject basic names into this type.  Other negative
+  indexes should not be used.
 
   \end{description}
 *}
 
 
-subsection {* Qualified names and name spaces *}
+subsection {* Long names \label{sec:long-name} *}
 
-text {*
-  A \emph{qualified name} consists of a non-empty sequence of basic
-  name components.  The packed representation uses a dot as separator,
-  as in ``@{text "A.b.c"}''.  The last component is called \emph{base}
-  name, the remaining prefix \emph{qualifier} (which may be empty).
-  The idea of qualified names is to encode nested structures by
-  recording the access paths as qualifiers.  For example, an item
-  named ``@{text "A.b.c"}'' may be understood as a local entity @{text
-  "c"}, within a local structure @{text "b"}, within a global
-  structure @{text "A"}.  Typically, name space hierarchies consist of
-  1--2 levels of qualification, but this need not be always so.
+text {* A \emph{long name} consists of a sequence of non-empty name
+  components.  The packed representation uses a dot as separator, as
+  in ``@{text "A.b.c"}''.  The last component is called \emph{base
+  name}, the remaining prefix is called \emph{qualifier} (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 ``@{text "A.b.c"}'' may be understood as
+  a local entity @{text "c"}, within a local structure @{text "b"},
+  within a global structure @{text "A"}.  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, whenever this makes any sense.  The basic operations on
-  qualified names are smart enough to pass through such improper names
-  unchanged.
-
-  \medskip A @{text "naming"} policy tells how to turn a name
-  specification into a fully qualified internal name (by the @{text
-  "full"} operation), and how fully qualified names may be accessed
-  externally.  For example, the default naming policy is to prefix an
-  implicit path: @{text "full x"} produces @{text "path.x"}, and the
-  standard accesses for @{text "path.x"} include both @{text "x"} and
-  @{text "path.x"}.  Normally, the naming is implicit in the theory or
-  proof context; there are separate versions of the corresponding.
-
-  \medskip A @{text "name space"} manages a collection of fully
-  internalized names, together with a mapping between external names
-  and internal names (in both directions).  The corresponding @{text
-  "intern"} and @{text "extern"} operations are mostly used for
-  parsing and printing only!  The @{text "declare"} operation augments
-  a name space according to the accesses determined by the naming
-  policy.
-
-  \medskip As a general principle, there is a separate name space for
-  each kind of formal entity, e.g.\ logical constant, type
-  constructor, type class, theorem.  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 @{text
-  "c"} may be used uniformly for a constant, type constructor, and
-  type class.
-
-  There are common schemes to name theorems systematically, according
-  to the name of the main logical entity involved, e.g.\ @{text
-  "c.intro"} for a canonical theorem related to constant @{text "c"}.
-  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 are better suffixed in
-  addition to the usual qualification, e.g.\ @{text "c_type.intro"}
-  and @{text "c_class.intro"} for theorems related to type @{text "c"}
-  and class @{text "c"}, respectively.
+  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.
 *}
 
 text %mlref {*
@@ -674,6 +803,85 @@
   @{index_ML Long_Name.implode: "string list -> string"} \\
   @{index_ML Long_Name.explode: "string -> string list"} \\
   \end{mldecls}
+
+  \begin{description}
+
+  \item @{ML Long_Name.base_name}~@{text "name"} returns the base name
+  of a long name.
+
+  \item @{ML Long_Name.qualifier}~@{text "name"} returns the qualifier
+  of a long name.
+
+  \item @{ML Long_Name.append}~@{text "name\<^isub>1 name\<^isub>2"} appends two long
+  names.
+
+  \item @{ML Long_Name.implode}~@{text "names"} and @{ML
+  Long_Name.explode}~@{text "name"} convert between the packed string
+  representation and the explicit list form of long names.
+
+  \end{description}
+*}
+
+
+subsection {* Name spaces \label{sec:name-space} *}
+
+text {* A @{text "name space"} 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 @{text "intern"} and @{text "extern"} operations
+  are mostly used for parsing and printing only!  The @{text
+  "declare"} operation augments a name space according to the accesses
+  determined by a given binding, and a naming policy from the context.
+
+  \medskip A @{text "binding"} 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 @{text "naming"} provides some additional details for
+  producing a long name from a binding.  Normally, the naming is
+  implicit in the theory or proof context.  The @{text "full"}
+  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:
+
+  \smallskip
+  \begin{tabular}{l}
+  @{text "binding + naming \<longrightarrow> long name + name space accesses"}
+  \end{tabular}
+  \smallskip
+
+  \medskip 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 @{text "c"} 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 @{text "c.intro"} for a canonical introduction rule related to
+  constant @{text "c"}.  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 are better suffixed in addition to the usual qualification,
+  e.g.\ @{text "c_type.intro"} and @{text "c_class.intro"} for
+  theorems related to type @{text "c"} and class @{text "c"},
+  respectively.
+*}
+
+text %mlref {*
+  \begin{mldecls}
+  @{index_ML_type binding} \\
+  @{index_ML Binding.empty: binding} \\
+  @{index_ML Binding.name: "string -> binding"} \\
+  @{index_ML Binding.qualify: "bool -> string -> binding -> binding"} \\
+  @{index_ML Binding.prefix: "bool -> string -> binding -> binding"} \\
+  @{index_ML Binding.conceal: "binding -> binding"} \\
+  @{index_ML Binding.str_of: "binding -> string"} \\
+  \end{mldecls}
   \begin{mldecls}
   @{index_ML_type Name_Space.naming} \\
   @{index_ML Name_Space.default_naming: Name_Space.naming} \\
@@ -688,22 +896,41 @@
   string * Name_Space.T"} \\
   @{index_ML Name_Space.intern: "Name_Space.T -> string -> string"} \\
   @{index_ML Name_Space.extern: "Name_Space.T -> string -> string"} \\
+  @{index_ML Name_Space.is_concealed: "Name_Space.T -> string -> bool"}
   \end{mldecls}
 
   \begin{description}
 
-  \item @{ML Long_Name.base_name}~@{text "name"} returns the base name of a
-  qualified name.
+  \item @{ML_type binding} represents the abstract concept of name
+  bindings.
+
+  \item @{ML Binding.empty} is the empty binding.
 
-  \item @{ML Long_Name.qualifier}~@{text "name"} returns the qualifier
-  of a qualified name.
+  \item @{ML Binding.name}~@{text "name"} produces a binding with base
+  name @{text "name"}.
+
+  \item @{ML Binding.qualify}~@{text "mandatory name binding"}
+  prefixes qualifier @{text "name"} to @{text "binding"}.  The @{text
+  "mandatory"} flag tells if this name component always needs to be
+  given in name space accesses --- this is mostly @{text "false"} in
+  practice.  Note that this part of qualification is typically used in
+  derived specification mechanisms.
 
-  \item @{ML Long_Name.append}~@{text "name\<^isub>1 name\<^isub>2"}
-  appends two qualified names.
+  \item @{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}).
 
-  \item @{ML Long_Name.implode}~@{text "names"} and @{ML
-  Long_Name.explode}~@{text "name"} convert between the packed string
-  representation and the explicit list form of qualified names.
+  \item @{ML Binding.conceal}~@{text "binding"} 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}).
+
+  \item @{ML Binding.str_of}~@{text "binding"} produces a string
+  representation for human-readable output, together with some formal
+  markup that might get used in GUI front-ends, for example.
 
   \item @{ML_type Name_Space.naming} represents the abstract concept of
   a naming policy.
@@ -747,6 +974,10 @@
   This operation is mostly for printing!  User code should not rely on
   the precise result too much.
 
+  \item @{ML Name_Space.is_concealed}~@{text "space name"} indicates
+  whether @{text "name"} refers to a strictly private entity that
+  other tools are supposed to ignore!
+
   \end{description}
 *}
 
--- a/doc-src/IsarImplementation/Thy/Proof.thy	Mon Feb 08 14:06:58 2010 +0100
+++ b/doc-src/IsarImplementation/Thy/Proof.thy	Mon Feb 08 14:08:32 2010 +0100
@@ -25,10 +25,10 @@
   categories for \emph{fixed variables} (e.g.\ @{text "x"}) vs.\
   \emph{schematic variables} (e.g.\ @{text "?x"}).  Incidently, a
   universal result @{text "\<turnstile> \<And>x. B(x)"} has the HHF normal form @{text
-  "\<turnstile> B(?x)"}, which represents its generality nicely without requiring
-  an explicit quantifier.  The same principle works for type
-  variables: @{text "\<turnstile> B(?\<alpha>)"} represents the idea of ``@{text "\<turnstile>
-  \<forall>\<alpha>. B(\<alpha>)"}'' without demanding a truly polymorphic framework.
+  "\<turnstile> B(?x)"}, which represents its generality without requiring an
+  explicit quantifier.  The same principle works for type variables:
+  @{text "\<turnstile> B(?\<alpha>)"} represents the idea of ``@{text "\<turnstile> \<forall>\<alpha>. B(\<alpha>)"}''
+  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
@@ -95,7 +95,8 @@
   @{index_ML Variable.polymorphic: "Proof.context -> term list -> term list"} \\
   @{index_ML Variable.import: "bool -> thm list -> Proof.context ->
   (((ctyp * ctyp) list * (cterm * cterm) list) * thm list) * Proof.context"} \\
-  @{index_ML Variable.focus: "cterm -> Proof.context -> ((string * cterm) list * cterm) * Proof.context"} \\
+  @{index_ML Variable.focus: "cterm -> Proof.context ->
+  ((string * cterm) list * cterm) * Proof.context"} \\
   \end{mldecls}
 
   \begin{description}
@@ -144,6 +145,72 @@
   \end{description}
 *}
 
+text %mlex {* The following example (in theory @{theory Pure}) shows
+  how to work with fixed term and type parameters work with
+  type-inference.
+*}
+
+typedecl foo  -- {* some basic type for testing purposes *}
+
+ML {*
+  (*static compile-time context -- for testing only*)
+  val ctxt0 = @{context};
+
+  (*locally fixed parameters -- no type assignment yet*)
+  val ([x, y], ctxt1) = ctxt0 |> Variable.add_fixes ["x", "y"];
+
+  (*t1: most general fixed type; t1': most general arbitrary type*)
+  val t1 = Syntax.read_term ctxt1 "x";
+  val t1' = singleton (Variable.polymorphic ctxt1) t1;
+
+  (*term u enforces specific type assignment*)
+  val u = Syntax.read_term ctxt1 "(x::foo) \<equiv> y";
+
+  (*official declaration of u -- propagates constraints etc.*)
+  val ctxt2 = ctxt1 |> Variable.declare_term u;
+  val t2 = Syntax.read_term ctxt2 "x";  (*x::foo is enforced*)
+*}
+
+text {* In the above example, the starting context had been derived
+  from the toplevel theory, which means that fixed variables are
+  internalized literally: @{verbatim "x"} is mapped again to
+  @{verbatim "x"}, and attempting to fix it again in the subsequent
+  context is an error.  Alternatively, fixed parameters can be renamed
+  explicitly as follows:
+*}
+
+ML {*
+  val ctxt0 = @{context};
+  val ([x1, x2, x3], ctxt1) =
+    ctxt0 |> Variable.variant_fixes ["x", "x", "x"];
+*}
+
+text {* \noindent Subsequent ML code can now work with the invented
+  names of @{verbatim x1}, @{verbatim x2}, @{verbatim x3}, 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:
+*}
+
+lemma "PROP XXX"
+proof -
+  ML_prf %"ML" {*
+    val ctxt0 = @{context};
+
+    val ([x1], ctxt1) = ctxt0 |> Variable.add_fixes ["x"];
+    val ([x2], ctxt2) = ctxt1 |> Variable.add_fixes ["x"];
+    val ([x3], ctxt3) = ctxt2 |> Variable.add_fixes ["x"];
+
+    val ([y1, y2], ctxt4) =
+      ctxt3 |> Variable.variant_fixes ["y", "y"];
+  *}
+  oops
+
+text {* \noindent 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.
+*}
+
 
 section {* Assumptions \label{sec:assumptions} *}
 
@@ -171,21 +238,21 @@
   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 contexts, not the facts being exported!  There is a
+  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 @{text "\<Longrightarrow>"} introduction rule:
   \[
-  \infer[(@{text "\<Longrightarrow>_intro"})]{@{text "\<Gamma> \\ A \<turnstile> A \<Longrightarrow> B"}}{@{text "\<Gamma> \<turnstile> B"}}
+  \infer[(@{text "\<Longrightarrow>\<dash>intro"})]{@{text "\<Gamma> - A \<turnstile> A \<Longrightarrow> B"}}{@{text "\<Gamma> \<turnstile> B"}}
   \]
 
   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[(@{text "#\<Longrightarrow>_intro"})]{@{text "\<Gamma> \\ A \<turnstile> #A \<Longrightarrow> B"}}{@{text "\<Gamma> \<turnstile> B"}}
+  \infer[(@{text "#\<Longrightarrow>\<dash>intro"})]{@{text "\<Gamma> - A \<turnstile> #A \<Longrightarrow> B"}}{@{text "\<Gamma> \<turnstile> B"}}
   \]
 
   \medskip Alternative versions of assumptions may perform arbitrary
@@ -194,7 +261,7 @@
   definition works by fixing @{text "x"} and assuming @{text "x \<equiv> t"},
   with the following export rule to reverse the effect:
   \[
-  \infer[(@{text "\<equiv>-expand"})]{@{text "\<Gamma> \\ x \<equiv> t \<turnstile> B t"}}{@{text "\<Gamma> \<turnstile> B x"}}
+  \infer[(@{text "\<equiv>\<dash>expand"})]{@{text "\<Gamma> - (x \<equiv> t) \<turnstile> B t"}}{@{text "\<Gamma> \<turnstile> B x"}}
   \]
   This works, because the assumption @{text "x \<equiv> t"} was introduced in
   a context with @{text "x"} being fresh, so @{text "x"} does not
@@ -222,8 +289,8 @@
   simultaneously.
 
   \item @{ML Assumption.assume}~@{text "A"} turns proposition @{text
-  "A"} into a raw assumption @{text "A \<turnstile> A'"}, where the conclusion
-  @{text "A'"} is in HHF normal form.
+  "A"} into a primitive assumption @{text "A \<turnstile> A'"}, where the
+  conclusion @{text "A'"} is in HHF normal form.
 
   \item @{ML Assumption.add_assms}~@{text "r As"} augments the context
   by assumptions @{text "As"} with export rule @{text "r"}.  The
@@ -232,7 +299,8 @@
 
   \item @{ML Assumption.add_assumes}~@{text "As"} is a special case of
   @{ML Assumption.add_assms} where the export rule performs @{text
-  "\<Longrightarrow>_intro"} or @{text "#\<Longrightarrow>_intro"}, depending on goal mode.
+  "\<Longrightarrow>\<dash>intro"} or @{text "#\<Longrightarrow>\<dash>intro"}, depending on goal
+  mode.
 
   \item @{ML Assumption.export}~@{text "is_goal inner outer thm"}
   exports result @{text "thm"} from the the @{text "inner"} context
@@ -244,8 +312,32 @@
   \end{description}
 *}
 
+text %mlex {* 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.
+*}
 
-section {* Results \label{sec:results} *}
+ML {*
+  (*static compile-time context -- for testing only*)
+  val ctxt0 = @{context};
+
+  val ([eq], ctxt1) =
+    ctxt0 |> Assumption.add_assumes [@{cprop "x \<equiv> y"}];
+  val eq' = Thm.symmetric eq;
+
+  (*back to original context -- discharges assumption*)
+  val r = Assumption.export false ctxt1 ctxt0 eq';
+*}
+
+text {* \noindent 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 "ProofContext.export"}).
+*}
+
+
+section {* Structured goals and results \label{sec:struct-goals} *}
 
 text {*
   Local results are established by monotonic reasoning from facts
@@ -263,6 +355,10 @@
   the tactic needs to solve the conclusion, but may use the premise as
   a local fact, for locally fixed variables.
 
+  The family of @{text "FOCUS"} combinators is similar to @{text
+  "SUBPROOF"}, but allows to retain schematic variables and pending
+  subgoals in the resulting goal state.
+
   The @{text "prove"} 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
@@ -275,7 +371,8 @@
   The @{text "obtain"} 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 certain fixed variables and assumptions.  See also
+  by parameters and assumptions, without affecting any conclusions
+  that do not mention these parameters.  See also
   \cite{isabelle-isar-ref} for the user-level @{text "\<OBTAIN>"} and
   @{text "\<GUESS>"} elements.  Final results, which may not refer to
   the parameters in the conclusion, need to exported explicitly into
@@ -285,7 +382,11 @@
 text %mlref {*
   \begin{mldecls}
   @{index_ML SUBPROOF: "(Subgoal.focus -> tactic) -> Proof.context -> int -> tactic"} \\
+  @{index_ML Subgoal.FOCUS: "(Subgoal.focus -> tactic) -> Proof.context -> int -> tactic"} \\
+  @{index_ML Subgoal.FOCUS_PREMS: "(Subgoal.focus -> tactic) -> Proof.context -> int -> tactic"} \\
+  @{index_ML Subgoal.FOCUS_PARAMS: "(Subgoal.focus -> tactic) -> Proof.context -> int -> tactic"} \\
   \end{mldecls}
+
   \begin{mldecls}
   @{index_ML Goal.prove: "Proof.context -> string list -> term list -> term ->
   ({prems: thm list, context: Proof.context} -> tactic) -> thm"} \\
@@ -305,6 +406,12 @@
   schematic parameters of the goal are imported into the context as
   fixed ones, which may not be instantiated in the sub-proof.
 
+  \item @{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.
+
   \item @{ML Goal.prove}~@{text "ctxt xs As C tac"} states goal @{text
   "C"} in the context augmented by fixed variables @{text "xs"} and
   assumptions @{text "As"}, and applies tactic @{text "tac"} to solve
--- a/doc-src/IsarImplementation/Thy/Syntax.thy	Mon Feb 08 14:06:58 2010 +0100
+++ b/doc-src/IsarImplementation/Thy/Syntax.thy	Mon Feb 08 14:08:32 2010 +0100
@@ -2,8 +2,14 @@
 imports Base
 begin
 
-chapter {* Syntax and type-checking *}
+chapter {* Concrete syntax and type-checking *}
 
 text FIXME
 
+section {* Parsing and printing *}
+
+text FIXME
+
+section {* Checking and unchecking \label{sec:term-check} *}
+
 end
--- a/doc-src/IsarImplementation/Thy/Tactic.thy	Mon Feb 08 14:06:58 2010 +0100
+++ b/doc-src/IsarImplementation/Thy/Tactic.thy	Mon Feb 08 14:08:32 2010 +0100
@@ -4,15 +4,13 @@
 
 chapter {* Tactical reasoning *}
 
-text {*
-  Tactical reasoning works by refining the initial claim in a
+text {* Tactical reasoning works by refining an initial claim in a
   backwards fashion, until a solved form is reached.  A @{text "goal"}
   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 @{text "tactic"} is a refinement operation that maps
   a goal to a lazy sequence of potential successors.  A @{text
-  "tactical"} is a combinator for composing tactics.
-*}
+  "tactical"} is a combinator for composing tactics.  *}
 
 
 section {* Goals \label{sec:tactical-goals} *}
@@ -40,8 +38,8 @@
 
   The main conclusion @{text C} is internally marked as a protected
   proposition, which is represented explicitly by the notation @{text
-  "#C"}.  This ensures that the decomposition into subgoals and main
-  conclusion is well-defined for arbitrarily structured claims.
+  "#C"} 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:
@@ -98,26 +96,27 @@
   supporting memoing.\footnote{The lack of memoing and the strict
   nature of SML requires some care when working with low-level
   sequence operations, to avoid duplicate or premature evaluation of
-  results.}
+  results.  It also means that modified runtime behavior, such as
+  timeout, is very hard to achieve for general tactics.}
 
   An \emph{empty result sequence} means that the tactic has failed: in
-  a compound tactic expressions other tactics might be tried instead,
+  a compound tactic expression other tactics might be tried instead,
   or the whole refinement step might fail outright, producing a
-  toplevel error message.  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.
+  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{multiple results}, 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 in \secref{sec:tactical-goals}, 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).
+  \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).
 
   Tactics with explicit \emph{subgoal addressing} are of the form
   @{text "int \<rightarrow> tactic"} and may be applied to a particular subgoal
@@ -139,7 +138,7 @@
 
   Tactics with internal subgoal addressing should expose the subgoal
   index as @{text "int"} argument in full generality; a hardwired
-  subgoal 1 inappropriate.
+  subgoal 1 is not acceptable.
   
   \medskip The main well-formedness conditions for proper tactics are
   summarized as follows.
@@ -161,11 +160,11 @@
   \end{itemize}
 
   Some of these conditions are checked by higher-level goal
-  infrastructure (\secref{sec:results}); others are not checked
+  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 disallow composition with
-  basic tacticals).
+  applicable only to singleton goals, or prevent composition via
+  standard tacticals).
 *}
 
 text %mlref {*
@@ -356,7 +355,7 @@
   "(?'a, \<tau>)"}.  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 type-inference as expected, explicit type
+  already performs simple type-inference, so explicit type
   instantiations are seldom necessary.
 *}
 
@@ -389,6 +388,12 @@
   names} (which need to be distinct indentifiers).
 
   \end{description}
+
+  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.
 *}
 
 
--- a/doc-src/IsarImplementation/Thy/document/Integration.tex	Mon Feb 08 14:06:58 2010 +0100
+++ b/doc-src/IsarImplementation/Thy/document/Integration.tex	Mon Feb 08 14:08:32 2010 +0100
@@ -106,8 +106,9 @@
   \item \verb|Toplevel.is_toplevel|~\isa{state} checks for an empty
   toplevel state.
 
-  \item \verb|Toplevel.theory_of|~\isa{state} selects the theory of
-  a theory or proof (!), otherwise raises \verb|Toplevel.UNDEF|.
+  \item \verb|Toplevel.theory_of|~\isa{state} selects the
+  background theory of \isa{state}, raises \verb|Toplevel.UNDEF|
+  for an empty toplevel state.
 
   \item \verb|Toplevel.proof_of|~\isa{state} selects the Isar proof
   state if available, otherwise raises \verb|Toplevel.UNDEF|.
@@ -150,16 +151,16 @@
   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} acts
+  alternative state transitions.  For example, \isakeyword{qed} works
   differently for a local proofs vs.\ the global ending of the main
   proof.
 
   Toplevel transitions are composed via transition transformers.
   Internally, Isar commands are put together from an empty transition
-  extended by name and source position (and optional source text).  It
-  is then left to the individual command parser to turn the given
-  concrete syntax into a suitable transition transformer that adjoins
-  actual operations on a theory or proof state etc.%
+  extended by name and source position.  It is then left to the
+  individual command parser to turn the given concrete syntax into a
+  suitable transition transformer that adjoins actual operations on a
+  theory or proof state etc.%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -240,7 +241,7 @@
   the toplevel itself, and only make sense in interactive mode.  Under
   normal circumstances, the user encounters these only implicitly as
   part of the protocol between the Isabelle/Isar system and a
-  user-interface such as ProofGeneral.
+  user-interface such as Proof~General.
 
   \begin{description}
 
@@ -391,9 +392,7 @@
   associated with each theory, by declaring these dependencies in the
   theory header as \isa{{\isasymUSES}}, and loading them consecutively
   within the theory context.  The system keeps track of incoming {\ML}
-  sources and associates them with the current theory.  The file
-  \isa{A}\verb,.ML, is loaded after a theory has been concluded, in
-  order to support legacy proof {\ML} proof scripts.
+  sources and associates them with the current theory.
 
   The basic internal actions of the theory database are \isa{update}, \isa{outdate}, and \isa{remove}:
 
@@ -458,12 +457,14 @@
 
   \item \verb|use_thy|~\isa{A} ensures that theory \isa{A} is fully
   up-to-date wrt.\ the external file store, reloading outdated
-  ancestors as required.
+  ancestors as required.  In batch mode, the simultaneous \verb|use_thys| should be used exclusively.
 
   \item \verb|use_thys| is similar to \verb|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, though.
+  result.  By loading a whole sub-graph of theories like that, the
+  intrinsic parallelism can be exploited by the system, to speedup
+  loading.
 
   \item \verb|ThyInfo.touch_thy|~\isa{A} performs and \isa{outdate} action
   on theory \isa{A} and all descendants.
@@ -472,7 +473,7 @@
   descendants from the theory database.
 
   \item \verb|ThyInfo.begin_theory| is the basic operation behind a
-  \isa{{\isasymTHEORY}} header declaration.  This is {\ML} functions is
+  \isa{{\isasymTHEORY}} header declaration.  This {\ML} function is
   normally not invoked directly.
 
   \item \verb|ThyInfo.end_theory| concludes the loading of a theory
--- a/doc-src/IsarImplementation/Thy/document/Local_Theory.tex	Mon Feb 08 14:06:58 2010 +0100
+++ b/doc-src/IsarImplementation/Thy/document/Local_Theory.tex	Mon Feb 08 14:08:32 2010 +0100
@@ -18,7 +18,7 @@
 %
 \endisadelimtheory
 %
-\isamarkupchapter{Local theory specifications%
+\isamarkupchapter{Local theory specifications \label{ch:local-theory}%
 }
 \isamarkuptrue%
 %
--- a/doc-src/IsarImplementation/Thy/document/Logic.tex	Mon Feb 08 14:06:58 2010 +0100
+++ b/doc-src/IsarImplementation/Thy/document/Logic.tex	Mon Feb 08 14:08:32 2010 +0100
@@ -39,7 +39,9 @@
   schematic polymorphism, which is strictly speaking outside the
   logic.\footnote{This is the deeper logical reason, why the theory
   context \isa{{\isasymTheta}} is separate from the proof context \isa{{\isasymGamma}}
-  of the core calculus.}%
+  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!}%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -57,16 +59,15 @@
   internally.  The resulting relation is an ordering: reflexive,
   transitive, and antisymmetric.
 
-  A \emph{sort} is a list of type classes written as \isa{s\ {\isacharequal}\ {\isacharbraceleft}c\isactrlisub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ c\isactrlisub m{\isacharbraceright}}, which represents symbolic
-  intersection.  Notationally, the curly braces are omitted for
-  singleton intersections, i.e.\ any class \isa{c} may be read as
-  a sort \isa{{\isacharbraceleft}c{\isacharbraceright}}.  The ordering on type classes is extended to
-  sorts according to the meaning of intersections: \isa{{\isacharbraceleft}c\isactrlisub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}\ c\isactrlisub m{\isacharbraceright}\ {\isasymsubseteq}\ {\isacharbraceleft}d\isactrlisub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ d\isactrlisub n{\isacharbraceright}} iff
-  \isa{{\isasymforall}j{\isachardot}\ {\isasymexists}i{\isachardot}\ c\isactrlisub i\ {\isasymsubseteq}\ d\isactrlisub j}.  The empty intersection
-  \isa{{\isacharbraceleft}{\isacharbraceright}} refers to the universal sort, which is the largest
-  element wrt.\ the sort order.  The intersections of all (finitely
-  many) classes declared in the current theory are the minimal
-  elements wrt.\ the sort order.
+  A \emph{sort} is a list of type classes written as \isa{s\ {\isacharequal}\ {\isacharbraceleft}c\isactrlisub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ c\isactrlisub m{\isacharbraceright}}, it represents symbolic intersection.  Notationally, the
+  curly braces are omitted for singleton intersections, i.e.\ any
+  class \isa{c} may be read as a sort \isa{{\isacharbraceleft}c{\isacharbraceright}}.  The ordering
+  on type classes is extended to sorts according to the meaning of
+  intersections: \isa{{\isacharbraceleft}c\isactrlisub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}\ c\isactrlisub m{\isacharbraceright}\ {\isasymsubseteq}\ {\isacharbraceleft}d\isactrlisub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ d\isactrlisub n{\isacharbraceright}} iff \isa{{\isasymforall}j{\isachardot}\ {\isasymexists}i{\isachardot}\ c\isactrlisub i\ {\isasymsubseteq}\ d\isactrlisub j}.  The empty intersection \isa{{\isacharbraceleft}{\isacharbraceright}} refers to
+  the universal sort, which is the largest element wrt.\ the sort
+  order.  Thus \isa{{\isacharbraceleft}{\isacharbraceright}} 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{fixed type variable} is a pair of a basic name
   (starting with a \isa{{\isacharprime}} character) and a sort constraint, e.g.\
@@ -76,10 +77,10 @@
   printed as \isa{{\isacharquery}{\isasymalpha}\isactrlisub s}.
 
   Note that \emph{all} syntactic components contribute to the identity
-  of type variables, including the sort constraint.  The core logic
-  handles type variables with the same name but different sorts as
-  different, although some outer layers of the system make it hard to
-  produce anything like this.
+  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{type constructor} \isa{{\isasymkappa}} is a \isa{k}-ary operator
   on types declared in the theory.  Type constructor application is
@@ -90,8 +91,8 @@
   Further notation is provided for specific constructors, notably the
   right-associative infix \isa{{\isasymalpha}\ {\isasymRightarrow}\ {\isasymbeta}} instead of \isa{{\isacharparenleft}{\isasymalpha}{\isacharcomma}\ {\isasymbeta}{\isacharparenright}fun}.
   
-  A \emph{type} is defined inductively over type variables and type
-  constructors as follows: \isa{{\isasymtau}\ {\isacharequal}\ {\isasymalpha}\isactrlisub s\ {\isacharbar}\ {\isacharquery}{\isasymalpha}\isactrlisub s\ {\isacharbar}\ {\isacharparenleft}{\isasymtau}\isactrlsub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ {\isasymtau}\isactrlsub k{\isacharparenright}{\isasymkappa}}.
+  The logical category \emph{type} is defined inductively over type
+  variables and type constructors as follows: \isa{{\isasymtau}\ {\isacharequal}\ {\isasymalpha}\isactrlisub s\ {\isacharbar}\ {\isacharquery}{\isasymalpha}\isactrlisub s\ {\isacharbar}\ {\isacharparenleft}{\isasymtau}\isactrlsub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ {\isasymtau}\isactrlsub k{\isacharparenright}{\isasymkappa}}.
 
   A \emph{type abbreviation} is a syntactic definition \isa{{\isacharparenleft}\isactrlvec {\isasymalpha}{\isacharparenright}{\isasymkappa}\ {\isacharequal}\ {\isasymtau}} of an arbitrary type expression \isa{{\isasymtau}} over
   variables \isa{\isactrlvec {\isasymalpha}}.  Type abbreviations appear as type
@@ -127,9 +128,9 @@
 %
 \begin{isamarkuptext}%
 \begin{mldecls}
-  \indexdef{}{ML type}{class}\verb|type class| \\
-  \indexdef{}{ML type}{sort}\verb|type sort| \\
-  \indexdef{}{ML type}{arity}\verb|type arity| \\
+  \indexdef{}{ML type}{class}\verb|type class = string| \\
+  \indexdef{}{ML type}{sort}\verb|type sort = class list| \\
+  \indexdef{}{ML type}{arity}\verb|type arity = string * sort list * sort| \\
   \indexdef{}{ML type}{typ}\verb|type typ| \\
   \indexdef{}{ML}{map\_atyps}\verb|map_atyps: (typ -> typ) -> typ -> typ| \\
   \indexdef{}{ML}{fold\_atyps}\verb|fold_atyps: (typ -> 'a -> 'a) -> typ -> 'a -> 'a| \\
@@ -147,14 +148,14 @@
 
   \begin{description}
 
-  \item \verb|class| represents type classes; this is an alias for
-  \verb|string|.
+  \item \verb|class| represents type classes.
 
-  \item \verb|sort| represents sorts; this is an alias for
-  \verb|class list|.
+  \item \verb|sort| represents sorts, i.e.\ finite intersections
+  of classes.  The empty list \verb|[]: sort| refers to the empty
+  class intersection, i.e.\ the ``full sort''.
 
-  \item \verb|arity| represents type arities; this is an alias for
-  triples of the form \isa{{\isacharparenleft}{\isasymkappa}{\isacharcomma}\ \isactrlvec s{\isacharcomma}\ s{\isacharparenright}} for \isa{{\isasymkappa}\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}\isactrlvec s{\isacharparenright}s} described above.
+  \item \verb|arity| represents type arities.  A triple \isa{{\isacharparenleft}{\isasymkappa}{\isacharcomma}\ \isactrlvec s{\isacharcomma}\ s{\isacharparenright}\ {\isacharcolon}\ arity} represents \isa{{\isasymkappa}\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}\isactrlvec s{\isacharparenright}s} as
+  described above.
 
   \item \verb|typ| represents types; this is a datatype with
   constructors \verb|TFree|, \verb|TVar|, \verb|Type|.
@@ -207,13 +208,13 @@
   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
-  are have an explicit name and type in each occurrence.
+  have an explicit name and type in each occurrence.
 
   \medskip A \emph{bound variable} is a natural number \isa{b},
   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 \isa{{\isasymlambda}\isactrlbsub nat\isactrlesub {\isachardot}\ {\isasymlambda}\isactrlbsub nat\isactrlesub {\isachardot}\ {\isadigit{1}}\ {\isacharplus}\ {\isadigit{0}}} would
-  correspond to \isa{{\isasymlambda}x\isactrlbsub nat\isactrlesub {\isachardot}\ {\isasymlambda}y\isactrlbsub nat\isactrlesub {\isachardot}\ x\ {\isacharplus}\ y} in a named
+  example, the de-Bruijn term \isa{{\isasymlambda}\isactrlbsub bool\isactrlesub {\isachardot}\ {\isasymlambda}\isactrlbsub bool\isactrlesub {\isachardot}\ {\isadigit{1}}\ {\isasymand}\ {\isadigit{0}}} would
+  correspond to \isa{{\isasymlambda}x\isactrlbsub bool\isactrlesub {\isachardot}\ {\isasymlambda}y\isactrlbsub bool\isactrlesub {\isachardot}\ x\ {\isasymand}\ y} 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.
@@ -225,19 +226,23 @@
   without any loose variables.
 
   A \emph{fixed variable} is a pair of a basic name and a type, e.g.\
-  \isa{{\isacharparenleft}x{\isacharcomma}\ {\isasymtau}{\isacharparenright}} which is usually printed \isa{x\isactrlisub {\isasymtau}}.  A
+  \isa{{\isacharparenleft}x{\isacharcomma}\ {\isasymtau}{\isacharparenright}} which is usually printed \isa{x\isactrlisub {\isasymtau}} here.  A
   \emph{schematic variable} is a pair of an indexname and a type,
-  e.g.\ \isa{{\isacharparenleft}{\isacharparenleft}x{\isacharcomma}\ {\isadigit{0}}{\isacharparenright}{\isacharcomma}\ {\isasymtau}{\isacharparenright}} which is usually printed as \isa{{\isacharquery}x\isactrlisub {\isasymtau}}.
+  e.g.\ \isa{{\isacharparenleft}{\isacharparenleft}x{\isacharcomma}\ {\isadigit{0}}{\isacharparenright}{\isacharcomma}\ {\isasymtau}{\isacharparenright}} which is likewise printed as \isa{{\isacharquery}x\isactrlisub {\isasymtau}}.
 
   \medskip A \emph{constant} is a pair of a basic name and a type,
-  e.g.\ \isa{{\isacharparenleft}c{\isacharcomma}\ {\isasymtau}{\isacharparenright}} which is usually printed as \isa{c\isactrlisub {\isasymtau}}.  Constants are declared in the context as polymorphic
-  families \isa{c\ {\isacharcolon}{\isacharcolon}\ {\isasymsigma}}, meaning that all substitution instances
-  \isa{c\isactrlisub {\isasymtau}} for \isa{{\isasymtau}\ {\isacharequal}\ {\isasymsigma}{\isasymvartheta}} are valid.
+  e.g.\ \isa{{\isacharparenleft}c{\isacharcomma}\ {\isasymtau}{\isacharparenright}} which is usually printed as \isa{c\isactrlisub {\isasymtau}}
+  here.  Constants are declared in the context as polymorphic families
+  \isa{c\ {\isacharcolon}{\isacharcolon}\ {\isasymsigma}}, meaning that all substitution instances \isa{c\isactrlisub {\isasymtau}} for \isa{{\isasymtau}\ {\isacharequal}\ {\isasymsigma}{\isasymvartheta}} are valid.
 
-  The vector of \emph{type arguments} of constant \isa{c\isactrlisub {\isasymtau}}
-  wrt.\ the declaration \isa{c\ {\isacharcolon}{\isacharcolon}\ {\isasymsigma}} is defined as the codomain of
-  the matcher \isa{{\isasymvartheta}\ {\isacharequal}\ {\isacharbraceleft}{\isacharquery}{\isasymalpha}\isactrlisub {\isadigit{1}}\ {\isasymmapsto}\ {\isasymtau}\isactrlisub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ {\isacharquery}{\isasymalpha}\isactrlisub n\ {\isasymmapsto}\ {\isasymtau}\isactrlisub n{\isacharbraceright}} presented in canonical order \isa{{\isacharparenleft}{\isasymtau}\isactrlisub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ {\isasymtau}\isactrlisub n{\isacharparenright}}.  Within a given theory context,
-  there is a one-to-one correspondence between any constant \isa{c\isactrlisub {\isasymtau}} and the application \isa{c{\isacharparenleft}{\isasymtau}\isactrlisub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ {\isasymtau}\isactrlisub n{\isacharparenright}} of its type arguments.  For example, with \isa{plus\ {\isacharcolon}{\isacharcolon}\ {\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}}, the instance \isa{plus\isactrlbsub nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ nat\isactrlesub } corresponds to \isa{plus{\isacharparenleft}nat{\isacharparenright}}.
+  The vector of \emph{type arguments} of constant \isa{c\isactrlisub {\isasymtau}} wrt.\
+  the declaration \isa{c\ {\isacharcolon}{\isacharcolon}\ {\isasymsigma}} is defined as the codomain of the
+  matcher \isa{{\isasymvartheta}\ {\isacharequal}\ {\isacharbraceleft}{\isacharquery}{\isasymalpha}\isactrlisub {\isadigit{1}}\ {\isasymmapsto}\ {\isasymtau}\isactrlisub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ {\isacharquery}{\isasymalpha}\isactrlisub n\ {\isasymmapsto}\ {\isasymtau}\isactrlisub n{\isacharbraceright}} presented in
+  canonical order \isa{{\isacharparenleft}{\isasymtau}\isactrlisub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ {\isasymtau}\isactrlisub n{\isacharparenright}}, corresponding to the
+  left-to-right occurrences of the \isa{{\isasymalpha}\isactrlisub i} in \isa{{\isasymsigma}}.
+  Within a given theory context, there is a one-to-one correspondence
+  between any constant \isa{c\isactrlisub {\isasymtau}} and the application \isa{c{\isacharparenleft}{\isasymtau}\isactrlisub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ {\isasymtau}\isactrlisub n{\isacharparenright}} of its type arguments.  For example, with \isa{plus\ {\isacharcolon}{\isacharcolon}\ {\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}}, the instance \isa{plus\isactrlbsub nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ nat\isactrlesub } corresponds to
+  \isa{plus{\isacharparenleft}nat{\isacharparenright}}.
 
   Constant declarations \isa{c\ {\isacharcolon}{\isacharcolon}\ {\isasymsigma}} may contain sort constraints
   for type variables in \isa{{\isasymsigma}}.  These are observed by
@@ -246,13 +251,12 @@
   polymorphic constants that the user-level type-checker would reject
   due to violation of type class restrictions.
 
-  \medskip An \emph{atomic} term is either a variable or constant.  A
-  \emph{term} is defined inductively over atomic terms, with
-  abstraction and application as follows: \isa{t\ {\isacharequal}\ b\ {\isacharbar}\ x\isactrlisub {\isasymtau}\ {\isacharbar}\ {\isacharquery}x\isactrlisub {\isasymtau}\ {\isacharbar}\ c\isactrlisub {\isasymtau}\ {\isacharbar}\ {\isasymlambda}\isactrlisub {\isasymtau}{\isachardot}\ t\ {\isacharbar}\ t\isactrlisub {\isadigit{1}}\ t\isactrlisub {\isadigit{2}}}.
-  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.
+  \medskip An \emph{atomic} term is either a variable or constant.
+  The logical category \emph{term} is defined inductively over atomic
+  terms, with abstraction and application as follows: \isa{t\ {\isacharequal}\ b\ {\isacharbar}\ x\isactrlisub {\isasymtau}\ {\isacharbar}\ {\isacharquery}x\isactrlisub {\isasymtau}\ {\isacharbar}\ c\isactrlisub {\isasymtau}\ {\isacharbar}\ {\isasymlambda}\isactrlisub {\isasymtau}{\isachardot}\ t\ {\isacharbar}\ t\isactrlisub {\isadigit{1}}\ t\isactrlisub {\isadigit{2}}}.  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 \isa{t\ {\isacharcolon}{\isacharcolon}\ {\isasymtau}} assigns a (unique) type to a
   term according to the structure of atomic terms, abstractions, and
@@ -273,16 +277,17 @@
   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 \isa{x\isactrlbsub {\isasymtau}\isactrlisub {\isadigit{1}}\isactrlesub } and \isa{x\isactrlbsub {\isasymtau}\isactrlisub {\isadigit{2}}\isactrlesub } may become the same after type
-  instantiation.  Some outer layers of the system make it hard to
-  produce variables of the same name, but different types.  In
-  contrast, mixed instances of polymorphic constants occur frequently.
+  component.  This means that different variables \isa{x\isactrlbsub {\isasymtau}\isactrlisub {\isadigit{1}}\isactrlesub } and \isa{x\isactrlbsub {\isasymtau}\isactrlisub {\isadigit{2}}\isactrlesub } 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{hidden polymorphism} of a term \isa{t\ {\isacharcolon}{\isacharcolon}\ {\isasymsigma}}
   is the set of type variables occurring in \isa{t}, but not in
-  \isa{{\isasymsigma}}.  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 \isa{t{\isasymvartheta}\ {\isacharcolon}{\isacharcolon}\ {\isasymsigma}} and \isa{t{\isasymvartheta}{\isacharprime}\ {\isacharcolon}{\isacharcolon}\ {\isasymsigma}} with the same type.  This slightly
+  its type \isa{{\isasymsigma}}.  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 \isa{t{\isasymvartheta}\ {\isacharcolon}{\isacharcolon}\ {\isasymsigma}} and
+  \isa{t{\isasymvartheta}{\isacharprime}\ {\isacharcolon}{\isacharcolon}\ {\isasymsigma}} with the same type.  This slightly
   pathological situation notoriously demands additional care.
 
   \medskip A \emph{term abbreviation} is a syntactic definition \isa{c\isactrlisub {\isasymsigma}\ {\isasymequiv}\ t} of a closed term \isa{t} of type \isa{{\isasymsigma}},
@@ -435,14 +440,14 @@
   \infer[\isa{{\isacharparenleft}assume{\isacharparenright}}]{\isa{A\ {\isasymturnstile}\ A}}{}
   \]
   \[
-  \infer[\isa{{\isacharparenleft}{\isasymAnd}{\isacharunderscore}intro{\isacharparenright}}]{\isa{{\isasymGamma}\ {\isasymturnstile}\ {\isasymAnd}x{\isachardot}\ b{\isacharbrackleft}x{\isacharbrackright}}}{\isa{{\isasymGamma}\ {\isasymturnstile}\ b{\isacharbrackleft}x{\isacharbrackright}} & \isa{x\ {\isasymnotin}\ {\isasymGamma}}}
+  \infer[\isa{{\isacharparenleft}{\isasymAnd}{\isasymdash}intro{\isacharparenright}}]{\isa{{\isasymGamma}\ {\isasymturnstile}\ {\isasymAnd}x{\isachardot}\ b{\isacharbrackleft}x{\isacharbrackright}}}{\isa{{\isasymGamma}\ {\isasymturnstile}\ b{\isacharbrackleft}x{\isacharbrackright}} & \isa{x\ {\isasymnotin}\ {\isasymGamma}}}
   \qquad
-  \infer[\isa{{\isacharparenleft}{\isasymAnd}{\isacharunderscore}elim{\isacharparenright}}]{\isa{{\isasymGamma}\ {\isasymturnstile}\ b{\isacharbrackleft}a{\isacharbrackright}}}{\isa{{\isasymGamma}\ {\isasymturnstile}\ {\isasymAnd}x{\isachardot}\ b{\isacharbrackleft}x{\isacharbrackright}}}
+  \infer[\isa{{\isacharparenleft}{\isasymAnd}{\isasymdash}elim{\isacharparenright}}]{\isa{{\isasymGamma}\ {\isasymturnstile}\ b{\isacharbrackleft}a{\isacharbrackright}}}{\isa{{\isasymGamma}\ {\isasymturnstile}\ {\isasymAnd}x{\isachardot}\ b{\isacharbrackleft}x{\isacharbrackright}}}
   \]
   \[
-  \infer[\isa{{\isacharparenleft}{\isasymLongrightarrow}{\isacharunderscore}intro{\isacharparenright}}]{\isa{{\isasymGamma}\ {\isacharminus}\ A\ {\isasymturnstile}\ A\ {\isasymLongrightarrow}\ B}}{\isa{{\isasymGamma}\ {\isasymturnstile}\ B}}
+  \infer[\isa{{\isacharparenleft}{\isasymLongrightarrow}{\isasymdash}intro{\isacharparenright}}]{\isa{{\isasymGamma}\ {\isacharminus}\ A\ {\isasymturnstile}\ A\ {\isasymLongrightarrow}\ B}}{\isa{{\isasymGamma}\ {\isasymturnstile}\ B}}
   \qquad
-  \infer[\isa{{\isacharparenleft}{\isasymLongrightarrow}{\isacharunderscore}elim{\isacharparenright}}]{\isa{{\isasymGamma}\isactrlsub {\isadigit{1}}\ {\isasymunion}\ {\isasymGamma}\isactrlsub {\isadigit{2}}\ {\isasymturnstile}\ B}}{\isa{{\isasymGamma}\isactrlsub {\isadigit{1}}\ {\isasymturnstile}\ A\ {\isasymLongrightarrow}\ B} & \isa{{\isasymGamma}\isactrlsub {\isadigit{2}}\ {\isasymturnstile}\ A}}
+  \infer[\isa{{\isacharparenleft}{\isasymLongrightarrow}{\isasymdash}elim{\isacharparenright}}]{\isa{{\isasymGamma}\isactrlsub {\isadigit{1}}\ {\isasymunion}\ {\isasymGamma}\isactrlsub {\isadigit{2}}\ {\isasymturnstile}\ B}}{\isa{{\isasymGamma}\isactrlsub {\isadigit{1}}\ {\isasymturnstile}\ A\ {\isasymLongrightarrow}\ B} & \isa{{\isasymGamma}\isactrlsub {\isadigit{2}}\ {\isasymturnstile}\ A}}
   \]
   \caption{Primitive inferences of Pure}\label{fig:prim-rules}
   \end{center}
@@ -469,17 +474,18 @@
   terms, and \isa{{\isasymAnd}{\isacharslash}{\isasymLongrightarrow}} for proofs (cf.\
   \cite{Berghofer-Nipkow:2000:TPHOL}).
 
-  Observe that locally fixed parameters (as in \isa{{\isasymAnd}{\isacharunderscore}intro}) need
-  not be recorded in the hypotheses, because the simple syntactic
-  types of Pure are always inhabitable.  ``Assumptions'' \isa{x\ {\isacharcolon}{\isacharcolon}\ {\isasymtau}} for type-membership are only present as long as some \isa{x\isactrlisub {\isasymtau}} occurs in the statement body.\footnote{This is the key
-  difference to ``\isa{{\isasymlambda}HOL}'' in the PTS framework
-  \cite{Barendregt-Geuvers:2001}, where hypotheses \isa{x\ {\isacharcolon}\ A} are
-  treated uniformly for propositions and types.}
+  Observe that locally fixed parameters (as in \isa{{\isasymAnd}{\isasymdash}intro}) need not be recorded in the hypotheses, because
+  the simple syntactic types of Pure are always inhabitable.
+  ``Assumptions'' \isa{x\ {\isacharcolon}{\isacharcolon}\ {\isasymtau}} for type-membership are only
+  present as long as some \isa{x\isactrlisub {\isasymtau}} occurs in the statement
+  body.\footnote{This is the key difference to ``\isa{{\isasymlambda}HOL}'' in
+  the PTS framework \cite{Barendregt-Geuvers:2001}, where hypotheses
+  \isa{x\ {\isacharcolon}\ A} are treated uniformly for propositions and types.}
 
   \medskip The axiomatization of a theory is implicitly closed by
   forming all instances of type and term variables: \isa{{\isasymturnstile}\ A{\isasymvartheta}} holds for any substitution instance of an axiom
   \isa{{\isasymturnstile}\ A}.  By pushing substitutions through derivations
-  inductively, we also get admissible \isa{generalize} and \isa{instance} rules as shown in \figref{fig:subst-rules}.
+  inductively, we also get admissible \isa{generalize} and \isa{instantiate} rules as shown in \figref{fig:subst-rules}.
 
   \begin{figure}[htb]
   \begin{center}
@@ -586,11 +592,11 @@
   Every \verb|thm| value contains a sliding back-reference to the
   enclosing theory, cf.\ \secref{sec:context-theory}.
 
-  \item \verb|proofs| determines the detail of proof recording within
+  \item \verb|proofs| specifies the detail of proof recording within
   \verb|thm| values: \verb|0| records only the names of oracles,
   \verb|1| records oracle names and propositions, \verb|2| additionally
   records full proof terms.  Officially named theorems that contribute
-  to a result are always recorded.
+  to a result are recorded in any case.
 
   \item \verb|Thm.assume|, \verb|Thm.forall_intr|, \verb|Thm.forall_elim|, \verb|Thm.implies_intr|, and \verb|Thm.implies_elim|
   correspond to the primitive inferences of \figref{fig:prim-rules}.
@@ -646,8 +652,8 @@
   \begin{figure}[htb]
   \begin{center}
   \begin{tabular}{ll}
-  \isa{conjunction\ {\isacharcolon}{\isacharcolon}\ prop\ {\isasymRightarrow}\ prop\ {\isasymRightarrow}\ prop} & (infix \isa{{\isacharampersand}}) \\
-  \isa{{\isasymturnstile}\ A\ {\isacharampersand}\ B\ {\isasymequiv}\ {\isacharparenleft}{\isasymAnd}C{\isachardot}\ {\isacharparenleft}A\ {\isasymLongrightarrow}\ B\ {\isasymLongrightarrow}\ C{\isacharparenright}\ {\isasymLongrightarrow}\ C{\isacharparenright}} \\[1ex]
+  \isa{conjunction\ {\isacharcolon}{\isacharcolon}\ prop\ {\isasymRightarrow}\ prop\ {\isasymRightarrow}\ prop} & (infix \isa{{\isacharampersand}{\isacharampersand}{\isacharampersand}}) \\
+  \isa{{\isasymturnstile}\ A\ {\isacharampersand}{\isacharampersand}{\isacharampersand}\ B\ {\isasymequiv}\ {\isacharparenleft}{\isasymAnd}C{\isachardot}\ {\isacharparenleft}A\ {\isasymLongrightarrow}\ B\ {\isasymLongrightarrow}\ C{\isacharparenright}\ {\isasymLongrightarrow}\ C{\isacharparenright}} \\[1ex]
   \isa{prop\ {\isacharcolon}{\isacharcolon}\ prop\ {\isasymRightarrow}\ prop} & (prefix \isa{{\isacharhash}}, suppressed) \\
   \isa{{\isacharhash}A\ {\isasymequiv}\ A} \\[1ex]
   \isa{term\ {\isacharcolon}{\isacharcolon}\ {\isasymalpha}\ {\isasymRightarrow}\ prop} & (prefix \isa{TERM}) \\
@@ -659,12 +665,15 @@
   \end{center}
   \end{figure}
 
-  Derived conjunction rules include introduction \isa{A\ {\isasymLongrightarrow}\ B\ {\isasymLongrightarrow}\ A\ {\isacharampersand}\ B}, and destructions \isa{A\ {\isacharampersand}\ B\ {\isasymLongrightarrow}\ A} and \isa{A\ {\isacharampersand}\ B\ {\isasymLongrightarrow}\ B}.
-  Conjunction allows to treat simultaneous assumptions and conclusions
-  uniformly.  For example, multiple claims are intermediately
-  represented as explicit conjunction, but this is refined into
-  separate sub-goals before the user continues the proof; the final
-  result is projected into a list of theorems (cf.\
+  The introduction \isa{A\ {\isasymLongrightarrow}\ B\ {\isasymLongrightarrow}\ A\ {\isacharampersand}{\isacharampersand}{\isacharampersand}\ B}, and eliminations
+  (projections) \isa{A\ {\isacharampersand}{\isacharampersand}{\isacharampersand}\ B\ {\isasymLongrightarrow}\ A} and \isa{A\ {\isacharampersand}{\isacharampersand}{\isacharampersand}\ B\ {\isasymLongrightarrow}\ B} are
+  available as derived rules.  Conjunction allows to treat
+  simultaneous assumptions and conclusions uniformly, e.g.\ consider
+  \isa{A\ {\isasymLongrightarrow}\ B\ {\isasymLongrightarrow}\ C\ {\isacharampersand}{\isacharampersand}{\isacharampersand}\ D}.  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 \isa{prop} marker (\isa{{\isacharhash}}) makes arbitrarily complex
@@ -680,7 +689,7 @@
   the unspecified type \isa{{\isasymalpha}\ itself}; it essentially injects the
   language of types into that of terms.  There is specific notation
   \isa{TYPE{\isacharparenleft}{\isasymtau}{\isacharparenright}} for \isa{TYPE\isactrlbsub {\isasymtau}\ itself\isactrlesub }.
-  Although being devoid of any particular meaning, the \isa{TYPE{\isacharparenleft}{\isasymtau}{\isacharparenright}} accounts for the type \isa{{\isasymtau}} within the term
+  Although being devoid of any particular meaning, the term \isa{TYPE{\isacharparenleft}{\isasymtau}{\isacharparenright}} accounts for the type \isa{{\isasymtau}} within the term
   language.  In particular, \isa{TYPE{\isacharparenleft}{\isasymalpha}{\isacharparenright}} may be used as formal
   argument in primitive definitions, in order to circumvent hidden
   polymorphism (cf.\ \secref{sec:terms}).  For example, \isa{c\ TYPE{\isacharparenleft}{\isasymalpha}{\isacharparenright}\ {\isasymequiv}\ A{\isacharbrackleft}{\isasymalpha}{\isacharbrackright}} defines \isa{c\ {\isacharcolon}{\isacharcolon}\ {\isasymalpha}\ itself\ {\isasymRightarrow}\ prop} in terms of
@@ -707,10 +716,10 @@
 
   \begin{description}
 
-  \item \verb|Conjunction.intr| derives \isa{A\ {\isacharampersand}\ B} from \isa{A} and \isa{B}.
+  \item \verb|Conjunction.intr| derives \isa{A\ {\isacharampersand}{\isacharampersand}{\isacharampersand}\ B} from \isa{A} and \isa{B}.
 
   \item \verb|Conjunction.elim| derives \isa{A} and \isa{B}
-  from \isa{A\ {\isacharampersand}\ B}.
+  from \isa{A\ {\isacharampersand}{\isacharampersand}{\isacharampersand}\ B}.
 
   \item \verb|Drule.mk_term| derives \isa{TERM\ t}.
 
@@ -792,7 +801,8 @@
   prefix of parameters and compound premises, concluding an atomic
   proposition.  Typical examples are \isa{{\isasymlongrightarrow}}-introduction \isa{{\isacharparenleft}A\ {\isasymLongrightarrow}\ B{\isacharparenright}\ {\isasymLongrightarrow}\ A\ {\isasymlongrightarrow}\ B} or mathematical induction \isa{P\ {\isadigit{0}}\ {\isasymLongrightarrow}\ {\isacharparenleft}{\isasymAnd}n{\isachardot}\ P\ n\ {\isasymLongrightarrow}\ P\ {\isacharparenleft}Suc\ n{\isacharparenright}{\isacharparenright}\ {\isasymLongrightarrow}\ P\ n}.  Even deeper nesting occurs in well-founded
   induction \isa{{\isacharparenleft}{\isasymAnd}x{\isachardot}\ {\isacharparenleft}{\isasymAnd}y{\isachardot}\ y\ {\isasymprec}\ x\ {\isasymLongrightarrow}\ P\ y{\isacharparenright}\ {\isasymLongrightarrow}\ P\ x{\isacharparenright}\ {\isasymLongrightarrow}\ P\ x}, but this
-  already marks the limit of rule complexity seen in practice.
+  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:
@@ -917,9 +927,9 @@
 
   \begin{description}
 
-  \item \isa{rule\isactrlsub {\isadigit{1}}\ RS\ rule\isactrlsub {\isadigit{2}}} resolves \isa{rule\isactrlsub {\isadigit{1}}} with \isa{rule\isactrlsub {\isadigit{2}}} according to the
-  \hyperlink{inference.resolution}{\mbox{\isa{resolution}}} principle explained above.  Note that the
-  corresponding attribute in the Isar language is called \hyperlink{attribute.THEN}{\mbox{\isa{THEN}}}.
+  \item \isa{rule\isactrlsub {\isadigit{1}}\ RS\ rule\isactrlsub {\isadigit{2}}} resolves \isa{rule\isactrlsub {\isadigit{1}}} with \isa{rule\isactrlsub {\isadigit{2}}} according to the \hyperlink{inference.resolution}{\mbox{\isa{resolution}}} principle
+  explained above.  Note that the corresponding rule attribute in the
+  Isar language is called \hyperlink{attribute.THEN}{\mbox{\isa{THEN}}}.
 
   \item \isa{rule\ OF\ rules} resolves a list of rules with the
   first rule, addressing its premises \isa{{\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ length\ rules}
--- a/doc-src/IsarImplementation/Thy/document/Prelim.tex	Mon Feb 08 14:06:58 2010 +0100
+++ b/doc-src/IsarImplementation/Thy/document/Prelim.tex	Mon Feb 08 14:08:32 2010 +0100
@@ -93,39 +93,46 @@
 \isamarkuptrue%
 %
 \begin{isamarkuptext}%
-A \emph{theory} is a data container with explicit name and unique
-  identifier.  Theories are related by a (nominal) sub-theory
+A \emph{theory} 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.
-
-  The \isa{merge} operation produces the least upper bound of two
-  theories, which actually degenerates into absorption of one theory
-  into the other (due to the nominal sub-theory relation).
+  ancestor theories.  To this end, the system maintains a set of
+  symbolic ``identification stamps'' within each theory.
 
-  The \isa{begin} operation starts a new theory by importing
-  several parent theories and entering a special \isa{draft} mode,
-  which is sustained until the final \isa{end} operation.  A draft
-  theory acts like a linear type, where updates invalidate earlier
-  versions.  An invalidated draft is called ``stale''.
+  In order to avoid the full-scale overhead of explicit sub-theory
+  identification of arbitrary intermediate stages, a theory is
+  switched into \isa{draft} mode under certain circumstances.  A
+  draft theory acts like a linear type, where updates invalidate
+  earlier versions.  An invalidated draft is called \emph{stale}.
 
-  The \isa{checkpoint} operation produces an intermediate stepping
-  stone that will survive the next update: both the original and the
-  changed theory remain valid and are related by the sub-theory
-  relation.  Checkpointing essentially recovers purely functional
-  theory values, at the expense of some extra internal bookkeeping.
+  The \isa{checkpoint} operation produces a safe stepping stone
+  that will survive the next update without becoming stale: both the
+  old and the new theory remain valid and are related by the
+  sub-theory relation.  Checkpointing essentially recovers purely
+  functional theory values, at the expense of some extra internal
+  bookkeeping.
 
   The \isa{copy} operation produces an auxiliary version that has
   the same data content, but is unrelated to the original: updates of
   the copy do not affect the original, neither does the sub-theory
   relation hold.
 
+  The \isa{merge} operation produces the least upper bound of two
+  theories, which actually degenerates into absorption of one theory
+  into the other (according to the nominal sub-theory relation).
+
+  The \isa{begin} operation starts a new theory by importing
+  several parent theories and entering a special mode of nameless
+  incremental updates, until the final \isa{end} operation is
+  performed.
+
   \medskip The example in \figref{fig:ex-theory} below shows a theory
   graph derived from \isa{Pure}, with theory \isa{Length}
   importing \isa{Nat} and \isa{List}.  The body of \isa{Length} consists of a sequence of updates, working mostly on
-  drafts.  Intermediate checkpoints may occur as well, due to the
-  history mechanism provided by the Isar top-level, cf.\
-  \secref{sec:isar-toplevel}.
+  drafts internally, while transaction boundaries of Isar top-level
+  commands (\secref{sec:isar-toplevel}) are guaranteed to be safe
+  checkpoints.
 
   \begin{figure}[htb]
   \begin{center}
@@ -172,9 +179,10 @@
 \begin{mldecls}
   \indexdef{}{ML type}{theory}\verb|type theory| \\
   \indexdef{}{ML}{Theory.subthy}\verb|Theory.subthy: theory * theory -> bool| \\
-  \indexdef{}{ML}{Theory.merge}\verb|Theory.merge: theory * theory -> theory| \\
   \indexdef{}{ML}{Theory.checkpoint}\verb|Theory.checkpoint: theory -> theory| \\
   \indexdef{}{ML}{Theory.copy}\verb|Theory.copy: theory -> theory| \\
+  \indexdef{}{ML}{Theory.merge}\verb|Theory.merge: theory * theory -> theory| \\
+  \indexdef{}{ML}{Theory.begin\_theory}\verb|Theory.begin_theory: string -> theory list -> theory| \\
   \end{mldecls}
   \begin{mldecls}
   \indexdef{}{ML type}{theory\_ref}\verb|type theory_ref| \\
@@ -185,24 +193,31 @@
   \begin{description}
 
   \item \verb|theory| represents theory contexts.  This is
-  essentially a linear type!  Most operations destroy the original
-  version, which then becomes ``stale''.
+  essentially a linear type, with explicit runtime checking!  Most
+  internal theory operations destroy the original version, which then
+  becomes ``stale''.
 
-  \item \verb|Theory.subthy|~\isa{{\isacharparenleft}thy\isactrlsub {\isadigit{1}}{\isacharcomma}\ thy\isactrlsub {\isadigit{2}}{\isacharparenright}}
-  compares theories according to the inherent graph structure of the
-  construction.  This sub-theory relation is a nominal approximation
-  of inclusion (\isa{{\isasymsubseteq}}) of the corresponding content.
-
-  \item \verb|Theory.merge|~\isa{{\isacharparenleft}thy\isactrlsub {\isadigit{1}}{\isacharcomma}\ thy\isactrlsub {\isadigit{2}}{\isacharparenright}}
-  absorbs one theory into the other.  This fails for unrelated
-  theories!
+  \item \verb|Theory.subthy|~\isa{{\isacharparenleft}thy\isactrlsub {\isadigit{1}}{\isacharcomma}\ thy\isactrlsub {\isadigit{2}}{\isacharparenright}} compares theories
+  according to the intrinsic graph structure of the construction.
+  This sub-theory relation is a nominal approximation of inclusion
+  (\isa{{\isasymsubseteq}}) of the corresponding content (according to the
+  semantics of the ML modules that implement the data).
 
   \item \verb|Theory.checkpoint|~\isa{thy} produces a safe
-  stepping stone in the linear development of \isa{thy}.  The next
-  update will result in two related, valid theories.
+  stepping stone in the linear development of \isa{thy}.  This
+  changes the old theory, but the next update will result in two
+  related, valid theories.
+
+  \item \verb|Theory.copy|~\isa{thy} produces a variant of \isa{thy} with the same data.  The copy is not related to the original,
+  but the original is unchanged.
 
-  \item \verb|Theory.copy|~\isa{thy} produces a variant of \isa{thy} with the same data.  The result is not related to the
-  original; the original is unchanged.
+  \item \verb|Theory.merge|~\isa{{\isacharparenleft}thy\isactrlsub {\isadigit{1}}{\isacharcomma}\ thy\isactrlsub {\isadigit{2}}{\isacharparenright}} absorbs one theory
+  into the other, without changing \isa{thy\isactrlsub {\isadigit{1}}} or \isa{thy\isactrlsub {\isadigit{2}}}.
+  This version of ad-hoc theory merge fails for unrelated theories!
+
+  \item \verb|Theory.begin_theory|~\isa{name\ parents} constructs
+  a new theory based on the given parents.  This {\ML} function is
+  normally not invoked directly.
 
   \item \verb|theory_ref| represents a sliding reference to an
   always valid theory; updates on the original are propagated
@@ -229,30 +244,32 @@
 \isamarkuptrue%
 %
 \begin{isamarkuptext}%
-A proof context is a container for pure data with a back-reference
-  to the theory it belongs to.  The \isa{init} operation creates a
-  proof context from a given theory.  Modifications to draft theories
-  are propagated to the proof context as usual, but there is also an
-  explicit \isa{transfer} operation to force resynchronization
-  with more substantial updates to the underlying theory.  The actual
-  context data does not require any special bookkeeping, thanks to the
-  lack of destructive features.
+A proof context is a container for pure data with a
+  back-reference to the theory it belongs to.  The \isa{init}
+  operation creates a proof context from a given theory.
+  Modifications to draft theories are propagated to the proof context
+  as usual, but there is also an explicit \isa{transfer} operation
+  to force resynchronization with more substantial updates to the
+  underlying theory.
 
-  Entities derived in a proof context need to record inherent logical
+  Entities derived in a proof context need to record logical
   requirements explicitly, since there is no separate context
-  identification as for theories.  For example, hypotheses used in
-  primitive derivations (cf.\ \secref{sec:thms}) are recorded
-  separately within the sequent \isa{{\isasymGamma}\ {\isasymturnstile}\ {\isasymphi}}, 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.
+  identification or symbolic inclusion as for theories.  For example,
+  hypotheses used in primitive derivations (cf.\ \secref{sec:thms})
+  are recorded separately within the sequent \isa{{\isasymGamma}\ {\isasymturnstile}\ {\isasymphi}}, 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 the Isar proof states model
+  into the original context.  Note that an Isar proof state models
   block-structured reasoning explicitly, using a stack of proof
-  contexts internally.%
+  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!%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -311,7 +328,8 @@
 
   Moreover, there are total operations \isa{theory{\isacharunderscore}of} and \isa{proof{\isacharunderscore}of} 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 \isa{init} operation.%
+  have to be constructed by an ad-hoc \isa{init} operation, which
+  incurs a small runtime overhead.%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -376,6 +394,15 @@
   \noindent The \isa{empty} value acts as initial default for
   \emph{any} theory that does not declare actual data content; \isa{extend} is acts like a unitary version of \isa{merge}.
 
+  Implementing \isa{merge} can be tricky.  The general idea is
+  that \isa{merge\ {\isacharparenleft}data\isactrlsub {\isadigit{1}}{\isacharcomma}\ data\isactrlsub {\isadigit{2}}{\isacharparenright}} inserts those parts of \isa{data\isactrlsub {\isadigit{2}}} into \isa{data\isactrlsub {\isadigit{1}}} that are not yet present, while
+  keeping the general order of things.  The \verb|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!
+
   \paragraph{Proof context data} declarations need to implement the
   following SML signature:
 
@@ -387,15 +414,18 @@
   \medskip
 
   \noindent The \isa{init} operation is supposed to produce a pure
-  value from the given background theory.
+  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 \isa{init}
+  operation of \emph{all} theory data slots ever declared.
 
   \paragraph{Generic data} provides a hybrid interface for both theory
   and proof data.  The \isa{init} operation for proof contexts is
   predefined to select the current data value from the background
   theory.
 
-  \bigskip A data declaration of type \isa{T} results in the
-  following interface:
+  \bigskip Any of these data declaration over type \isa{T} result
+  in an ML structure with the following signature:
 
   \medskip
   \begin{tabular}{ll}
@@ -405,12 +435,12 @@
   \end{tabular}
   \medskip
 
-  \noindent These other operations provide access for the particular
-  kind of context (theory, proof, or generic context).  Note that this
-  is a safe interface: there is no other way to access the
-  corresponding data slot of a context.  By keeping these operations
-  private, a component may maintain abstract values authentically,
-  without other components interfering.%
+  \noindent These other operations provide exclusive access for the
+  particular kind of context (theory, proof, or generic context).
+  This interface fully 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.%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -451,16 +481,157 @@
 %
 \endisadelimmlref
 %
+\isadelimmlex
+%
+\endisadelimmlex
+%
+\isatagmlex
+%
+\begin{isamarkuptext}%
+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:%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\endisatagmlex
+{\isafoldmlex}%
+%
+\isadelimmlex
+%
+\endisadelimmlex
+%
+\isadelimML
+%
+\endisadelimML
+%
+\isatagML
+\isacommand{ML}\isamarkupfalse%
+\ {\isacharverbatimopen}\isanewline
+\ \ signature\ WELLFORMED{\isacharunderscore}TERMS\ {\isacharequal}\isanewline
+\ \ sig\isanewline
+\ \ \ \ val\ get{\isacharcolon}\ theory\ {\isacharminus}{\isachargreater}\ term\ list\isanewline
+\ \ \ \ val\ add{\isacharcolon}\ term\ {\isacharminus}{\isachargreater}\ theory\ {\isacharminus}{\isachargreater}\ theory\isanewline
+\ \ end{\isacharsemicolon}\isanewline
+{\isacharverbatimclose}%
+\endisatagML
+{\isafoldML}%
+%
+\isadelimML
+%
+\endisadelimML
+%
+\begin{isamarkuptext}%
+\noindent The implementation uses private theory data
+  internally, and only exposes an operation that involves explicit
+  argument checking wrt.\ the given theory.%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\isadelimML
+%
+\endisadelimML
+%
+\isatagML
+\isacommand{ML}\isamarkupfalse%
+\ {\isacharverbatimopen}\isanewline
+\ \ structure\ Wellformed{\isacharunderscore}Terms{\isacharcolon}\ WELLFORMED{\isacharunderscore}TERMS\ {\isacharequal}\isanewline
+\ \ struct\isanewline
+\isanewline
+\ \ structure\ Terms\ {\isacharequal}\ Theory{\isacharunderscore}Data\isanewline
+\ \ {\isacharparenleft}\isanewline
+\ \ \ \ type\ T\ {\isacharequal}\ term\ OrdList{\isachardot}T{\isacharsemicolon}\isanewline
+\ \ \ \ val\ empty\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}{\isacharsemicolon}\isanewline
+\ \ \ \ val\ extend\ {\isacharequal}\ I{\isacharsemicolon}\isanewline
+\ \ \ \ fun\ merge\ {\isacharparenleft}ts{\isadigit{1}}{\isacharcomma}\ ts{\isadigit{2}}{\isacharparenright}\ {\isacharequal}\isanewline
+\ \ \ \ \ \ OrdList{\isachardot}union\ TermOrd{\isachardot}fast{\isacharunderscore}term{\isacharunderscore}ord\ ts{\isadigit{1}}\ ts{\isadigit{2}}{\isacharsemicolon}\isanewline
+\ \ {\isacharparenright}\isanewline
+\isanewline
+\ \ val\ get\ {\isacharequal}\ Terms{\isachardot}get{\isacharsemicolon}\isanewline
+\isanewline
+\ \ fun\ add\ raw{\isacharunderscore}t\ thy\ {\isacharequal}\isanewline
+\ \ \ \ let\ val\ t\ {\isacharequal}\ Sign{\isachardot}cert{\isacharunderscore}term\ thy\ raw{\isacharunderscore}t\isanewline
+\ \ \ \ in\ Terms{\isachardot}map\ {\isacharparenleft}OrdList{\isachardot}insert\ TermOrd{\isachardot}fast{\isacharunderscore}term{\isacharunderscore}ord\ t{\isacharparenright}\ thy\ end{\isacharsemicolon}\isanewline
+\isanewline
+\ \ end{\isacharsemicolon}\isanewline
+{\isacharverbatimclose}%
+\endisatagML
+{\isafoldML}%
+%
+\isadelimML
+%
+\endisadelimML
+%
+\begin{isamarkuptext}%
+We use \verb|term OrdList.T| for reasonably efficient
+  representation of a set of terms: all operations are linear in the
+  number of stored elements.  Here we assume that our users do not
+  care about the declaration order, since that data structure forces
+  its own arrangement of elements.
+
+  Observe how the \verb|merge| operation joins the data slots of
+  the two constituents: \verb|OrdList.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:
+  \begin{enumerate}
+
+  \item \verb|Wellformed_Terms.add| only admits terms that have passed
+  the \verb|Sign.cert_term| check of the given theory at that point.
+
+  \item Wellformedness in the sense of \verb|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.
+
+  \end{enumerate}
+
+  Note that all basic operations of the inference kernel (which
+  includes \verb|Sign.cert_term|) observe this monotonicity principle,
+  but other user-space tools don't.  For example, fully-featured
+  type-inference via \verb|Syntax.check_term| (cf.\
+  \secref{sec:term-check}) is not necessarily monotonic wrt.\ the
+  background theory, since constraints of term constants can be
+  strengthened 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.%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
 \isamarkupsection{Names \label{sec:names}%
 }
 \isamarkuptrue%
 %
 \begin{isamarkuptext}%
 In principle, a name is just a string, but there are various
-  convention for encoding additional structure.  For example, ``\isa{Foo{\isachardot}bar{\isachardot}baz}'' is considered as a qualified name consisting of
-  three basic name components.  The individual constituents of a name
-  may have further substructure, e.g.\ the string
-  ``\verb,\,\verb,<alpha>,'' encodes as a single symbol.%
+  conventions for representing additional structure.  For example,
+  ``\isa{Foo{\isachardot}bar{\isachardot}baz}'' is considered as a long name consisting of
+  qualifier \isa{Foo{\isachardot}bar} and base name \isa{baz}.  The
+  individual constituents of a name may have further substructure,
+  e.g.\ the string ``\verb,\,\verb,<alpha>,'' encodes as a single
+  symbol.
+
+  \medskip Subsequently, we shall introduce specific categories of
+  names.  Roughly speaking these correspond to logical entities as
+  follows:
+  \begin{itemize}
+
+  \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}).
+
+  \end{itemize}%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -469,16 +640,16 @@
 \isamarkuptrue%
 %
 \begin{isamarkuptext}%
-A \emph{symbol} constitutes the smallest textual unit in Isabelle
-  --- raw characters are normally not encountered at all.  Isabelle
-  strings consist of a sequence of symbols, represented as a packed
-  string or a list of strings.  Each symbol is in itself a small
-  string, which has either one of the following forms:
+A \emph{symbol} constitutes the smallest textual unit in
+  Isabelle --- 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:
 
   \begin{enumerate}
 
-  \item a single ASCII character ``\isa{c}'', for example
-  ``\verb,a,'',
+  \item a single ASCII character ``\isa{c}'' or raw byte in the
+  range of 128\dots 255, for example ``\verb,a,'',
 
   \item a regular symbol ``\verb,\,\verb,<,\isa{ident}\verb,>,'',
   for example ``\verb,\,\verb,<alpha>,'',
@@ -487,7 +658,7 @@
   for example ``\verb,\,\verb,<^bold>,'',
 
   \item a raw symbol ``\verb,\,\verb,<^raw:,\isa{text}\verb,>,''
-  where \isa{text} constists of printable characters excluding
+  where \isa{text} consists of printable characters excluding
   ``\verb,.,'' and ``\verb,>,'', for example
   ``\verb,\,\verb,<^raw:$\sum_{i = 1}^n$>,'',
 
@@ -503,16 +674,25 @@
   may occur within regular Isabelle identifiers.
 
   Since the character set underlying Isabelle symbols is 7-bit ASCII
-  and 8-bit characters are passed through transparently, Isabelle may
-  also process Unicode/UCS data in UTF-8 encoding.  Unicode provides
-  its own collection of mathematical symbols, but there is no built-in
-  link to the standard collection of Isabelle.
+  and 8-bit characters are passed through transparently, Isabelle can
+  also process Unicode/UCS data in UTF-8 encoding.\footnote{When
+  counting precise source positions internally, bytes in the range of
+  128\dots 191 are ignored.  In UTF-8 encoding, this interval covers
+  the additional trailer bytes, so Isabelle happens to count Unicode
+  characters here, not bytes in memory.  In ISO-Latin encoding, the
+  ignored range merely includes some extra punctuation characters that
+  even have replacements within the standard collection of Isabelle
+  symbols; the accented letters range is counted properly.} Unicode
+  provides its own collection of mathematical symbols, but within the
+  core Isabelle/ML world there is no link to the standard collection
+  of Isabelle regular symbols.
 
   \medskip Output of Isabelle symbols depends on the print mode
   (\secref{print-mode}).  For example, the standard {\LaTeX} setup of
   the Isabelle document preparation system would present
   ``\verb,\,\verb,<alpha>,'' as \isa{{\isasymalpha}}, and
-  ``\verb,\,\verb,<^bold>,\verb,\,\verb,<alpha>,'' as \isa{\isactrlbold {\isasymalpha}}.%
+  ``\verb,\,\verb,<^bold>,\verb,\,\verb,<alpha>,'' as \isa{\isactrlbold {\isasymalpha}}.  On-screen rendering usually works by mapping a finite
+  subset of Isabelle symbols to suitable Unicode characters.%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -524,7 +704,7 @@
 %
 \begin{isamarkuptext}%
 \begin{mldecls}
-  \indexdef{}{ML type}{Symbol.symbol}\verb|type Symbol.symbol| \\
+  \indexdef{}{ML type}{Symbol.symbol}\verb|type Symbol.symbol = string| \\
   \indexdef{}{ML}{Symbol.explode}\verb|Symbol.explode: string -> Symbol.symbol list| \\
   \indexdef{}{ML}{Symbol.is\_letter}\verb|Symbol.is_letter: Symbol.symbol -> bool| \\
   \indexdef{}{ML}{Symbol.is\_digit}\verb|Symbol.is_digit: Symbol.symbol -> bool| \\
@@ -539,11 +719,14 @@
   \begin{description}
 
   \item \verb|Symbol.symbol| represents individual Isabelle
-  symbols; this is an alias for \verb|string|.
+  symbols.
 
   \item \verb|Symbol.explode|~\isa{str} produces a symbol list
   from the packed form.  This function supercedes \verb|String.explode| for virtually all purposes of manipulating text in
-  Isabelle!
+  Isabelle!\footnote{The runtime overhead for exploded strings is
+  mainly that of the list structure: individual symbols that happen to
+  be a singleton string --- which is the most common case --- do not
+  require extra memory in Poly/ML.}
 
   \item \verb|Symbol.is_letter|, \verb|Symbol.is_digit|, \verb|Symbol.is_quasi|, \verb|Symbol.is_blank| classify standard
   symbols according to fixed syntactic conventions of Isabelle, cf.\
@@ -555,7 +738,16 @@
   \item \verb|Symbol.decode| converts the string representation of a
   symbol into the datatype version.
 
-  \end{description}%
+  \end{description}
+
+  \paragraph{Historical note.} In the original SML90 standard the
+  primitive ML type \verb|char| did not exists, and the basic \verb|explode: string -> string list| operation would produce a list of
+  singleton strings as in Isabelle/ML today.  When SML97 came out,
+  Isabelle did not adopt its slightly anachronistic 8-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.%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -566,7 +758,7 @@
 %
 \endisadelimmlref
 %
-\isamarkupsubsection{Basic names \label{sec:basic-names}%
+\isamarkupsubsection{Basic names \label{sec:basic-name}%
 }
 \isamarkuptrue%
 %
@@ -583,7 +775,8 @@
   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 name references like \isa{xaa} to appear in the text.
+  as internal, which prevents mysterious names like \isa{xaa} to
+  appear in human-readable text.
 
   \medskip Manipulating binding scopes often requires on-the-fly
   renamings.  A \emph{name context} contains a collection of already
@@ -618,6 +811,9 @@
   \indexdef{}{ML}{Name.invents}\verb|Name.invents: Name.context -> string -> int -> string list| \\
   \indexdef{}{ML}{Name.variants}\verb|Name.variants: string list -> Name.context -> string list * Name.context| \\
   \end{mldecls}
+  \begin{mldecls}
+  \indexdef{}{ML}{Variable.names\_of}\verb|Variable.names_of: Proof.context -> Name.context| \\
+  \end{mldecls}
 
   \begin{description}
 
@@ -638,6 +834,14 @@
   \item \verb|Name.variants|~\isa{names\ context} produces fresh
   variants of \isa{names}; the result is entered into the context.
 
+  \item \verb|Variable.names_of|~\isa{ctxt} 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 \verb|Variable|, which is also able to provide an official status of
+  ``locally fixed variable'' within the logical environment (cf.\
+  \secref{sec:variables}).
+
   \end{description}%
 \end{isamarkuptext}%
 \isamarkuptrue%
@@ -649,7 +853,7 @@
 %
 \endisadelimmlref
 %
-\isamarkupsubsection{Indexed names%
+\isamarkupsubsection{Indexed names \label{sec:indexname}%
 }
 \isamarkuptrue%
 %
@@ -663,9 +867,9 @@
   \isa{maxidx\ {\isacharplus}\ {\isadigit{1}}}; the maximum index of an empty collection is
   \isa{{\isacharminus}{\isadigit{1}}}.
 
-  Occasionally, basic names and indexed names are injected into the
-  same pair type: the (improper) indexname \isa{{\isacharparenleft}x{\isacharcomma}\ {\isacharminus}{\isadigit{1}}{\isacharparenright}} is used
-  to encode basic names.
+  Occasionally, basic names are injected into the same pair type of
+  indexed names: then \isa{{\isacharparenleft}x{\isacharcomma}\ {\isacharminus}{\isadigit{1}}{\isacharparenright}} is used to encode the basic
+  name \isa{x}.
 
   \medskip Isabelle syntax observes the following rules for
   representing an indexname \isa{{\isacharparenleft}x{\isacharcomma}\ i{\isacharparenright}} as a packed string:
@@ -680,11 +884,12 @@
 
   \end{itemize}
 
-  Indexnames may acquire large index numbers over time.  Results are
-  normalized towards \isa{{\isadigit{0}}} 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
-  \isa{{\isacharquery}x{\isadigit{1}}{\isacharcomma}\ {\isacharquery}x{\isadigit{7}}{\isacharcomma}\ {\isacharquery}x{\isadigit{4}}{\isadigit{2}}} becomes \isa{{\isacharquery}x{\isacharcomma}\ {\isacharquery}xa{\isacharcomma}\ {\isacharquery}xb}.%
+  Indexnames may acquire large index numbers after several maxidx
+  shifts have been applied.  Results are usually normalized towards
+  \isa{{\isadigit{0}}} 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 \isa{{\isacharquery}x{\isadigit{1}}{\isacharcomma}\ {\isacharquery}x{\isadigit{7}}{\isacharcomma}\ {\isacharquery}x{\isadigit{4}}{\isadigit{2}}}
+  becomes \isa{{\isacharquery}x{\isacharcomma}\ {\isacharquery}xa{\isacharcomma}\ {\isacharquery}xb}.%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -704,7 +909,8 @@
   \item \verb|indexname| represents indexed names.  This is an
   abbreviation for \verb|string * int|.  The second component is
   usually non-negative, except for situations where \isa{{\isacharparenleft}x{\isacharcomma}\ {\isacharminus}{\isadigit{1}}{\isacharparenright}}
-  is used to embed basic names into this type.
+  is used to inject basic names into this type.  Other negative
+  indexes should not be used.
 
   \end{description}%
 \end{isamarkuptext}%
@@ -717,56 +923,31 @@
 %
 \endisadelimmlref
 %
-\isamarkupsubsection{Qualified names and name spaces%
+\isamarkupsubsection{Long names \label{sec:long-name}%
 }
 \isamarkuptrue%
 %
 \begin{isamarkuptext}%
-A \emph{qualified name} consists of a non-empty sequence of basic
-  name components.  The packed representation uses a dot as separator,
-  as in ``\isa{A{\isachardot}b{\isachardot}c}''.  The last component is called \emph{base}
-  name, the remaining prefix \emph{qualifier} (which may be empty).
-  The idea of qualified names is to encode nested structures by
-  recording the access paths as qualifiers.  For example, an item
-  named ``\isa{A{\isachardot}b{\isachardot}c}'' may be understood as a local entity \isa{c}, within a local structure \isa{b}, within a global
-  structure \isa{A}.  Typically, name space hierarchies consist of
-  1--2 levels of qualification, but this need not be always so.
+A \emph{long name} consists of a sequence of non-empty name
+  components.  The packed representation uses a dot as separator, as
+  in ``\isa{A{\isachardot}b{\isachardot}c}''.  The last component is called \emph{base
+  name}, the remaining prefix is called \emph{qualifier} (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 ``\isa{A{\isachardot}b{\isachardot}c}'' may be understood as
+  a local entity \isa{c}, within a local structure \isa{b},
+  within a global structure \isa{A}.  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, whenever this makes any sense.  The basic operations on
-  qualified names are smart enough to pass through such improper names
-  unchanged.
-
-  \medskip A \isa{naming} policy tells how to turn a name
-  specification into a fully qualified internal name (by the \isa{full} operation), and how fully qualified names may be accessed
-  externally.  For example, the default naming policy is to prefix an
-  implicit path: \isa{full\ x} produces \isa{path{\isachardot}x}, and the
-  standard accesses for \isa{path{\isachardot}x} include both \isa{x} and
-  \isa{path{\isachardot}x}.  Normally, the naming is implicit in the theory or
-  proof context; there are separate versions of the corresponding.
-
-  \medskip A \isa{name\ space} manages a collection of fully
-  internalized names, together with a mapping between external names
-  and internal names (in both directions).  The corresponding \isa{intern} and \isa{extern} operations are mostly used for
-  parsing and printing only!  The \isa{declare} operation augments
-  a name space according to the accesses determined by the naming
-  policy.
-
-  \medskip As a general principle, there is a separate name space for
-  each kind of formal entity, e.g.\ logical constant, type
-  constructor, type class, theorem.  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 \isa{c} may be used uniformly for a constant, type constructor, and
-  type class.
-
-  There are common schemes to name theorems systematically, according
-  to the name of the main logical entity involved, e.g.\ \isa{c{\isachardot}intro} for a canonical theorem related to constant \isa{c}.
-  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 are better suffixed in
-  addition to the usual qualification, e.g.\ \isa{c{\isacharunderscore}type{\isachardot}intro}
-  and \isa{c{\isacharunderscore}class{\isachardot}intro} for theorems related to type \isa{c}
-  and class \isa{c}, respectively.%
+  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.%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -784,6 +965,100 @@
   \indexdef{}{ML}{Long\_Name.implode}\verb|Long_Name.implode: string list -> string| \\
   \indexdef{}{ML}{Long\_Name.explode}\verb|Long_Name.explode: string -> string list| \\
   \end{mldecls}
+
+  \begin{description}
+
+  \item \verb|Long_Name.base_name|~\isa{name} returns the base name
+  of a long name.
+
+  \item \verb|Long_Name.qualifier|~\isa{name} returns the qualifier
+  of a long name.
+
+  \item \verb|Long_Name.append|~\isa{name\isactrlisub {\isadigit{1}}\ name\isactrlisub {\isadigit{2}}} appends two long
+  names.
+
+  \item \verb|Long_Name.implode|~\isa{names} and \verb|Long_Name.explode|~\isa{name} convert between the packed string
+  representation and the explicit list form of long names.
+
+  \end{description}%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\endisatagmlref
+{\isafoldmlref}%
+%
+\isadelimmlref
+%
+\endisadelimmlref
+%
+\isamarkupsubsection{Name spaces \label{sec:name-space}%
+}
+\isamarkuptrue%
+%
+\begin{isamarkuptext}%
+A \isa{name\ space} 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 \isa{intern} and \isa{extern} operations
+  are mostly used for parsing and printing only!  The \isa{declare} operation augments a name space according to the accesses
+  determined by a given binding, and a naming policy from the context.
+
+  \medskip A \isa{binding} 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 \isa{naming} provides some additional details for
+  producing a long name from a binding.  Normally, the naming is
+  implicit in the theory or proof context.  The \isa{full}
+  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:
+
+  \smallskip
+  \begin{tabular}{l}
+  \isa{binding\ {\isacharplus}\ naming\ {\isasymlongrightarrow}\ long\ name\ {\isacharplus}\ name\ space\ accesses}
+  \end{tabular}
+  \smallskip
+
+  \medskip 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 \isa{c} 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 \isa{c{\isachardot}intro} for a canonical introduction rule related to
+  constant \isa{c}.  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 are better suffixed in addition to the usual qualification,
+  e.g.\ \isa{c{\isacharunderscore}type{\isachardot}intro} and \isa{c{\isacharunderscore}class{\isachardot}intro} for
+  theorems related to type \isa{c} and class \isa{c},
+  respectively.%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\isadelimmlref
+%
+\endisadelimmlref
+%
+\isatagmlref
+%
+\begin{isamarkuptext}%
+\begin{mldecls}
+  \indexdef{}{ML type}{binding}\verb|type binding| \\
+  \indexdef{}{ML}{Binding.empty}\verb|Binding.empty: binding| \\
+  \indexdef{}{ML}{Binding.name}\verb|Binding.name: string -> binding| \\
+  \indexdef{}{ML}{Binding.qualify}\verb|Binding.qualify: bool -> string -> binding -> binding| \\
+  \indexdef{}{ML}{Binding.prefix}\verb|Binding.prefix: bool -> string -> binding -> binding| \\
+  \indexdef{}{ML}{Binding.conceal}\verb|Binding.conceal: binding -> binding| \\
+  \indexdef{}{ML}{Binding.str\_of}\verb|Binding.str_of: binding -> string| \\
+  \end{mldecls}
   \begin{mldecls}
   \indexdef{}{ML type}{Name\_Space.naming}\verb|type Name_Space.naming| \\
   \indexdef{}{ML}{Name\_Space.default\_naming}\verb|Name_Space.default_naming: Name_Space.naming| \\
@@ -798,21 +1073,39 @@
 \verb|  string * Name_Space.T| \\
   \indexdef{}{ML}{Name\_Space.intern}\verb|Name_Space.intern: Name_Space.T -> string -> string| \\
   \indexdef{}{ML}{Name\_Space.extern}\verb|Name_Space.extern: Name_Space.T -> string -> string| \\
+  \indexdef{}{ML}{Name\_Space.is\_concealed}\verb|Name_Space.is_concealed: Name_Space.T -> string -> bool|
   \end{mldecls}
 
   \begin{description}
 
-  \item \verb|Long_Name.base_name|~\isa{name} returns the base name of a
-  qualified name.
+  \item \verb|binding| represents the abstract concept of name
+  bindings.
+
+  \item \verb|Binding.empty| is the empty binding.
 
-  \item \verb|Long_Name.qualifier|~\isa{name} returns the qualifier
-  of a qualified name.
+  \item \verb|Binding.name|~\isa{name} produces a binding with base
+  name \isa{name}.
+
+  \item \verb|Binding.qualify|~\isa{mandatory\ name\ binding}
+  prefixes qualifier \isa{name} to \isa{binding}.  The \isa{mandatory} flag tells if this name component always needs to be
+  given in name space accesses --- this is mostly \isa{false} in
+  practice.  Note that this part of qualification is typically used in
+  derived specification mechanisms.
 
-  \item \verb|Long_Name.append|~\isa{name\isactrlisub {\isadigit{1}}\ name\isactrlisub {\isadigit{2}}}
-  appends two qualified names.
+  \item \verb|Binding.prefix| is similar to \verb|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}).
 
-  \item \verb|Long_Name.implode|~\isa{names} and \verb|Long_Name.explode|~\isa{name} convert between the packed string
-  representation and the explicit list form of qualified names.
+  \item \verb|Binding.conceal|~\isa{binding} 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.\ \verb|Name_Space.is_concealed|).
+
+  \item \verb|Binding.str_of|~\isa{binding} produces a string
+  representation for human-readable output, together with some formal
+  markup that might get used in GUI front-ends, for example.
 
   \item \verb|Name_Space.naming| represents the abstract concept of
   a naming policy.
@@ -853,6 +1146,10 @@
   This operation is mostly for printing!  User code should not rely on
   the precise result too much.
 
+  \item \verb|Name_Space.is_concealed|~\isa{space\ name} indicates
+  whether \isa{name} refers to a strictly private entity that
+  other tools are supposed to ignore!
+
   \end{description}%
 \end{isamarkuptext}%
 \isamarkuptrue%
--- a/doc-src/IsarImplementation/Thy/document/Proof.tex	Mon Feb 08 14:06:58 2010 +0100
+++ b/doc-src/IsarImplementation/Thy/document/Proof.tex	Mon Feb 08 14:08:32 2010 +0100
@@ -42,9 +42,10 @@
   or outside the current scope by providing separate syntactic
   categories for \emph{fixed variables} (e.g.\ \isa{x}) vs.\
   \emph{schematic variables} (e.g.\ \isa{{\isacharquery}x}).  Incidently, a
-  universal result \isa{{\isasymturnstile}\ {\isasymAnd}x{\isachardot}\ B{\isacharparenleft}x{\isacharparenright}} has the HHF normal form \isa{{\isasymturnstile}\ B{\isacharparenleft}{\isacharquery}x{\isacharparenright}}, which represents its generality nicely without requiring
-  an explicit quantifier.  The same principle works for type
-  variables: \isa{{\isasymturnstile}\ B{\isacharparenleft}{\isacharquery}{\isasymalpha}{\isacharparenright}} represents the idea of ``\isa{{\isasymturnstile}\ {\isasymforall}{\isasymalpha}{\isachardot}\ B{\isacharparenleft}{\isasymalpha}{\isacharparenright}}'' without demanding a truly polymorphic framework.
+  universal result \isa{{\isasymturnstile}\ {\isasymAnd}x{\isachardot}\ B{\isacharparenleft}x{\isacharparenright}} has the HHF normal form \isa{{\isasymturnstile}\ B{\isacharparenleft}{\isacharquery}x{\isacharparenright}}, which represents its generality without requiring an
+  explicit quantifier.  The same principle works for type variables:
+  \isa{{\isasymturnstile}\ B{\isacharparenleft}{\isacharquery}{\isasymalpha}{\isacharparenright}} represents the idea of ``\isa{{\isasymturnstile}\ {\isasymforall}{\isasymalpha}{\isachardot}\ B{\isacharparenleft}{\isasymalpha}{\isacharparenright}}''
+  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
@@ -113,7 +114,8 @@
   \indexdef{}{ML}{Variable.polymorphic}\verb|Variable.polymorphic: Proof.context -> term list -> term list| \\
   \indexdef{}{ML}{Variable.import}\verb|Variable.import: bool -> thm list -> Proof.context ->|\isasep\isanewline%
 \verb|  (((ctyp * ctyp) list * (cterm * cterm) list) * thm list) * Proof.context| \\
-  \indexdef{}{ML}{Variable.focus}\verb|Variable.focus: cterm -> Proof.context -> ((string * cterm) list * cterm) * Proof.context| \\
+  \indexdef{}{ML}{Variable.focus}\verb|Variable.focus: cterm -> Proof.context ->|\isasep\isanewline%
+\verb|  ((string * cterm) list * cterm) * Proof.context| \\
   \end{mldecls}
 
   \begin{description}
@@ -167,6 +169,160 @@
 %
 \endisadelimmlref
 %
+\isadelimmlex
+%
+\endisadelimmlex
+%
+\isatagmlex
+%
+\begin{isamarkuptext}%
+The following example (in theory \hyperlink{theory.Pure}{\mbox{\isa{Pure}}}) shows
+  how to work with fixed term and type parameters work with
+  type-inference.%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\endisatagmlex
+{\isafoldmlex}%
+%
+\isadelimmlex
+%
+\endisadelimmlex
+\isacommand{typedecl}\isamarkupfalse%
+\ foo\ \ %
+\isamarkupcmt{some basic type for testing purposes%
+}
+\isanewline
+%
+\isadelimML
+\isanewline
+%
+\endisadelimML
+%
+\isatagML
+\isacommand{ML}\isamarkupfalse%
+\ {\isacharverbatimopen}\isanewline
+\ \ {\isacharparenleft}{\isacharasterisk}static\ compile{\isacharminus}time\ context\ {\isacharminus}{\isacharminus}\ for\ testing\ only{\isacharasterisk}{\isacharparenright}\isanewline
+\ \ val\ ctxt{\isadigit{0}}\ {\isacharequal}\ %
+\isaantiq
+context%
+\endisaantiq
+{\isacharsemicolon}\isanewline
+\isanewline
+\ \ {\isacharparenleft}{\isacharasterisk}locally\ fixed\ parameters\ {\isacharminus}{\isacharminus}\ no\ type\ assignment\ yet{\isacharasterisk}{\isacharparenright}\isanewline
+\ \ val\ {\isacharparenleft}{\isacharbrackleft}x{\isacharcomma}\ y{\isacharbrackright}{\isacharcomma}\ ctxt{\isadigit{1}}{\isacharparenright}\ {\isacharequal}\ ctxt{\isadigit{0}}\ {\isacharbar}{\isachargreater}\ Variable{\isachardot}add{\isacharunderscore}fixes\ {\isacharbrackleft}{\isachardoublequote}x{\isachardoublequote}{\isacharcomma}\ {\isachardoublequote}y{\isachardoublequote}{\isacharbrackright}{\isacharsemicolon}\isanewline
+\isanewline
+\ \ {\isacharparenleft}{\isacharasterisk}t{\isadigit{1}}{\isacharcolon}\ most\ general\ fixed\ type{\isacharsemicolon}\ t{\isadigit{1}}{\isacharprime}{\isacharcolon}\ most\ general\ arbitrary\ type{\isacharasterisk}{\isacharparenright}\isanewline
+\ \ val\ t{\isadigit{1}}\ {\isacharequal}\ Syntax{\isachardot}read{\isacharunderscore}term\ ctxt{\isadigit{1}}\ {\isachardoublequote}x{\isachardoublequote}{\isacharsemicolon}\isanewline
+\ \ val\ t{\isadigit{1}}{\isacharprime}\ {\isacharequal}\ singleton\ {\isacharparenleft}Variable{\isachardot}polymorphic\ ctxt{\isadigit{1}}{\isacharparenright}\ t{\isadigit{1}}{\isacharsemicolon}\isanewline
+\isanewline
+\ \ {\isacharparenleft}{\isacharasterisk}term\ u\ enforces\ specific\ type\ assignment{\isacharasterisk}{\isacharparenright}\isanewline
+\ \ val\ u\ {\isacharequal}\ Syntax{\isachardot}read{\isacharunderscore}term\ ctxt{\isadigit{1}}\ {\isachardoublequote}{\isacharparenleft}x{\isacharcolon}{\isacharcolon}foo{\isacharparenright}\ {\isasymequiv}\ y{\isachardoublequote}{\isacharsemicolon}\isanewline
+\isanewline
+\ \ {\isacharparenleft}{\isacharasterisk}official\ declaration\ of\ u\ {\isacharminus}{\isacharminus}\ propagates\ constraints\ etc{\isachardot}{\isacharasterisk}{\isacharparenright}\isanewline
+\ \ val\ ctxt{\isadigit{2}}\ {\isacharequal}\ ctxt{\isadigit{1}}\ {\isacharbar}{\isachargreater}\ Variable{\isachardot}declare{\isacharunderscore}term\ u{\isacharsemicolon}\isanewline
+\ \ val\ t{\isadigit{2}}\ {\isacharequal}\ Syntax{\isachardot}read{\isacharunderscore}term\ ctxt{\isadigit{2}}\ {\isachardoublequote}x{\isachardoublequote}{\isacharsemicolon}\ \ {\isacharparenleft}{\isacharasterisk}x{\isacharcolon}{\isacharcolon}foo\ is\ enforced{\isacharasterisk}{\isacharparenright}\isanewline
+{\isacharverbatimclose}%
+\endisatagML
+{\isafoldML}%
+%
+\isadelimML
+%
+\endisadelimML
+%
+\begin{isamarkuptext}%
+In the above example, the starting context had been derived
+  from the toplevel theory, which means that fixed variables are
+  internalized literally: \verb|x| is mapped again to
+  \verb|x|, and attempting to fix it again in the subsequent
+  context is an error.  Alternatively, fixed parameters can be renamed
+  explicitly as follows:%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\isadelimML
+%
+\endisadelimML
+%
+\isatagML
+\isacommand{ML}\isamarkupfalse%
+\ {\isacharverbatimopen}\isanewline
+\ \ val\ ctxt{\isadigit{0}}\ {\isacharequal}\ %
+\isaantiq
+context%
+\endisaantiq
+{\isacharsemicolon}\isanewline
+\ \ val\ {\isacharparenleft}{\isacharbrackleft}x{\isadigit{1}}{\isacharcomma}\ x{\isadigit{2}}{\isacharcomma}\ x{\isadigit{3}}{\isacharbrackright}{\isacharcomma}\ ctxt{\isadigit{1}}{\isacharparenright}\ {\isacharequal}\isanewline
+\ \ \ \ ctxt{\isadigit{0}}\ {\isacharbar}{\isachargreater}\ Variable{\isachardot}variant{\isacharunderscore}fixes\ {\isacharbrackleft}{\isachardoublequote}x{\isachardoublequote}{\isacharcomma}\ {\isachardoublequote}x{\isachardoublequote}{\isacharcomma}\ {\isachardoublequote}x{\isachardoublequote}{\isacharbrackright}{\isacharsemicolon}\isanewline
+{\isacharverbatimclose}%
+\endisatagML
+{\isafoldML}%
+%
+\isadelimML
+%
+\endisadelimML
+%
+\begin{isamarkuptext}%
+\noindent Subsequent ML code can now work with the invented
+  names of \verb|x1|, \verb|x2|, \verb|x3|, 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:%
+\end{isamarkuptext}%
+\isamarkuptrue%
+\isacommand{lemma}\isamarkupfalse%
+\ {\isachardoublequoteopen}PROP\ XXX{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+%
+\endisadelimproof
+%
+\isatagproof
+\isacommand{proof}\isamarkupfalse%
+\ {\isacharminus}%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+%
+\isadelimML
+\ \ %
+\endisadelimML
+%
+\isatagML
+\isacommand{ML{\isacharunderscore}prf}\isamarkupfalse%
+\ {\isacharverbatimopen}\isanewline
+\ \ \ \ val\ ctxt{\isadigit{0}}\ {\isacharequal}\ %
+\isaantiq
+context%
+\endisaantiq
+{\isacharsemicolon}\isanewline
+\isanewline
+\ \ \ \ val\ {\isacharparenleft}{\isacharbrackleft}x{\isadigit{1}}{\isacharbrackright}{\isacharcomma}\ ctxt{\isadigit{1}}{\isacharparenright}\ {\isacharequal}\ ctxt{\isadigit{0}}\ {\isacharbar}{\isachargreater}\ Variable{\isachardot}add{\isacharunderscore}fixes\ {\isacharbrackleft}{\isachardoublequote}x{\isachardoublequote}{\isacharbrackright}{\isacharsemicolon}\isanewline
+\ \ \ \ val\ {\isacharparenleft}{\isacharbrackleft}x{\isadigit{2}}{\isacharbrackright}{\isacharcomma}\ ctxt{\isadigit{2}}{\isacharparenright}\ {\isacharequal}\ ctxt{\isadigit{1}}\ {\isacharbar}{\isachargreater}\ Variable{\isachardot}add{\isacharunderscore}fixes\ {\isacharbrackleft}{\isachardoublequote}x{\isachardoublequote}{\isacharbrackright}{\isacharsemicolon}\isanewline
+\ \ \ \ val\ {\isacharparenleft}{\isacharbrackleft}x{\isadigit{3}}{\isacharbrackright}{\isacharcomma}\ ctxt{\isadigit{3}}{\isacharparenright}\ {\isacharequal}\ ctxt{\isadigit{2}}\ {\isacharbar}{\isachargreater}\ Variable{\isachardot}add{\isacharunderscore}fixes\ {\isacharbrackleft}{\isachardoublequote}x{\isachardoublequote}{\isacharbrackright}{\isacharsemicolon}\isanewline
+\isanewline
+\ \ \ \ val\ {\isacharparenleft}{\isacharbrackleft}y{\isadigit{1}}{\isacharcomma}\ y{\isadigit{2}}{\isacharbrackright}{\isacharcomma}\ ctxt{\isadigit{4}}{\isacharparenright}\ {\isacharequal}\isanewline
+\ \ \ \ \ \ ctxt{\isadigit{3}}\ {\isacharbar}{\isachargreater}\ Variable{\isachardot}variant{\isacharunderscore}fixes\ {\isacharbrackleft}{\isachardoublequote}y{\isachardoublequote}{\isacharcomma}\ {\isachardoublequote}y{\isachardoublequote}{\isacharbrackright}{\isacharsemicolon}\isanewline
+\ \ {\isacharverbatimclose}\isanewline
+\ \ \isacommand{oops}\isamarkupfalse%
+%
+\endisatagML
+{\isafoldML}%
+%
+\isadelimML
+%
+\endisadelimML
+%
+\begin{isamarkuptext}%
+\noindent In this situation \verb|Variable.add_fixes| and \verb|Variable.variant_fixes| are very similar, but identical name
+  proposals given in a row are only accepted by the second version.%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
 \isamarkupsection{Assumptions \label{sec:assumptions}%
 }
 \isamarkuptrue%
@@ -192,21 +348,21 @@
   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 contexts, not the facts being exported!  There is a
+  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 \isa{{\isasymLongrightarrow}} introduction rule:
   \[
-  \infer[(\isa{{\isasymLongrightarrow}{\isacharunderscore}intro})]{\isa{{\isasymGamma}\ {\isacharbackslash}\ A\ {\isasymturnstile}\ A\ {\isasymLongrightarrow}\ B}}{\isa{{\isasymGamma}\ {\isasymturnstile}\ B}}
+  \infer[(\isa{{\isasymLongrightarrow}{\isasymdash}intro})]{\isa{{\isasymGamma}\ {\isacharminus}\ A\ {\isasymturnstile}\ A\ {\isasymLongrightarrow}\ B}}{\isa{{\isasymGamma}\ {\isasymturnstile}\ B}}
   \]
 
   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[(\isa{{\isacharhash}{\isasymLongrightarrow}{\isacharunderscore}intro})]{\isa{{\isasymGamma}\ {\isacharbackslash}\ A\ {\isasymturnstile}\ {\isacharhash}A\ {\isasymLongrightarrow}\ B}}{\isa{{\isasymGamma}\ {\isasymturnstile}\ B}}
+  \infer[(\isa{{\isacharhash}{\isasymLongrightarrow}{\isasymdash}intro})]{\isa{{\isasymGamma}\ {\isacharminus}\ A\ {\isasymturnstile}\ {\isacharhash}A\ {\isasymLongrightarrow}\ B}}{\isa{{\isasymGamma}\ {\isasymturnstile}\ B}}
   \]
 
   \medskip Alternative versions of assumptions may perform arbitrary
@@ -215,7 +371,7 @@
   definition works by fixing \isa{x} and assuming \isa{x\ {\isasymequiv}\ t},
   with the following export rule to reverse the effect:
   \[
-  \infer[(\isa{{\isasymequiv}{\isacharminus}expand})]{\isa{{\isasymGamma}\ {\isacharbackslash}\ x\ {\isasymequiv}\ t\ {\isasymturnstile}\ B\ t}}{\isa{{\isasymGamma}\ {\isasymturnstile}\ B\ x}}
+  \infer[(\isa{{\isasymequiv}{\isasymdash}expand})]{\isa{{\isasymGamma}\ {\isacharminus}\ {\isacharparenleft}x\ {\isasymequiv}\ t{\isacharparenright}\ {\isasymturnstile}\ B\ t}}{\isa{{\isasymGamma}\ {\isasymturnstile}\ B\ x}}
   \]
   This works, because the assumption \isa{x\ {\isasymequiv}\ t} was introduced in
   a context with \isa{x} being fresh, so \isa{x} does not
@@ -247,8 +403,8 @@
   where the \verb|bool| indicates goal mode, and the \verb|cterm list| the collection of assumptions to be discharged
   simultaneously.
 
-  \item \verb|Assumption.assume|~\isa{A} turns proposition \isa{A} into a raw assumption \isa{A\ {\isasymturnstile}\ A{\isacharprime}}, where the conclusion
-  \isa{A{\isacharprime}} is in HHF normal form.
+  \item \verb|Assumption.assume|~\isa{A} turns proposition \isa{A} into a primitive assumption \isa{A\ {\isasymturnstile}\ A{\isacharprime}}, where the
+  conclusion \isa{A{\isacharprime}} is in HHF normal form.
 
   \item \verb|Assumption.add_assms|~\isa{r\ As} augments the context
   by assumptions \isa{As} with export rule \isa{r}.  The
@@ -256,7 +412,8 @@
   \verb|Assumption.assume|.
 
   \item \verb|Assumption.add_assumes|~\isa{As} is a special case of
-  \verb|Assumption.add_assms| where the export rule performs \isa{{\isasymLongrightarrow}{\isacharunderscore}intro} or \isa{{\isacharhash}{\isasymLongrightarrow}{\isacharunderscore}intro}, depending on goal mode.
+  \verb|Assumption.add_assms| where the export rule performs \isa{{\isasymLongrightarrow}{\isasymdash}intro} or \isa{{\isacharhash}{\isasymLongrightarrow}{\isasymdash}intro}, depending on goal
+  mode.
 
   \item \verb|Assumption.export|~\isa{is{\isacharunderscore}goal\ inner\ outer\ thm}
   exports result \isa{thm} from the the \isa{inner} context
@@ -276,7 +433,68 @@
 %
 \endisadelimmlref
 %
-\isamarkupsection{Results \label{sec:results}%
+\isadelimmlex
+%
+\endisadelimmlex
+%
+\isatagmlex
+%
+\begin{isamarkuptext}%
+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 \hyperlink{theory.Pure}{\mbox{\isa{Pure}}} equality
+  here for testing purposes.%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\endisatagmlex
+{\isafoldmlex}%
+%
+\isadelimmlex
+%
+\endisadelimmlex
+%
+\isadelimML
+%
+\endisadelimML
+%
+\isatagML
+\isacommand{ML}\isamarkupfalse%
+\ {\isacharverbatimopen}\isanewline
+\ \ {\isacharparenleft}{\isacharasterisk}static\ compile{\isacharminus}time\ context\ {\isacharminus}{\isacharminus}\ for\ testing\ only{\isacharasterisk}{\isacharparenright}\isanewline
+\ \ val\ ctxt{\isadigit{0}}\ {\isacharequal}\ %
+\isaantiq
+context%
+\endisaantiq
+{\isacharsemicolon}\isanewline
+\isanewline
+\ \ val\ {\isacharparenleft}{\isacharbrackleft}eq{\isacharbrackright}{\isacharcomma}\ ctxt{\isadigit{1}}{\isacharparenright}\ {\isacharequal}\isanewline
+\ \ \ \ ctxt{\isadigit{0}}\ {\isacharbar}{\isachargreater}\ Assumption{\isachardot}add{\isacharunderscore}assumes\ {\isacharbrackleft}%
+\isaantiq
+cprop\ {\isachardoublequote}x\ {\isasymequiv}\ y{\isachardoublequote}%
+\endisaantiq
+{\isacharbrackright}{\isacharsemicolon}\isanewline
+\ \ val\ eq{\isacharprime}\ {\isacharequal}\ Thm{\isachardot}symmetric\ eq{\isacharsemicolon}\isanewline
+\isanewline
+\ \ {\isacharparenleft}{\isacharasterisk}back\ to\ original\ context\ {\isacharminus}{\isacharminus}\ discharges\ assumption{\isacharasterisk}{\isacharparenright}\isanewline
+\ \ val\ r\ {\isacharequal}\ Assumption{\isachardot}export\ false\ ctxt{\isadigit{1}}\ ctxt{\isadigit{0}}\ eq{\isacharprime}{\isacharsemicolon}\isanewline
+{\isacharverbatimclose}%
+\endisatagML
+{\isafoldML}%
+%
+\isadelimML
+%
+\endisadelimML
+%
+\begin{isamarkuptext}%
+\noindent 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.\
+  \verb|Variable.export| or the combined \verb|ProofContext.export|).%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\isamarkupsection{Structured goals and results \label{sec:struct-goals}%
 }
 \isamarkuptrue%
 %
@@ -296,6 +514,9 @@
   the tactic needs to solve the conclusion, but may use the premise as
   a local fact, for locally fixed variables.
 
+  The family of \isa{FOCUS} combinators is similar to \isa{SUBPROOF}, but allows to retain schematic variables and pending
+  subgoals in the resulting goal state.
+
   The \isa{prove} 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
@@ -308,7 +529,8 @@
   The \isa{obtain} 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 certain fixed variables and assumptions.  See also
+  by parameters and assumptions, without affecting any conclusions
+  that do not mention these parameters.  See also
   \cite{isabelle-isar-ref} for the user-level \isa{{\isasymOBTAIN}} and
   \isa{{\isasymGUESS}} elements.  Final results, which may not refer to
   the parameters in the conclusion, need to exported explicitly into
@@ -325,7 +547,11 @@
 \begin{isamarkuptext}%
 \begin{mldecls}
   \indexdef{}{ML}{SUBPROOF}\verb|SUBPROOF: (Subgoal.focus -> tactic) -> Proof.context -> int -> tactic| \\
+  \indexdef{}{ML}{Subgoal.FOCUS}\verb|Subgoal.FOCUS: (Subgoal.focus -> tactic) -> Proof.context -> int -> tactic| \\
+  \indexdef{}{ML}{Subgoal.FOCUS\_PREMS}\verb|Subgoal.FOCUS_PREMS: (Subgoal.focus -> tactic) -> Proof.context -> int -> tactic| \\
+  \indexdef{}{ML}{Subgoal.FOCUS\_PARAMS}\verb|Subgoal.FOCUS_PARAMS: (Subgoal.focus -> tactic) -> Proof.context -> int -> tactic| \\
   \end{mldecls}
+
   \begin{mldecls}
   \indexdef{}{ML}{Goal.prove}\verb|Goal.prove: Proof.context -> string list -> term list -> term ->|\isasep\isanewline%
 \verb|  ({prems: thm list, context: Proof.context} -> tactic) -> thm| \\
@@ -345,6 +571,11 @@
   schematic parameters of the goal are imported into the context as
   fixed ones, which may not be instantiated in the sub-proof.
 
+  \item \verb|Subgoal.FOCUS|, \verb|Subgoal.FOCUS_PREMS|, and \verb|Subgoal.FOCUS_PARAMS| are similar to \verb|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.
+
   \item \verb|Goal.prove|~\isa{ctxt\ xs\ As\ C\ tac} states goal \isa{C} in the context augmented by fixed variables \isa{xs} and
   assumptions \isa{As}, and applies tactic \isa{tac} to solve
   it.  The latter may depend on the local assumptions being presented
--- a/doc-src/IsarImplementation/Thy/document/Syntax.tex	Mon Feb 08 14:06:58 2010 +0100
+++ b/doc-src/IsarImplementation/Thy/document/Syntax.tex	Mon Feb 08 14:08:32 2010 +0100
@@ -18,7 +18,7 @@
 %
 \endisadelimtheory
 %
-\isamarkupchapter{Syntax and type-checking%
+\isamarkupchapter{Concrete syntax and type-checking%
 }
 \isamarkuptrue%
 %
@@ -27,6 +27,19 @@
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
+\isamarkupsection{Parsing and printing%
+}
+\isamarkuptrue%
+%
+\begin{isamarkuptext}%
+FIXME%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\isamarkupsection{Checking and unchecking \label{sec:term-check}%
+}
+\isamarkuptrue%
+%
 \isadelimtheory
 %
 \endisadelimtheory
--- a/doc-src/IsarImplementation/Thy/document/Tactic.tex	Mon Feb 08 14:06:58 2010 +0100
+++ b/doc-src/IsarImplementation/Thy/document/Tactic.tex	Mon Feb 08 14:08:32 2010 +0100
@@ -23,7 +23,7 @@
 \isamarkuptrue%
 %
 \begin{isamarkuptext}%
-Tactical reasoning works by refining the initial claim in a
+Tactical reasoning works by refining an initial claim in a
   backwards fashion, until a solved form is reached.  A \isa{goal}
   consists of several subgoals that need to be solved in order to
   achieve the main statement; zero subgoals means that the proof may
@@ -54,8 +54,8 @@
   nesting rarely exceeds 1--2 in practice.
 
   The main conclusion \isa{C} is internally marked as a protected
-  proposition, which is represented explicitly by the notation \isa{{\isacharhash}C}.  This ensures that the decomposition into subgoals and main
-  conclusion is well-defined for arbitrarily structured claims.
+  proposition, which is represented explicitly by the notation \isa{{\isacharhash}C} 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:
@@ -129,26 +129,27 @@
   supporting memoing.\footnote{The lack of memoing and the strict
   nature of SML requires some care when working with low-level
   sequence operations, to avoid duplicate or premature evaluation of
-  results.}
+  results.  It also means that modified runtime behavior, such as
+  timeout, is very hard to achieve for general tactics.}
 
   An \emph{empty result sequence} means that the tactic has failed: in
-  a compound tactic expressions other tactics might be tried instead,
+  a compound tactic expression other tactics might be tried instead,
   or the whole refinement step might fail outright, producing a
-  toplevel error message.  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.
+  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{multiple results}, 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 in \secref{sec:tactical-goals}, 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).
+  \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).
 
   Tactics with explicit \emph{subgoal addressing} are of the form
   \isa{int\ {\isasymrightarrow}\ tactic} and may be applied to a particular subgoal
@@ -170,7 +171,7 @@
 
   Tactics with internal subgoal addressing should expose the subgoal
   index as \isa{int} argument in full generality; a hardwired
-  subgoal 1 inappropriate.
+  subgoal 1 is not acceptable.
   
   \medskip The main well-formedness conditions for proper tactics are
   summarized as follows.
@@ -192,11 +193,11 @@
   \end{itemize}
 
   Some of these conditions are checked by higher-level goal
-  infrastructure (\secref{sec:results}); others are not checked
+  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 disallow composition with
-  basic tacticals).%
+  applicable only to singleton goals, or prevent composition via
+  standard tacticals).%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -414,7 +415,7 @@
   Type instantiations may be given as well, via pairs like \isa{{\isacharparenleft}{\isacharquery}{\isacharprime}a{\isacharcomma}\ {\isasymtau}{\isacharparenright}}.  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 type-inference as expected, explicit type
+  already performs simple type-inference, so explicit type
   instantiations are seldom necessary.%
 \end{isamarkuptext}%
 \isamarkuptrue%
@@ -452,7 +453,13 @@
   \item \verb|rename_tac|~\isa{names\ i} renames the innermost
   parameters of subgoal \isa{i} according to the provided \isa{names} (which need to be distinct indentifiers).
 
-  \end{description}%
+  \end{description}
+
+  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 \verb|Subgoal.FOCUS| combinator
+  of \secref{sec:struct-goals} allows to refer to internal goal
+  structure with explicit context management.%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
--- a/doc-src/IsarImplementation/implementation.tex	Mon Feb 08 14:06:58 2010 +0100
+++ b/doc-src/IsarImplementation/implementation.tex	Mon Feb 08 14:08:32 2010 +0100
@@ -25,7 +25,7 @@
 
 \begin{document}
 
-\maketitle 
+\maketitle
 
 \begin{abstract}
   We describe the key concepts underlying the Isabelle/Isar
@@ -37,7 +37,7 @@
 
 \vspace*{2.5cm}
 \begin{quote}
-  
+
   {\small\em Isabelle was not designed; it evolved.  Not everyone
     likes this idea.  Specification experts rightly abhor
     trial-and-error programming.  They suggest that no one should
@@ -45,17 +45,28 @@
     specification. But university departments are not software houses.
     Programs like Isabelle are not products: when they have served
     their purpose, they are discarded.}
-  
+
   Lawrence C. Paulson, ``Isabelle: The Next 700 Theorem Provers''
 
   \vspace*{1cm}
-  
+
   {\small\em As I did 20 years ago, I still fervently believe that the
     only way to make software secure, reliable, and fast is to make it
     small.  Fight features.}
-  
+
   Andrew S. Tanenbaum
 
+  \vspace*{1cm}
+
+  {\small\em One thing that UNIX does not need is more features. It is
+    successful in part because it has a small number of good ideas
+    that work well together. Merely adding features does not make it
+    easier for users to do things --- it just makes the manual
+    thicker. The right solution in the right place is always more
+    effective than haphazard hacking.}
+
+  Rob Pike and Brian W. Kernighan
+
 \end{quote}
 
 \thispagestyle{empty}\clearpage
@@ -89,7 +100,7 @@
 \end{document}
 
 
-%%% Local Variables: 
+%%% Local Variables:
 %%% mode: latex
 %%% TeX-master: t
-%%% End: 
+%%% End:
--- a/doc-src/IsarImplementation/style.sty	Mon Feb 08 14:06:58 2010 +0100
+++ b/doc-src/IsarImplementation/style.sty	Mon Feb 08 14:08:32 2010 +0100
@@ -19,7 +19,6 @@
 \pagestyle{headings}
 \sloppy
 \binperiod
-\underscoreon
 
 \renewcommand{\isadigit}[1]{\isamath{#1}}
 
@@ -30,6 +29,13 @@
 \renewcommand{\isatagmlref}{\subsection*{\makebox[0pt][r]{\fbox{\ML}~~}Reference}\begingroup\def\isastyletext{\rm}\small}
 \renewcommand{\endisatagmlref}{\endgroup}
 
+\isakeeptag{mlex}
+\renewcommand{\isatagmlex}{\subsection*{\makebox[0pt][r]{\fbox{\ML}~~}Examples}}
+\renewcommand{\endisatagmlex}{}
+
+\renewcommand{\isatagML}{\begingroup\isabellestyle{default}\isastyle\def\isadigit##1{##1}}
+\renewcommand{\endisatagML}{\endgroup}
+
 \newcommand{\minorcmd}[1]{{\sf #1}}
 \newcommand{\isasymtype}{\minorcmd{type}}
 \newcommand{\isasymval}{\minorcmd{val}}
@@ -49,9 +55,7 @@
 \newcommand{\isasymDEFINITION}{\isakeyword{definition}}
 
 \isabellestyle{it}
-
-
-%%% Local Variables: 
-%%% mode: latex
-%%% TeX-master: "implementation"
-%%% End: 
+\underscoreon
+\renewcommand{\isacharunderscore}{\_}
+\renewcommand{\isacharunderscorekeyword}{\_}
+\newcommand{\isasymdash}{\mbox{-}}
--- a/doc-src/Main/Docs/Main_Doc.thy	Mon Feb 08 14:06:58 2010 +0100
+++ b/doc-src/Main/Docs/Main_Doc.thy	Mon Feb 08 14:08:32 2010 +0100
@@ -48,8 +48,8 @@
 \smallskip
 
 \begin{supertabular}{@ {} l @ {~::~} l l @ {}}
-@{const HOL.less_eq} & @{typeof HOL.less_eq} & (\verb$<=$)\\
-@{const HOL.less} & @{typeof HOL.less}\\
+@{const Algebras.less_eq} & @{typeof Algebras.less_eq} & (\verb$<=$)\\
+@{const Algebras.less} & @{typeof Algebras.less}\\
 @{const Orderings.Least} & @{typeof Orderings.Least}\\
 @{const Orderings.min} & @{typeof Orderings.min}\\
 @{const Orderings.max} & @{typeof Orderings.max}\\
--- a/etc/settings	Mon Feb 08 14:06:58 2010 +0100
+++ b/etc/settings	Mon Feb 08 14:08:32 2010 +0100
@@ -50,12 +50,6 @@
 #ML_PLATFORM=$(eval $("$ML_HOME/.arch-n-opsys" 2>/dev/null); echo "$HEAP_SUFFIX")
 #SMLNJ_CYGWIN_RUNTIME=1
 
-# Moscow ML 2.00 (experimental!)
-#ML_SYSTEM=mosml
-#ML_HOME="/usr/local/mosml/bin"
-#ML_OPTIONS=""
-#ML_PLATFORM=""
-
 
 ###
 ### JVM components (Scala or Java)
--- a/lib/Tools/keywords	Mon Feb 08 14:06:58 2010 +0100
+++ b/lib/Tools/keywords	Mon Feb 08 14:08:32 2010 +0100
@@ -66,5 +66,4 @@
     gzip -dc "$LOG"
   fi
   echo
-done | \
-perl -w "$ISABELLE_HOME/lib/scripts/keywords.pl" "$KEYWORDS_NAME" "$SESSIONS"
+done | "$ISABELLE_HOME/lib/scripts/keywords" "$KEYWORDS_NAME" "$SESSIONS"
--- a/lib/Tools/unsymbolize	Mon Feb 08 14:06:58 2010 +0100
+++ b/lib/Tools/unsymbolize	Mon Feb 08 14:08:32 2010 +0100
@@ -34,4 +34,4 @@
 ## main
 
 find $SPECS \( -name \*.ML -o -name \*.thy \) -print | \
-  xargs perl -w "$ISABELLE_HOME/lib/scripts/unsymbolize.pl"
+  xargs "$ISABELLE_HOME/lib/scripts/unsymbolize"
--- a/lib/Tools/yxml	Mon Feb 08 14:06:58 2010 +0100
+++ b/lib/Tools/yxml	Mon Feb 08 14:08:32 2010 +0100
@@ -31,4 +31,4 @@
 
 ## main
 
-exec perl -w "$ISABELLE_HOME/lib/scripts/yxml.pl"
+exec "$ISABELLE_HOME/lib/scripts/yxml"
--- a/lib/fonts/IsabelleText.sfd	Mon Feb 08 14:06:58 2010 +0100
+++ b/lib/fonts/IsabelleText.sfd	Mon Feb 08 14:08:32 2010 +0100
@@ -13,13 +13,14 @@
 LayerCount: 2
 Layer: 0 1 "Back"  1
 Layer: 1 1 "Fore"  0
+NeedsXUIDChange: 1
 XUID: [1021 906 1711068302 4288927]
 FSType: 4
 OS2Version: 1
 OS2_WeightWidthSlopeOnly: 0
 OS2_UseTypoMetrics: 1
 CreationTime: 1050361371
-ModificationTime: 1263327886
+ModificationTime: 1265397211
 PfmFamily: 17
 TTFWeight: 400
 TTFWidth: 5
@@ -2240,7 +2241,7 @@
 DisplaySize: -48
 AntiAlias: 1
 FitToEm: 1
-WinInfo: 160 16 10
+WinInfo: 8800 16 10
 TeXData: 1 0 0 631296 315648 210432 572416 -1048576 210432 783286 444596 497025 792723 393216 433062 380633 303038 157286 324010 404750 52429 2506097 1059062 262144
 BeginChars: 1114189 648
 
@@ -13869,61 +13870,67 @@
 LayerCount: 2
 Fore
 SplineSet
-1401 1505 m 6,0,-1
- 402 -440 l 2,1,2
- 382 -480 382 -480 353 -480 c 0,3,4
- 332 -480 332 -480 320.5 -465 c 128,-1,5
- 309 -450 309 -450 309 -435 c 0,6,7
- 309 -426 309 -426 324 -395 c 2,8,-1
- 1323 1550 l 2,9,10
- 1343 1590 1343 1590 1372 1590 c 0,11,12
- 1393 1590 1393 1590 1404.5 1575 c 128,-1,13
- 1416 1560 1416 1560 1416 1545 c 0,14,15
- 1417 1536 1417 1536 1401 1505 c 6,0,-1
-1525 726 m 2,16,-1
- 200 726 l 2,17,18
- 186 726 186 726 178.5 726.5 c 128,-1,19
- 171 727 171 727 159 729 c 128,-1,20
- 147 731 147 731 140.5 735.5 c 128,-1,21
- 134 740 134 740 129 749 c 128,-1,22
- 124 758 124 758 124 770 c 0,23,24
- 124 786 124 786 130.5 796 c 128,-1,25
- 137 806 137 806 150 809.5 c 128,-1,26
- 163 813 163 813 172 814 c 128,-1,27
- 181 815 181 815 198 815 c 2,28,-1
- 1527 815 l 2,29,30
- 1543 815 1543 815 1552.5 814 c 128,-1,31
- 1562 813 1562 813 1575 809.5 c 128,-1,32
- 1588 806 1588 806 1594.5 796 c 128,-1,33
- 1601 786 1601 786 1601 770 c 0,34,35
- 1601 757 1601 757 1595.5 748.5 c 128,-1,36
- 1590 740 1590 740 1584 735.5 c 128,-1,37
- 1578 731 1578 731 1566 729 c 128,-1,38
- 1554 727 1554 727 1546.5 726.5 c 128,-1,39
- 1539 726 1539 726 1525 726 c 2,16,-1
-1527 295 m 2,40,-1
- 198 295 l 2,41,42
- 182 295 182 295 172.5 296 c 128,-1,43
- 163 297 163 297 150 300.5 c 128,-1,44
- 137 304 137 304 130.5 314 c 128,-1,45
- 124 324 124 324 124 340 c 0,46,47
- 124 353 124 353 129 361.5 c 128,-1,48
- 134 370 134 370 140.5 374.5 c 128,-1,49
- 147 379 147 379 159 381 c 128,-1,50
- 171 383 171 383 178.5 383.5 c 128,-1,51
- 186 384 186 384 200 384 c 2,52,-1
- 1525 384 l 2,53,54
- 1539 384 1539 384 1546.5 383.5 c 128,-1,55
- 1554 383 1554 383 1566 381 c 128,-1,56
- 1578 379 1578 379 1584 374.5 c 128,-1,57
- 1590 370 1590 370 1595.5 361 c 128,-1,58
- 1601 352 1601 352 1601 340 c 0,59,60
- 1601 324 1601 324 1594.5 314 c 128,-1,61
- 1588 304 1588 304 1575 300.5 c 128,-1,62
- 1562 297 1562 297 1552.5 296 c 128,-1,63
- 1543 295 1543 295 1527 295 c 2,40,-1
-EndSplineSet
-Validated: 5
+945 815 m 1,0,-1
+ 1323 1550 l 2,1,2
+ 1343 1590 1343 1590 1372 1590 c 0,3,4
+ 1393 1590 1393 1590 1404.5 1575 c 128,-1,5
+ 1416 1560 1416 1560 1416 1545 c 1,6,7
+ 1417 1536 1417 1536 1401 1505 c 2,8,-1
+ 1047 815 l 1,9,-1
+ 1527 815 l 2,10,11
+ 1543 815 1543 815 1552.5 814 c 128,-1,12
+ 1562 813 1562 813 1575 809.5 c 128,-1,13
+ 1588 806 1588 806 1594.5 796 c 128,-1,14
+ 1601 786 1601 786 1601 770 c 0,15,16
+ 1601 757 1601 757 1595.5 748.5 c 128,-1,17
+ 1590 740 1590 740 1584 735.5 c 128,-1,18
+ 1578 731 1578 731 1566 729 c 128,-1,19
+ 1554 727 1554 727 1546.5 726.5 c 128,-1,20
+ 1539 726 1539 726 1525 726 c 2,21,-1
+ 1001 726 l 1,22,-1
+ 825 384 l 1,23,-1
+ 1525 384 l 2,24,25
+ 1539 384 1539 384 1546.5 383.5 c 128,-1,26
+ 1554 383 1554 383 1566 381 c 128,-1,27
+ 1578 379 1578 379 1584 374.5 c 128,-1,28
+ 1590 370 1590 370 1595.5 361 c 128,-1,29
+ 1601 352 1601 352 1601 340 c 0,30,31
+ 1601 324 1601 324 1594.5 314 c 128,-1,32
+ 1588 304 1588 304 1575 300.5 c 128,-1,33
+ 1562 297 1562 297 1552.5 296 c 128,-1,34
+ 1543 295 1543 295 1527 295 c 2,35,-1
+ 780 295 l 1,36,-1
+ 402 -440 l 2,37,38
+ 382 -480 382 -480 353 -480 c 0,39,40
+ 332 -480 332 -480 320.5 -465 c 128,-1,41
+ 309 -450 309 -450 309 -435 c 0,42,43
+ 309 -426 309 -426 324 -395 c 2,44,-1
+ 678 295 l 1,45,-1
+ 198 295 l 2,46,47
+ 182 295 182 295 172.5 296 c 128,-1,48
+ 163 297 163 297 150 300.5 c 128,-1,49
+ 137 304 137 304 130.5 314 c 128,-1,50
+ 124 324 124 324 124 340 c 0,51,52
+ 124 353 124 353 129 361.5 c 128,-1,53
+ 134 370 134 370 140.5 374.5 c 128,-1,54
+ 147 379 147 379 159 381 c 128,-1,55
+ 171 383 171 383 178.5 383.5 c 128,-1,56
+ 186 384 186 384 200 384 c 2,57,-1
+ 724 384 l 1,58,-1
+ 900 726 l 1,59,-1
+ 200 726 l 2,60,61
+ 186 726 186 726 178.5 726.5 c 128,-1,62
+ 171 727 171 727 159 729 c 128,-1,63
+ 147 731 147 731 140.5 735.5 c 128,-1,64
+ 134 740 134 740 129 749 c 128,-1,65
+ 124 758 124 758 124 770 c 0,66,67
+ 124 786 124 786 130 796 c 0,68,69
+ 137 806 137 806 150 809.5 c 128,-1,70
+ 163 813 163 813 172 814 c 128,-1,71
+ 181 815 181 815 198 815 c 2,72,-1
+ 945 815 l 1,0,-1
+EndSplineSet
+Validated: 1
 EndChar
 
 StartChar: AE
@@ -14371,6 +14378,7 @@
  1544 695 1544 695 1527 695 c 2,55,-1
  908 695 l 1,0,-1
 EndSplineSet
+Validated: 1
 EndChar
 
 StartChar: lessequal
@@ -17238,6 +17246,7 @@
  163 598 163 598 172.5 599 c 128,-1,43
  182 600 182 600 198 600 c 2,24,-1
 EndSplineSet
+Validated: 33
 EndChar
 
 StartChar: lozenge
@@ -17930,6 +17939,7 @@
  357 672 357 672 391.5 637.5 c 128,-1,0
  426 603 426 603 426 555 c 128,-1,1
 EndSplineSet
+Validated: 33
 EndChar
 
 StartChar: quotesinglbase
@@ -18068,7 +18078,7 @@
  2098 665 2098 665 2080 634.5 c 128,-1,13
  2062 604 2062 604 2049 563.5 c 128,-1,14
  2036 523 2036 523 2027.5 460.5 c 128,-1,15
- 2019 398 2019 398 2019 322 c 0,16,17
+ 2019 398 2019 398 2019 322 c 128,-1,17
  2019 246 2019 246 2027.5 184 c 128,-1,18
  2036 122 2036 122 2049 81.5 c 128,-1,19
  2062 41 2062 41 2080.5 11 c 128,-1,20
@@ -18145,7 +18155,7 @@
  353 1560 353 1560 335 1529.5 c 128,-1,110
  317 1499 317 1499 304 1458.5 c 128,-1,111
  291 1418 291 1418 282.5 1355.5 c 128,-1,112
- 274 1293 274 1293 274 1217 c 0,113,114
+ 274 1293 274 1293 274 1217 c 128,-1,114
  274 1141 274 1141 282.5 1079 c 128,-1,115
  291 1017 291 1017 304 976.5 c 128,-1,116
  317 936 317 936 335.5 905.5 c 128,-1,117
@@ -18166,7 +18176,7 @@
  1337 665 1337 665 1318.5 634.5 c 128,-1,134
  1300 604 1300 604 1287 563.5 c 128,-1,135
  1274 523 1274 523 1266 460.5 c 128,-1,136
- 1258 398 1258 398 1258 322 c 0,137,138
+ 1258 398 1258 398 1258 322 c 128,-1,138
  1258 246 1258 246 1266 184 c 128,-1,139
  1274 122 1274 122 1287.5 81.5 c 128,-1,140
  1301 41 1301 41 1319 11 c 128,-1,141
@@ -20091,6 +20101,7 @@
  1389 1018 1389 1018 1370 999 c 2,41,-1
  928 555 l 1,0,-1
 EndSplineSet
+Validated: 1
 EndChar
 
 StartChar: uni00B9
@@ -28736,75 +28747,63 @@
 LayerCount: 2
 Fore
 SplineSet
-213 511 m 5,0,-1
- 213 215 l 2,1,2
- 213 199 213 199 212 190 c 128,-1,3
- 211 181 211 181 207.5 168 c 128,-1,4
- 204 155 204 155 194 148.5 c 128,-1,5
- 184 142 184 142 168.5 142 c 128,-1,6
- 153 142 153 142 143 148.5 c 128,-1,7
- 133 155 133 155 129.5 168 c 128,-1,8
- 126 181 126 181 125 190 c 128,-1,9
- 124 199 124 199 124 215 c 2,10,-1
- 124 895 l 2,11,12
- 124 911 124 911 125 920 c 128,-1,13
- 126 929 126 929 129.5 942 c 128,-1,14
- 133 955 133 955 143 961.5 c 128,-1,15
- 153 968 153 968 168.5 968 c 128,-1,16
- 184 968 184 968 194 961.5 c 128,-1,17
- 204 955 204 955 207.5 942 c 128,-1,18
- 211 929 211 929 212 920 c 128,-1,19
- 213 911 213 911 213 895 c 2,20,-1
- 213 599 l 1,21,22
- 275 599 275 599 275 555 c 128,-1,23
- 275 511 275 511 213 511 c 5,0,-1
-1905 511 m 5,24,-1
- 198 511 l 2,25,26
- 184 511 184 511 177 511 c 128,-1,27
- 170 511 170 511 158 513.5 c 128,-1,28
- 146 516 146 516 140 520.5 c 128,-1,29
- 134 525 134 525 129 533.5 c 128,-1,30
- 124 542 124 542 124 555 c 128,-1,31
- 124 568 124 568 129 576.5 c 128,-1,32
- 134 585 134 585 140 589.5 c 128,-1,33
- 146 594 146 594 158 596.5 c 128,-1,34
- 170 599 170 599 177 599 c 128,-1,35
- 184 599 184 599 198 599 c 2,36,-1
- 1905 599 l 1,37,38
- 1847 641 1847 641 1798 698 c 128,-1,39
- 1749 755 1749 755 1722 802 c 128,-1,40
- 1695 849 1695 849 1680 882 c 128,-1,41
- 1665 915 1665 915 1665 924 c 0,42,43
- 1665 938 1665 938 1675 944 c 128,-1,44
- 1685 950 1685 950 1698 950 c 0,45,46
- 1715 950 1715 950 1721.5 943.5 c 128,-1,47
- 1728 937 1728 937 1736 919 c 0,48,49
- 1769 846 1769 846 1809.5 788.5 c 128,-1,50
- 1850 731 1850 731 1896.5 691.5 c 128,-1,51
- 1943 652 1943 652 1983.5 626.5 c 128,-1,52
- 2024 601 2024 601 2076 577 c 1,53,54
- 2079 577 2079 577 2085 571 c 128,-1,55
- 2091 565 2091 565 2091 555 c 0,56,57
- 2091 544 2091 544 2086 539 c 128,-1,58
- 2081 534 2081 534 2060 524 c 0,59,60
- 2014 502 2014 502 1973.5 475.5 c 128,-1,61
- 1933 449 1933 449 1904 424.5 c 128,-1,62
- 1875 400 1875 400 1848.5 369 c 128,-1,63
- 1822 338 1822 338 1806.5 316 c 128,-1,64
- 1791 294 1791 294 1774.5 264 c 128,-1,65
- 1758 234 1758 234 1751 220 c 128,-1,66
- 1744 206 1744 206 1734 182 c 0,67,68
- 1730 174 1730 174 1728.5 171.5 c 128,-1,69
- 1727 169 1727 169 1719 164.5 c 128,-1,70
- 1711 160 1711 160 1698 160 c 0,71,72
- 1684 160 1684 160 1674.5 166.5 c 128,-1,73
- 1665 173 1665 173 1665 186 c 0,74,75
- 1665 195 1665 195 1680 228 c 128,-1,76
- 1695 261 1695 261 1722 308 c 128,-1,77
- 1749 355 1749 355 1798 412 c 128,-1,78
- 1847 469 1847 469 1905 511 c 5,24,-1
-EndSplineSet
-Validated: 5
+213 599 m 1,0,-1
+ 1905 599 l 1,1,2
+ 1847 641 1847 641 1798 698 c 128,-1,3
+ 1749 755 1749 755 1722 802 c 128,-1,4
+ 1695 849 1695 849 1680 882 c 128,-1,5
+ 1665 915 1665 915 1665 924 c 0,6,7
+ 1665 938 1665 938 1675 944 c 128,-1,8
+ 1685 950 1685 950 1698 950 c 0,9,10
+ 1715 950 1715 950 1722 944 c 0,11,12
+ 1728 937 1728 937 1736 919 c 0,13,14
+ 1767 848 1767 848 1810 788 c 0,15,16
+ 1850 731 1850 731 1896.5 691.5 c 128,-1,17
+ 1943 652 1943 652 1983.5 626.5 c 128,-1,18
+ 2024 601 2024 601 2076 577 c 1,19,20
+ 2079 577 2079 577 2085 571 c 128,-1,21
+ 2091 565 2091 565 2091 555 c 0,22,23
+ 2091 544 2091 544 2086 539 c 128,-1,24
+ 2081 534 2081 534 2060 524 c 0,25,26
+ 2014 502 2014 502 1973.5 475.5 c 128,-1,27
+ 1933 449 1933 449 1904 424.5 c 128,-1,28
+ 1875 400 1875 400 1848.5 369 c 128,-1,29
+ 1822 338 1822 338 1806.5 316 c 128,-1,30
+ 1791 294 1791 294 1774.5 264 c 128,-1,31
+ 1758 234 1758 234 1751 220 c 128,-1,32
+ 1744 206 1744 206 1734 182 c 0,33,34
+ 1730 174 1730 174 1728.5 171.5 c 128,-1,35
+ 1727 169 1727 169 1719 164.5 c 128,-1,36
+ 1711 160 1711 160 1698 160 c 0,37,38
+ 1684 160 1684 160 1674 166 c 0,39,40
+ 1665 173 1665 173 1665 186 c 0,41,42
+ 1665 195 1665 195 1680 228 c 128,-1,43
+ 1695 261 1695 261 1722 308 c 128,-1,44
+ 1749 355 1749 355 1798 412 c 128,-1,45
+ 1847 469 1847 469 1905 511 c 1,46,-1
+ 213 511 l 1,47,-1
+ 213 215 l 2,48,49
+ 213 199 213 199 212 190 c 128,-1,50
+ 211 181 211 181 207.5 168 c 128,-1,51
+ 204 155 204 155 194 148.5 c 128,-1,52
+ 184 142 184 142 168.5 142 c 128,-1,53
+ 153 142 153 142 143 148.5 c 128,-1,54
+ 133 155 133 155 129.5 168 c 128,-1,55
+ 126 181 126 181 125 190 c 128,-1,56
+ 124 199 124 199 124 215 c 2,57,-1
+ 124 555 l 1,58,-1
+ 124 895 l 2,59,60
+ 124 911 124 911 125 920 c 128,-1,61
+ 126 929 126 929 129.5 942 c 128,-1,62
+ 133 955 133 955 143 961.5 c 128,-1,63
+ 153 968 153 968 168.5 968 c 128,-1,64
+ 184 968 184 968 194 961.5 c 128,-1,65
+ 204 955 204 955 207.5 942 c 128,-1,66
+ 211 929 211 929 212 920 c 128,-1,67
+ 213 911 213 911 213 895 c 2,68,-1
+ 213 599 l 1,0,-1
+EndSplineSet
+Validated: 1
 EndChar
 
 StartChar: uni21A9
@@ -28817,78 +28816,68 @@
 LayerCount: 2
 Fore
 SplineSet
-2045 511 m 6,0,-1
- 337 511 l 1,1,2
- 395 469 395 469 444.5 412 c 128,-1,3
- 494 355 494 355 521 308 c 128,-1,4
- 548 261 548 261 562.5 228 c 128,-1,5
- 577 195 577 195 577 186 c 0,6,7
- 577 172 577 172 567 166 c 128,-1,8
- 557 160 557 160 544 160 c 0,9,10
- 527 160 527 160 520.5 166.5 c 128,-1,11
- 514 173 514 173 506 191 c 0,12,13
- 455 301 455 301 380 383.5 c 128,-1,14
- 305 466 305 466 178 526 c 0,15,16
- 172 529 172 529 167 532 c 128,-1,17
- 162 535 162 535 158.5 538 c 128,-1,18
- 155 541 155 541 153 545 c 128,-1,19
- 151 549 151 549 151 555 c 0,20,21
- 151 561 151 561 151.5 564 c 128,-1,22
- 152 567 152 567 158 571.5 c 128,-1,23
- 164 576 164 576 166 577 c 128,-1,24
- 168 578 168 578 182 586 c 0,25,26
- 245 616 245 616 298 655.5 c 128,-1,27
- 351 695 351 695 383 728.5 c 128,-1,28
- 415 762 415 762 443 805 c 128,-1,29
- 471 848 471 848 482.5 871.5 c 128,-1,30
- 494 895 494 895 508 928 c 0,31,32
- 512 936 512 936 514 938.5 c 128,-1,33
- 516 941 516 941 524 945.5 c 128,-1,34
- 532 950 532 950 544 950 c 0,35,36
- 558 950 558 950 567.5 943.5 c 128,-1,37
- 577 937 577 937 577 924 c 0,38,39
- 577 915 577 915 562.5 882 c 128,-1,40
- 548 849 548 849 521 802 c 128,-1,41
- 494 755 494 755 444.5 698.5 c 128,-1,42
- 395 642 395 642 337 599 c 1,43,-1
- 2045 599 l 2,44,45
- 2059 599 2059 599 2066 599 c 128,-1,46
- 2073 599 2073 599 2084.5 596.5 c 128,-1,47
- 2096 594 2096 594 2102 589.5 c 128,-1,48
- 2108 585 2108 585 2113 576.5 c 128,-1,49
- 2118 568 2118 568 2118 555 c 128,-1,50
- 2118 542 2118 542 2113 533.5 c 128,-1,51
- 2108 525 2108 525 2102 520.5 c 128,-1,52
- 2096 516 2096 516 2084.5 513.5 c 128,-1,53
- 2073 511 2073 511 2066 511 c 128,-1,54
- 2059 511 2059 511 2045 511 c 6,0,-1
-2402 770.5 m 132,-1,56
- 2402 720 2402 720 2382 676.5 c 128,-1,57
- 2362 633 2362 633 2329.5 603 c 128,-1,58
- 2297 573 2297 573 2255.5 551.5 c 128,-1,59
- 2214 530 2214 530 2171 520.5 c 128,-1,60
- 2128 511 2128 511 2087 511 c 0,61,62
- 2031 511 2031 511 2031 555 c 0,63,64
- 2031 558 2031 558 2032 563 c 128,-1,65
- 2033 568 2033 568 2037 577 c 128,-1,66
- 2041 586 2041 586 2051.5 592.5 c 128,-1,67
- 2062 599 2062 599 2078 599 c 0,68,69
- 2194 604 2194 604 2253.5 653.5 c 128,-1,70
- 2313 703 2313 703 2313 770 c 0,71,72
- 2313 798 2313 798 2302 824.5 c 128,-1,73
- 2291 851 2291 851 2266 877 c 128,-1,74
- 2241 903 2241 903 2192 920.5 c 128,-1,75
- 2143 938 2143 938 2076 941 c 0,76,77
- 2054 943 2054 943 2042.5 956 c 128,-1,78
- 2031 969 2031 969 2031 986 c 0,79,80
- 2031 1030 2031 1030 2087 1030 c 0,81,82
- 2128 1030 2128 1030 2171 1020 c 128,-1,83
- 2214 1010 2214 1010 2255.5 989 c 128,-1,84
- 2297 968 2297 968 2329.5 938 c 128,-1,85
- 2362 908 2362 908 2382 864.5 c 128,-1,55
- 2402 821 2402 821 2402 770.5 c 132,-1,56
-EndSplineSet
-Validated: 5
+2402 770.5 m 128,-1,1
+ 2402 720 2402 720 2382 676.5 c 128,-1,2
+ 2362 633 2362 633 2329.5 603 c 128,-1,3
+ 2297 573 2297 573 2256 552 c 0,4,5
+ 2211 528 2211 528 2171 520 c 0,6,7
+ 2126 511 2126 511 2087 511 c 0,8,9
+ 2080 511 2080 511 2074 512 c 1,10,11
+ 2070 511 2070 511 2066 511 c 2,12,-1
+ 2045 511 l 1,13,-1
+ 337 511 l 1,14,15
+ 395 469 395 469 444.5 412 c 128,-1,16
+ 494 355 494 355 521 308 c 128,-1,17
+ 548 261 548 261 562.5 228 c 128,-1,18
+ 577 195 577 195 577 186 c 0,19,20
+ 577 172 577 172 567 166 c 128,-1,21
+ 557 160 557 160 544 160 c 0,22,23
+ 527 160 527 160 520.5 166.5 c 128,-1,24
+ 514 173 514 173 506 191 c 0,25,26
+ 455 301 455 301 380 383.5 c 128,-1,27
+ 305 466 305 466 178 526 c 0,28,29
+ 172 529 172 529 167 532 c 128,-1,30
+ 162 535 162 535 158.5 538 c 128,-1,31
+ 155 541 155 541 153 545 c 128,-1,32
+ 151 549 151 549 151 555 c 128,-1,33
+ 151 561 151 561 151.5 564 c 128,-1,34
+ 152 567 152 567 158 571.5 c 128,-1,35
+ 164 576 164 576 166 577 c 128,-1,36
+ 168 578 168 578 182 586 c 1,37,38
+ 245 616 245 616 298 655.5 c 128,-1,39
+ 351 695 351 695 383 728 c 0,40,41
+ 415 762 415 762 443 805 c 128,-1,42
+ 471 848 471 848 482.5 871.5 c 128,-1,43
+ 494 895 494 895 508 928 c 0,44,45
+ 512 936 512 936 514 938.5 c 128,-1,46
+ 516 941 516 941 524 945.5 c 128,-1,47
+ 532 950 532 950 544 950 c 0,48,49
+ 558 950 558 950 567.5 943.5 c 128,-1,50
+ 577 937 577 937 577 924 c 0,51,52
+ 577 915 577 915 562.5 882 c 128,-1,53
+ 548 849 548 849 521 802 c 128,-1,54
+ 494 755 494 755 444.5 698.5 c 128,-1,55
+ 395 642 395 642 337 599 c 1,56,-1
+ 2045 599 l 1,57,-1
+ 2066 599 l 2,58,59
+ 2068 599 2068 599 2071 599 c 128,-1,60
+ 2074 599 2074 599 2078 599 c 0,61,62
+ 2188 599 2188 599 2254 654 c 0,63,64
+ 2313 703 2313 703 2313 770 c 0,65,66
+ 2313 798 2313 798 2302 824.5 c 128,-1,67
+ 2291 851 2291 851 2266 877 c 128,-1,68
+ 2241 903 2241 903 2192 920.5 c 128,-1,69
+ 2143 938 2143 938 2076 941 c 1,70,71
+ 2054 943 2054 943 2042 956 c 0,72,73
+ 2031 969 2031 969 2031 986 c 0,74,75
+ 2031 1030 2031 1030 2087 1030 c 0,76,77
+ 2128 1030 2128 1030 2171 1020 c 0,78,79
+ 2216 1010 2216 1010 2256 989 c 0,80,81
+ 2297 968 2297 968 2330 938 c 0,82,83
+ 2362 908 2362 908 2382 864.5 c 128,-1,0
+ 2402 821 2402 821 2402 770.5 c 128,-1,1
+EndSplineSet
+Validated: 1
 EndChar
 
 StartChar: uni21AA
@@ -28901,78 +28890,69 @@
 LayerCount: 2
 Fore
 SplineSet
-2211 511 m 5,0,-1
- 504 511 l 2,1,2
- 490 511 490 511 483 511 c 128,-1,3
- 476 511 476 511 464.5 513.5 c 128,-1,4
- 453 516 453 516 446.5 520.5 c 128,-1,5
- 440 525 440 525 435.5 533.5 c 128,-1,6
- 431 542 431 542 431 555 c 128,-1,7
- 431 568 431 568 435.5 576.5 c 128,-1,8
- 440 585 440 585 446.5 589.5 c 128,-1,9
- 453 594 453 594 464.5 596.5 c 128,-1,10
- 476 599 476 599 483 599 c 128,-1,11
- 490 599 490 599 504 599 c 2,12,-1
- 2211 599 l 1,13,14
- 2153 641 2153 641 2104 698 c 128,-1,15
- 2055 755 2055 755 2028 802 c 128,-1,16
- 2001 849 2001 849 1986 882 c 128,-1,17
- 1971 915 1971 915 1971 924 c 0,18,19
- 1971 938 1971 938 1981 944 c 128,-1,20
- 1991 950 1991 950 2005 950 c 0,21,22
- 2022 950 2022 950 2028 943.5 c 128,-1,23
- 2034 937 2034 937 2042 919 c 0,24,25
- 2075 846 2075 846 2115.5 788.5 c 128,-1,26
- 2156 731 2156 731 2202.5 691.5 c 128,-1,27
- 2249 652 2249 652 2290 626.5 c 128,-1,28
- 2331 601 2331 601 2382 577 c 1,29,30
- 2385 577 2385 577 2391.5 571 c 128,-1,31
- 2398 565 2398 565 2398 555 c 0,32,33
- 2398 544 2398 544 2393 539 c 128,-1,34
- 2388 534 2388 534 2367 524 c 0,35,36
- 2321 502 2321 502 2280.5 475.5 c 128,-1,37
- 2240 449 2240 449 2211 424.5 c 128,-1,38
- 2182 400 2182 400 2155.5 369 c 128,-1,39
- 2129 338 2129 338 2113 316 c 128,-1,40
- 2097 294 2097 294 2080.5 264 c 128,-1,41
- 2064 234 2064 234 2057.5 220 c 128,-1,42
- 2051 206 2051 206 2040 182 c 0,43,44
- 2036 174 2036 174 2034.5 171.5 c 128,-1,45
- 2033 169 2033 169 2025 164.5 c 128,-1,46
- 2017 160 2017 160 2005 160 c 0,47,48
- 1991 160 1991 160 1981 166.5 c 128,-1,49
- 1971 173 1971 173 1971 186 c 0,50,51
- 1971 195 1971 195 1986 228 c 128,-1,52
- 2001 261 2001 261 2028 308 c 128,-1,53
- 2055 355 2055 355 2104 412 c 128,-1,54
- 2153 469 2153 469 2211 511 c 5,0,-1
-526 555 m 4,55,56
- 526 511 526 511 471 511 c 0,57,58
- 430 511 430 511 387 520.5 c 128,-1,59
- 344 530 344 530 302.5 551.5 c 128,-1,60
- 261 573 261 573 228 603 c 128,-1,61
- 195 633 195 633 175 676.5 c 128,-1,62
- 155 720 155 720 155 770.5 c 128,-1,63
- 155 821 155 821 175 864.5 c 128,-1,64
- 195 908 195 908 228 938 c 128,-1,65
- 261 968 261 968 302.5 989 c 128,-1,66
- 344 1010 344 1010 387 1020 c 128,-1,67
- 430 1030 430 1030 471 1030 c 0,68,69
- 527 1030 527 1030 526 986 c 0,70,71
- 526 983 526 983 525 978 c 128,-1,72
- 524 973 524 973 520.5 964 c 128,-1,73
- 517 955 517 955 506.5 948.5 c 128,-1,74
- 496 942 496 942 480 941 c 0,75,76
- 364 936 364 936 304 887 c 128,-1,77
- 244 838 244 838 244 770 c 0,78,79
- 244 742 244 742 255.5 716 c 128,-1,80
- 267 690 267 690 292 663.5 c 128,-1,81
- 317 637 317 637 366 620 c 128,-1,82
- 415 603 415 603 482 599 c 0,83,84
- 504 597 504 597 515 584.5 c 128,-1,85
- 526 572 526 572 526 555 c 4,55,56
-EndSplineSet
-Validated: 37
+2211 511 m 1,0,-1
+ 504 511 l 1,1,-1
+ 483 511 l 2,2,3
+ 481 511 481 511 478 511 c 128,-1,4
+ 475 511 475 511 471 511 c 0,5,6
+ 430 511 430 511 387 520.5 c 128,-1,7
+ 344 530 344 530 302.5 551.5 c 128,-1,8
+ 261 573 261 573 228 603 c 128,-1,9
+ 195 633 195 633 175 676.5 c 128,-1,10
+ 155 720 155 720 155 770.5 c 128,-1,11
+ 155 821 155 821 175 864.5 c 128,-1,12
+ 195 908 195 908 228 938 c 128,-1,13
+ 261 968 261 968 302.5 989 c 128,-1,14
+ 344 1010 344 1010 387 1020 c 128,-1,15
+ 430 1030 430 1030 471 1030 c 0,16,17
+ 527 1030 527 1030 526 986 c 0,18,19
+ 526 983 526 983 525 978 c 128,-1,20
+ 524 973 524 973 520.5 964 c 128,-1,21
+ 517 955 517 955 506.5 948.5 c 128,-1,22
+ 496 942 496 942 480 941 c 0,23,24
+ 364 936 364 936 304 887 c 128,-1,25
+ 244 838 244 838 244 770 c 0,26,27
+ 244 742 244 742 255.5 716 c 128,-1,28
+ 267 690 267 690 292 663.5 c 128,-1,29
+ 317 637 317 637 366 620 c 128,-1,30
+ 415 603 415 603 482 599 c 0,31,32
+ 483 599 483 599 483 599 c 2,33,-1
+ 504 599 l 1,34,-1
+ 2211 599 l 1,35,36
+ 2153 641 2153 641 2104 698 c 128,-1,37
+ 2055 755 2055 755 2028 802 c 128,-1,38
+ 2001 849 2001 849 1986 882 c 128,-1,39
+ 1971 915 1971 915 1971 924 c 0,40,41
+ 1971 938 1971 938 1981 944 c 128,-1,42
+ 1991 950 1991 950 2005 950 c 0,43,44
+ 2022 950 2022 950 2028 944 c 0,45,46
+ 2034 937 2034 937 2042 919 c 0,47,48
+ 2073 848 2073 848 2116 788 c 0,49,50
+ 2156 731 2156 731 2202.5 691.5 c 128,-1,51
+ 2249 652 2249 652 2290 626.5 c 128,-1,52
+ 2331 601 2331 601 2382 577 c 1,53,54
+ 2385 577 2385 577 2391.5 571 c 128,-1,55
+ 2398 565 2398 565 2398 555 c 0,56,57
+ 2398 544 2398 544 2393 539 c 128,-1,58
+ 2388 534 2388 534 2367 524 c 0,59,60
+ 2321 502 2321 502 2280.5 475.5 c 128,-1,61
+ 2240 449 2240 449 2211 424.5 c 128,-1,62
+ 2182 400 2182 400 2155.5 369 c 128,-1,63
+ 2129 338 2129 338 2113 316 c 128,-1,64
+ 2097 294 2097 294 2080.5 264 c 128,-1,65
+ 2064 234 2064 234 2058 220 c 2,66,-1
+ 2040 182 l 2,67,68
+ 2036 174 2036 174 2034.5 171.5 c 128,-1,69
+ 2033 169 2033 169 2025 164.5 c 128,-1,70
+ 2017 160 2017 160 2005 160 c 0,71,72
+ 1991 160 1991 160 1981 166.5 c 128,-1,73
+ 1971 173 1971 173 1971 186 c 0,74,75
+ 1971 195 1971 195 1986 228 c 128,-1,76
+ 2001 261 2001 261 2028 308 c 128,-1,77
+ 2055 355 2055 355 2104 412 c 128,-1,78
+ 2153 469 2153 469 2211 511 c 1,0,-1
+EndSplineSet
+Validated: 33
 EndChar
 
 StartChar: carriagereturn
@@ -30183,65 +30163,78 @@
 LayerCount: 2
 Fore
 SplineSet
-1221 511 m 2,0,-1
- 275 511 l 1,1,2
- 288 365 288 365 370.5 247.5 c 128,-1,3
- 453 130 453 130 583.5 65 c 128,-1,4
- 714 0 714 0 868 0 c 2,5,-1
- 1221 0 l 2,6,7
- 1237 0 1237 0 1246 -1 c 128,-1,8
- 1255 -2 1255 -2 1268.5 -5.5 c 128,-1,9
- 1282 -9 1282 -9 1288 -19 c 128,-1,10
- 1294 -29 1294 -29 1294 -44.5 c 128,-1,11
- 1294 -60 1294 -60 1288 -70 c 128,-1,12
- 1282 -80 1282 -80 1268.5 -83.5 c 128,-1,13
- 1255 -87 1255 -87 1246 -88 c 128,-1,14
- 1237 -89 1237 -89 1221 -89 c 2,15,-1
- 861 -89 l 2,16,17
- 727 -89 727 -89 603 -39.5 c 128,-1,18
- 479 10 479 10 387 94.5 c 128,-1,19
- 295 179 295 179 239.5 299.5 c 128,-1,20
- 184 420 184 420 184 555 c 0,21,22
- 184 691 184 691 240 812 c 128,-1,23
- 296 933 296 933 388.5 1017 c 128,-1,24
- 481 1101 481 1101 604 1150 c 128,-1,25
- 727 1199 727 1199 859 1199 c 2,26,-1
- 1221 1199 l 2,27,28
- 1237 1199 1237 1199 1246 1198 c 128,-1,29
- 1255 1197 1255 1197 1268.5 1193.5 c 128,-1,30
- 1282 1190 1282 1190 1288 1180 c 128,-1,31
- 1294 1170 1294 1170 1294 1154.5 c 128,-1,32
- 1294 1139 1294 1139 1288 1129 c 128,-1,33
- 1282 1119 1282 1119 1268.5 1115.5 c 128,-1,34
- 1255 1112 1255 1112 1246 1111 c 128,-1,35
- 1237 1110 1237 1110 1221 1110 c 2,36,-1
- 866 1110 l 2,37,38
- 754 1110 754 1110 650.5 1072.5 c 128,-1,39
- 547 1035 547 1035 467.5 968.5 c 128,-1,40
- 388 902 388 902 336.5 806 c 128,-1,41
- 285 710 285 710 275 599 c 1,42,-1
- 1221 599 l 2,43,44
- 1237 599 1237 599 1246 598.5 c 128,-1,45
- 1255 598 1255 598 1268.5 594 c 128,-1,46
- 1282 590 1282 590 1288 580.5 c 128,-1,47
- 1294 571 1294 571 1294 555 c 128,-1,48
- 1294 539 1294 539 1288 529.5 c 128,-1,49
- 1282 520 1282 520 1268.5 516 c 128,-1,50
- 1255 512 1255 512 1246 511.5 c 128,-1,51
- 1237 511 1237 511 1221 511 c 2,0,-1
-1281 1505 m 2,52,-1
- 282 -440 l 2,53,54
- 262 -480 262 -480 233 -480 c 0,55,56
- 212 -480 212 -480 200.5 -465 c 128,-1,57
- 189 -450 189 -450 189 -435 c 0,58,59
- 189 -426 189 -426 204 -395 c 2,60,-1
+715 599 m 1,0,-1
+ 977 1110 l 1,1,-1
+ 866 1110 l 2,2,3
+ 754 1110 754 1110 650.5 1072.5 c 128,-1,4
+ 547 1035 547 1035 467.5 968.5 c 128,-1,5
+ 388 902 388 902 336.5 806 c 128,-1,6
+ 285 710 285 710 275 599 c 1,7,-1
+ 715 599 l 1,0,-1
+816 599 m 1,8,-1
+ 1221 599 l 2,9,10
+ 1237 599 1237 599 1246 598.5 c 128,-1,11
+ 1255 598 1255 598 1268.5 594 c 128,-1,12
+ 1282 590 1282 590 1288 580.5 c 128,-1,13
+ 1294 571 1294 571 1294 555 c 128,-1,14
+ 1294 539 1294 539 1288 529.5 c 128,-1,15
+ 1282 520 1282 520 1268 516 c 0,16,17
+ 1255 512 1255 512 1246 511.5 c 128,-1,18
+ 1237 511 1237 511 1221 511 c 2,19,-1
+ 770 511 l 1,20,-1
+ 550 83 l 1,21,22
+ 567 73 567 73 584 65 c 0,23,24
+ 718 0 718 0 868 0 c 2,25,-1
+ 1221 0 l 2,26,27
+ 1237 0 1237 0 1246 -1 c 128,-1,28
+ 1255 -2 1255 -2 1268.5 -5.5 c 128,-1,29
+ 1282 -9 1282 -9 1288 -19 c 128,-1,30
+ 1294 -29 1294 -29 1294 -44.5 c 128,-1,31
+ 1294 -60 1294 -60 1288 -70 c 128,-1,32
+ 1282 -80 1282 -80 1268.5 -83.5 c 128,-1,33
+ 1255 -87 1255 -87 1246 -88 c 128,-1,34
+ 1237 -89 1237 -89 1221 -89 c 2,35,-1
+ 861 -89 l 2,36,37
+ 723 -89 723 -89 603 -40 c 0,38,39
+ 554 -20 554 -20 510 5 c 1,40,-1
+ 282 -440 l 2,41,42
+ 262 -480 262 -480 233 -480 c 0,43,44
+ 212 -480 212 -480 200.5 -465 c 128,-1,45
+ 189 -450 189 -450 189 -435 c 0,46,47
+ 189 -426 189 -426 204 -395 c 2,48,-1
+ 435 55 l 1,49,50
+ 407 75 407 75 387 94 c 0,51,52
+ 294 182 294 182 240 300 c 0,53,54
+ 184 419 184 419 184 555 c 128,-1,55
+ 184 691 184 691 240 812 c 128,-1,56
+ 296 933 296 933 388.5 1017 c 128,-1,57
+ 481 1101 481 1101 604 1150 c 128,-1,58
+ 727 1199 727 1199 859 1199 c 2,59,-1
+ 1023 1199 l 1,60,-1
  1203 1550 l 2,61,62
  1223 1590 1223 1590 1252 1590 c 0,63,64
  1273 1590 1273 1590 1284.5 1575 c 128,-1,65
- 1296 1560 1296 1560 1296 1545 c 0,66,67
- 1297 1536 1297 1536 1281 1505 c 2,52,-1
-EndSplineSet
-Validated: 5
+ 1296 1560 1296 1560 1296 1545 c 1,66,67
+ 1297 1536 1297 1536 1281 1505 c 2,68,-1
+ 1124 1199 l 1,69,-1
+ 1221 1199 l 2,70,71
+ 1237 1199 1237 1199 1246 1198 c 128,-1,72
+ 1255 1197 1255 1197 1268.5 1193.5 c 128,-1,73
+ 1282 1190 1282 1190 1288 1180 c 128,-1,74
+ 1294 1170 1294 1170 1294 1154.5 c 128,-1,75
+ 1294 1139 1294 1139 1288 1129 c 128,-1,76
+ 1282 1119 1282 1119 1268.5 1115.5 c 128,-1,77
+ 1255 1112 1255 1112 1246 1111 c 128,-1,78
+ 1237 1110 1237 1110 1221 1110 c 2,79,-1
+ 1078 1110 l 1,80,-1
+ 816 599 l 1,8,-1
+476 134 m 1,81,-1
+ 669 511 l 1,82,-1
+ 275 511 l 1,83,84
+ 288 364 288 364 370 248 c 0,85,86
+ 416 183 416 183 476 134 c 1,81,-1
+EndSplineSet
+Validated: 1
 EndChar
 
 StartChar: uni220D
@@ -32320,56 +32313,62 @@
 LayerCount: 2
 Fore
 SplineSet
-1401 1505 m 2,0,-1
- 402 -440 l 2,1,2
- 382 -480 382 -480 353 -480 c 0,3,4
- 332 -480 332 -480 320.5 -465 c 128,-1,5
- 309 -450 309 -450 309 -435 c 0,6,7
- 309 -426 309 -426 324 -395 c 2,8,-1
- 1323 1550 l 2,9,10
- 1343 1590 1343 1590 1372 1590 c 0,11,12
- 1393 1590 1393 1590 1404.5 1575 c 128,-1,13
- 1416 1560 1416 1560 1416 1545 c 0,14,15
- 1417 1536 1417 1536 1401 1505 c 2,0,-1
-1467 1110 m 2,16,-1
- 866 1110 l 2,17,18
- 743 1110 743 1110 633 1066.5 c 128,-1,19
- 523 1023 523 1023 444 948.5 c 128,-1,20
- 365 874 365 874 319 771.5 c 128,-1,21
- 273 669 273 669 273 555 c 128,-1,22
- 273 441 273 441 319.5 338.5 c 128,-1,23
- 366 236 366 236 445 161.5 c 128,-1,24
- 524 87 524 87 634 43.5 c 128,-1,25
- 744 0 744 0 866 0 c 2,26,-1
- 1467 0 l 2,27,28
- 1483 0 1483 0 1492.5 -1 c 128,-1,29
- 1502 -2 1502 -2 1515 -5.5 c 128,-1,30
- 1528 -9 1528 -9 1534.5 -19 c 128,-1,31
- 1541 -29 1541 -29 1541 -44.5 c 128,-1,32
- 1541 -60 1541 -60 1534.5 -70 c 128,-1,33
- 1528 -80 1528 -80 1515 -83.5 c 128,-1,34
- 1502 -87 1502 -87 1493 -88 c 128,-1,35
- 1484 -89 1484 -89 1467 -89 c 2,36,-1
- 859 -89 l 2,37,38
- 727 -89 727 -89 604 -40 c 128,-1,39
- 481 9 481 9 388.5 93 c 128,-1,40
- 296 177 296 177 240 298 c 128,-1,41
- 184 419 184 419 184 555 c 128,-1,42
- 184 691 184 691 240 812 c 128,-1,43
- 296 933 296 933 388.5 1017 c 128,-1,44
- 481 1101 481 1101 604 1150 c 128,-1,45
- 727 1199 727 1199 859 1199 c 2,46,-1
- 1467 1199 l 2,47,48
- 1483 1199 1483 1199 1492.5 1198 c 128,-1,49
- 1502 1197 1502 1197 1515 1193.5 c 128,-1,50
- 1528 1190 1528 1190 1534.5 1180 c 128,-1,51
- 1541 1170 1541 1170 1541 1154.5 c 128,-1,52
- 1541 1139 1541 1139 1534.5 1129 c 128,-1,53
- 1528 1119 1528 1119 1515 1115.5 c 128,-1,54
- 1502 1112 1502 1112 1492.5 1111 c 128,-1,55
- 1483 1110 1483 1110 1467 1110 c 2,16,-1
-EndSplineSet
-Validated: 5
+1143 1199 m 1,0,-1
+ 1323 1550 l 2,1,2
+ 1343 1590 1343 1590 1372 1590 c 0,3,4
+ 1393 1590 1393 1590 1404.5 1575 c 128,-1,5
+ 1416 1560 1416 1560 1416 1545 c 1,6,7
+ 1417 1536 1417 1536 1401 1505 c 2,8,-1
+ 1244 1199 l 1,9,-1
+ 1467 1199 l 2,10,11
+ 1483 1199 1483 1199 1492.5 1198 c 128,-1,12
+ 1502 1197 1502 1197 1515 1193.5 c 128,-1,13
+ 1528 1190 1528 1190 1534.5 1180 c 128,-1,14
+ 1541 1170 1541 1170 1541 1154.5 c 128,-1,15
+ 1541 1139 1541 1139 1534.5 1129 c 128,-1,16
+ 1528 1119 1528 1119 1515 1115.5 c 128,-1,17
+ 1502 1112 1502 1112 1492.5 1111 c 128,-1,18
+ 1483 1110 1483 1110 1467 1110 c 2,19,-1
+ 1198 1110 l 1,20,-1
+ 648 38 l 1,21,22
+ 752 0 752 0 866 0 c 2,23,-1
+ 1467 0 l 2,24,25
+ 1483 0 1483 0 1492.5 -1 c 128,-1,26
+ 1502 -2 1502 -2 1515 -5.5 c 128,-1,27
+ 1528 -9 1528 -9 1534.5 -19 c 128,-1,28
+ 1541 -29 1541 -29 1541 -44.5 c 128,-1,29
+ 1541 -60 1541 -60 1534.5 -70 c 128,-1,30
+ 1528 -80 1528 -80 1515 -83.5 c 128,-1,31
+ 1502 -87 1502 -87 1493 -88 c 128,-1,32
+ 1484 -89 1484 -89 1467 -89 c 2,33,-1
+ 859 -89 l 2,34,35
+ 729 -89 729 -89 607 -41 c 1,36,-1
+ 402 -440 l 2,37,38
+ 382 -480 382 -480 353 -480 c 0,39,40
+ 332 -480 332 -480 320.5 -465 c 128,-1,41
+ 309 -450 309 -450 309 -435 c 0,42,43
+ 309 -426 309 -426 324 -395 c 2,44,-1
+ 525 -3 l 1,45,46
+ 450 37 450 37 388 93 c 0,47,48
+ 296 177 296 177 240 298 c 128,-1,49
+ 184 419 184 419 184 555 c 128,-1,50
+ 184 691 184 691 240 812 c 128,-1,51
+ 296 933 296 933 388.5 1017 c 128,-1,52
+ 481 1101 481 1101 604 1150 c 128,-1,53
+ 727 1199 727 1199 859 1199 c 2,54,-1
+ 1143 1199 l 1,0,-1
+566 75 m 1,55,-1
+ 1097 1110 l 1,56,-1
+ 866 1110 l 2,57,58
+ 743 1110 743 1110 633 1066.5 c 128,-1,59
+ 523 1023 523 1023 444 948.5 c 128,-1,60
+ 365 874 365 874 319 771.5 c 128,-1,61
+ 273 669 273 669 273 555 c 0,62,63
+ 273 443 273 443 320 338 c 0,64,65
+ 365 237 365 237 445 162 c 0,66,67
+ 499 110 499 110 566 75 c 1,55,-1
+EndSplineSet
+Validated: 1
 EndChar
 
 StartChar: reflexsubset
@@ -33613,77 +33612,60 @@
 Width: 1971
 VWidth: 2220
 Flags: W
-HStem: 295 89<294 1679> 726 89<294 1679>
-LayerCount: 2
-Fore
-SplineSet
-1657 726 m 6,0,-1
- 332 726 l 2,1,2
- 318 726 318 726 310.5 726.5 c 128,-1,3
- 303 727 303 727 291 729 c 128,-1,4
- 279 731 279 731 273 735.5 c 128,-1,5
- 267 740 267 740 261.5 749 c 128,-1,6
- 256 758 256 758 256 770 c 0,7,8
- 256 786 256 786 262.5 796 c 128,-1,9
- 269 806 269 806 282 809.5 c 128,-1,10
- 295 813 295 813 304 814 c 128,-1,11
- 313 815 313 815 330 815 c 2,12,-1
- 1659 815 l 2,13,14
- 1675 815 1675 815 1684.5 814 c 128,-1,15
- 1694 813 1694 813 1707 809.5 c 128,-1,16
- 1720 806 1720 806 1726.5 796 c 128,-1,17
- 1733 786 1733 786 1733 770 c 0,18,19
- 1733 757 1733 757 1728 748.5 c 128,-1,20
- 1723 740 1723 740 1716.5 735.5 c 128,-1,21
- 1710 731 1710 731 1698 729 c 128,-1,22
- 1686 727 1686 727 1678.5 726.5 c 128,-1,23
- 1671 726 1671 726 1657 726 c 6,0,-1
-1659 295 m 6,24,-1
- 330 295 l 2,25,26
- 314 295 314 295 304.5 296 c 128,-1,27
- 295 297 295 297 282 300.5 c 128,-1,28
- 269 304 269 304 262.5 314 c 128,-1,29
- 256 324 256 324 256 340 c 0,30,31
- 256 353 256 353 261.5 361.5 c 128,-1,32
- 267 370 267 370 273 374.5 c 128,-1,33
- 279 379 279 379 291 381 c 128,-1,34
- 303 383 303 383 310.5 383.5 c 128,-1,35
- 318 384 318 384 332 384 c 2,36,-1
- 1657 384 l 2,37,38
- 1671 384 1671 384 1678.5 383.5 c 128,-1,39
- 1686 383 1686 383 1698 381 c 128,-1,40
- 1710 379 1710 379 1716.5 374.5 c 128,-1,41
- 1723 370 1723 370 1728 361 c 128,-1,42
- 1733 352 1733 352 1733 340 c 0,43,44
- 1733 324 1733 324 1726.5 314 c 128,-1,45
- 1720 304 1720 304 1707 300.5 c 128,-1,46
- 1694 297 1694 297 1684.5 296 c 128,-1,47
- 1675 295 1675 295 1659 295 c 6,24,-1
-374 1567 m 2,48,-1
- 374 -457 l 2,49,50
- 374 -472 374 -472 373.5 -481 c 128,-1,51
- 373 -490 373 -490 369.5 -505.5 c 128,-1,52
- 366 -521 366 -521 359.5 -530.5 c 128,-1,53
- 353 -540 353 -540 339 -547.5 c 128,-1,54
- 325 -555 325 -555 305 -555 c 0,55,56
- 286 -555 286 -555 272.5 -547.5 c 128,-1,57
- 259 -540 259 -540 252.5 -530.5 c 128,-1,58
- 246 -521 246 -521 243 -505 c 128,-1,59
- 240 -489 240 -489 239.5 -480 c 128,-1,60
- 239 -471 239 -471 239 -457 c 2,61,-1
- 239 1567 l 2,62,63
- 239 1582 239 1582 239.5 1591 c 128,-1,64
- 240 1600 240 1600 243.5 1615.5 c 128,-1,65
- 247 1631 247 1631 253.5 1640.5 c 128,-1,66
- 260 1650 260 1650 274 1657.5 c 128,-1,67
- 288 1665 288 1665 307 1665 c 0,68,69
- 326 1665 326 1665 339.5 1657.5 c 128,-1,70
- 353 1650 353 1650 359.5 1640.5 c 128,-1,71
- 366 1631 366 1631 369.5 1615 c 128,-1,72
- 373 1599 373 1599 373.5 1590 c 128,-1,73
- 374 1581 374 1581 374 1567 c 2,48,-1
-EndSplineSet
-Validated: 5
+HStem: 295 89<374 1726.3> 726 89<374 1726.3>
+VStem: 239 135<-551.614 295 384 726 815 1661.61>
+LayerCount: 2
+Fore
+SplineSet
+374 815 m 1,0,-1
+ 1659 815 l 2,1,2
+ 1675 815 1675 815 1684.5 814 c 128,-1,3
+ 1694 813 1694 813 1707 809.5 c 128,-1,4
+ 1720 806 1720 806 1726.5 796 c 128,-1,5
+ 1733 786 1733 786 1733 770 c 0,6,7
+ 1733 757 1733 757 1728 748.5 c 128,-1,8
+ 1723 740 1723 740 1716.5 735.5 c 128,-1,9
+ 1710 731 1710 731 1698 729 c 128,-1,10
+ 1686 727 1686 727 1678.5 726.5 c 128,-1,11
+ 1671 726 1671 726 1657 726 c 2,12,-1
+ 374 726 l 1,13,-1
+ 374 384 l 1,14,-1
+ 1657 384 l 2,15,16
+ 1671 384 1671 384 1678.5 383.5 c 128,-1,17
+ 1686 383 1686 383 1698 381 c 128,-1,18
+ 1710 379 1710 379 1716.5 374.5 c 128,-1,19
+ 1723 370 1723 370 1728 361 c 128,-1,20
+ 1733 352 1733 352 1733 340 c 0,21,22
+ 1733 324 1733 324 1726.5 314 c 128,-1,23
+ 1720 304 1720 304 1707 300.5 c 128,-1,24
+ 1694 297 1694 297 1684.5 296 c 128,-1,25
+ 1675 295 1675 295 1659 295 c 2,26,-1
+ 374 295 l 1,27,-1
+ 374 -457 l 2,28,29
+ 374 -472 374 -472 373.5 -481 c 128,-1,30
+ 373 -490 373 -490 369.5 -505.5 c 128,-1,31
+ 366 -521 366 -521 359.5 -530.5 c 128,-1,32
+ 353 -540 353 -540 339 -547.5 c 128,-1,33
+ 325 -555 325 -555 305 -555 c 0,34,35
+ 286 -555 286 -555 272.5 -547.5 c 128,-1,36
+ 259 -540 259 -540 252.5 -530.5 c 128,-1,37
+ 246 -521 246 -521 243 -505 c 128,-1,38
+ 240 -489 240 -489 239.5 -480 c 128,-1,39
+ 239 -471 239 -471 239 -457 c 2,40,-1
+ 239 1567 l 2,41,42
+ 239 1582 239 1582 239.5 1591 c 128,-1,43
+ 240 1600 240 1600 243.5 1615.5 c 128,-1,44
+ 247 1631 247 1631 254 1640 c 0,45,46
+ 260 1650 260 1650 274 1657.5 c 128,-1,47
+ 288 1665 288 1665 307 1665 c 128,-1,48
+ 326 1665 326 1665 339.5 1657.5 c 128,-1,49
+ 353 1650 353 1650 359.5 1640.5 c 128,-1,50
+ 366 1631 366 1631 369.5 1615 c 128,-1,51
+ 373 1599 373 1599 373.5 1590 c 128,-1,52
+ 374 1581 374 1581 374 1567 c 2,53,-1
+ 374 815 l 1,0,-1
+EndSplineSet
+Validated: 1
 EndChar
 
 StartChar: uni22A9
@@ -33691,73 +33673,67 @@
 Width: 1445
 VWidth: 2220
 Flags: W
-HStem: 0 47G<147 191 369 413> 726 89<213 435 435 435 435 1431> 1496 44G<147 191 369 413>
-VStem: 124 89<22 726 815 1519> 346 89<22 726 815 1519>
-LayerCount: 2
-Fore
-SplineSet
-1157 726 m 6,0,-1
- 213 726 l 1,1,-1
- 213 73 l 2,2,3
- 213 57 213 57 212 48 c 128,-1,4
- 211 39 211 39 207.5 26 c 128,-1,5
- 204 13 204 13 194 6.5 c 128,-1,6
- 184 0 184 0 168.5 0 c 128,-1,7
- 153 0 153 0 143 6.5 c 128,-1,8
- 133 13 133 13 129.5 26.5 c 128,-1,9
- 126 40 126 40 125 49.5 c 128,-1,10
- 124 59 124 59 124 75 c 2,11,-1
- 124 1465 l 2,12,13
- 124 1481 124 1481 125 1490.5 c 128,-1,14
- 126 1500 126 1500 129.5 1513.5 c 128,-1,15
- 133 1527 133 1527 143 1534 c 128,-1,16
- 153 1541 153 1541 168.5 1541 c 128,-1,17
- 184 1541 184 1541 194 1534.5 c 128,-1,18
- 204 1528 204 1528 207.5 1515 c 128,-1,19
- 211 1502 211 1502 212 1493 c 128,-1,20
- 213 1484 213 1484 213 1467 c 2,21,-1
- 213 815 l 1,22,-1
- 1157 815 l 2,23,24
- 1173 815 1173 815 1182 814 c 128,-1,25
- 1191 813 1191 813 1204 809.5 c 128,-1,26
- 1217 806 1217 806 1223.5 796 c 128,-1,27
- 1230 786 1230 786 1230 770.5 c 128,-1,28
- 1230 755 1230 755 1223.5 745 c 128,-1,29
- 1217 735 1217 735 1204 731.5 c 128,-1,30
- 1191 728 1191 728 1182 727 c 128,-1,31
- 1173 726 1173 726 1157 726 c 6,0,-1
-1379 726 m 2,32,-1
- 435 726 l 1,33,-1
- 435 73 l 2,34,35
- 435 57 435 57 434 48 c 128,-1,36
- 433 39 433 39 429.5 26 c 128,-1,37
- 426 13 426 13 416 6.5 c 128,-1,38
- 406 0 406 0 390.5 0 c 128,-1,39
- 375 0 375 0 365 6.5 c 128,-1,40
- 355 13 355 13 351.5 26.5 c 128,-1,41
- 348 40 348 40 347 49.5 c 128,-1,42
- 346 59 346 59 346 75 c 2,43,-1
- 346 1465 l 2,44,45
- 346 1481 346 1481 347 1490.5 c 128,-1,46
- 348 1500 348 1500 351.5 1513.5 c 128,-1,47
- 355 1527 355 1527 365 1534 c 128,-1,48
- 375 1541 375 1541 390.5 1541 c 128,-1,49
- 406 1541 406 1541 416 1534.5 c 128,-1,50
- 426 1528 426 1528 429.5 1515 c 128,-1,51
- 433 1502 433 1502 434 1493 c 128,-1,52
- 435 1484 435 1484 435 1467 c 2,53,-1
- 435 815 l 1,54,-1
- 1379 815 l 2,55,56
- 1395 815 1395 815 1404 814 c 128,-1,57
- 1413 813 1413 813 1426 809.5 c 128,-1,58
- 1439 806 1439 806 1445.5 796 c 128,-1,59
- 1452 786 1452 786 1452 770.5 c 128,-1,60
- 1452 755 1452 755 1445.5 745 c 128,-1,61
- 1439 735 1439 735 1426 731.5 c 128,-1,62
- 1413 728 1413 728 1404 727 c 128,-1,63
- 1395 726 1395 726 1379 726 c 2,32,-1
-EndSplineSet
-Validated: 5
+HStem: 0 21G<160.75 176.25 382.75 398.25> 726 89<213 346 435 1452> 1521 20G<160.75 176.25 382.75 398.25>
+VStem: 124 89<6.70483 726 815 1533.78> 346 89<6.70483 726 815 1533.78>
+LayerCount: 2
+Fore
+SplineSet
+1157 815 m 1,0,-1
+ 1379 815 l 2,1,2
+ 1395 815 1395 815 1404 814 c 128,-1,3
+ 1413 813 1413 813 1426 809.5 c 128,-1,4
+ 1439 806 1439 806 1445.5 796 c 128,-1,5
+ 1452 786 1452 786 1452 770.5 c 128,-1,6
+ 1452 755 1452 755 1445.5 745 c 128,-1,7
+ 1439 735 1439 735 1426 731.5 c 128,-1,8
+ 1413 728 1413 728 1404 727 c 128,-1,9
+ 1395 726 1395 726 1379 726 c 2,10,-1
+ 1157 726 l 1,11,-1
+ 435 726 l 1,12,-1
+ 435 73 l 2,13,14
+ 435 57 435 57 434 48 c 0,15,16
+ 432 34 432 34 430 26 c 0,17,18
+ 426 13 426 13 416 6 c 0,19,20
+ 406 0 406 0 390.5 0 c 128,-1,21
+ 375 0 375 0 365 6 c 0,22,23
+ 355 13 355 13 351.5 26.5 c 128,-1,24
+ 348 40 348 40 347 49.5 c 128,-1,25
+ 346 59 346 59 346 75 c 2,26,-1
+ 346 726 l 1,27,-1
+ 213 726 l 1,28,-1
+ 213 73 l 2,29,30
+ 213 57 213 57 212 48 c 0,31,32
+ 210 34 210 34 208 26 c 0,33,34
+ 204 13 204 13 194 6 c 0,35,36
+ 184 0 184 0 168.5 0 c 128,-1,37
+ 153 0 153 0 143 6 c 0,38,39
+ 133 13 133 13 129.5 26.5 c 128,-1,40
+ 126 40 126 40 125 49.5 c 128,-1,41
+ 124 59 124 59 124 75 c 2,42,-1
+ 124 1465 l 2,43,44
+ 124 1481 124 1481 125 1490.5 c 128,-1,45
+ 126 1500 126 1500 129.5 1513.5 c 128,-1,46
+ 133 1527 133 1527 143 1534 c 128,-1,47
+ 153 1541 153 1541 168.5 1541 c 128,-1,48
+ 184 1541 184 1541 194 1534.5 c 128,-1,49
+ 204 1528 204 1528 207.5 1515 c 128,-1,50
+ 211 1502 211 1502 212 1493 c 128,-1,51
+ 213 1484 213 1484 213 1467 c 2,52,-1
+ 213 815 l 1,53,-1
+ 346 815 l 1,54,-1
+ 346 1465 l 2,55,56
+ 346 1481 346 1481 347 1490.5 c 128,-1,57
+ 348 1500 348 1500 351.5 1513.5 c 128,-1,58
+ 355 1527 355 1527 365 1534 c 128,-1,59
+ 375 1541 375 1541 390.5 1541 c 128,-1,60
+ 406 1541 406 1541 416 1534.5 c 128,-1,61
+ 426 1528 426 1528 429.5 1515 c 128,-1,62
+ 433 1502 433 1502 434 1493 c 128,-1,63
+ 435 1484 435 1484 435 1467 c 2,64,-1
+ 435 815 l 1,65,-1
+ 1157 815 l 1,0,-1
+EndSplineSet
+Validated: 1
 EndChar
 
 StartChar: uni22AB
@@ -33765,12 +33741,12 @@
 Width: 2148
 VWidth: 2220
 Flags: W
-HStem: 295 89<565 1978> 726 89<565 1978>
-VStem: 264 89<-534 1644> 523 118<295 384 726 815>
-LayerCount: 2
-Fore
-SplineSet
-353 1592 m 6,0,-1
+HStem: 295 89<640 1992.3> 726 89<640 1992.3>
+VStem: 264 89<-554.998 1665> 505 135<-551.614 295 384 726 815 1661.61>
+LayerCount: 2
+Fore
+SplineSet
+353 1592 m 2,0,-1
  353 -482 l 2,1,2
  353 -498 353 -498 352 -507 c 128,-1,3
  351 -516 351 -516 347.5 -529 c 128,-1,4
@@ -33788,74 +33764,56 @@
  324 1665 324 1665 334 1658.5 c 128,-1,17
  344 1652 344 1652 347.5 1639 c 128,-1,18
  351 1626 351 1626 352 1617 c 128,-1,19
- 353 1608 353 1608 353 1592 c 6,0,-1
-1924 726 m 2,20,-1
- 598 726 l 2,21,22
- 584 726 584 726 577 726.5 c 128,-1,23
- 570 727 570 727 558 729 c 128,-1,24
- 546 731 546 731 539.5 735.5 c 128,-1,25
- 533 740 533 740 528 749 c 128,-1,26
- 523 758 523 758 523 770 c 0,27,28
- 523 786 523 786 529 796 c 128,-1,29
- 535 806 535 806 548.5 809.5 c 128,-1,30
- 562 813 562 813 571 814 c 128,-1,31
- 580 815 580 815 596 815 c 2,32,-1
- 1926 815 l 2,33,34
- 1942 815 1942 815 1951 814 c 128,-1,35
- 1960 813 1960 813 1973 809.5 c 128,-1,36
- 1986 806 1986 806 1992.5 796 c 128,-1,37
- 1999 786 1999 786 1999 770 c 0,38,39
- 1999 757 1999 757 1994 748.5 c 128,-1,40
- 1989 740 1989 740 1982.5 735.5 c 128,-1,41
- 1976 731 1976 731 1964 729 c 128,-1,42
- 1952 727 1952 727 1945 726.5 c 128,-1,43
- 1938 726 1938 726 1924 726 c 2,20,-1
-1926 295 m 2,44,-1
- 596 295 l 2,45,46
- 580 295 580 295 571 296 c 128,-1,47
- 562 297 562 297 548.5 300.5 c 128,-1,48
- 535 304 535 304 529 314 c 128,-1,49
- 523 324 523 324 523 340 c 0,50,51
- 523 353 523 353 528 361.5 c 128,-1,52
- 533 370 533 370 539.5 374.5 c 128,-1,53
- 546 379 546 379 558 381 c 128,-1,54
- 570 383 570 383 577.5 383.5 c 128,-1,55
- 585 384 585 384 598 384 c 2,56,-1
- 1924 384 l 2,57,58
- 1938 384 1938 384 1945 383.5 c 128,-1,59
- 1952 383 1952 383 1964 381 c 128,-1,60
- 1976 379 1976 379 1982.5 374.5 c 128,-1,61
- 1989 370 1989 370 1994 361 c 128,-1,62
- 1999 352 1999 352 1999 340 c 0,63,64
- 1999 324 1999 324 1992.5 314 c 128,-1,65
- 1986 304 1986 304 1973 300.5 c 128,-1,66
- 1960 297 1960 297 1951 296 c 128,-1,67
- 1942 295 1942 295 1926 295 c 2,44,-1
-640 1567 m 2,68,-1
- 640 -457 l 2,69,70
- 640 -472 640 -472 640 -481 c 128,-1,71
- 640 -490 640 -490 636 -505.5 c 128,-1,72
- 632 -521 632 -521 625.5 -530.5 c 128,-1,73
- 619 -540 619 -540 605 -547.5 c 128,-1,74
- 591 -555 591 -555 572 -555 c 0,75,76
- 553 -555 553 -555 539.5 -547.5 c 128,-1,77
- 526 -540 526 -540 519.5 -530.5 c 128,-1,78
- 513 -521 513 -521 509.5 -505 c 128,-1,79
- 506 -489 506 -489 505.5 -480 c 128,-1,80
- 505 -471 505 -471 505 -457 c 2,81,-1
- 505 1567 l 2,82,83
- 505 1582 505 1582 505.5 1591 c 128,-1,84
- 506 1600 506 1600 509.5 1615.5 c 128,-1,85
- 513 1631 513 1631 519.5 1640.5 c 128,-1,86
- 526 1650 526 1650 540 1657.5 c 128,-1,87
- 554 1665 554 1665 574 1665 c 0,88,89
- 593 1665 593 1665 606.5 1657.5 c 128,-1,90
- 620 1650 620 1650 626.5 1640.5 c 128,-1,91
- 633 1631 633 1631 636.5 1615 c 128,-1,92
- 640 1599 640 1599 640 1590 c 128,-1,93
- 640 1581 640 1581 640 1567 c 2,68,-1
-EndSplineSet
-Validated: 1029
+ 353 1608 353 1608 353 1592 c 2,0,-1
+640 384 m 1,20,-1
+ 1924 384 l 2,21,22
+ 1938 384 1938 384 1945 383.5 c 128,-1,23
+ 1952 383 1952 383 1964 381 c 128,-1,24
+ 1976 379 1976 379 1982.5 374.5 c 128,-1,25
+ 1989 370 1989 370 1994 361 c 128,-1,26
+ 1999 352 1999 352 1999 340 c 0,27,28
+ 1999 324 1999 324 1992.5 314 c 128,-1,29
+ 1986 304 1986 304 1973 300.5 c 128,-1,30
+ 1960 297 1960 297 1951 296 c 128,-1,31
+ 1942 295 1942 295 1926 295 c 2,32,-1
+ 640 295 l 1,33,-1
+ 640 -457 l 1,34,-1
+ 640 -481 l 2,35,36
+ 640 -490 640 -490 636 -505.5 c 128,-1,37
+ 632 -521 632 -521 625.5 -530.5 c 128,-1,38
+ 619 -540 619 -540 605 -547.5 c 128,-1,39
+ 591 -555 591 -555 572 -555 c 128,-1,40
+ 553 -555 553 -555 539.5 -547.5 c 128,-1,41
+ 526 -540 526 -540 519.5 -530.5 c 128,-1,42
+ 513 -521 513 -521 509.5 -505 c 128,-1,43
+ 506 -489 506 -489 505.5 -480 c 128,-1,44
+ 505 -471 505 -471 505 -457 c 2,45,-1
+ 505 1567 l 2,46,47
+ 505 1582 505 1582 505.5 1591 c 128,-1,48
+ 506 1600 506 1600 509.5 1615.5 c 128,-1,49
+ 513 1631 513 1631 519.5 1640.5 c 128,-1,50
+ 526 1650 526 1650 540 1657.5 c 128,-1,51
+ 554 1665 554 1665 574 1665 c 0,52,53
+ 593 1665 593 1665 606.5 1657.5 c 128,-1,54
+ 620 1650 620 1650 626.5 1640.5 c 128,-1,55
+ 633 1631 633 1631 636.5 1615 c 128,-1,56
+ 640 1599 640 1599 640 1590 c 2,57,-1
+ 640 1567 l 1,58,-1
+ 640 815 l 1,59,-1
+ 1926 815 l 2,60,61
+ 1942 815 1942 815 1951 814 c 128,-1,62
+ 1960 813 1960 813 1973 809.5 c 128,-1,63
+ 1986 806 1986 806 1992.5 796 c 128,-1,64
+ 1999 786 1999 786 1999 770 c 0,65,66
+ 1999 757 1999 757 1994 748.5 c 128,-1,67
+ 1989 740 1989 740 1982.5 735.5 c 128,-1,68
+ 1976 731 1976 731 1964 729 c 128,-1,69
+ 1952 727 1952 727 1945 726.5 c 128,-1,70
+ 1938 726 1938 726 1924 726 c 2,71,-1
+ 640 726 l 1,72,-1
+ 640 384 l 1,20,-1
+EndSplineSet
+Validated: 1
 EndChar
 
 StartChar: uni22B2
@@ -35915,72 +35873,61 @@
 LayerCount: 2
 Fore
 SplineSet
-2101 511 m 6,0,-1
- 394 511 l 1,1,2
- 452 469 452 469 501 412 c 128,-1,3
- 550 355 550 355 577 308 c 128,-1,4
- 604 261 604 261 619 228 c 128,-1,5
- 634 195 634 195 634 186 c 0,6,7
- 634 172 634 172 624 166 c 128,-1,8
- 614 160 614 160 601 160 c 0,9,10
- 584 160 584 160 577.5 166.5 c 128,-1,11
- 571 173 571 173 563 191 c 0,12,13
- 512 301 512 301 437 383.5 c 128,-1,14
- 362 466 362 466 234 526 c 0,15,16
- 228 529 228 529 223.5 532 c 128,-1,17
- 219 535 219 535 215.5 538 c 128,-1,18
- 212 541 212 541 210 545 c 128,-1,19
- 208 549 208 549 208 555 c 0,20,21
- 208 561 208 561 208.5 564 c 128,-1,22
- 209 567 209 567 215 571.5 c 128,-1,23
- 221 576 221 576 222.5 577 c 128,-1,24
- 224 578 224 578 239 586 c 0,25,26
- 302 616 302 616 355 655.5 c 128,-1,27
- 408 695 408 695 440 728.5 c 128,-1,28
- 472 762 472 762 500 805 c 128,-1,29
- 528 848 528 848 539.5 871.5 c 128,-1,30
- 551 895 551 895 565 928 c 0,31,32
- 569 936 569 936 570.5 938.5 c 128,-1,33
- 572 941 572 941 580 945.5 c 128,-1,34
- 588 950 588 950 601 950 c 0,35,36
- 615 950 615 950 624.5 943.5 c 128,-1,37
- 634 937 634 937 634 924 c 0,38,39
- 634 915 634 915 619 882 c 128,-1,40
- 604 849 604 849 577 802 c 128,-1,41
- 550 755 550 755 501 698.5 c 128,-1,42
- 452 642 452 642 394 599 c 1,43,-1
- 2101 599 l 2,44,45
- 2115 599 2115 599 2122 599 c 128,-1,46
- 2129 599 2129 599 2141 596.5 c 128,-1,47
- 2153 594 2153 594 2159 589.5 c 128,-1,48
- 2165 585 2165 585 2169.5 576.5 c 128,-1,49
- 2174 568 2174 568 2174 555 c 128,-1,50
- 2174 542 2174 542 2169.5 533.5 c 128,-1,51
- 2165 525 2165 525 2159 520.5 c 128,-1,52
- 2153 516 2153 516 2141 513.5 c 128,-1,53
- 2129 511 2129 511 2122 511 c 128,-1,54
- 2115 511 2115 511 2101 511 c 6,0,-1
-3382 511 m 6,55,-1
- 2172 511 l 2,56,57
- 2156 511 2156 511 2147 511.5 c 128,-1,58
- 2138 512 2138 512 2125 516 c 128,-1,59
- 2112 520 2112 520 2105.5 529.5 c 128,-1,60
- 2099 539 2099 539 2099 555 c 128,-1,61
- 2099 571 2099 571 2105.5 580.5 c 128,-1,62
- 2112 590 2112 590 2125 594 c 128,-1,63
- 2138 598 2138 598 2147 598.5 c 128,-1,64
- 2156 599 2156 599 2172 599 c 2,65,-1
- 3382 599 l 2,66,67
- 3398 599 3398 599 3407.5 598.5 c 128,-1,68
- 3417 598 3417 598 3430 594 c 128,-1,69
- 3443 590 3443 590 3449 580.5 c 128,-1,70
- 3455 571 3455 571 3455 555 c 128,-1,71
- 3455 539 3455 539 3449 529.5 c 128,-1,72
- 3443 520 3443 520 3430 516 c 128,-1,73
- 3417 512 3417 512 3407.5 511.5 c 128,-1,74
- 3398 511 3398 511 3382 511 c 6,55,-1
-EndSplineSet
-Validated: 5
+3382 511 m 2,0,-1
+ 2172 511 l 2,1,2
+ 2156 511 2156 511 2147 512 c 0,3,4
+ 2143 512 2143 512 2137 513 c 1,5,6
+ 2128 511 2128 511 2122 511 c 2,7,-1
+ 2101 511 l 1,8,-1
+ 394 511 l 1,9,10
+ 452 469 452 469 501 412 c 128,-1,11
+ 550 355 550 355 577 308 c 128,-1,12
+ 604 261 604 261 619 228 c 128,-1,13
+ 634 195 634 195 634 186 c 0,14,15
+ 634 172 634 172 624 166 c 128,-1,16
+ 614 160 614 160 601 160 c 0,17,18
+ 584 160 584 160 577.5 166.5 c 128,-1,19
+ 571 173 571 173 563 191 c 0,20,21
+ 512 301 512 301 437 383.5 c 128,-1,22
+ 362 466 362 466 234 526 c 0,23,24
+ 228 529 228 529 223.5 532 c 128,-1,25
+ 219 535 219 535 215.5 538 c 128,-1,26
+ 212 541 212 541 210 545 c 128,-1,27
+ 208 549 208 549 208 555 c 128,-1,28
+ 208 561 208 561 208.5 564 c 128,-1,29
+ 209 567 209 567 215 571.5 c 128,-1,30
+ 221 576 221 576 222.5 577 c 128,-1,31
+ 224 578 224 578 239 586 c 0,32,33
+ 302 616 302 616 355 655.5 c 128,-1,34
+ 408 695 408 695 440 728 c 0,35,36
+ 472 762 472 762 500 805 c 128,-1,37
+ 528 848 528 848 539.5 871.5 c 128,-1,38
+ 551 895 551 895 565 928 c 0,39,40
+ 569 936 569 936 570.5 938.5 c 128,-1,41
+ 572 941 572 941 580 945.5 c 128,-1,42
+ 588 950 588 950 601 950 c 0,43,44
+ 615 950 615 950 624.5 943.5 c 128,-1,45
+ 634 937 634 937 634 924 c 0,46,47
+ 634 915 634 915 619 882 c 128,-1,48
+ 604 849 604 849 577 802 c 128,-1,49
+ 550 755 550 755 501 698.5 c 128,-1,50
+ 452 642 452 642 394 599 c 1,51,-1
+ 2101 599 l 1,52,-1
+ 2122 599 l 2,53,54
+ 2128 599 2128 599 2137 597 c 1,55,56
+ 2143 598 2143 598 2147 598.5 c 128,-1,57
+ 2151 599 2151 599 2172 599 c 2,58,-1
+ 3382 599 l 2,59,60
+ 3398 599 3398 599 3407.5 598.5 c 128,-1,61
+ 3417 598 3417 598 3430 594 c 128,-1,62
+ 3443 590 3443 590 3449 580.5 c 128,-1,63
+ 3455 571 3455 571 3455 555 c 128,-1,64
+ 3455 539 3455 539 3449 529.5 c 128,-1,65
+ 3443 520 3443 520 3430 516 c 128,-1,66
+ 3417 512 3417 512 3407.5 511.5 c 128,-1,67
+ 3398 511 3398 511 3382 511 c 2,0,-1
+EndSplineSet
+Validated: 1
 EndChar
 
 StartChar: uni27F6
@@ -35988,77 +35935,66 @@
 Width: 3663
 VWidth: 2220
 Flags: W
-HStem: 511 89<233 1525 1525 1525 1525 3265>
-VStem: 1484 84<511 599>
-LayerCount: 2
-Fore
-SplineSet
-3265 511 m 5,0,-1
- 1557 511 l 2,1,2
- 1543 511 1543 511 1536 511 c 128,-1,3
- 1529 511 1529 511 1517.5 513.5 c 128,-1,4
- 1506 516 1506 516 1500 520.5 c 128,-1,5
- 1494 525 1494 525 1489 533.5 c 128,-1,6
- 1484 542 1484 542 1484 555 c 128,-1,7
- 1484 568 1484 568 1489 576.5 c 128,-1,8
- 1494 585 1494 585 1500 589.5 c 128,-1,9
- 1506 594 1506 594 1517.5 596.5 c 128,-1,10
- 1529 599 1529 599 1536.5 599 c 128,-1,11
- 1544 599 1544 599 1557 599 c 2,12,-1
- 3265 599 l 1,13,14
- 3207 641 3207 641 3157.5 698 c 128,-1,15
- 3108 755 3108 755 3081 802 c 128,-1,16
- 3054 849 3054 849 3039.5 882 c 128,-1,17
- 3025 915 3025 915 3025 924 c 0,18,19
- 3025 938 3025 938 3035 944 c 128,-1,20
- 3045 950 3045 950 3058 950 c 0,21,22
- 3075 950 3075 950 3081.5 943.5 c 128,-1,23
- 3088 937 3088 937 3096 919 c 0,24,25
- 3129 846 3129 846 3169.5 788.5 c 128,-1,26
- 3210 731 3210 731 3256 691.5 c 128,-1,27
- 3302 652 3302 652 3343 626.5 c 128,-1,28
- 3384 601 3384 601 3435 577 c 1,29,30
- 3438 577 3438 577 3444.5 571 c 128,-1,31
- 3451 565 3451 565 3451 555 c 0,32,33
- 3451 544 3451 544 3446 539 c 128,-1,34
- 3441 534 3441 534 3420 524 c 0,35,36
- 3374 502 3374 502 3333.5 475.5 c 128,-1,37
- 3293 449 3293 449 3264 424.5 c 128,-1,38
- 3235 400 3235 400 3208.5 369 c 128,-1,39
- 3182 338 3182 338 3166 316 c 128,-1,40
- 3150 294 3150 294 3133.5 264 c 128,-1,41
- 3117 234 3117 234 3110.5 220 c 128,-1,42
- 3104 206 3104 206 3094 182 c 0,43,44
- 3090 174 3090 174 3088 171.5 c 128,-1,45
- 3086 169 3086 169 3078 164.5 c 128,-1,46
- 3070 160 3070 160 3058 160 c 0,47,48
- 3044 160 3044 160 3034.5 166.5 c 128,-1,49
- 3025 173 3025 173 3025 186 c 0,50,51
- 3025 195 3025 195 3039.5 228 c 128,-1,52
- 3054 261 3054 261 3081 308 c 128,-1,53
- 3108 355 3108 355 3157.5 412 c 128,-1,54
- 3207 469 3207 469 3265 511 c 5,0,-1
-1495 511 m 6,55,-1
- 285 511 l 2,56,57
- 269 511 269 511 260 511.5 c 128,-1,58
- 251 512 251 512 238 516 c 128,-1,59
- 225 520 225 520 218.5 529.5 c 128,-1,60
- 212 539 212 539 212 555 c 128,-1,61
- 212 571 212 571 218.5 580.5 c 128,-1,62
- 225 590 225 590 238 594 c 128,-1,63
- 251 598 251 598 260 598.5 c 128,-1,64
- 269 599 269 599 285 599 c 2,65,-1
- 1495 599 l 2,66,67
- 1511 599 1511 599 1520.5 598.5 c 128,-1,68
- 1530 598 1530 598 1543 594 c 128,-1,69
- 1556 590 1556 590 1562 580.5 c 128,-1,70
- 1568 571 1568 571 1568 555 c 128,-1,71
- 1568 539 1568 539 1562 529.5 c 128,-1,72
- 1556 520 1556 520 1543 516 c 128,-1,73
- 1530 512 1530 512 1520.5 511.5 c 128,-1,74
- 1511 511 1511 511 1495 511 c 6,55,-1
-EndSplineSet
-Validated: 5
+HStem: 511 88<218.705 3265>
+VStem: 3025 71<849.875 943.302>
+LayerCount: 2
+Fore
+SplineSet
+1495 511 m 2,0,-1
+ 285 511 l 2,1,2
+ 269 511 269 511 260 511.5 c 128,-1,3
+ 251 512 251 512 238 516 c 128,-1,4
+ 225 520 225 520 218.5 529.5 c 128,-1,5
+ 212 539 212 539 212 555 c 128,-1,6
+ 212 571 212 571 218.5 580.5 c 128,-1,7
+ 225 590 225 590 238 594 c 128,-1,8
+ 251 598 251 598 260 598.5 c 128,-1,9
+ 269 599 269 599 285 599 c 2,10,-1
+ 1495 599 l 2,11,12
+ 1511 599 1511 599 1520 598 c 0,13,14
+ 1526 598 1526 598 1526 598 c 0,15,16
+ 1532 599 1532 599 1536 599 c 2,17,-1
+ 1557 599 l 1,18,-1
+ 3265 599 l 1,19,20
+ 3207 641 3207 641 3157.5 698 c 128,-1,21
+ 3108 755 3108 755 3081 802 c 128,-1,22
+ 3054 849 3054 849 3039.5 882 c 128,-1,23
+ 3025 915 3025 915 3025 924 c 0,24,25
+ 3025 938 3025 938 3035 944 c 128,-1,26
+ 3045 950 3045 950 3058 950 c 0,27,28
+ 3075 950 3075 950 3082 944 c 0,29,30
+ 3088 937 3088 937 3096 919 c 0,31,32
+ 3127 848 3127 848 3170 788 c 0,33,34
+ 3210 731 3210 731 3256 691.5 c 128,-1,35
+ 3302 652 3302 652 3343 626.5 c 128,-1,36
+ 3384 601 3384 601 3435 577 c 1,37,38
+ 3438 577 3438 577 3444.5 571 c 128,-1,39
+ 3451 565 3451 565 3451 555 c 0,40,41
+ 3451 544 3451 544 3446 539 c 128,-1,42
+ 3441 534 3441 534 3420 524 c 0,43,44
+ 3374 502 3374 502 3333.5 475.5 c 128,-1,45
+ 3293 449 3293 449 3264 424.5 c 128,-1,46
+ 3235 400 3235 400 3208.5 369 c 128,-1,47
+ 3182 338 3182 338 3166 316 c 128,-1,48
+ 3150 294 3150 294 3133.5 264 c 128,-1,49
+ 3117 234 3117 234 3110.5 220 c 128,-1,50
+ 3104 206 3104 206 3094 182 c 0,51,52
+ 3090 174 3090 174 3088 171.5 c 128,-1,53
+ 3086 169 3086 169 3078 164 c 0,54,55
+ 3070 160 3070 160 3058 160 c 0,56,57
+ 3044 160 3044 160 3034.5 166.5 c 128,-1,58
+ 3025 173 3025 173 3025 186 c 0,59,60
+ 3025 195 3025 195 3039.5 228 c 128,-1,61
+ 3054 261 3054 261 3081 308 c 128,-1,62
+ 3108 355 3108 355 3157.5 412 c 128,-1,63
+ 3207 469 3207 469 3265 511 c 1,64,-1
+ 1557 511 l 1,65,-1
+ 1536 511 l 2,66,67
+ 1532 511 1532 511 1526 512 c 0,68,69
+ 1523 512 1523 512 1520.5 511.5 c 128,-1,71
+ 1518 511 1518 511 1495 511 c 2,0,-1
+EndSplineSet
+Validated: 1
 EndChar
 
 StartChar: uni27F7
@@ -36066,102 +36002,88 @@
 Width: 3663
 VWidth: 2220
 Flags: W
-HStem: 511 89<313 2047 1583 3326>
-LayerCount: 2
-Fore
-SplineSet
-2020 511 m 6,0,-1
- 313 511 l 1,1,2
- 371 469 371 469 420 412 c 128,-1,3
- 469 355 469 355 496 308 c 128,-1,4
- 523 261 523 261 538 228 c 128,-1,5
- 553 195 553 195 553 186 c 0,6,7
- 553 172 553 172 543 166 c 128,-1,8
- 533 160 533 160 519 160 c 0,9,10
- 502 160 502 160 496 166.5 c 128,-1,11
- 490 173 490 173 482 191 c 0,12,13
- 431 301 431 301 356 383.5 c 128,-1,14
- 281 466 281 466 153 526 c 0,15,16
- 147 529 147 529 142.5 532 c 128,-1,17
- 138 535 138 535 134 538 c 128,-1,18
- 130 541 130 541 128.5 545 c 128,-1,19
- 127 549 127 549 127 555 c 0,20,21
- 127 561 127 561 127.5 564 c 128,-1,22
- 128 567 128 567 134 571.5 c 128,-1,23
- 140 576 140 576 141.5 577 c 128,-1,24
- 143 578 143 578 158 586 c 0,25,26
- 221 616 221 616 274 655.5 c 128,-1,27
- 327 695 327 695 359 728.5 c 128,-1,28
- 391 762 391 762 419 805 c 128,-1,29
- 447 848 447 848 458.5 871.5 c 128,-1,30
- 470 895 470 895 484 928 c 0,31,32
- 488 936 488 936 489.5 938.5 c 128,-1,33
- 491 941 491 941 499 945.5 c 128,-1,34
- 507 950 507 950 519 950 c 0,35,36
- 533 950 533 950 543 943.5 c 128,-1,37
- 553 937 553 937 553 924 c 0,38,39
- 553 915 553 915 538 882 c 128,-1,40
- 523 849 523 849 496 802 c 128,-1,41
- 469 755 469 755 420 698.5 c 128,-1,42
- 371 642 371 642 313 599 c 1,43,-1
- 2020 599 l 2,44,45
- 2034 599 2034 599 2041 599 c 128,-1,46
- 2048 599 2048 599 2060 596.5 c 128,-1,47
- 2072 594 2072 594 2078 589.5 c 128,-1,48
- 2084 585 2084 585 2088.5 576.5 c 128,-1,49
- 2093 568 2093 568 2093 555 c 128,-1,50
- 2093 542 2093 542 2088.5 533.5 c 128,-1,51
- 2084 525 2084 525 2078 520.5 c 128,-1,52
- 2072 516 2072 516 2060 513.5 c 128,-1,53
- 2048 511 2048 511 2041 511 c 128,-1,54
- 2034 511 2034 511 2020 511 c 6,0,-1
-3328 511 m 5,55,-1
- 1621 511 l 2,56,57
- 1607 511 1607 511 1600 511 c 128,-1,58
- 1593 511 1593 511 1581 513.5 c 128,-1,59
- 1569 516 1569 516 1563 520.5 c 128,-1,60
- 1557 525 1557 525 1552 533.5 c 128,-1,61
- 1547 542 1547 542 1547 555 c 128,-1,62
- 1547 568 1547 568 1552 576.5 c 128,-1,63
- 1557 585 1557 585 1563 589.5 c 128,-1,64
- 1569 594 1569 594 1581 596.5 c 128,-1,65
- 1593 599 1593 599 1600 599 c 128,-1,66
- 1607 599 1607 599 1621 599 c 2,67,-1
- 3328 599 l 1,68,69
- 3270 641 3270 641 3221 698 c 128,-1,70
- 3172 755 3172 755 3145 802 c 128,-1,71
- 3118 849 3118 849 3103 882 c 128,-1,72
- 3088 915 3088 915 3088 924 c 0,73,74
- 3088 938 3088 938 3098 944 c 128,-1,75
- 3108 950 3108 950 3121 950 c 0,76,77
- 3138 950 3138 950 3144.5 943.5 c 128,-1,78
- 3151 937 3151 937 3159 919 c 0,79,80
- 3192 846 3192 846 3232.5 788.5 c 128,-1,81
- 3273 731 3273 731 3319.5 691.5 c 128,-1,82
- 3366 652 3366 652 3406.5 626.5 c 128,-1,83
- 3447 601 3447 601 3499 577 c 1,84,85
- 3502 577 3502 577 3508 571 c 128,-1,86
- 3514 565 3514 565 3514 555 c 0,87,88
- 3514 544 3514 544 3509 539 c 128,-1,89
- 3504 534 3504 534 3483 524 c 0,90,91
- 3437 502 3437 502 3396.5 475.5 c 128,-1,92
- 3356 449 3356 449 3327 424.5 c 128,-1,93
- 3298 400 3298 400 3271.5 369 c 128,-1,94
- 3245 338 3245 338 3229.5 316 c 128,-1,95
- 3214 294 3214 294 3197.5 264 c 128,-1,96
- 3181 234 3181 234 3174 220 c 128,-1,97
- 3167 206 3167 206 3157 182 c 0,98,99
- 3153 174 3153 174 3151.5 171.5 c 128,-1,100
- 3150 169 3150 169 3142 164.5 c 128,-1,101
- 3134 160 3134 160 3121 160 c 0,102,103
- 3107 160 3107 160 3097.5 166.5 c 128,-1,104
- 3088 173 3088 173 3088 186 c 0,105,106
- 3088 195 3088 195 3103 228 c 128,-1,107
- 3118 261 3118 261 3145 308 c 128,-1,108
- 3172 355 3172 355 3221 412 c 128,-1,109
- 3270 469 3270 469 3328 511 c 5,55,-1
-EndSplineSet
-Validated: 1029
+HStem: 511 88<313 3328>
+VStem: 482 71<166.706 257.064> 3088 71<849.875 943.302>
+LayerCount: 2
+Fore
+SplineSet
+2041 511 m 1,0,-1
+ 2020 511 l 1,1,-1
+ 1621 511 l 1,2,-1
+ 1600 511 l 1,3,-1
+ 313 511 l 1,4,5
+ 371 469 371 469 420 412 c 128,-1,6
+ 469 355 469 355 496 308 c 128,-1,7
+ 523 261 523 261 538 228 c 128,-1,8
+ 553 195 553 195 553 186 c 0,9,10
+ 553 172 553 172 543 166 c 128,-1,11
+ 533 160 533 160 519 160 c 0,12,13
+ 502 160 502 160 496 166.5 c 128,-1,14
+ 490 173 490 173 482 191 c 0,15,16
+ 431 301 431 301 356 383.5 c 128,-1,17
+ 281 466 281 466 153 526 c 0,18,19
+ 147 529 147 529 142.5 532 c 128,-1,20
+ 138 535 138 535 134 538 c 128,-1,21
+ 130 541 130 541 128.5 545 c 128,-1,22
+ 127 549 127 549 127 555 c 128,-1,23
+ 127 561 127 561 127.5 564 c 128,-1,24
+ 128 567 128 567 134 571.5 c 128,-1,25
+ 140 576 140 576 141.5 577 c 128,-1,26
+ 143 578 143 578 158 586 c 0,27,28
+ 221 616 221 616 274 655.5 c 128,-1,29
+ 327 695 327 695 359 728 c 0,30,31
+ 391 762 391 762 419 805 c 128,-1,32
+ 447 848 447 848 458.5 871.5 c 128,-1,33
+ 470 895 470 895 484 928 c 0,34,35
+ 488 936 488 936 489.5 938.5 c 128,-1,36
+ 491 941 491 941 499 945.5 c 128,-1,37
+ 507 950 507 950 519 950 c 0,38,39
+ 533 950 533 950 543 943.5 c 128,-1,40
+ 553 937 553 937 553 924 c 0,41,42
+ 553 915 553 915 538 882 c 128,-1,43
+ 523 849 523 849 496 802 c 128,-1,44
+ 469 755 469 755 420 698.5 c 128,-1,45
+ 371 642 371 642 313 599 c 1,46,-1
+ 1600 599 l 1,47,-1
+ 1621 599 l 1,48,-1
+ 2020 599 l 1,49,-1
+ 2041 599 l 1,50,-1
+ 3328 599 l 1,51,52
+ 3270 641 3270 641 3221 698 c 128,-1,53
+ 3172 755 3172 755 3145 802 c 128,-1,54
+ 3118 849 3118 849 3103 882 c 128,-1,55
+ 3088 915 3088 915 3088 924 c 0,56,57
+ 3088 938 3088 938 3098 944 c 128,-1,58
+ 3108 950 3108 950 3121 950 c 0,59,60
+ 3138 950 3138 950 3144 944 c 0,61,62
+ 3151 937 3151 937 3159 919 c 0,63,64
+ 3190 848 3190 848 3232 788 c 0,65,66
+ 3273 731 3273 731 3319.5 691.5 c 128,-1,67
+ 3366 652 3366 652 3406.5 626.5 c 128,-1,68
+ 3447 601 3447 601 3499 577 c 1,69,70
+ 3502 577 3502 577 3508 571 c 128,-1,71
+ 3514 565 3514 565 3514 555 c 0,72,73
+ 3514 544 3514 544 3509 539 c 128,-1,74
+ 3504 534 3504 534 3483 524 c 0,75,76
+ 3437 502 3437 502 3396.5 475.5 c 128,-1,77
+ 3356 449 3356 449 3327 424.5 c 128,-1,78
+ 3298 400 3298 400 3271.5 369 c 128,-1,79
+ 3245 338 3245 338 3229.5 316 c 128,-1,80
+ 3214 294 3214 294 3197.5 264 c 128,-1,81
+ 3181 234 3181 234 3174 220 c 128,-1,82
+ 3167 206 3167 206 3157 182 c 0,83,84
+ 3153 174 3153 174 3151.5 171.5 c 128,-1,85
+ 3150 169 3150 169 3142 164.5 c 128,-1,86
+ 3134 160 3134 160 3121 160 c 0,87,88
+ 3107 160 3107 160 3097.5 166.5 c 128,-1,89
+ 3088 173 3088 173 3088 186 c 0,90,91
+ 3088 195 3088 195 3103 228 c 128,-1,92
+ 3118 261 3118 261 3145 308 c 128,-1,93
+ 3172 355 3172 355 3221 412 c 128,-1,94
+ 3270 469 3270 469 3328 511 c 1,95,-1
+ 2041 511 l 1,0,-1
+EndSplineSet
+Validated: 1
 EndChar
 
 StartChar: uni27F8
@@ -36169,117 +36091,95 @@
 Width: 3663
 VWidth: 2220
 Flags: W
-HStem: 295 89<677 2073 2071 3457> 726 89<737 2073 2071 3457>
-LayerCount: 2
-Fore
-SplineSet
-2047 295 m 2,0,-1
- 737 295 l 1,1,2
- 799 227 799 227 846.5 163 c 128,-1,3
- 894 99 894 99 916 59.5 c 128,-1,4
- 938 20 938 20 948.5 -2.5 c 128,-1,5
- 959 -25 959 -25 959 -29 c 0,6,7
- 959 -41 959 -41 951 -47.5 c 128,-1,8
- 943 -54 943 -54 935.5 -55 c 128,-1,9
- 928 -56 928 -56 915 -56 c 0,10,11
- 891 -56 891 -56 881 -49 c 0,12,13
- 879 -48 879 -48 878 -46 c 128,-1,14
- 877 -44 877 -44 875 -40 c 128,-1,15
- 873 -36 873 -36 866.5 -24.5 c 128,-1,16
- 860 -13 860 -13 852 0 c 0,17,18
- 817 64 817 64 775.5 123 c 128,-1,19
- 734 182 734 182 674 243.5 c 128,-1,20
- 614 305 614 305 544.5 356 c 128,-1,21
- 475 407 475 407 383.5 452 c 128,-1,22
- 292 497 292 497 189 526 c 0,23,24
- 169 532 169 532 161 537.5 c 128,-1,25
- 153 543 153 543 153 555.5 c 128,-1,26
- 153 568 153 568 161 573 c 128,-1,27
- 169 578 169 578 191 584 c 0,28,29
- 318 620 318 620 426.5 679 c 128,-1,30
- 535 738 535 738 615.5 811.5 c 128,-1,31
- 696 885 696 885 755 961 c 128,-1,32
- 814 1037 814 1037 859 1123 c 0,33,34
- 876 1154 876 1154 883 1160 c 128,-1,35
- 890 1166 890 1166 915 1166 c 0,36,37
- 929 1166 929 1166 936 1165 c 128,-1,38
- 943 1164 943 1164 951 1157.5 c 128,-1,39
- 959 1151 959 1151 959 1139 c 0,40,41
- 959 1134 959 1134 948.5 1111.5 c 128,-1,42
- 938 1089 938 1089 916 1049 c 128,-1,43
- 894 1009 894 1009 846.5 945.5 c 128,-1,44
- 799 882 799 882 737 815 c 1,45,-1
- 2047 815 l 2,46,47
- 2063 815 2063 815 2072 814 c 128,-1,48
- 2081 813 2081 813 2094 809.5 c 128,-1,49
- 2107 806 2107 806 2113.5 796 c 128,-1,50
- 2120 786 2120 786 2120 770 c 0,51,52
- 2120 757 2120 757 2115 748.5 c 128,-1,53
- 2110 740 2110 740 2103.5 735.5 c 128,-1,54
- 2097 731 2097 731 2085 729 c 128,-1,55
- 2073 727 2073 727 2065.5 726.5 c 128,-1,56
- 2058 726 2058 726 2045 726 c 2,57,-1
- 677 726 l 2,58,59
- 649 726 649 726 639.5 723 c 128,-1,60
- 630 720 630 720 610 704 c 0,61,62
- 479 609 479 609 355 555 c 1,63,64
- 490 494 490 494 608 406 c 0,65,66
- 629 390 629 390 639 387 c 128,-1,67
- 649 384 649 384 677 384 c 2,68,-1
- 2045 384 l 2,69,70
- 2059 384 2059 384 2066 383.5 c 128,-1,71
- 2073 383 2073 383 2085 381 c 128,-1,72
- 2097 379 2097 379 2103.5 374.5 c 128,-1,73
- 2110 370 2110 370 2115 361 c 128,-1,74
- 2120 352 2120 352 2120 340 c 0,75,76
- 2120 324 2120 324 2113.5 314 c 128,-1,77
- 2107 304 2107 304 2094 300.5 c 128,-1,78
- 2081 297 2081 297 2072 296 c 128,-1,79
- 2063 295 2063 295 2047 295 c 2,0,-1
-3434 726 m 2,80,-1
+HStem: 295 89<737 3503.3> 726 89<737 3503.3>
+LayerCount: 2
+Fore
+SplineSet
+3437 295 m 2,0,-1
+ 2107 295 l 2,1,2
+ 2091 295 2091 295 2082 296 c 0,3,4
+ 2079 296 2079 296 2077 297 c 0,5,6
+ 2074 296 2074 296 2072 296 c 0,7,8
+ 2063 295 2063 295 2047 295 c 2,9,-1
+ 737 295 l 1,10,11
+ 793 233 793 233 846 163 c 0,12,13
+ 894 99 894 99 916 59 c 0,14,15
+ 938 20 938 20 948.5 -2.5 c 128,-1,16
+ 959 -25 959 -25 959 -29 c 0,17,18
+ 959 -41 959 -41 951 -47.5 c 128,-1,19
+ 943 -54 943 -54 935.5 -55 c 128,-1,20
+ 928 -56 928 -56 915 -56 c 0,21,22
+ 891 -56 891 -56 881 -49 c 0,23,24
+ 879 -48 879 -48 878 -46 c 2,25,-1
+ 875 -40 l 2,26,27
+ 873 -36 873 -36 866 -24 c 2,28,-1
+ 852 0 l 1,29,30
+ 817 64 817 64 775.5 123 c 128,-1,31
+ 734 182 734 182 674 243.5 c 128,-1,32
+ 614 305 614 305 544.5 356 c 128,-1,33
+ 475 407 475 407 383.5 452 c 128,-1,34
+ 292 497 292 497 189 526 c 0,35,36
+ 169 532 169 532 161 537.5 c 128,-1,37
+ 153 543 153 543 153 555.5 c 128,-1,38
+ 153 568 153 568 161 573 c 0,39,40
+ 168 578 168 578 191 584 c 0,41,42
+ 318 620 318 620 426.5 679 c 128,-1,43
+ 535 738 535 738 615.5 811.5 c 128,-1,44
+ 696 885 696 885 755 961 c 128,-1,45
+ 814 1037 814 1037 859 1123 c 0,46,47
+ 876 1154 876 1154 883 1160 c 128,-1,48
+ 890 1166 890 1166 915 1166 c 0,49,50
+ 929 1166 929 1166 936 1165 c 128,-1,51
+ 943 1164 943 1164 951 1157.5 c 128,-1,52
+ 959 1151 959 1151 959 1139 c 0,53,54
+ 959 1134 959 1134 948.5 1111.5 c 128,-1,55
+ 938 1089 938 1089 916 1049 c 128,-1,56
+ 894 1009 894 1009 846.5 945.5 c 128,-1,57
+ 799 882 799 882 737 815 c 1,58,-1
+ 2047 815 l 2,59,60
+ 2063 815 2063 815 2072 814 c 0,61,62
+ 2074 814 2074 814 2077 813 c 0,63,64
+ 2079 814 2079 814 2082 814 c 0,65,66
+ 2091 815 2091 815 2107 815 c 2,67,-1
+ 3437 815 l 2,68,69
+ 3453 815 3453 815 3462 814 c 128,-1,70
+ 3471 813 3471 813 3484 809.5 c 128,-1,71
+ 3497 806 3497 806 3503.5 796 c 128,-1,72
+ 3510 786 3510 786 3510 770 c 0,73,74
+ 3510 757 3510 757 3505 748.5 c 128,-1,75
+ 3500 740 3500 740 3494 736 c 0,76,77
+ 3487 731 3487 731 3475 729 c 128,-1,78
+ 3463 727 3463 727 3455.5 726.5 c 128,-1,79
+ 3448 726 3448 726 3434 726 c 2,80,-1
  2109 726 l 2,81,82
- 2095 726 2095 726 2088 726.5 c 128,-1,83
- 2081 727 2081 727 2068.5 729 c 128,-1,84
- 2056 731 2056 731 2050 735.5 c 128,-1,85
- 2044 740 2044 740 2039 749 c 128,-1,86
- 2034 758 2034 758 2034 770 c 0,87,88
- 2034 786 2034 786 2040 796 c 128,-1,89
- 2046 806 2046 806 2059 809.5 c 128,-1,90
- 2072 813 2072 813 2081.5 814 c 128,-1,91
- 2091 815 2091 815 2107 815 c 2,92,-1
- 3437 815 l 2,93,94
- 3453 815 3453 815 3462 814 c 128,-1,95
- 3471 813 3471 813 3484 809.5 c 128,-1,96
- 3497 806 3497 806 3503.5 796 c 128,-1,97
- 3510 786 3510 786 3510 770 c 0,98,99
- 3510 757 3510 757 3505 748.5 c 128,-1,100
- 3500 740 3500 740 3493.5 735.5 c 128,-1,101
- 3487 731 3487 731 3475 729 c 128,-1,102
- 3463 727 3463 727 3455.5 726.5 c 128,-1,103
- 3448 726 3448 726 3434 726 c 2,80,-1
-3437 295 m 2,104,-1
- 2107 295 l 2,105,106
- 2091 295 2091 295 2081.5 296 c 128,-1,107
- 2072 297 2072 297 2059 300.5 c 128,-1,108
- 2046 304 2046 304 2040 314 c 128,-1,109
- 2034 324 2034 324 2034 340 c 0,110,111
- 2034 353 2034 353 2039 361.5 c 128,-1,112
- 2044 370 2044 370 2050 374.5 c 128,-1,113
- 2056 379 2056 379 2068.5 381 c 128,-1,114
- 2081 383 2081 383 2088 383.5 c 128,-1,115
- 2095 384 2095 384 2109 384 c 2,116,-1
- 3434 384 l 2,117,118
- 3448 384 3448 384 3455.5 383.5 c 128,-1,119
- 3463 383 3463 383 3475 381 c 128,-1,120
- 3487 379 3487 379 3493.5 374.5 c 128,-1,121
- 3500 370 3500 370 3505 361 c 128,-1,122
- 3510 352 3510 352 3510 340 c 0,123,124
- 3510 324 3510 324 3503.5 314 c 128,-1,125
- 3497 304 3497 304 3484 300.5 c 128,-1,126
- 3471 297 3471 297 3462 296 c 128,-1,127
- 3453 295 3453 295 3437 295 c 2,104,-1
-EndSplineSet
-Validated: 1029
+ 2088 726 2088 726 2088 726 c 0,83,84
+ 2084 727 2084 727 2077 728 c 1,85,86
+ 2070 727 2070 727 2065.5 726.5 c 128,-1,87
+ 2061 726 2061 726 2045 726 c 2,88,-1
+ 677 726 l 2,89,90
+ 649 726 649 726 639.5 723 c 128,-1,91
+ 630 720 630 720 610 704 c 1,92,93
+ 479 609 479 609 355 555 c 1,94,95
+ 490 494 490 494 608 406 c 0,96,97
+ 630 390 630 390 639 387 c 0,98,99
+ 649 384 649 384 677 384 c 2,100,-1
+ 2045 384 l 2,101,102
+ 2066 384 2066 384 2066 384 c 0,103,104
+ 2070 383 2070 383 2077 382 c 1,105,106
+ 2084 383 2084 383 2088 383.5 c 128,-1,108
+ 2092 384 2092 384 2109 384 c 2,109,-1
+ 3434 384 l 2,110,111
+ 3448 384 3448 384 3455.5 383.5 c 128,-1,112
+ 3463 383 3463 383 3475 381 c 128,-1,113
+ 3487 379 3487 379 3493.5 374.5 c 128,-1,114
+ 3500 370 3500 370 3505 361 c 128,-1,115
+ 3510 352 3510 352 3510 340 c 0,116,117
+ 3510 324 3510 324 3503.5 314 c 128,-1,118
+ 3497 304 3497 304 3484 300.5 c 128,-1,119
+ 3471 297 3471 297 3462 296 c 128,-1,120
+ 3453 295 3453 295 3437 295 c 2,0,-1
+EndSplineSet
+Validated: 1025
 EndChar
 
 StartChar: uni27F9
@@ -36287,117 +36187,91 @@
 Width: 3663
 VWidth: 2220
 Flags: W
-HStem: 295 89<195 1581 1576 2922> 726 89<195 1581 1576 2922>
-LayerCount: 2
-Fore
-SplineSet
-2981 726 m 6,0,-1
- 1614 726 l 2,1,2
- 1600 726 1600 726 1592.5 726.5 c 128,-1,3
- 1585 727 1585 727 1573 729 c 128,-1,4
- 1561 731 1561 731 1555 735.5 c 128,-1,5
- 1549 740 1549 740 1543.5 749 c 128,-1,6
- 1538 758 1538 758 1538 770 c 0,7,8
- 1538 786 1538 786 1544.5 796 c 128,-1,9
- 1551 806 1551 806 1564 809.5 c 128,-1,10
- 1577 813 1577 813 1586 814 c 128,-1,11
- 1595 815 1595 815 1612 815 c 2,12,-1
- 2922 815 l 1,13,14
- 2860 883 2860 883 2812 947 c 128,-1,15
- 2764 1011 2764 1011 2742 1050.5 c 128,-1,16
- 2720 1090 2720 1090 2710 1112.5 c 128,-1,17
- 2700 1135 2700 1135 2700 1139 c 0,18,19
- 2700 1151 2700 1151 2707.5 1157.5 c 128,-1,20
- 2715 1164 2715 1164 2722.5 1165 c 128,-1,21
- 2730 1166 2730 1166 2744 1166 c 0,22,23
- 2768 1166 2768 1166 2777 1159 c 0,24,25
- 2779 1158 2779 1158 2780.5 1156 c 128,-1,26
- 2782 1154 2782 1154 2784 1150 c 128,-1,27
- 2786 1146 2786 1146 2792 1134.5 c 128,-1,28
- 2798 1123 2798 1123 2806 1110 c 0,29,30
- 2841 1046 2841 1046 2883 987 c 128,-1,31
- 2925 928 2925 928 2985 866.5 c 128,-1,32
- 3045 805 3045 805 3114 754 c 128,-1,33
- 3183 703 3183 703 3275 658 c 128,-1,34
- 3367 613 3367 613 3470 584 c 0,35,36
- 3490 578 3490 578 3497.5 572.5 c 128,-1,37
- 3505 567 3505 567 3505 554.5 c 128,-1,38
- 3505 542 3505 542 3497 537 c 128,-1,39
- 3489 532 3489 532 3468 526 c 0,40,41
- 3341 490 3341 490 3232.5 431 c 128,-1,42
- 3124 372 3124 372 3043 298.5 c 128,-1,43
- 2962 225 2962 225 2903.5 149 c 128,-1,44
- 2845 73 2845 73 2799 -13 c 0,45,46
- 2782 -44 2782 -44 2775 -50 c 128,-1,47
- 2768 -56 2768 -56 2744 -56 c 0,48,49
- 2730 -56 2730 -56 2722.5 -55 c 128,-1,50
- 2715 -54 2715 -54 2707.5 -47.5 c 128,-1,51
- 2700 -41 2700 -41 2700 -29 c 0,52,53
- 2700 -24 2700 -24 2710 -1.5 c 128,-1,54
- 2720 21 2720 21 2742 61 c 128,-1,55
- 2764 101 2764 101 2811.5 164.5 c 128,-1,56
- 2859 228 2859 228 2922 295 c 1,57,-1
- 1612 295 l 2,58,59
- 1596 295 1596 295 1586.5 296 c 128,-1,60
- 1577 297 1577 297 1564 300.5 c 128,-1,61
- 1551 304 1551 304 1544.5 314 c 128,-1,62
- 1538 324 1538 324 1538 340 c 0,63,64
- 1538 353 1538 353 1543.5 361.5 c 128,-1,65
- 1549 370 1549 370 1555 374.5 c 128,-1,66
- 1561 379 1561 379 1573 381 c 128,-1,67
- 1585 383 1585 383 1592.5 383.5 c 128,-1,68
- 1600 384 1600 384 1614 384 c 2,69,-1
- 2981 384 l 2,70,71
- 3009 384 3009 384 3019 387 c 128,-1,72
- 3029 390 3029 390 3048 406 c 0,73,74
- 3179 501 3179 501 3303 555 c 1,75,76
- 3168 616 3168 616 3050 704 c 0,77,78
- 3029 720 3029 720 3019 723 c 128,-1,79
- 3009 726 3009 726 2981 726 c 6,0,-1
-1558 726 m 6,80,-1
- 233 726 l 2,81,82
- 219 726 219 726 212 726.5 c 128,-1,83
- 205 727 205 727 193 729 c 128,-1,84
- 181 731 181 731 174.5 735.5 c 128,-1,85
- 168 740 168 740 163 749 c 128,-1,86
- 158 758 158 758 158 770 c 0,87,88
- 158 786 158 786 164 796 c 128,-1,89
- 170 806 170 806 183 809.5 c 128,-1,90
- 196 813 196 813 205.5 814 c 128,-1,91
- 215 815 215 815 231 815 c 2,92,-1
- 1561 815 l 2,93,94
- 1577 815 1577 815 1586 814 c 128,-1,95
- 1595 813 1595 813 1608 809.5 c 128,-1,96
- 1621 806 1621 806 1627.5 796 c 128,-1,97
- 1634 786 1634 786 1634 770 c 0,98,99
- 1634 757 1634 757 1629 748.5 c 128,-1,100
- 1624 740 1624 740 1617.5 735.5 c 128,-1,101
- 1611 731 1611 731 1599 729 c 128,-1,102
- 1587 727 1587 727 1579.5 726.5 c 128,-1,103
- 1572 726 1572 726 1558 726 c 6,80,-1
-1561 295 m 6,104,-1
- 231 295 l 2,105,106
- 215 295 215 295 205.5 296 c 128,-1,107
- 196 297 196 297 183 300.5 c 128,-1,108
- 170 304 170 304 164 314 c 128,-1,109
- 158 324 158 324 158 340 c 0,110,111
- 158 353 158 353 163 361.5 c 128,-1,112
- 168 370 168 370 174.5 374.5 c 128,-1,113
- 181 379 181 379 193 381 c 128,-1,114
- 205 383 205 383 212 383.5 c 128,-1,115
- 219 384 219 384 233 384 c 2,116,-1
- 1558 384 l 2,117,118
- 1572 384 1572 384 1579.5 383.5 c 128,-1,119
- 1587 383 1587 383 1599 381 c 128,-1,120
- 1611 379 1611 379 1617.5 374.5 c 128,-1,121
- 1624 370 1624 370 1629 361 c 128,-1,122
- 1634 352 1634 352 1634 340 c 0,123,124
- 1634 324 1634 324 1627.5 314 c 128,-1,125
- 1621 304 1621 304 1608 300.5 c 128,-1,126
- 1595 297 1595 297 1586 296 c 128,-1,127
- 1577 295 1577 295 1561 295 c 6,104,-1
-EndSplineSet
-Validated: 1029
+HStem: 295 89<164.189 2922> 726.5 88.5<164.189 2922>
+LayerCount: 2
+Fore
+SplineSet
+1561 295 m 2,0,-1
+ 231 295 l 2,1,2
+ 215 295 215 295 205.5 296 c 128,-1,3
+ 196 297 196 297 183 300.5 c 128,-1,4
+ 170 304 170 304 164 314 c 128,-1,5
+ 158 324 158 324 158 340 c 0,6,7
+ 158 353 158 353 163 361.5 c 128,-1,8
+ 168 370 168 370 174.5 374.5 c 128,-1,9
+ 181 379 181 379 193 381 c 128,-1,10
+ 205 383 205 383 212 383.5 c 128,-1,11
+ 219 384 219 384 233 384 c 2,12,-1
+ 1558 384 l 1,13,-1
+ 1580 384 l 1,14,15
+ 1582 383 1582 383 1586 383 c 128,-1,16
+ 1590 383 1590 383 1592 384 c 0,17,18
+ 1595 384 1595 384 1614 384 c 2,19,-1
+ 2981 384 l 2,20,21
+ 3009 384 3009 384 3019 387 c 128,-1,22
+ 3029 390 3029 390 3048 406 c 1,23,24
+ 3179 501 3179 501 3303 555 c 1,25,26
+ 3168 616 3168 616 3050 704 c 0,27,28
+ 3029 720 3029 720 3019 723 c 128,-1,29
+ 3009 726 3009 726 2981 726 c 2,30,-1
+ 1614 726 l 1,31,-1
+ 1592 726 l 1,32,33
+ 1590 727 1590 727 1586 727 c 128,-1,34
+ 1582 727 1582 727 1579.5 726.5 c 128,-1,35
+ 1577 726 1577 726 1558 726 c 2,36,-1
+ 233 726 l 2,37,38
+ 219 726 219 726 212 726.5 c 128,-1,39
+ 205 727 205 727 193 729 c 128,-1,40
+ 181 731 181 731 174.5 735.5 c 128,-1,41
+ 168 740 168 740 163 749 c 128,-1,42
+ 158 758 158 758 158 770 c 0,43,44
+ 158 786 158 786 164 796 c 128,-1,45
+ 170 806 170 806 183 809.5 c 128,-1,46
+ 196 813 196 813 205.5 814 c 128,-1,47
+ 215 815 215 815 231 815 c 2,48,-1
+ 1561 815 l 2,49,50
+ 1577 815 1577 815 1586 814 c 1,51,52
+ 1595 815 1595 815 1612 815 c 2,53,-1
+ 2922 815 l 1,54,55
+ 2860 883 2860 883 2812 947 c 128,-1,56
+ 2764 1011 2764 1011 2742 1050.5 c 128,-1,57
+ 2720 1090 2720 1090 2710 1112.5 c 128,-1,58
+ 2700 1135 2700 1135 2700 1139 c 0,59,60
+ 2700 1151 2700 1151 2707.5 1157.5 c 128,-1,61
+ 2715 1164 2715 1164 2722.5 1165 c 128,-1,62
+ 2730 1166 2730 1166 2744 1166 c 0,63,64
+ 2768 1166 2768 1166 2777 1159 c 0,65,66
+ 2779 1158 2779 1158 2780.5 1156 c 128,-1,67
+ 2782 1154 2782 1154 2784 1150 c 2,68,-1
+ 2792 1134 l 2,69,70
+ 2798 1123 2798 1123 2806 1110 c 0,71,72
+ 2873 1001 2873 1001 2883 987 c 0,73,74
+ 2926 927 2926 927 2985 866 c 0,75,76
+ 3045 805 3045 805 3114 754 c 128,-1,77
+ 3183 703 3183 703 3275 658 c 128,-1,78
+ 3367 613 3367 613 3470 584 c 0,79,80
+ 3490 578 3490 578 3497.5 572.5 c 128,-1,81
+ 3505 567 3505 567 3505 554.5 c 128,-1,82
+ 3505 542 3505 542 3497 537 c 128,-1,83
+ 3489 532 3489 532 3468 526 c 0,84,85
+ 3340 489 3340 489 3232 431 c 0,86,87
+ 3129 375 3129 375 3043 298 c 0,88,89
+ 2961 224 2961 224 2904 149 c 0,90,91
+ 2845 73 2845 73 2799 -13 c 0,92,93
+ 2783 -43 2783 -43 2775 -50 c 0,94,95
+ 2768 -56 2768 -56 2744 -56 c 0,96,97
+ 2730 -56 2730 -56 2722.5 -55 c 128,-1,98
+ 2715 -54 2715 -54 2707.5 -47.5 c 128,-1,99
+ 2700 -41 2700 -41 2700 -29 c 0,100,101
+ 2700 -24 2700 -24 2710 -1.5 c 128,-1,102
+ 2720 21 2720 21 2742 61 c 128,-1,103
+ 2764 101 2764 101 2811.5 164.5 c 128,-1,104
+ 2859 228 2859 228 2922 295 c 1,105,-1
+ 1612 295 l 2,106,107
+ 1596 295 1596 295 1586 296 c 1,108,109
+ 1577 295 1577 295 1561 295 c 2,0,-1
+EndSplineSet
+Validated: 1025
 EndChar
 
 StartChar: uni27FA
@@ -36405,138 +36279,106 @@
 Width: 3663
 VWidth: 2220
 Flags: W
-HStem: 295 89<665 2061 1593 2938> 726 89<725 2061 1593 2938>
-LayerCount: 2
-Fore
-SplineSet
-2035 295 m 6,0,-1
- 725 295 l 1,1,2
- 787 227 787 227 834.5 163 c 128,-1,3
- 882 99 882 99 904 59.5 c 128,-1,4
- 926 20 926 20 936.5 -2.5 c 128,-1,5
- 947 -25 947 -25 947 -29 c 0,6,7
- 947 -41 947 -41 939 -47.5 c 128,-1,8
- 931 -54 931 -54 923.5 -55 c 128,-1,9
- 916 -56 916 -56 902 -56 c 0,10,11
- 878 -56 878 -56 869 -49 c 0,12,13
- 867 -48 867 -48 865.5 -46 c 128,-1,14
- 864 -44 864 -44 862 -40 c 128,-1,15
- 860 -36 860 -36 854 -24.5 c 128,-1,16
- 848 -13 848 -13 840 0 c 0,17,18
- 805 64 805 64 763.5 123 c 128,-1,19
- 722 182 722 182 662 243.5 c 128,-1,20
- 602 305 602 305 532.5 356 c 128,-1,21
- 463 407 463 407 371.5 452 c 128,-1,22
- 280 497 280 497 176 526 c 0,23,24
- 156 532 156 532 148.5 537.5 c 128,-1,25
- 141 543 141 543 141 555.5 c 128,-1,26
- 141 568 141 568 149 573 c 128,-1,27
- 157 578 157 578 179 584 c 0,28,29
- 306 620 306 620 414 679 c 128,-1,30
- 522 738 522 738 603 811.5 c 128,-1,31
- 684 885 684 885 743 961 c 128,-1,32
- 802 1037 802 1037 847 1123 c 0,33,34
- 864 1154 864 1154 871 1160 c 128,-1,35
- 878 1166 878 1166 902 1166 c 0,36,37
- 916 1166 916 1166 923.5 1165 c 128,-1,38
- 931 1164 931 1164 939 1157.5 c 128,-1,39
- 947 1151 947 1151 947 1139 c 0,40,41
- 947 1134 947 1134 936.5 1111.5 c 128,-1,42
- 926 1089 926 1089 904 1049 c 128,-1,43
- 882 1009 882 1009 834.5 945.5 c 128,-1,44
- 787 882 787 882 725 815 c 1,45,-1
- 2035 815 l 2,46,47
- 2051 815 2051 815 2060 814 c 128,-1,48
- 2069 813 2069 813 2082 809.5 c 128,-1,49
- 2095 806 2095 806 2101.5 796 c 128,-1,50
- 2108 786 2108 786 2108 770 c 0,51,52
- 2108 757 2108 757 2103 748.5 c 128,-1,53
- 2098 740 2098 740 2091.5 735.5 c 128,-1,54
- 2085 731 2085 731 2073 729 c 128,-1,55
- 2061 727 2061 727 2053.5 726.5 c 128,-1,56
- 2046 726 2046 726 2032 726 c 2,57,-1
- 665 726 l 2,58,59
- 637 726 637 726 627.5 723 c 128,-1,60
- 618 720 618 720 598 704 c 0,61,62
- 467 609 467 609 343 555 c 1,63,64
- 478 494 478 494 596 406 c 0,65,66
- 617 390 617 390 627 387 c 128,-1,67
- 637 384 637 384 665 384 c 2,68,-1
- 2032 384 l 2,69,70
- 2046 384 2046 384 2053.5 383.5 c 128,-1,71
- 2061 383 2061 383 2073 381 c 128,-1,72
- 2085 379 2085 379 2091.5 374.5 c 128,-1,73
- 2098 370 2098 370 2103 361 c 128,-1,74
- 2108 352 2108 352 2108 340 c 0,75,76
- 2108 324 2108 324 2101.5 314 c 128,-1,77
- 2095 304 2095 304 2082 300.5 c 128,-1,78
- 2069 297 2069 297 2060 296 c 128,-1,79
- 2051 295 2051 295 2035 295 c 6,0,-1
-2998 726 m 6,80,-1
- 1631 726 l 2,81,82
- 1617 726 1617 726 1609.5 726.5 c 128,-1,83
- 1602 727 1602 727 1590 729 c 128,-1,84
- 1578 731 1578 731 1571.5 735.5 c 128,-1,85
- 1565 740 1565 740 1560 749 c 128,-1,86
- 1555 758 1555 758 1555 770 c 0,87,88
- 1555 786 1555 786 1561.5 796 c 128,-1,89
- 1568 806 1568 806 1581 809.5 c 128,-1,90
- 1594 813 1594 813 1603 814 c 128,-1,91
- 1612 815 1612 815 1628 815 c 2,92,-1
- 2938 815 l 1,93,94
- 2876 883 2876 883 2828.5 947 c 128,-1,95
- 2781 1011 2781 1011 2759 1050.5 c 128,-1,96
- 2737 1090 2737 1090 2726.5 1112.5 c 128,-1,97
- 2716 1135 2716 1135 2716 1139 c 0,98,99
- 2716 1151 2716 1151 2724 1157.5 c 128,-1,100
- 2732 1164 2732 1164 2739.5 1165 c 128,-1,101
- 2747 1166 2747 1166 2761 1166 c 0,102,103
- 2785 1166 2785 1166 2794 1159 c 0,104,105
- 2796 1158 2796 1158 2797.5 1156 c 128,-1,106
- 2799 1154 2799 1154 2801 1150 c 128,-1,107
- 2803 1146 2803 1146 2809 1134.5 c 128,-1,108
- 2815 1123 2815 1123 2823 1110 c 0,109,110
- 2858 1046 2858 1046 2899.5 987 c 128,-1,111
- 2941 928 2941 928 3001 866.5 c 128,-1,112
- 3061 805 3061 805 3130.5 754 c 128,-1,113
- 3200 703 3200 703 3291.5 658 c 128,-1,114
- 3383 613 3383 613 3487 584 c 0,115,116
- 3507 578 3507 578 3514.5 572.5 c 128,-1,117
- 3522 567 3522 567 3522 554.5 c 128,-1,118
- 3522 542 3522 542 3514 537 c 128,-1,119
- 3506 532 3506 532 3484 526 c 0,120,121
- 3357 490 3357 490 3249 431 c 128,-1,122
- 3141 372 3141 372 3060 298.5 c 128,-1,123
- 2979 225 2979 225 2920 149 c 128,-1,124
- 2861 73 2861 73 2816 -13 c 0,125,126
- 2799 -44 2799 -44 2792 -50 c 128,-1,127
- 2785 -56 2785 -56 2761 -56 c 0,128,129
- 2747 -56 2747 -56 2739.5 -55 c 128,-1,130
- 2732 -54 2732 -54 2724 -47.5 c 128,-1,131
- 2716 -41 2716 -41 2716 -29 c 0,132,133
- 2716 -24 2716 -24 2726.5 -1.5 c 128,-1,134
- 2737 21 2737 21 2759 61 c 128,-1,135
- 2781 101 2781 101 2828.5 164.5 c 128,-1,136
- 2876 228 2876 228 2938 295 c 1,137,-1
- 1628 295 l 2,138,139
- 1612 295 1612 295 1603 296 c 128,-1,140
- 1594 297 1594 297 1581 300.5 c 128,-1,141
- 1568 304 1568 304 1561.5 314 c 128,-1,142
- 1555 324 1555 324 1555 340 c 0,143,144
- 1555 353 1555 353 1560 361.5 c 128,-1,145
- 1565 370 1565 370 1571.5 374.5 c 128,-1,146
- 1578 379 1578 379 1590 381 c 128,-1,147
- 1602 383 1602 383 1609.5 383.5 c 128,-1,148
- 1617 384 1617 384 1631 384 c 2,149,-1
- 2998 384 l 2,150,151
- 3026 384 3026 384 3035.5 387 c 128,-1,152
- 3045 390 3045 390 3065 406 c 0,153,154
- 3196 501 3196 501 3320 555 c 1,155,156
- 3185 616 3185 616 3067 704 c 0,157,158
- 3046 720 3046 720 3036 723 c 128,-1,159
- 3026 726 3026 726 2998 726 c 6,80,-1
-EndSplineSet
-Validated: 1029
+HStem: 295 89<725 2938> 726 89<725 2938>
+LayerCount: 2
+Fore
+SplineSet
+2032 384 m 1,0,-1
+ 2998 384 l 2,1,2
+ 3026 384 3026 384 3035.5 387 c 128,-1,3
+ 3045 390 3045 390 3065 406 c 1,4,5
+ 3196 501 3196 501 3320 555 c 1,6,7
+ 3185 616 3185 616 3067 704 c 0,8,9
+ 3046 720 3046 720 3036 723 c 128,-1,10
+ 3026 726 3026 726 2998 726 c 2,11,-1
+ 2032 726 l 1,12,-1
+ 1631 726 l 1,13,-1
+ 665 726 l 2,14,15
+ 637 726 637 726 627.5 723 c 128,-1,16
+ 618 720 618 720 598 704 c 1,17,18
+ 467 609 467 609 343 555 c 1,19,20
+ 478 494 478 494 596 406 c 0,21,22
+ 618 390 618 390 627 387 c 0,23,24
+ 637 384 637 384 665 384 c 2,25,-1
+ 1631 384 l 1,26,-1
+ 2032 384 l 1,0,-1
+2035 815 m 1,27,-1
+ 2938 815 l 1,28,29
+ 2876 883 2876 883 2828.5 947 c 128,-1,30
+ 2781 1011 2781 1011 2759 1050.5 c 128,-1,31
+ 2737 1090 2737 1090 2726.5 1112.5 c 128,-1,32
+ 2716 1135 2716 1135 2716 1139 c 0,33,34
+ 2716 1151 2716 1151 2724 1157.5 c 128,-1,35
+ 2732 1164 2732 1164 2739.5 1165 c 128,-1,36
+ 2747 1166 2747 1166 2761 1166 c 0,37,38
+ 2785 1166 2785 1166 2794 1159 c 0,39,40
+ 2796 1158 2796 1158 2797.5 1156 c 128,-1,41
+ 2799 1154 2799 1154 2801 1150 c 2,42,-1
+ 2809 1134 l 2,43,44
+ 2815 1123 2815 1123 2823 1110 c 0,45,46
+ 2893 996 2893 996 2900 987 c 0,47,48
+ 2941 928 2941 928 3001 866.5 c 128,-1,49
+ 3061 805 3061 805 3130.5 754 c 128,-1,50
+ 3200 703 3200 703 3291.5 658 c 128,-1,51
+ 3383 613 3383 613 3487 584 c 0,52,53
+ 3507 578 3507 578 3514.5 572.5 c 128,-1,54
+ 3522 567 3522 567 3522 554.5 c 128,-1,55
+ 3522 542 3522 542 3514 537 c 0,56,57
+ 3507 532 3507 532 3484 526 c 0,58,59
+ 3357 490 3357 490 3249 431 c 128,-1,60
+ 3141 372 3141 372 3060 298.5 c 128,-1,61
+ 2979 225 2979 225 2920 149 c 128,-1,62
+ 2861 73 2861 73 2816 -13 c 0,63,64
+ 2800 -43 2800 -43 2792 -50 c 0,65,66
+ 2785 -56 2785 -56 2761 -56 c 0,67,68
+ 2747 -56 2747 -56 2739.5 -55 c 128,-1,69
+ 2732 -54 2732 -54 2724 -47.5 c 128,-1,70
+ 2716 -41 2716 -41 2716 -29 c 0,71,72
+ 2716 -24 2716 -24 2726.5 -1.5 c 128,-1,73
+ 2737 21 2737 21 2759 61 c 128,-1,74
+ 2781 101 2781 101 2828.5 164.5 c 128,-1,75
+ 2876 228 2876 228 2938 295 c 1,76,-1
+ 2035 295 l 1,77,-1
+ 1628 295 l 1,78,-1
+ 725 295 l 1,79,80
+ 781 233 781 233 834 163 c 0,81,82
+ 882 99 882 99 904 59 c 0,83,84
+ 926 20 926 20 936.5 -2.5 c 128,-1,85
+ 947 -25 947 -25 947 -29 c 0,86,87
+ 947 -41 947 -41 939 -47.5 c 128,-1,88
+ 931 -54 931 -54 923.5 -55 c 128,-1,89
+ 916 -56 916 -56 902 -56 c 0,90,91
+ 878 -56 878 -56 869 -49 c 0,92,93
+ 867 -48 867 -48 865.5 -46 c 128,-1,94
+ 864 -44 864 -44 862 -40 c 2,95,-1
+ 854 -24 l 2,96,97
+ 848 -13 848 -13 840 0 c 0,98,99
+ 770 114 770 114 764 123 c 0,100,101
+ 722 182 722 182 662 243.5 c 128,-1,102
+ 602 305 602 305 532.5 356 c 128,-1,103
+ 463 407 463 407 371.5 452 c 128,-1,104
+ 280 497 280 497 176 526 c 0,105,106
+ 156 532 156 532 148.5 537.5 c 128,-1,107
+ 141 543 141 543 141 555.5 c 128,-1,108
+ 141 568 141 568 149 573 c 0,109,110
+ 156 578 156 578 179 584 c 0,111,112
+ 306 620 306 620 414 679 c 128,-1,113
+ 522 738 522 738 603 811.5 c 128,-1,114
+ 684 885 684 885 743 961 c 128,-1,115
+ 802 1037 802 1037 847 1123 c 0,116,117
+ 864 1154 864 1154 871 1160 c 128,-1,118
+ 878 1166 878 1166 902 1166 c 0,119,120
+ 916 1166 916 1166 923.5 1165 c 128,-1,121
+ 931 1164 931 1164 939 1157.5 c 128,-1,122
+ 947 1151 947 1151 947 1139 c 0,123,124
+ 947 1134 947 1134 936.5 1111.5 c 128,-1,125
+ 926 1089 926 1089 904 1049 c 128,-1,126
+ 882 1009 882 1009 834.5 945.5 c 128,-1,127
+ 787 882 787 882 725 815 c 1,128,-1
+ 1628 815 l 1,129,-1
+ 2035 815 l 1,27,-1
+EndSplineSet
+Validated: 1025
 EndChar
 
 StartChar: uni27FC
@@ -36544,99 +36386,83 @@
 Width: 3663
 VWidth: 2220
 Flags: W
-HStem: 511 89<257 273 320 1569 1569 1569 1569 3308> 924 44G<190 235>
-VStem: 168 89<163 511 599 947> 256 63<511 599> 1528 84<511 599>
-LayerCount: 2
-Fore
-SplineSet
-1539 511 m 6,0,-1
- 329 511 l 2,1,2
- 313 511 313 511 304 511.5 c 128,-1,3
- 295 512 295 512 281.5 516 c 128,-1,4
- 268 520 268 520 262 529.5 c 128,-1,5
- 256 539 256 539 256 555 c 128,-1,6
- 256 571 256 571 262 580.5 c 128,-1,7
- 268 590 268 590 281.5 594 c 128,-1,8
- 295 598 295 598 304 598.5 c 128,-1,9
- 313 599 313 599 329 599 c 2,10,-1
- 1539 599 l 2,11,12
- 1555 599 1555 599 1564 598.5 c 128,-1,13
- 1573 598 1573 598 1586.5 594 c 128,-1,14
- 1600 590 1600 590 1606 580.5 c 128,-1,15
- 1612 571 1612 571 1612 555 c 128,-1,16
- 1612 539 1612 539 1606 529.5 c 128,-1,17
- 1600 520 1600 520 1586.5 516 c 128,-1,18
- 1573 512 1573 512 1564 511.5 c 128,-1,19
- 1555 511 1555 511 1539 511 c 6,0,-1
-3308 511 m 5,20,-1
- 1601 511 l 2,21,22
- 1587 511 1587 511 1580 511 c 128,-1,23
- 1573 511 1573 511 1561.5 513.5 c 128,-1,24
- 1550 516 1550 516 1544 520.5 c 128,-1,25
- 1538 525 1538 525 1533 533.5 c 128,-1,26
- 1528 542 1528 542 1528 555 c 128,-1,27
- 1528 568 1528 568 1533 576.5 c 128,-1,28
- 1538 585 1538 585 1544 589.5 c 128,-1,29
- 1550 594 1550 594 1561.5 596.5 c 128,-1,30
- 1573 599 1573 599 1580.5 599 c 128,-1,31
- 1588 599 1588 599 1601 599 c 2,32,-1
- 3308 599 l 1,33,34
- 3250 641 3250 641 3201 698 c 128,-1,35
- 3152 755 3152 755 3125 802 c 128,-1,36
- 3098 849 3098 849 3083.5 882 c 128,-1,37
- 3069 915 3069 915 3069 924 c 0,38,39
- 3069 938 3069 938 3078.5 944 c 128,-1,40
- 3088 950 3088 950 3102 950 c 0,41,42
- 3119 950 3119 950 3125.5 943.5 c 128,-1,43
- 3132 937 3132 937 3140 919 c 0,44,45
- 3173 846 3173 846 3213.5 788.5 c 128,-1,46
- 3254 731 3254 731 3300 691.5 c 128,-1,47
- 3346 652 3346 652 3387 626.5 c 128,-1,48
- 3428 601 3428 601 3479 577 c 1,49,50
- 3482 577 3482 577 3488.5 571 c 128,-1,51
- 3495 565 3495 565 3495 555 c 0,52,53
- 3495 544 3495 544 3490 539 c 128,-1,54
- 3485 534 3485 534 3464 524 c 0,55,56
- 3418 502 3418 502 3377.5 475.5 c 128,-1,57
- 3337 449 3337 449 3308 424.5 c 128,-1,58
- 3279 400 3279 400 3252.5 369 c 128,-1,59
- 3226 338 3226 338 3210 316 c 128,-1,60
- 3194 294 3194 294 3177.5 264 c 128,-1,61
- 3161 234 3161 234 3154.5 220 c 128,-1,62
- 3148 206 3148 206 3137 182 c 0,63,64
- 3133 174 3133 174 3131.5 171.5 c 128,-1,65
- 3130 169 3130 169 3122 164.5 c 128,-1,66
- 3114 160 3114 160 3102 160 c 0,67,68
- 3088 160 3088 160 3078.5 166.5 c 128,-1,69
- 3069 173 3069 173 3069 186 c 0,70,71
- 3069 195 3069 195 3083.5 228 c 128,-1,72
- 3098 261 3098 261 3125 308 c 128,-1,73
- 3152 355 3152 355 3201 412 c 128,-1,74
- 3250 469 3250 469 3308 511 c 5,20,-1
-257 511 m 5,75,-1
- 257 215 l 2,76,77
- 257 199 257 199 256 190 c 128,-1,78
- 255 181 255 181 251.5 168 c 128,-1,79
- 248 155 248 155 238 148.5 c 128,-1,80
- 228 142 228 142 212.5 142 c 128,-1,81
- 197 142 197 142 187 148.5 c 128,-1,82
- 177 155 177 155 173.5 168 c 128,-1,83
- 170 181 170 181 169 190 c 128,-1,84
- 168 199 168 199 168 215 c 2,85,-1
- 168 895 l 2,86,87
- 168 911 168 911 169 920 c 128,-1,88
- 170 929 170 929 173.5 942 c 128,-1,89
- 177 955 177 955 187 961.5 c 128,-1,90
- 197 968 197 968 212.5 968 c 128,-1,91
- 228 968 228 968 238 961.5 c 128,-1,92
- 248 955 248 955 251.5 942 c 128,-1,93
- 255 929 255 929 256 920 c 128,-1,94
- 257 911 257 911 257 895 c 2,95,-1
- 257 599 l 1,96,97
- 319 599 319 599 319 555 c 128,-1,98
- 319 511 319 511 257 511 c 5,75,-1
-EndSplineSet
-Validated: 1029
+HStem: 511 88<257 286.567 287.269 3308>
+VStem: 168 89<142.002 511 599 967.998> 3069 71<849.875 943.302>
+LayerCount: 2
+Fore
+SplineSet
+257 511 m 1,0,-1
+ 257 215 l 2,1,2
+ 257 199 257 199 256 190 c 128,-1,3
+ 255 181 255 181 251.5 168 c 128,-1,4
+ 248 155 248 155 238 148.5 c 128,-1,5
+ 228 142 228 142 212.5 142 c 128,-1,6
+ 197 142 197 142 187 148.5 c 128,-1,7
+ 177 155 177 155 173.5 168 c 128,-1,8
+ 170 181 170 181 169 190 c 128,-1,9
+ 168 199 168 199 168 215 c 2,10,-1
+ 168 895 l 2,11,12
+ 168 911 168 911 169 920 c 128,-1,13
+ 170 929 170 929 173.5 942 c 128,-1,14
+ 177 955 177 955 187 961.5 c 128,-1,15
+ 197 968 197 968 212.5 968 c 128,-1,16
+ 228 968 228 968 238 961.5 c 128,-1,17
+ 248 955 248 955 251.5 942 c 128,-1,18
+ 255 929 255 929 256 920 c 128,-1,19
+ 257 911 257 911 257 895 c 2,20,-1
+ 257 599 l 1,21,22
+ 274 599 274 599 287 596 c 1,23,24
+ 297 598 297 598 304 598.5 c 128,-1,25
+ 311 599 311 599 329 599 c 2,26,-1
+ 1539 599 l 2,27,28
+ 1555 599 1555 599 1564 598 c 0,29,30
+ 1569 598 1569 598 1569 598 c 0,31,32
+ 1576 599 1576 599 1580 599 c 2,33,-1
+ 1601 599 l 1,34,-1
+ 3308 599 l 1,35,36
+ 3250 641 3250 641 3201 698 c 128,-1,37
+ 3152 755 3152 755 3125 802 c 128,-1,38
+ 3098 849 3098 849 3083.5 882 c 128,-1,39
+ 3069 915 3069 915 3069 924 c 0,40,41
+ 3069 938 3069 938 3078.5 944 c 128,-1,42
+ 3088 950 3088 950 3102 950 c 0,43,44
+ 3119 950 3119 950 3126 944 c 0,45,46
+ 3132 937 3132 937 3140 919 c 0,47,48
+ 3171 848 3171 848 3214 788 c 0,49,50
+ 3254 731 3254 731 3300 691.5 c 128,-1,51
+ 3346 652 3346 652 3387 626.5 c 128,-1,52
+ 3428 601 3428 601 3479 577 c 1,53,54
+ 3482 577 3482 577 3488.5 571 c 128,-1,55
+ 3495 565 3495 565 3495 555 c 0,56,57
+ 3495 544 3495 544 3490 539 c 128,-1,58
+ 3485 534 3485 534 3464 524 c 0,59,60
+ 3418 502 3418 502 3377.5 475.5 c 128,-1,61
+ 3337 449 3337 449 3308 424.5 c 128,-1,62
+ 3279 400 3279 400 3252.5 369 c 128,-1,63
+ 3226 338 3226 338 3210 316 c 128,-1,64
+ 3194 294 3194 294 3177.5 264 c 128,-1,65
+ 3161 234 3161 234 3154 220 c 2,66,-1
+ 3137 182 l 2,67,68
+ 3133 174 3133 174 3131.5 171.5 c 128,-1,69
+ 3130 169 3130 169 3122 164.5 c 128,-1,70
+ 3114 160 3114 160 3102 160 c 0,71,72
+ 3088 160 3088 160 3078.5 166.5 c 128,-1,73
+ 3069 173 3069 173 3069 186 c 0,74,75
+ 3069 195 3069 195 3083.5 228 c 128,-1,76
+ 3098 261 3098 261 3125 308 c 128,-1,77
+ 3152 355 3152 355 3201 412 c 128,-1,78
+ 3250 469 3250 469 3308 511 c 1,79,-1
+ 1601 511 l 1,80,-1
+ 1580 511 l 2,81,82
+ 1576 511 1576 511 1569 512 c 0,83,84
+ 1566 512 1566 512 1564 511.5 c 128,-1,85
+ 1562 511 1562 511 1539 511 c 2,86,-1
+ 329 511 l 2,87,88
+ 313 511 313 511 304 512 c 0,89,90
+ 297 512 297 512 287 514 c 1,91,92
+ 274 511 274 511 257 511 c 1,0,-1
+EndSplineSet
+Validated: 1
 EndChar
 
 StartChar: uni2983
@@ -37269,61 +37095,54 @@
 Width: 2220
 VWidth: 2220
 Flags: W
-HStem: -11 47G<211 236 1981 2009>
+HStem: -11 21G<1982 1993.5>
 VStem: 189 89<111 999> 1942 89<111 999>
 LayerCount: 2
 Fore
 SplineSet
-1143 506 m 6,0,-1
- 273 4 l 2,1,2
- 246 -12 246 -12 233 -11 c 0,3,4
- 217 -11 217 -11 207.5 -4 c 128,-1,5
- 198 3 198 3 194 17.5 c 128,-1,6
- 190 32 190 32 189.5 41.5 c 128,-1,7
- 189 51 189 51 189 69 c 2,8,-1
- 189 1041 l 2,9,10
- 189 1058 189 1058 189.5 1068 c 128,-1,11
- 190 1078 190 1078 194 1092.5 c 128,-1,12
- 198 1107 198 1107 207.5 1114 c 128,-1,13
- 217 1121 217 1121 233 1121 c 0,14,15
- 250 1121 250 1121 278 1103 c 2,16,-1
- 1143 604 l 2,17,18
- 1160 595 1160 595 1169 585 c 128,-1,19
- 1178 575 1178 575 1179.5 569 c 128,-1,20
- 1181 563 1181 563 1181 555 c 128,-1,21
- 1181 547 1181 547 1179.5 541 c 128,-1,22
- 1178 535 1178 535 1169 525 c 128,-1,23
- 1160 515 1160 515 1143 506 c 6,0,-1
-1048 555 m 5,24,-1
- 278 999 l 1,25,-1
- 278 111 l 1,26,-1
- 1048 555 l 5,24,-1
-1086 606 m 2,27,-1
- 1949 1106 l 2,28,29
- 1976 1122 1976 1122 1987 1121 c 0,30,31
- 2000 1121 2000 1121 2008.5 1116 c 128,-1,32
- 2017 1111 2017 1111 2021.5 1105 c 128,-1,33
- 2026 1099 2026 1099 2028.5 1086 c 128,-1,34
- 2031 1073 2031 1073 2031 1065 c 128,-1,35
- 2031 1057 2031 1057 2031 1041 c 2,36,-1
- 2031 69 l 2,37,38
- 2031 53 2031 53 2031 45 c 128,-1,39
- 2031 37 2031 37 2028.5 24 c 128,-1,40
- 2026 11 2026 11 2021.5 5 c 128,-1,41
- 2017 -1 2017 -1 2008.5 -6 c 128,-1,42
- 2000 -11 2000 -11 1987 -11 c 0,43,44
- 1977 -11 1977 -11 1945 7 c 2,45,-1
- 1081 504 l 2,46,47
- 1039 526 1039 526 1039 555 c 0,48,49
- 1039 567 1039 567 1045 576 c 128,-1,50
- 1051 585 1051 585 1059 590 c 128,-1,51
- 1067 595 1067 595 1086 606 c 2,27,-1
-1172 555 m 1,52,-1
- 1942 111 l 1,53,-1
- 1942 999 l 1,54,-1
- 1172 555 l 1,52,-1
-EndSplineSet
-Validated: 5
+1039 550 m 1,0,1
+ 1039 552 1039 552 1039 555 c 128,-1,2
+ 1039 558 1039 558 1039 560 c 1,3,-1
+ 278 999 l 1,4,-1
+ 278 111 l 1,5,-1
+ 1039 550 l 1,0,1
+1181 550 m 1,6,-1
+ 1942 111 l 1,7,-1
+ 1942 999 l 1,8,-1
+ 1181 560 l 1,9,10
+ 1181 558 1181 558 1181 555 c 128,-1,11
+ 1181 552 1181 552 1181 550 c 1,6,-1
+1113 621 m 1,12,-1
+ 1949 1106 l 2,13,14
+ 1977 1122 1977 1122 1987 1121 c 1,15,16
+ 2000 1121 2000 1121 2008.5 1116 c 128,-1,17
+ 2017 1111 2017 1111 2021.5 1105 c 128,-1,18
+ 2026 1099 2026 1099 2028.5 1086 c 128,-1,19
+ 2031 1073 2031 1073 2031 1065 c 2,20,-1
+ 2031 1041 l 1,21,-1
+ 2031 69 l 1,22,-1
+ 2031 45 l 2,23,24
+ 2031 37 2031 37 2028.5 24 c 128,-1,25
+ 2026 11 2026 11 2021.5 5 c 128,-1,26
+ 2017 -1 2017 -1 2008.5 -6 c 128,-1,27
+ 2000 -11 2000 -11 1987 -11 c 0,28,29
+ 1977 -11 1977 -11 1945 7 c 2,30,-1
+ 1110 487 l 1,31,-1
+ 273 4 l 2,32,33
+ 246 -12 246 -12 233 -11 c 1,34,35
+ 217 -11 217 -11 208 -4 c 0,36,37
+ 198 3 198 3 194 17.5 c 128,-1,38
+ 190 32 190 32 189.5 41.5 c 128,-1,39
+ 189 51 189 51 189 69 c 2,40,-1
+ 189 1041 l 2,41,42
+ 189 1058 189 1058 189.5 1068 c 128,-1,43
+ 190 1078 190 1078 194 1092.5 c 128,-1,44
+ 198 1107 198 1107 207.5 1114 c 128,-1,45
+ 217 1121 217 1121 233 1121 c 0,46,47
+ 250 1121 250 1121 278 1103 c 1,48,-1
+ 1113 621 l 1,12,-1
+EndSplineSet
+Validated: 1
 EndChar
 
 StartChar: uni2A3F
Binary file lib/fonts/IsabelleText.ttf has changed
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/lib/scripts/bash	Mon Feb 08 14:08:32 2010 +0100
@@ -0,0 +1,31 @@
+#!/usr/bin/env perl
+#
+# Author: Makarius
+#
+# bash - invoke shell command line (with robust signal handling)
+#
+
+use warnings;
+use strict;
+
+
+# args
+
+my ($group, $script_name, $pid_name, $output_name) = @ARGV;
+
+
+# process id
+
+if ($group eq "group") {
+  use POSIX "setsid";
+  POSIX::setsid || die $!;
+}
+
+open (PID_FILE, ">", $pid_name) || die $!;
+print PID_FILE "$$";
+close PID_FILE;
+
+
+# exec script
+
+exec qq/exec bash '$script_name' > '$output_name'/;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/lib/scripts/keywords	Mon Feb 08 14:08:32 2010 +0100
@@ -0,0 +1,124 @@
+#!/usr/bin/env perl
+#
+# Author: Makarius
+#
+# keywords.pl - generate outer syntax keyword files from session logs
+#
+
+use warnings;
+use strict;
+
+
+## arguments
+
+my ($keywords_name, $sessions) = @ARGV;
+
+
+## keywords
+
+my %keywords;
+
+sub set_keyword {
+  my ($name, $kind) = @_;
+  if (defined $keywords{$name} and $keywords{$name} ne $kind and $keywords{$name} ne "minor") {
+    if ($kind ne "minor") {
+      print STDERR "### Inconsistent declaration of keyword \"${name}\": $keywords{$name} vs ${kind}\n";
+      $keywords{$name} = $kind;
+    }
+  } else {
+    $keywords{$name} = $kind;
+  }
+}
+
+sub collect_keywords {
+  while(<STDIN>) {
+    if (m/^Outer syntax keyword:\s*"(.*)"/) {
+      my $name = $1;
+      &set_keyword($name, "minor");
+    }
+    elsif (m/^Outer syntax command:\s*"(.*)"\s*\((.*)\)/) {
+      my $name = $1;
+      my $kind = $2;
+      &set_keyword($name, $kind);
+    }
+  }
+}
+
+
+## Emacs output
+
+sub emacs_output {
+  my @kinds = (
+    "major",
+    "minor",
+    "control",
+    "diag",
+    "theory-begin",
+    "theory-switch",
+    "theory-end",
+    "theory-heading",
+    "theory-decl",
+    "theory-script",
+    "theory-goal",
+    "qed",
+    "qed-block",
+    "qed-global",
+    "proof-heading",
+    "proof-goal",
+    "proof-block",
+    "proof-open",
+    "proof-close",
+    "proof-chain",
+    "proof-decl",
+    "proof-asm",
+    "proof-asm-goal",
+    "proof-script"
+  );
+  my $file = $keywords_name eq "" ? "isar-keywords.el" : "isar-keywords-${keywords_name}.el";
+  open (OUTPUT, "> ${file}") || die "$!";
+  select OUTPUT;
+
+  print ";;\n";
+  print ";; Keyword classification tables for Isabelle/Isar.\n";
+  print ";; Generated from ${sessions}.\n";
+  print ";; *** DO NOT EDIT *** DO NOT EDIT *** DO NOT EDIT ***\n";
+  print ";;\n";
+
+  for my $kind (@kinds) {
+    my @names;
+    for my $name (keys(%keywords)) {
+      if ($kind eq "major" ? $keywords{$name} ne "minor" : $keywords{$name} eq $kind) {
+        if ($kind ne "minor" or $name =~ m/^[A-Za-z0-9_]+$/) {
+          push @names, $name;
+        }
+      }
+    }
+    @names = sort(@names);
+
+    print "\n(defconst isar-keywords-${kind}";
+    print "\n  '(";
+    my $first = 1;
+    for my $name (@names) {
+      $name =~ s/([\.\*\+\?\[\]\^\$])/\\\\$1/g;
+      if ($first) {
+        print "\"${name}\"";
+        $first = 0;
+      }
+      else {
+        print "\n    \"${name}\"";
+      }
+    }
+    print "))\n";
+  }
+  print "\n(provide 'isar-keywords)\n";
+
+  close OUTPUT;
+  select;
+  print STDERR "${file}\n";
+}
+
+
+## main
+
+&collect_keywords();
+&emacs_output();
--- a/lib/scripts/keywords.pl	Mon Feb 08 14:06:58 2010 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,120 +0,0 @@
-#
-# Author: Makarius
-#
-# keywords.pl - generate outer syntax keyword files from session logs
-#
-
-## arguments
-
-my ($keywords_name, $sessions) = @ARGV;
-
-
-## keywords
-
-my %keywords;
-
-sub set_keyword {
-  my ($name, $kind) = @_;
-  if (defined $keywords{$name} and $keywords{$name} ne $kind and $keywords{$name} ne "minor") {
-    if ($kind ne "minor") {
-      print STDERR "### Inconsistent declaration of keyword \"${name}\": $keywords{$name} vs ${kind}\n";
-      $keywords{$name} = $kind;
-    }
-  } else {
-    $keywords{$name} = $kind;
-  }
-}
-
-sub collect_keywords {
-  while(<STDIN>) {
-    if (m/^Outer syntax keyword:\s*"(.*)"/) {
-      my $name = $1;
-      &set_keyword($name, "minor");
-    }
-    elsif (m/^Outer syntax command:\s*"(.*)"\s*\((.*)\)/) {
-      my $name = $1;
-      my $kind = $2;
-      &set_keyword($name, $kind);
-    }
-  }
-}
-
-
-## Emacs output
-
-sub emacs_output {
-  my @kinds = (
-    "major",
-    "minor",
-    "control",
-    "diag",
-    "theory-begin",
-    "theory-switch",
-    "theory-end",
-    "theory-heading",
-    "theory-decl",
-    "theory-script",
-    "theory-goal",
-    "qed",
-    "qed-block",
-    "qed-global",
-    "proof-heading",
-    "proof-goal",
-    "proof-block",
-    "proof-open",
-    "proof-close",
-    "proof-chain",
-    "proof-decl",
-    "proof-asm",
-    "proof-asm-goal",
-    "proof-script"
-  );
-  my $file = $keywords_name eq "" ? "isar-keywords.el" : "isar-keywords-${keywords_name}.el";
-  open (OUTPUT, "> ${file}") || die "$!";
-  select OUTPUT;
-
-  print ";;\n";
-  print ";; Keyword classification tables for Isabelle/Isar.\n";
-  print ";; Generated from ${sessions}.\n";
-  print ";; *** DO NOT EDIT *** DO NOT EDIT *** DO NOT EDIT ***\n";
-  print ";;\n";
-
-  for my $kind (@kinds) {
-    my @names;
-    for my $name (keys(%keywords)) {
-      if ($kind eq "major" ? $keywords{$name} ne "minor" : $keywords{$name} eq $kind) {
-        if ($kind ne "minor" or $name =~ m/^[A-Za-z0-9_]+$/) {
-          push @names, $name;
-        }
-      }
-    }
-    @names = sort(@names);
-
-    print "\n(defconst isar-keywords-${kind}";
-    print "\n  '(";
-    my $first = 1;
-    for my $name (@names) {
-      $name =~ s/([\.\*\+\?\[\]\^\$])/\\\\$1/g;
-      if ($first) {
-        print "\"${name}\"";
-        $first = 0;
-      }
-      else {
-        print "\n    \"${name}\"";
-      }
-    }
-    print "))\n";
-  }
-  print "\n(provide 'isar-keywords)\n";
-
-  close OUTPUT;
-  select;
-  print STDERR "${file}\n";
-}
-
-
-## main
-
-&collect_keywords();
-&emacs_output();
-
--- a/lib/scripts/run-mosml	Mon Feb 08 14:06:58 2010 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,57 +0,0 @@
-#!/usr/bin/env bash
-#
-# Author: Markus Wenzel, TU Muenchen
-#
-# Moscow ML 2.00 startup script
-
-export -n INFILE OUTFILE MLTEXT TERMINATE NOWRITE
-
-
-## diagnostics
-
-function fail_out()
-{
-  echo "Unable to create output heap file: \"$OUTFILE\"" >&2
-  exit 2
-}
-
-
-## prepare databases
-
-MOSML="mosml -P sml90"
-
-if [ -z "$INFILE" ]; then
-  EXIT='load "OS"; fun exit 0 = (OS.Process.exit OS.Process.success): unit | exit _ = OS.Process.exit OS.Process.failure;'
-else
-  EXIT=""
-  echo "Cannot load images yet!" >&2
-  exit 2
-fi
-
-if [ -z "$OUTFILE" ]; then
-  COMMIT='fun commit () = (TextIO.output (TextIO.stdErr, "Error - Database is not opened for writing.\n"); false);'
-else
-  COMMIT="fun commit () = true;"
-  echo "WARNING: cannot save images yet!" >&2
-  [ -f "$OUTFILE" ] && { chmod +w "$OUTFILE" || fail_out; }
-fi
-
-
-## run it!
-
-MLTEXT="$EXIT $COMMIT $MLTEXT"
-MLEXIT="commit();"
-
-if [ -z "$TERMINATE" ]; then
-  FEEDER_OPTS=""
-else
-  FEEDER_OPTS="-q"
-fi
-
-"$ISABELLE_HOME/lib/scripts/feeder" -p -h "$MLTEXT" -t "$MLEXIT" $FEEDER_OPTS | \
-  { read FPID; "$ML_HOME"/$MOSML $ML_OPTIONS 2>&1; RC="$?"; kill -HUP "$FPID"; exit "$RC"; }
-RC="$?"
-
-[ -n "$OUTFILE" -a -f "$OUTFILE" -a -n "$NOWRITE" ] && chmod -w "$OUTFILE"
-
-exit "$RC"
--- a/lib/scripts/system.pl	Mon Feb 08 14:06:58 2010 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,28 +0,0 @@
-#
-# Author: Makarius
-#
-# system.pl - invoke shell command line (with robust signal handling)
-#
-
-# args
-
-($group, $script_name, $pid_name, $output_name) = @ARGV;
-
-
-# process id
-
-if ($group eq "group") {
-  use POSIX "setsid";
-  POSIX::setsid || die $!;
-}
-
-open (PID_FILE, ">", $pid_name) || die $!;
-print PID_FILE "$$";
-close PID_FILE;
-
-
-# exec script
-
-$SIG{'INT'} = "DEFAULT";   #paranoia setting, required for Cygwin
-exec qq/exec bash '$script_name' > '$output_name'/;
-
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/lib/scripts/unsymbolize	Mon Feb 08 14:08:32 2010 +0100
@@ -0,0 +1,65 @@
+#!/usr/bin/env perl
+#
+# Author: Markus Wenzel, TU Muenchen
+#
+# unsymbolize.pl - remove unreadable symbol names from sources
+#
+
+use warnings;
+use strict;
+
+sub unsymbolize {
+    my ($file) = @_;
+
+    open (FILE, $file) || die $!;
+    undef $/; my $text = <FILE>; $/ = "\n";         # slurp whole file
+    close FILE || die $!;
+
+    $_ = $text;
+
+    # Pure
+    s/\\?\\<And>/!!/g;
+    s/\\?\\<Colon>/::/g;
+    s/\\?\\<Longrightarrow>/==>/g;
+    s/\\?\\<Midarrow>\\?\\<Rightarrow>/==>/g;
+    s/\\?\\<Rightarrow>/=>/g;
+    s/\\?\\<equiv>/==/g;
+    s/\\?\\<dots>/.../g;
+    s/\\?\\<lbrakk> ?/[| /g;
+    s/\\?\\ ?<rbrakk>/ |]/g;
+    s/\\?\\<lparr> ?/(| /g;
+    s/\\?\\ ?<rparr>/ |)/g;
+    # HOL
+    s/\\?\\<longleftrightarrow>/<->/g;
+    s/\\?\\<longrightarrow>/-->/g;
+    s/\\?\\<midarrow>\\?\\<rightarrow>/-->/g;
+    s/\\?\\<rightarrow>/->/g;
+    s/\\?\\<not>/~/g;
+    s/\\?\\<notin>/~:/g;
+    s/\\?\\<noteq>/~=/g;
+    s/\\?\\<some> ?/SOME /g;
+    # outer syntax
+    s/\\?\\<rightleftharpoons>/==/g;
+    s/\\?\\<rightharpoonup>/=>/g;
+    s/\\?\\<leftharpoondown>/<=/g;
+
+    my $result = $_;
+
+    if ($text ne $result) {
+	print STDERR "fixing $file\n";
+        if (! -f "$file~~") {
+	    rename $file, "$file~~" || die $!;
+        }
+	open (FILE, "> $file") || die $!;
+	print FILE $result;
+	close FILE || die $!;
+    }
+}
+
+
+## main
+
+foreach my $file (@ARGV) {
+  eval { &unsymbolize($file); };
+  if ($@) { print STDERR "*** unsymbolize $file: ", $@, "\n"; }
+}
--- a/lib/scripts/unsymbolize.pl	Mon Feb 08 14:06:58 2010 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,61 +0,0 @@
-#
-# Author: Markus Wenzel, TU Muenchen
-#
-# unsymbolize.pl - remove unreadable symbol names from sources
-#
-
-sub unsymbolize {
-    my ($file) = @_;
-
-    open (FILE, $file) || die $!;
-    undef $/; $text = <FILE>; $/ = "\n";         # slurp whole file
-    close FILE || die $!;
-
-    $_ = $text;
-
-    # Pure
-    s/\\?\\<And>/!!/g;
-    s/\\?\\<Colon>/::/g;
-    s/\\?\\<Longrightarrow>/==>/g;
-    s/\\?\\<Midarrow>\\?\\<Rightarrow>/==>/g;
-    s/\\?\\<Rightarrow>/=>/g;
-    s/\\?\\<equiv>/==/g;
-    s/\\?\\<dots>/.../g;
-    s/\\?\\<lbrakk> ?/[| /g;
-    s/\\?\\ ?<rbrakk>/ |]/g;
-    s/\\?\\<lparr> ?/(| /g;
-    s/\\?\\ ?<rparr>/ |)/g;
-    # HOL
-    s/\\?\\<longleftrightarrow>/<->/g;
-    s/\\?\\<longrightarrow>/-->/g;
-    s/\\?\\<midarrow>\\?\\<rightarrow>/-->/g;
-    s/\\?\\<rightarrow>/->/g;
-    s/\\?\\<not>/~/g;
-    s/\\?\\<notin>/~:/g;
-    s/\\?\\<noteq>/~=/g;
-    s/\\?\\<some> ?/SOME /g;
-    # outer syntax
-    s/\\?\\<rightleftharpoons>/==/g;
-    s/\\?\\<rightharpoonup>/=>/g;
-    s/\\?\\<leftharpoondown>/<=/g;
-
-    $result = $_;
-
-    if ($text ne $result) {
-	print STDERR "fixing $file\n";
-        if (! -f "$file~~") {
-	    rename $file, "$file~~" || die $!;
-        }
-	open (FILE, "> $file") || die $!;
-	print FILE $result;
-	close FILE || die $!;
-    }
-}
-
-
-## main
-
-foreach $file (@ARGV) {
-  eval { &unsymbolize($file); };
-  if ($@) { print STDERR "*** unsymbolize $file: ", $@, "\n"; }
-}
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/lib/scripts/yxml	Mon Feb 08 14:08:32 2010 +0100
@@ -0,0 +1,37 @@
+#!/usr/bin/env perl
+#
+# Author: Makarius
+#
+# yxml.pl - simple XML to YXML converter
+#
+
+use warnings;
+use strict;
+
+use XML::Parser;
+
+binmode(STDOUT, ":utf8");
+
+sub handle_start {
+  print chr(5), chr(6), $_[1];
+  for (my $i = 2; $i <= $#_; $i++) {
+    print ($i % 2 == 0 ? chr(6) : "=");
+    print $_[$i];
+  }
+  print chr(5);
+}
+
+sub handle_end {
+  print chr(5), chr(6), chr(5);
+}
+
+sub handle_char {
+  print $_[1];
+}
+
+my $parser = new XML::Parser(Handlers =>
+  {Start => \&handle_start,
+    End => \&handle_end,
+    Char => \&handle_char});
+
+$parser->parse(*STDIN) or die $!;
--- a/lib/scripts/yxml.pl	Mon Feb 08 14:06:58 2010 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,35 +0,0 @@
-#
-# Author: Makarius
-#
-# yxml.pl - simple XML to YXML converter
-#
-
-use strict;
-use XML::Parser;
-
-binmode(STDOUT, ":utf8");
-
-sub handle_start {
-  print chr(5), chr(6), $_[1];
-  for (my $i = 2; $i <= $#_; $i++) {
-    print ($i % 2 == 0 ? chr(6) : "=");
-    print $_[$i];
-  }
-  print chr(5);
-}
-
-sub handle_end {
-  print chr(5), chr(6), chr(5);
-}
-
-sub handle_char {
-  print $_[1];
-}
-
-my $parser = new XML::Parser(Handlers =>
-  {Start => \&handle_start,
-    End => \&handle_end,
-    Char => \&handle_char});
-
-$parser->parse(*STDIN) or die $!;
-
--- a/src/FOL/simpdata.ML	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/FOL/simpdata.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -27,9 +27,9 @@
 
 (*Congruence rules for = or <-> (instead of ==)*)
 fun mk_meta_cong rl =
-  Drule.standard (mk_meta_eq (mk_meta_prems rl))
-  handle THM _ =>
-  error("Premises and conclusion of congruence rules must use =-equality or <->");
+  Drule.export_without_context (mk_meta_eq (mk_meta_prems rl))
+    handle THM _ =>
+      error("Premises and conclusion of congruence rules must use =-equality or <->");
 
 val mksimps_pairs =
   [("op -->", [@{thm mp}]), ("op &", [@{thm conjunct1}, @{thm conjunct2}]),
--- a/src/FOLP/simp.ML	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/FOLP/simp.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -519,7 +519,7 @@
 (* Compute Congruence rules for individual constants using the substition
    rules *)
 
-val subst_thms = map Drule.standard subst_thms;
+val subst_thms = map Drule.export_without_context subst_thms;
 
 
 fun exp_app(0,t) = t
--- a/src/HOL/Import/hol4rews.ML	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/HOL/Import/hol4rews.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -371,7 +371,7 @@
     fun process ((bthy,bthm), hth as (_,thm)) thy =
       let
         val thm1 = rewrite_rule (map (Thm.transfer thy) rews) (Thm.transfer thy thm);
-        val thm2 = Drule.standard thm1;
+        val thm2 = Drule.export_without_context thm1;
       in
         thy
         |> PureThy.store_thm (Binding.name bthm, thm2)
--- a/src/HOL/Import/shuffler.ML	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/HOL/Import/shuffler.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -100,7 +100,7 @@
         val th4 = implies_elim_list (assume cPPQ) [th3,th3]
                                     |> implies_intr_list [cPPQ,cP]
     in
-        equal_intr th4 th1 |> Drule.standard
+        equal_intr th4 th1 |> Drule.export_without_context
     end
 
 val imp_comm =
@@ -120,7 +120,7 @@
         val th2 = implies_elim_list (assume cQPR) [assume cQ,assume cP]
                                     |> implies_intr_list [cQPR,cP,cQ]
     in
-        equal_intr th1 th2 |> Drule.standard
+        equal_intr th1 th2 |> Drule.export_without_context
     end
 
 val def_norm =
@@ -147,7 +147,7 @@
                               |> forall_intr cv
                               |> implies_intr cPQ
     in
-        equal_intr th1 th2 |> Drule.standard
+        equal_intr th1 th2 |> Drule.export_without_context
     end
 
 val all_comm =
@@ -173,7 +173,7 @@
                          |> forall_intr_list [cy,cx]
                          |> implies_intr cl
     in
-        equal_intr th1 th2 |> Drule.standard
+        equal_intr th1 th2 |> Drule.export_without_context
     end
 
 val equiv_comm =
@@ -187,7 +187,7 @@
         val th1  = assume ctu |> symmetric |> implies_intr ctu
         val th2  = assume cut |> symmetric |> implies_intr cut
     in
-        equal_intr th1 th2 |> Drule.standard
+        equal_intr th1 th2 |> Drule.export_without_context
     end
 
 (* This simplification procedure rewrites !!x y. P x y
--- a/src/HOL/IsaMakefile	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/HOL/IsaMakefile	Mon Feb 08 14:08:32 2010 +0100
@@ -395,7 +395,7 @@
   Library/Library/document/root.tex Library/Library/document/root.bib	\
   Library/Transitive_Closure_Table.thy Library/While_Combinator.thy	\
   Library/Product_ord.thy Library/Char_nat.thy				\
-  Library/Char_ord.thy Library/Option_ord.thy				\
+  Library/Structure_Syntax.thy						\
   Library/Sublist_Order.thy Library/List_lexord.thy			\
   Library/Coinductive_List.thy Library/AssocList.thy			\
   Library/Formal_Power_Series.thy Library/Binomial.thy			\
--- a/src/HOL/Library/Library.thy	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/HOL/Library/Library.thy	Mon Feb 08 14:08:32 2010 +0100
@@ -51,6 +51,7 @@
   RBT
   SML_Quickcheck
   State_Monad
+  Structure_Syntax
   Sum_Of_Squares
   Transitive_Closure_Table
   Univ_Poly
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Library/Structure_Syntax.thy	Mon Feb 08 14:08:32 2010 +0100
@@ -0,0 +1,14 @@
+(* Author: Florian Haftmann, TU Muenchen *)
+
+header {* Index syntax for structures *}
+
+theory Structure_Syntax
+imports Pure
+begin
+
+syntax
+  "_index1"  :: index    ("\<^sub>1")
+translations
+  (index) "\<^sub>1" => (index) "\<^bsub>\<struct>\<^esub>"
+
+end
--- a/src/HOL/Library/Sum_Of_Squares/sos_wrapper.ML	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/HOL/Library/Sum_Of_Squares/sos_wrapper.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -54,10 +54,12 @@
 
     (* call solver *)
     val output_file = filename dir "sos_out"
-    val (output, rv) = system_out (
-      if File.exists cmd then space_implode " "
-        [File.shell_path cmd, File.shell_path input_file, File.shell_path output_file]
-      else error ("Bad executable: " ^ File.platform_path cmd))
+    val (output, rv) =
+      bash_output
+       (if File.exists cmd then
+          space_implode " "
+            [File.shell_path cmd, File.shell_path input_file, File.shell_path output_file]
+        else error ("Bad executable: " ^ File.platform_path cmd))
 
     (* read and analyze output *)
     val (res, res_msg) = find_failure rv
--- a/src/HOL/Matrix/cplex/Cplex_tools.ML	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/HOL/Matrix/cplex/Cplex_tools.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -1145,7 +1145,7 @@
     val cplex_path = getenv "GLPK_PATH"
     val cplex = if cplex_path = "" then "glpsol" else cplex_path
     val command = (wrap cplex)^" --lpt "^(wrap lpname)^" --output "^(wrap resultname)
-    val answer = #1 (system_out command)
+    val answer = #1 (bash_output command)
     in
     let
         val result = load_glpkResult resultname
@@ -1178,7 +1178,7 @@
     val cplex = if cplex_path = "" then "cplex" else cplex_path
     val _ = write_script scriptname lpname resultname
     val command = (wrap cplex)^" < "^(wrap scriptname)^" > /dev/null"
-    val answer = "return code "^(Int.toString (system command))
+    val answer = "return code "^(Int.toString (bash command))
     in
     let
         val result = load_cplexResult resultname
--- a/src/HOL/Matrix/cplex/matrixlp.ML	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/HOL/Matrix/cplex/matrixlp.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -18,7 +18,7 @@
 
 fun inst_real thm =
   let val certT = ctyp_of (Thm.theory_of_thm thm) in
-    Drule.standard (Thm.instantiate
+    Drule.export_without_context (Thm.instantiate
       ([(certT (TVar (hd (OldTerm.term_tvars (prop_of thm)))), certT HOLogic.realT)], []) thm)
   end
 
@@ -59,7 +59,7 @@
         val ord = prod_ord (prod_ord string_ord int_ord) (list_ord string_ord)
         val v = TVar (hd (sort ord (OldTerm.term_tvars (prop_of thm))))
     in
-        Drule.standard (Thm.instantiate ([(certT v, certT ty)], []) thm)
+        Drule.export_without_context (Thm.instantiate ([(certT v, certT ty)], []) thm)
     end
 
 fun inst_tvars [] thms = thms
--- a/src/HOL/Modelcheck/EindhovenSyn.thy	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/HOL/Modelcheck/EindhovenSyn.thy	Mon Feb 08 14:08:32 2010 +0100
@@ -40,7 +40,7 @@
       val pmu =
         if eindhoven_home = "" then error "Environment variable EINDHOVEN_HOME not set"
         else eindhoven_home ^ "/pmu";
-    in #1 (system_out ("echo \"" ^ s ^ "\" | " ^ pmu ^ " -w")) end;
+    in #1 (bash_output ("echo \"" ^ s ^ "\" | " ^ pmu ^ " -w")) end;
 in
   fn cgoal =>
     let
--- a/src/HOL/Modelcheck/mucke_oracle.ML	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/HOL/Modelcheck/mucke_oracle.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -498,7 +498,7 @@
       else mucke_home ^ "/mucke";
     val mucke_input_file = File.tmp_path (Path.basic "tmp.mu");
     val _ = File.write mucke_input_file s;
-    val (result, _) = system_out (mucke ^ " -nb -res " ^ File.shell_path mucke_input_file);
+    val (result, _) = bash_output (mucke ^ " -nb -res " ^ File.shell_path mucke_input_file);
   in
     if not (!trace_mc) then (File.rm mucke_input_file) else (); 
     result
--- a/src/HOL/Nominal/nominal_datatype.ML	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/HOL/Nominal/nominal_datatype.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -152,7 +152,7 @@
 
 fun projections rule =
   Project_Rule.projections (ProofContext.init (Thm.theory_of_thm rule)) rule
-  |> map (Drule.standard #> Rule_Cases.save rule);
+  |> map (Drule.export_without_context #> Rule_Cases.save rule);
 
 val supp_prod = thm "supp_prod";
 val fresh_prod = thm "fresh_prod";
@@ -312,7 +312,7 @@
 
     val unfolded_perm_eq_thms =
       if length descr = length new_type_names then []
-      else map Drule.standard (List.drop (split_conj_thm
+      else map Drule.export_without_context (List.drop (split_conj_thm
         (Goal.prove_global thy2 [] []
           (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
             (map (fn (c as (s, T), x) =>
@@ -332,7 +332,7 @@
 
     val perm_empty_thms = maps (fn a =>
       let val permT = mk_permT (Type (a, []))
-      in map Drule.standard (List.take (split_conj_thm
+      in map Drule.export_without_context (List.take (split_conj_thm
         (Goal.prove_global thy2 [] []
           (augment_sort thy2 [pt_class_of thy2 a]
             (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
@@ -364,7 +364,7 @@
         val pt_inst = pt_inst_of thy2 a;
         val pt2' = pt_inst RS pt2;
         val pt2_ax = PureThy.get_thm thy2 (Long_Name.map_base_name (fn s => "pt_" ^ s ^ "2") a);
-      in List.take (map Drule.standard (split_conj_thm
+      in List.take (map Drule.export_without_context (split_conj_thm
         (Goal.prove_global thy2 [] []
            (augment_sort thy2 [pt_class_of thy2 a]
              (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
@@ -399,7 +399,7 @@
         val pt3' = pt_inst RS pt3;
         val pt3_rev' = at_inst RS (pt_inst RS pt3_rev);
         val pt3_ax = PureThy.get_thm thy2 (Long_Name.map_base_name (fn s => "pt_" ^ s ^ "3") a);
-      in List.take (map Drule.standard (split_conj_thm
+      in List.take (map Drule.export_without_context (split_conj_thm
         (Goal.prove_global thy2 [] []
           (augment_sort thy2 [pt_class_of thy2 a] (Logic.mk_implies
              (HOLogic.mk_Trueprop (Const ("Nominal.prm_eq",
@@ -586,7 +586,7 @@
       (fn (x, (_, ("Nominal.noption", _, _))) => NONE | (x, _) => SOME x)
       (perm_indnames ~~ descr);
 
-    fun mk_perm_closed name = map (fn th => Drule.standard (th RS mp))
+    fun mk_perm_closed name = map (fn th => Drule.export_without_context (th RS mp))
       (List.take (split_conj_thm (Goal.prove_global thy4 [] []
         (augment_sort thy4
           (pt_class_of thy4 name :: map (cp_class_of thy4 name) (remove (op =) name dt_atoms))
@@ -812,7 +812,8 @@
         val rep_const = cterm_of thy
           (Const (Sign.intern_const thy ("Rep_" ^ tname), T --> T'));
         val dist =
-          Drule.standard (cterm_instantiate [(cterm_of thy distinct_f, rep_const)] distinct_lemma);
+          Drule.export_without_context
+            (cterm_instantiate [(cterm_of thy distinct_f, rep_const)] distinct_lemma);
         val (thy', defs', eqns') = fold (make_constr_def tname T T')
           (constrs ~~ constrs' ~~ constr_syntax) (Sign.add_path tname thy, defs, [])
       in
@@ -877,8 +878,9 @@
           let
             val dist_thm = Goal.prove_global thy8 [] [] t (fn _ =>
               simp_tac (global_simpset_of thy8 addsimps (dist_lemma :: rep_thms)) 1)
-          in dist_thm :: Drule.standard (dist_thm RS not_sym) ::
-            prove_distinct_thms p (k, ts)
+          in
+            dist_thm :: Drule.export_without_context (dist_thm RS not_sym) ::
+              prove_distinct_thms p (k, ts)
           end;
 
     val distinct_thms = map2 prove_distinct_thms
@@ -1092,7 +1094,7 @@
 
     val finite_supp_thms = map (fn atom =>
       let val atomT = Type (atom, [])
-      in map Drule.standard (List.take
+      in map Drule.export_without_context (List.take
         (split_conj_thm (Goal.prove_global thy8 [] []
            (augment_sort thy8 (fs_class_of thy8 atom :: pt_cp_sort)
              (HOLogic.mk_Trueprop
@@ -1540,7 +1542,7 @@
           in
             (R $ x $ y, R' $ mk_perm [] pi x $ mk_perm [] pi y)
           end) (recTs ~~ rec_result_Ts ~~ rec_sets ~~ rec_sets_pi ~~ (1 upto length recTs));
-        val ths = map (fn th => Drule.standard (th RS mp)) (split_conj_thm
+        val ths = map (fn th => Drule.export_without_context (th RS mp)) (split_conj_thm
           (Goal.prove_global thy11 [] []
             (augment_sort thy1 pt_cp_sort
               (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map HOLogic.mk_imp ps))))
@@ -1572,7 +1574,7 @@
           (finite $ (Const ("Nominal.supp", T --> aset) $ f)))
             (rec_fns ~~ rec_fn_Ts)
       in
-        map (fn th => Drule.standard (th RS mp)) (split_conj_thm
+        map (fn th => Drule.export_without_context (th RS mp)) (split_conj_thm
           (Goal.prove_global thy11 []
             (map (augment_sort thy11 fs_cp_sort) fins)
             (augment_sort thy11 fs_cp_sort
@@ -1615,7 +1617,7 @@
             val y = Free ("y", U);
             val y' = Free ("y'", U)
           in
-            Drule.standard (Goal.prove (ProofContext.init thy11) []
+            Drule.export_without_context (Goal.prove (ProofContext.init thy11) []
               (map (augment_sort thy11 fs_cp_sort)
                 (finite_prems @
                    [HOLogic.mk_Trueprop (R $ x $ y),
@@ -2060,7 +2062,7 @@
          ((Binding.name "rec_equiv'", flat rec_equiv_thms'), []),
          ((Binding.name "rec_fin_supp", flat rec_fin_supp_thms), []),
          ((Binding.name "rec_fresh", flat rec_fresh_thms), []),
-         ((Binding.name "rec_unique", map Drule.standard rec_unique_thms), []),
+         ((Binding.name "rec_unique", map Drule.export_without_context rec_unique_thms), []),
          ((Binding.name "recs", rec_thms), [])] ||>
       Sign.parent_path ||>
       map_nominal_datatypes (fold Symtab.update dt_infos);
--- a/src/HOL/SMT/Tools/smt_solver.ML	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/HOL/SMT/Tools/smt_solver.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -156,7 +156,7 @@
       else (["certificate"], "Using certificate from " ^ quote certs' ^ " ...")
     val _ = tracing msg
   in
-    system_out (space_implode " " ("perl -w" ::
+    bash_output (space_implode " " (
       File.shell_path (Path.explode (getenv "RUN_SMT_SOLVER")) :: certs' ::
       map File.shell_quote (solver @ args) @
       map File.shell_path [problem_path, proof_path]) ^ " 2>&1")
--- a/src/HOL/SMT/Tools/z3_proof_rules.ML	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/HOL/SMT/Tools/z3_proof_rules.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -83,7 +83,7 @@
 val rule_of_string = Symtab.lookup rule_names
 fun string_of_rule r =
   let fun fit (s, r') = if r = r' then SOME s else NONE 
-  in the (Symtab.get_first NONE fit rule_names) end
+  in the (Symtab.get_first fit rule_names) end
 
 
 (* proof representation *)
@@ -545,7 +545,7 @@
 
   fun derive conj t lits idx ptab =
     let
-      val (l, lit) = the (Termtab.get_first NONE (get_lit conj t) lits)
+      val (l, lit) = the (Termtab.get_first (get_lit conj t) lits)
       val ls = explode_thm conj false false [t] lit
       val lits' = fold (Termtab.update o ` prop_of) ls (Termtab.delete l lits)
       fun upd (Sequent {hyps, vars, thm}) =
@@ -1231,7 +1231,7 @@
       (case Termtab.lookup tab @{term False} of
         SOME rs => extract_lit thm rs
       | NONE =>
-          pairself (extract_lit thm) (the (Termtab.get_first NONE pnlits tab))
+          pairself (extract_lit thm) (the (Termtab.get_first pnlits tab))
           |> (fn (nlit, plit) => nlit COMP (plit COMP contra_rule)))
     end
   val falseE_v = Thm.dest_arg (Thm.dest_arg (Thm.cprop_of @{thm FalseE}))
--- a/src/HOL/SMT/etc/settings	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/HOL/SMT/etc/settings	Mon Feb 08 14:08:32 2010 +0100
@@ -1,6 +1,7 @@
 ISABELLE_SMT="$COMPONENT"
 
-RUN_SMT_SOLVER="$ISABELLE_SMT/lib/scripts/run_smt_solver.pl"
+RUN_SMT_SOLVER="$ISABELLE_SMT/lib/scripts/run_smt_solver"
+REMOTE_SMT="$ISABELLE_SMT/lib/scripts/remote_smt"
 
 REMOTE_SMT_URL="http://smt.in.tum.de/smt"
 
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SMT/lib/scripts/remote_smt	Mon Feb 08 14:08:32 2010 +0100
@@ -0,0 +1,31 @@
+#!/usr/bin/env perl
+#
+# Author: Sascha Boehme, TU Muenchen
+#
+# Invoke remote SMT solvers.
+
+use strict;
+use warnings;
+use LWP;
+
+
+# arguments
+
+my $solver = $ARGV[0];
+my @options = @ARGV[1 .. ($#ARGV - 1)];
+my $problem_file = $ARGV[-1];
+
+
+# call solver
+
+my $agent = LWP::UserAgent->new;
+$agent->agent("SMT-Request");
+$agent->timeout(180);
+my $response = $agent->post($ENV{"REMOTE_SMT_URL"}, [
+  "Solver" => $solver,
+  "Options" => join(" ", @options),
+  "Problem" => [$problem_file] ],
+  "Content_Type" => "form-data");
+if (not $response->is_success) { die "HTTP-Error: " . $response->message; }
+else { print $response->content; }
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SMT/lib/scripts/run_smt_solver	Mon Feb 08 14:08:32 2010 +0100
@@ -0,0 +1,93 @@
+#!/usr/bin/env perl
+#
+# Author: Sascha Boehme, TU Muenchen
+#
+# Cache for prover results, based on discriminating problems using MD5.
+
+use strict;
+use warnings;
+use Digest::MD5;
+
+
+# arguments
+
+my $certs_file = shift @ARGV;
+my $location = shift @ARGV;
+my $result_file = pop @ARGV;
+my $problem_file = $ARGV[-1];
+
+my $use_certs = not ($certs_file eq "-");
+
+
+# create MD5 problem digest
+
+my $problem_digest = "";
+if ($use_certs) {
+  my $md5 = Digest::MD5->new;
+  open FILE, "<$problem_file" or
+    die "ERROR: Failed to open '$problem_file' ($!)";
+  $md5->addfile(*FILE);
+  close FILE;
+  $problem_digest = $md5->b64digest;
+}
+
+
+# lookup problem digest among existing certificates and
+# return a cached result (if there is a certificate for the problem)
+
+if ($use_certs and -e $certs_file) {
+  my $cached = 0;
+  open CERTS, "<$certs_file" or die "ERROR: Failed to open '$certs_file' ($!)";
+  while (<CERTS>) {
+    if (m/(\S+) (\d+)/) {
+      if ($1 eq $problem_digest) {
+        my $start = tell CERTS;
+        open FILE, ">$result_file" or
+          die "ERROR: Failed to open '$result_file ($!)";
+        while ((tell CERTS) - $start < $2) {
+          my $line = <CERTS>;
+          print FILE $line;
+        }
+        close FILE;
+        $cached = 1;
+        last;
+      }
+      else { seek CERTS, $2, 1; }
+    }
+    else { die "ERROR: Invalid file format in '$certs_file'"; }
+  }
+  close CERTS;
+  if ($cached) { exit 0; }
+}
+
+
+# invoke (remote or local) prover
+
+my $result;
+
+if ($location eq "remote") {
+  $result = system "$ENV{'REMOTE_SMT'} @ARGV >$result_file 2>&1";
+}
+elsif ($location eq "local") {
+  $result = system "@ARGV >$result_file 2>&1";
+}
+else { die "ERROR: No SMT solver invoked"; }
+
+
+# cache result
+
+if ($use_certs) {
+  open CERTS, ">>$certs_file" or
+    die "ERROR: Failed to open '$certs_file' ($!)";
+  print CERTS
+    ("$problem_digest " . ((-s $result_file) + (length "\n")) . "\n");
+
+  open FILE, "<$result_file" or
+    die "ERROR: Failed to open '$result_file' ($!)";
+  foreach (<FILE>) { print CERTS $_; }
+  close FILE; 
+
+  print CERTS "\n";
+  close CERTS;
+}
+
--- a/src/HOL/SMT/lib/scripts/run_smt_solver.pl	Mon Feb 08 14:06:58 2010 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,110 +0,0 @@
-#
-# Author: Sascha Boehme, TU Muenchen
-#
-# Cache for prover results, based on discriminating problems using MD5.
-
-use strict;
-use warnings;
-use Digest::MD5;
-use LWP;
-
-
-# arguments
-
-my $certs_file = shift @ARGV;
-my $location = shift @ARGV;
-my $result_file = pop @ARGV;
-my $problem_file = $ARGV[-1];
-
-my $use_certs = not ($certs_file eq "-");
-
-
-# create MD5 problem digest
-
-my $problem_digest = "";
-if ($use_certs) {
-  my $md5 = Digest::MD5->new;
-  open FILE, "<$problem_file" or
-    die "ERROR: Failed to open '$problem_file' ($!)";
-  $md5->addfile(*FILE);
-  close FILE;
-  $problem_digest = $md5->b64digest;
-}
-
-
-# lookup problem digest among existing certificates and
-# return a cached result (if there is a certificate for the problem)
-
-if ($use_certs and -e $certs_file) {
-  my $cached = 0;
-  open CERTS, "<$certs_file" or die "ERROR: Failed to open '$certs_file' ($!)";
-  while (<CERTS>) {
-    if (m/(\S+) (\d+)/) {
-      if ($1 eq $problem_digest) {
-        my $start = tell CERTS;
-        open FILE, ">$result_file" or
-          die "ERROR: Failed to open '$result_file ($!)";
-        while ((tell CERTS) - $start < $2) {
-          my $line = <CERTS>;
-          print FILE $line;
-        }
-        close FILE;
-        $cached = 1;
-        last;
-      }
-      else { seek CERTS, $2, 1; }
-    }
-    else { die "ERROR: Invalid file format in '$certs_file'"; }
-  }
-  close CERTS;
-  if ($cached) { exit 0; }
-}
-
-
-# invoke (remote or local) prover
-
-if ($location eq "remote") {
-  # arguments
-  my $solver = $ARGV[0];
-  my @options = @ARGV[1 .. ($#ARGV - 1)];
-
-  # call solver
-  my $agent = LWP::UserAgent->new;
-  $agent->agent("SMT-Request");
-  $agent->timeout(180);
-  my $response = $agent->post($ENV{"REMOTE_SMT_URL"}, [
-    "Solver" => $solver,
-    "Options" => join(" ", @options),
-    "Problem" => [$problem_file] ],
-    "Content_Type" => "form-data");
-  if (not $response->is_success) { die "HTTP-Error: " . $response->message; }
-  else {
-    open FILE, ">$result_file" or
-      die "ERROR: Failed to create '$result_file' ($!)";
-    print FILE $response->content;
-    close FILE;
-  }
-}
-elsif ($location eq "local") {
-  system "@ARGV >$result_file 2>&1";
-}
-else { die "ERROR: No SMT solver invoked"; }
-
-
-# cache result
-
-if ($use_certs) {
-  open CERTS, ">>$certs_file" or
-    die "ERROR: Failed to open '$certs_file' ($!)";
-  print CERTS
-    ("$problem_digest " . ((-s $result_file) + (length "\n")) . "\n");
-
-  open FILE, "<$result_file" or
-    die "ERROR: Failed to open '$result_file' ($!)";
-  foreach (<FILE>) { print CERTS $_; }
-  close FILE; 
-
-  print CERTS "\n";
-  close CERTS;
-}
-
--- a/src/HOL/Statespace/state_fun.ML	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/HOL/Statespace/state_fun.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -310,7 +310,7 @@
                   val prop = list_all ([("n",nT),("x",eT)],
                               Logic.mk_equals (Const ("Ex",Tex)$Abs(s,T,eq),
                                                HOLogic.true_const));
-                  val thm = Drule.standard (prove prop);
+                  val thm = Drule.export_without_context (prove prop);
                   val thm' = if swap then swap_ex_eq OF [thm] else thm
              in SOME thm' end
              handle TERM _ => NONE)
--- a/src/HOL/Tools/ATP_Manager/atp_wrapper.ML	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/HOL/Tools/ATP_Manager/atp_wrapper.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -168,7 +168,7 @@
     fun run_on probfile =
       if File.exists cmd then
         write probfile clauses
-        |> pair (apfst split_time' (system_out (cmd_line probfile)))
+        |> pair (apfst split_time' (bash_output (cmd_line probfile)))
       else error ("Bad executable: " ^ Path.implode cmd);
 
     (* if problemfile has not been exported, delete problemfile; otherwise export proof, too *)
@@ -306,7 +306,7 @@
 
 fun get_systems () =
   let
-    val (answer, rc) = system_out ("\"$ISABELLE_ATP_MANAGER/SystemOnTPTP\" -w")
+    val (answer, rc) = bash_output ("\"$ISABELLE_ATP_MANAGER/SystemOnTPTP\" -w")
   in
     if rc <> 0 then error ("Failed to get available systems from SystemOnTPTP:\n" ^ answer)
     else split_lines answer
--- a/src/HOL/Tools/Datatype/datatype.ML	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/HOL/Tools/Datatype/datatype.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -253,9 +253,11 @@
         val rep_const = cterm_of thy
           (Const (Sign.intern_const thy ("Rep_" ^ tname), T --> Univ_elT));
         val cong' =
-          Drule.standard (cterm_instantiate [(cterm_of thy cong_f, rep_const)] arg_cong);
+          Drule.export_without_context
+            (cterm_instantiate [(cterm_of thy cong_f, rep_const)] arg_cong);
         val dist =
-          Drule.standard (cterm_instantiate [(cterm_of thy distinct_f, rep_const)] distinct_lemma);
+          Drule.export_without_context
+            (cterm_instantiate [(cterm_of thy distinct_f, rep_const)] distinct_lemma);
         val (thy', defs', eqns', _) = fold ((make_constr_def tname T) (length constrs))
           (constrs ~~ constr_syntax) (Sign.add_path tname thy, defs, [], 1);
       in
@@ -532,7 +534,7 @@
               let
                 val dist_thm = Skip_Proof.prove_global thy5 [] [] t (fn _ =>
                   EVERY [simp_tac (HOL_ss addsimps dist_rewrites') 1])
-              in dist_thm :: Drule.standard (dist_thm RS not_sym) :: prove ts end;
+              in dist_thm :: Drule.export_without_context (dist_thm RS not_sym) :: prove ts end;
       in prove ts end;
 
     val distinct_thms = map2 (prove_distinct_thms)
--- a/src/HOL/Tools/Function/size.ML	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/HOL/Tools/Function/size.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -197,7 +197,7 @@
 
     fun prove_size_eqs p size_fns size_ofp simpset =
       maps (fn (((_, (_, _, constrs)), size_const), T) =>
-        map (fn constr => Drule.standard (Skip_Proof.prove ctxt [] []
+        map (fn constr => Drule.export_without_context (Skip_Proof.prove ctxt [] []
           (gen_mk_size_eq p (AList.lookup op = (new_type_names ~~ size_fns))
              size_ofp size_const T constr)
           (fn _ => simp_tac simpset 1))) constrs)
--- a/src/HOL/Tools/Nitpick/kodkod.ML	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/HOL/Tools/Nitpick/kodkod.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -1004,7 +1004,7 @@
   read_next_problems (Substring.full (File.read path), [], []) |>> rev ||> rev
 
 (* The fudge term below is to account for Kodkodi's slow start-up time, which
-   is partly due to the JVM and partly due to the ML "system" function. *)
+   is partly due to the JVM and partly due to the ML "bash" function. *)
 val fudge_ms = 250
 
 (* bool -> Time.time option -> int -> int -> problem list -> outcome *)
@@ -1053,7 +1053,7 @@
           val outcome =
             let
               val code =
-                system ("cd " ^ temp_dir ^ ";\n" ^
+                bash ("cd " ^ File.shell_quote temp_dir ^ ";\n" ^
                         "env CLASSPATH=\"$KODKODI_CLASSPATH:$CLASSPATH\" \
                         \JAVA_LIBRARY_PATH=\"$KODKODI_JAVA_LIBRARY_PATH:\
                         \$JAVA_LIBRARY_PATH\" \
@@ -1068,9 +1068,9 @@
                            " -max-threads " ^ string_of_int max_threads
                          else
                            "") ^
-                        " < " ^ Path.implode in_path ^
-                        " > " ^ Path.implode out_path ^
-                        " 2> " ^ Path.implode err_path)
+                        " < " ^ File.shell_path in_path ^
+                        " > " ^ File.shell_path out_path ^
+                        " 2> " ^ File.shell_path err_path)
               val (ps, nontriv_js) = read_output_file out_path
                                      |>> map (apfst reindex) ||> map reindex
               val js = triv_js @ nontriv_js
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -577,7 +577,7 @@
         (*val elim = singleton (Inductive_Set.codegen_preproc thy) (preprocess_elim thy nparams 
           (expand_tuples_elim pre_elim))*)
         val elim =
-          (Drule.standard o Skip_Proof.make_thm thy)
+          (Drule.export_without_context o Skip_Proof.make_thm thy)
           (mk_casesrule (ProofContext.init thy) pred intros)
       in
         mk_pred_data ((intros, SOME elim), no_compilation)
@@ -641,7 +641,7 @@
       else ()
     val pred = Const (constname, T)
     val pre_elim = 
-      (Drule.standard o Skip_Proof.make_thm thy)
+      (Drule.export_without_context o Skip_Proof.make_thm thy)
       (mk_casesrule (ProofContext.init thy) pred pre_intros)
   in register_predicate (constname, pre_intros, pre_elim) thy end
 
@@ -2178,7 +2178,8 @@
     (join_preds_modes moded_clauses compiled_terms)
 
 fun prove_by_skip options thy _ _ _ _ compiled_terms =
-  map_preds_modes (fn pred => fn mode => fn t => Drule.standard (Skip_Proof.make_thm thy t))
+  map_preds_modes
+    (fn pred => fn mode => fn t => Drule.export_without_context (Skip_Proof.make_thm thy t))
     compiled_terms
 
 (* preparation of introduction rules into special datastructures *)
@@ -2269,7 +2270,7 @@
         val elim = the_elim_of thy predname
         val nparams = nparams_of thy predname
         val elim' =
-          (Drule.standard o (Skip_Proof.make_thm thy))
+          (Drule.export_without_context o Skip_Proof.make_thm thy)
           (mk_casesrule (ProofContext.init thy) nparams intros)
       in
         if not (Thm.equiv_thm (elim, elim')) then
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_fun.ML	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_fun.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -391,7 +391,7 @@
         |> map (fn (concl'::conclprems, _) =>
           Logic.list_implies ((flat prems') @ conclprems, concl')))
   in
-    map (Drule.standard o (Skip_Proof.make_thm thy)) intro_ts'
+    map (Drule.export_without_context o Skip_Proof.make_thm thy) intro_ts'
   end
 
 end;
--- a/src/HOL/Tools/TFL/post.ML	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/HOL/Tools/TFL/post.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -129,7 +129,7 @@
 
 (*lcp: put a theorem into Isabelle form, using meta-level connectives*)
 val meta_outer =
-  curry_rule o Drule.standard o
+  curry_rule o Drule.export_without_context o
   rule_by_tactic (REPEAT (FIRSTGOAL (resolve_tac [allI, impI, conjI] ORELSE' etac conjE)));
 
 (*Strip off the outer !P*)
@@ -151,7 +151,7 @@
                {f = f, R = R, rules = rules,
                 full_pats_TCs = full_pats_TCs,
                 TCs = TCs}
-       val rules' = map (Drule.standard o ObjectLogic.rulify_no_asm)
+       val rules' = map (Drule.export_without_context o ObjectLogic.rulify_no_asm)
                         (R.CONJUNCTS rules)
          in  {induct = meta_outer (ObjectLogic.rulify_no_asm (induction RS spec')),
         rules = ListPair.zip(rules', rows),
@@ -180,7 +180,7 @@
     | solve_eq (th, [a], i) = [(a, i)]
     | solve_eq (th, splitths as (_ :: _), i) = 
       (writeln "Proving unsplit equation...";
-      [((Drule.standard o ObjectLogic.rulify_no_asm)
+      [((Drule.export_without_context o ObjectLogic.rulify_no_asm)
           (CaseSplit.splitto splitths th), i)])
       (* if there's an error, pretend nothing happened with this definition 
          We should probably print something out so that the user knows...? *)
@@ -236,7 +236,7 @@
  in (theory,
      (*return the conjoined induction rule and recursion equations,
        with assumptions remaining to discharge*)
-     Drule.standard (induction RS (rules RS conjI)))
+     Drule.export_without_context (induction RS (rules RS conjI)))
  end
 
 fun defer thy congs fid seqs =
--- a/src/HOL/Tools/arith_data.ML	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/HOL/Tools/arith_data.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -110,8 +110,8 @@
 
 fun prove_conv tacs ctxt (_: thm list) = prove_conv_nohyps tacs ctxt;
 
-fun prove_conv2 expand_tac norm_tac ss tu = (*FIXME avoid standard*)
-  mk_meta_eq (Drule.standard (Goal.prove (Simplifier.the_context ss) [] []
+fun prove_conv2 expand_tac norm_tac ss tu = (* FIXME avoid Drule.export_without_context *)
+  mk_meta_eq (Drule.export_without_context (Goal.prove (Simplifier.the_context ss) [] []
       (HOLogic.mk_Trueprop (HOLogic.mk_eq tu))
     (K (EVERY [expand_tac, norm_tac ss]))));
 
--- a/src/HOL/Tools/choice_specification.ML	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/HOL/Tools/choice_specification.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -75,7 +75,7 @@
 fun add_specification axiomatic cos arg =
     arg |> apsnd Thm.freezeT
         |> proc_exprop axiomatic cos
-        |> apsnd Drule.standard
+        |> apsnd Drule.export_without_context
 
 
 (* Collect all intances of constants in term *)
@@ -189,7 +189,7 @@
                     in
                         args |> apsnd (remove_alls frees)
                              |> apsnd undo_imps
-                             |> apsnd Drule.standard
+                             |> apsnd Drule.export_without_context
                              |> Thm.theory_attributes (map (Attrib.attribute thy) atts)
                              |> add_final
                              |> Library.swap
--- a/src/HOL/Tools/inductive_codegen.ML	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/HOL/Tools/inductive_codegen.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -544,7 +544,7 @@
       (fn NONE => "X" | SOME k' => string_of_int k')
         (ks @ [SOME k]))) arities));
 
-fun prep_intrs intrs = map (rename_term o #prop o rep_thm o Drule.standard) intrs;
+fun prep_intrs intrs = map (rename_term o #prop o rep_thm o Drule.export_without_context) intrs;
 
 fun constrain cs [] = []
   | constrain cs ((s, xs) :: ys) =
--- a/src/HOL/Tools/numeral_simprocs.ML	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/HOL/Tools/numeral_simprocs.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -181,9 +181,8 @@
 (*To let us treat division as multiplication*)
 val divide_simps = [@{thm divide_inverse}, @{thm inverse_mult_distrib}, @{thm inverse_inverse_eq}];
 
-(*push the unary minus down: - x * y = x * - y *)
-val minus_mult_eq_1_to_2 =
-    [@{thm mult_minus_left}, @{thm minus_mult_right}] MRS trans |> Drule.standard;
+(*push the unary minus down*)
+val minus_mult_eq_1_to_2 = @{lemma "- (a::'a::ring) * b = a * - b" by simp};
 
 (*to extract again any uncancelled minuses*)
 val minus_from_mult_simps =
--- a/src/HOL/Tools/record.ML	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/HOL/Tools/record.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -1014,7 +1014,7 @@
     else if Goal.future_enabled () then
       Goal.future_result (ProofContext.init thy) (Future.fork_pri ~1 prf) prop
     else prf ()
-  in Drule.standard thm end;
+  in Drule.export_without_context thm end;
 
 fun prove_common immediate stndrd thy asms prop tac =
   let
@@ -1023,7 +1023,7 @@
       else if immediate orelse not (Goal.future_enabled ()) then Goal.prove
       else Goal.prove_future;
     val prf = prv (ProofContext.init thy) [] asms prop tac;
-  in if stndrd then Drule.standard prf else prf end;
+  in if stndrd then Drule.export_without_context prf else prf end;
 
 val prove_future_global = prove_common false;
 val prove_global = prove_common true;
@@ -1072,7 +1072,7 @@
           if is_sel_upd_pair thy acc upd
           then o_eq_dest
           else o_eq_id_dest;
-      in Drule.standard (othm RS dest) end;
+      in Drule.export_without_context (othm RS dest) end;
   in map get_simp upd_funs end;
 
 fun get_updupd_simp thy defset u u' comp =
@@ -1092,7 +1092,7 @@
           REPEAT_DETERM (Iso_Tuple_Support.iso_tuple_intros_tac 1) THEN
           TRY (simp_tac (HOL_ss addsimps [id_apply]) 1));
     val dest = if comp then o_eq_dest_lhs else o_eq_dest;
-  in Drule.standard (othm RS dest) end;
+  in Drule.export_without_context (othm RS dest) end;
 
 fun get_updupd_simps thy term defset =
   let
@@ -1279,8 +1279,8 @@
             val ss = get_sel_upd_defs thy;
             val uathm = get_upd_acc_cong_thm upd acc thy ss;
           in
-           [Drule.standard (uathm RS updacc_noopE),
-            Drule.standard (uathm RS updacc_noop_compE)]
+           [Drule.export_without_context (uathm RS updacc_noopE),
+            Drule.export_without_context (uathm RS updacc_noop_compE)]
           end;
 
         (*If f is constant then (f o g) = f.  We know that K_skeleton
@@ -1721,8 +1721,8 @@
       to prove other theorems. We haven't given names to the accessors
       f, g etc yet however, so we generate an ext structure with
       free variables as all arguments and allow the introduction tactic to
-      operate on it as far as it can. We then use Drule.standard to convert
-      the free variables into unifiable variables and unify them with
+      operate on it as far as it can. We then use Drule.export_without_context
+      to convert the free variables into unifiable variables and unify them with
       (roughly) the definition of the accessor.*)
     fun surject_prf () =
       let
@@ -1733,7 +1733,7 @@
           REPEAT_ALL_NEW Iso_Tuple_Support.iso_tuple_intros_tac 1;
         val tactic2 = REPEAT (rtac surject_assistI 1 THEN rtac refl 1);
         val [halfway] = Seq.list_of (tactic1 start);
-        val [surject] = Seq.list_of (tactic2 (Drule.standard halfway));
+        val [surject] = Seq.list_of (tactic2 (Drule.export_without_context halfway));
       in
         surject
       end;
@@ -2136,14 +2136,14 @@
     fun sel_convs_prf () =
       map (prove_simp false ss (sel_defs @ accessor_thms)) sel_conv_props;
     val sel_convs = timeit_msg "record sel_convs proof:" sel_convs_prf;
-    fun sel_convs_standard_prf () = map Drule.standard sel_convs
+    fun sel_convs_standard_prf () = map Drule.export_without_context sel_convs;
     val sel_convs_standard =
       timeit_msg "record sel_convs_standard proof:" sel_convs_standard_prf;
 
     fun upd_convs_prf () =
       map (prove_simp false ss (upd_defs @ updator_thms)) upd_conv_props;
     val upd_convs = timeit_msg "record upd_convs proof:" upd_convs_prf;
-    fun upd_convs_standard_prf () = map Drule.standard upd_convs
+    fun upd_convs_standard_prf () = map Drule.export_without_context upd_convs;
     val upd_convs_standard =
       timeit_msg "record upd_convs_standard proof:" upd_convs_standard_prf;
 
@@ -2151,7 +2151,7 @@
       let
         val symdefs = map symmetric (sel_defs @ upd_defs);
         val fold_ss = HOL_basic_ss addsimps symdefs;
-        val ua_congs = map (Drule.standard o simplify fold_ss) upd_acc_cong_assists;
+        val ua_congs = map (Drule.export_without_context o simplify fold_ss) upd_acc_cong_assists;
       in (ua_congs RL [updacc_foldE], ua_congs RL [updacc_unfoldE]) end;
     val (fold_congs, unfold_congs) =
       timeit_msg "record upd fold/unfold congs:" get_upd_acc_congs;
@@ -2221,7 +2221,7 @@
             rtac (prop_subst OF [surjective]),
             REPEAT o etac meta_allE, atac]);
     val split_meta = timeit_msg "record split_meta proof:" split_meta_prf;
-    fun split_meta_standardise () = Drule.standard split_meta;
+    fun split_meta_standardise () = Drule.export_without_context split_meta;
     val split_meta_standard =
       timeit_msg "record split_meta standard:" split_meta_standardise;
 
--- a/src/HOL/Tools/sat_solver.ML	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/HOL/Tools/sat_solver.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -282,7 +282,7 @@
   (* string -> (PropLogic.prop_formula -> unit) -> (unit -> result) -> solver *)
 
   fun make_external_solver cmd writefn readfn fm =
-    (writefn fm; system cmd; readfn ());
+    (writefn fm; bash cmd; readfn ());
 
 (* ------------------------------------------------------------------------- *)
 (* read_dimacs_cnf_file: returns a propositional formula that corresponds to *)
@@ -586,7 +586,7 @@
     val inpath     = File.tmp_path (Path.explode ("isabelle" ^ serial_str ^ ".cnf"))
     val outpath    = File.tmp_path (Path.explode ("result" ^ serial_str))
     val proofpath  = File.tmp_path (Path.explode ("result" ^ serial_str ^ ".prf"))
-    val cmd        = (getenv "MINISAT_HOME") ^ "/minisat " ^ (Path.implode inpath) ^ " -r " ^ (Path.implode outpath) ^ " -t " ^ (Path.implode proofpath) ^ "> /dev/null"
+    val cmd        = getenv "MINISAT_HOME" ^ "/minisat " ^ File.shell_path inpath ^ " -r " ^ File.shell_path outpath ^ " -t " ^ File.shell_path proofpath ^ "> /dev/null"
     fun writefn fm = SatSolver.write_dimacs_cnf_file inpath fm
     fun readfn ()  = SatSolver.read_std_result_file outpath ("SAT", "", "UNSAT")
     val _          = if File.exists inpath then warning ("overwriting existing file " ^ quote (Path.implode inpath)) else ()
@@ -767,11 +767,11 @@
 let
   fun minisat fm =
   let
-    val _          = if (getenv "MINISAT_HOME") = "" then raise SatSolver.NOT_CONFIGURED else ()
+    val _          = if getenv "MINISAT_HOME" = "" then raise SatSolver.NOT_CONFIGURED else ()
     val serial_str = serial_string ()
     val inpath     = File.tmp_path (Path.explode ("isabelle" ^ serial_str ^ ".cnf"))
     val outpath    = File.tmp_path (Path.explode ("result" ^ serial_str))
-    val cmd        = (getenv "MINISAT_HOME") ^ "/minisat " ^ (Path.implode inpath) ^ " -r " ^ (Path.implode outpath) ^ " > /dev/null"
+    val cmd        = getenv "MINISAT_HOME" ^ "/minisat " ^ File.shell_path inpath ^ " -r " ^ File.shell_path outpath ^ " > /dev/null"
     fun writefn fm = SatSolver.write_dimacs_cnf_file inpath (PropLogic.defcnf fm)
     fun readfn ()  = SatSolver.read_std_result_file outpath ("SAT", "", "UNSAT")
     val _          = if File.exists inpath then warning ("overwriting existing file " ^ quote (Path.implode inpath)) else ()
@@ -926,11 +926,11 @@
 let
   fun zchaff fm =
   let
-    val _          = if (getenv "ZCHAFF_HOME") = "" then raise SatSolver.NOT_CONFIGURED else ()
+    val _          = if getenv "ZCHAFF_HOME" = "" then raise SatSolver.NOT_CONFIGURED else ()
     val serial_str = serial_string ()
     val inpath     = File.tmp_path (Path.explode ("isabelle" ^ serial_str ^ ".cnf"))
     val outpath    = File.tmp_path (Path.explode ("result" ^ serial_str))
-    val cmd        = (getenv "ZCHAFF_HOME") ^ "/zchaff " ^ (Path.implode inpath) ^ " > " ^ (Path.implode outpath)
+    val cmd        = getenv "ZCHAFF_HOME" ^ "/zchaff " ^ File.shell_path inpath ^ " > " ^ File.shell_path outpath
     fun writefn fm = SatSolver.write_dimacs_cnf_file inpath (PropLogic.defcnf fm)
     fun readfn ()  = SatSolver.read_std_result_file outpath ("Instance Satisfiable", "", "Instance Unsatisfiable")
     val _          = if File.exists inpath then warning ("overwriting existing file " ^ quote (Path.implode inpath)) else ()
@@ -957,7 +957,7 @@
     val inpath     = File.tmp_path (Path.explode ("isabelle" ^ serial_str ^ ".cnf"))
     val outpath    = File.tmp_path (Path.explode ("result" ^ serial_str))
     val exec       = getenv "BERKMIN_EXE"
-    val cmd        = (getenv "BERKMIN_HOME") ^ "/" ^ (if exec = "" then "BerkMin561" else exec) ^ " " ^ (Path.implode inpath) ^ " > " ^ (Path.implode outpath)
+    val cmd        = getenv "BERKMIN_HOME" ^ "/" ^ (if exec = "" then "BerkMin561" else exec) ^ " " ^ File.shell_path inpath ^ " > " ^ File.shell_path outpath
     fun writefn fm = SatSolver.write_dimacs_cnf_file inpath (PropLogic.defcnf fm)
     fun readfn ()  = SatSolver.read_std_result_file outpath ("Satisfiable          !!", "solution =", "UNSATISFIABLE          !!")
     val _          = if File.exists inpath then warning ("overwriting existing file " ^ quote (Path.implode inpath)) else ()
@@ -983,7 +983,7 @@
     val serial_str = serial_string ()
     val inpath     = File.tmp_path (Path.explode ("isabelle" ^ serial_str ^ ".cnf"))
     val outpath    = File.tmp_path (Path.explode ("result" ^ serial_str))
-    val cmd        = (getenv "JERUSAT_HOME") ^ "/Jerusat1.3 " ^ (Path.implode inpath) ^ " > " ^ (Path.implode outpath)
+    val cmd        = getenv "JERUSAT_HOME" ^ "/Jerusat1.3 " ^ File.shell_path inpath ^ " > " ^ File.shell_path outpath
     fun writefn fm = SatSolver.write_dimacs_cnf_file inpath (PropLogic.defcnf fm)
     fun readfn ()  = SatSolver.read_std_result_file outpath ("s SATISFIABLE", "v ", "s UNSATISFIABLE")
     val _          = if File.exists inpath then warning ("overwriting existing file " ^ quote (Path.implode inpath)) else ()
--- a/src/HOL/Tools/split_rule.ML	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/HOL/Tools/split_rule.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -100,13 +100,13 @@
       | (t, ts) => fold collect_vars ts);
 
 
-val split_rule_var = (Drule.standard o remove_internal_split) oo split_rule_var';
+val split_rule_var = (Drule.export_without_context o remove_internal_split) oo split_rule_var';
 
 (*curries ALL function variables occurring in a rule's conclusion*)
 fun split_rule rl =
   fold_rev split_rule_var' (OldTerm.term_vars (concl_of rl)) rl
   |> remove_internal_split
-  |> Drule.standard;
+  |> Drule.export_without_context;
 
 (*curries ALL function variables*)
 fun complete_split_rule rl =
@@ -117,7 +117,7 @@
   in
     fst (fold_rev complete_split_rule_var vars (rl, xs))
     |> remove_internal_split
-    |> Drule.standard
+    |> Drule.export_without_context
     |> Rule_Cases.save rl
   end;
 
@@ -137,7 +137,7 @@
     rl
     |> fold_index one_goal xss
     |> Simplifier.full_simplify split_rule_ss
-    |> Drule.standard
+    |> Drule.export_without_context
     |> Rule_Cases.save rl
   end;
 
--- a/src/HOL/Tools/typedef.ML	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/HOL/Tools/typedef.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -122,7 +122,7 @@
           let
             val cert = Thm.cterm_of (Thm.theory_of_thm def_eq);
             val goal_eq = MetaSimplifier.rewrite true [def_eq] (cert goal');
-          in Drule.standard (Drule.equal_elim_rule2 OF [goal_eq, th]) end;
+          in Drule.export_without_context (Drule.equal_elim_rule2 OF [goal_eq, th]) end;
 
     fun typedef_result inhabited =
       ObjectLogic.typedecl (t, vs, mx)
@@ -139,7 +139,7 @@
       ##> Theory.add_deps "" (dest_Const AbsC) typedef_deps
       #-> (fn ([type_definition], set_def) => fn thy1 =>
         let
-          fun make th = Drule.standard (th OF [type_definition]);
+          fun make th = Drule.export_without_context (th OF [type_definition]);
           val ([Rep, Rep_inverse, Abs_inverse, Rep_inject, Abs_inject,
               Rep_cases, Abs_cases, Rep_induct, Abs_induct], thy2) =
             thy1
--- a/src/HOL/ex/svc_funcs.ML	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/HOL/ex/svc_funcs.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -62,11 +62,11 @@
       val _ = if !trace then tracing ("Calling SVC:\n" ^ svc_input) else ()
       val svc_input_file  = File.tmp_path (Path.basic "SVM_in");
       val svc_output_file = File.tmp_path (Path.basic "SVM_out");
-      val _ = (File.write svc_input_file svc_input;
-               #1 (system_out (check_valid ^ " -dump-result " ^
-                        File.shell_path svc_output_file ^
-                        " " ^ File.shell_path svc_input_file ^
-                        ">/dev/null 2>&1")))
+      val _ = File.write svc_input_file svc_input;
+      val _ =
+        bash_output (check_valid ^ " -dump-result " ^
+          File.shell_path svc_output_file ^ " " ^ File.shell_path svc_input_file ^
+          ">/dev/null 2>&1")
       val svc_output =
         (case try File.read svc_output_file of
           SOME out => out
--- a/src/HOLCF/Tools/Domain/domain_isomorphism.ML	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/HOLCF/Tools/Domain/domain_isomorphism.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -251,7 +251,7 @@
 
     (* register unfold theorems *)
     val (unfold_thms, thy) =
-      (PureThy.add_thms o map (Thm.no_attributes o apsnd Drule.standard))
+      (PureThy.add_thms o map (Thm.no_attributes o apsnd Drule.export_without_context))
         (mk_unfold_thms unfold_binds tuple_unfold_thm) thy;
   in
     ((proj_thms, unfold_thms), thy)
@@ -446,7 +446,7 @@
     (* prove isomorphism and isodefl rules *)
     fun mk_iso_thms ((tbind, REP_eq), (rep_def, abs_def)) thy =
       let
-        fun make thm = Drule.standard (thm OF [REP_eq, abs_def, rep_def]);
+        fun make thm = Drule.export_without_context (thm OF [REP_eq, abs_def, rep_def]);
         val rep_iso_thm = make @{thm domain_rep_iso};
         val abs_iso_thm = make @{thm domain_abs_iso};
         val isodefl_thm = make @{thm isodefl_abs_rep};
@@ -545,7 +545,7 @@
           val thmR = thm RS @{thm conjunct2};
         in (n, thmL):: conjuncts ns thmR end;
     val (isodefl_thms, thy) = thy |>
-      (PureThy.add_thms o map (Thm.no_attributes o apsnd Drule.standard))
+      (PureThy.add_thms o map (Thm.no_attributes o apsnd Drule.export_without_context))
         (conjuncts isodefl_binds isodefl_thm);
     val thy = IsodeflData.map (fold Thm.add_thm isodefl_thms) thy;
 
--- a/src/HOLCF/Tools/Domain/domain_theorems.ML	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/HOLCF/Tools/Domain/domain_theorems.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -180,7 +180,7 @@
 val rep_strict = ax_abs_iso RS (allI RS retraction_strict);
 val abs_defin' = iso_locale RS iso_abs_defin';
 val rep_defin' = iso_locale RS iso_rep_defin';
-val iso_rews = map Drule.standard [ax_abs_iso,ax_rep_iso,abs_strict,rep_strict];
+val iso_rews = map Drule.export_without_context [ax_abs_iso, ax_rep_iso, abs_strict, rep_strict];
 
 (* ----- generating beta reduction rules from definitions-------------------- *)
 
@@ -263,7 +263,7 @@
   val exhaust = pg con_appls (mk_trp exh) (K tacs);
   val _ = trace " Proving casedist...";
   val casedist =
-    Drule.standard (rewrite_rule exh_casedists (exhaust RS exh_casedist0));
+    Drule.export_without_context (rewrite_rule exh_casedists (exhaust RS exh_casedist0));
 end;
 
 local 
@@ -554,7 +554,7 @@
         flat
           (ListPair.map (distinct c) ((map #1 cs),leqs)) @
         distincts cs;
-  in map Drule.standard (distincts (cons ~~ distincts_le)) end;
+  in map Drule.export_without_context (distincts (cons ~~ distincts_le)) end;
 
 local 
   fun pgterm rel con args =
@@ -755,7 +755,7 @@
         maps (eq_tacs ctxt) eqs;
     in pg copy_take_defs goal tacs end;
 in
-  val take_rews = map Drule.standard
+  val take_rews = map Drule.export_without_context
     (atomize global_ctxt take_stricts @ take_0s @ atomize global_ctxt take_apps);
 end; (* local *)
 
--- a/src/HOLCF/Tools/pcpodef.ML	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/HOLCF/Tools/pcpodef.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -89,7 +89,7 @@
           (Tactic.rtac (@{thm typedef_cpo} OF cpo_thms) 1);
     (* transfer thms so that they will know about the new cpo instance *)
     val cpo_thms' = map (Thm.transfer thy2) cpo_thms;
-    fun make thm = Drule.standard (thm OF cpo_thms');
+    fun make thm = Drule.export_without_context (thm OF cpo_thms');
     val ([adm, cont_Rep, cont_Abs, lub, thelub, compact], thy3) =
       thy2
       |> Sign.add_path (Binding.name_of name)
@@ -127,7 +127,7 @@
       |> AxClass.prove_arity (full_tname, lhs_sorts, @{sort pcpo})
         (Tactic.rtac (@{thm typedef_pcpo} OF pcpo_thms) 1);
     val pcpo_thms' = map (Thm.transfer thy2) pcpo_thms;
-    fun make thm = Drule.standard (thm OF pcpo_thms');
+    fun make thm = Drule.export_without_context (thm OF pcpo_thms');
     val ([Rep_strict, Abs_strict, Rep_strict_iff, Abs_strict_iff,
           Rep_defined, Abs_defined], thy3) =
       thy2
--- a/src/HOLCF/Tools/repdef.ML	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/HOLCF/Tools/repdef.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -139,7 +139,7 @@
       |> Sign.add_path (Binding.name_of name)
       |> PureThy.add_thms
         [((Binding.prefix_name "REP_" name,
-          Drule.standard (@{thm typedef_REP} OF typedef_thms')), [])]
+          Drule.export_without_context (@{thm typedef_REP} OF typedef_thms')), [])]
       ||> Sign.parent_path;
 
     val rep_info =
--- a/src/Provers/hypsubst.ML	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/Provers/hypsubst.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -165,7 +165,7 @@
 
 end;
 
-val ssubst = Drule.standard (Data.sym RS Data.subst);
+val ssubst = Drule.export_without_context (Data.sym RS Data.subst);
 
 fun inst_subst_tac b rl = CSUBGOAL (fn (cBi, i) =>
   case try (Logic.strip_assums_hyp #> hd #>
--- a/src/Provers/typedsimp.ML	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/Provers/typedsimp.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -43,11 +43,11 @@
 (*For simplifying both sides of an equation:
       [| a=c; b=c |] ==> b=a
   Can use resolve_tac [split_eqn] to prepare an equation for simplification. *)
-val split_eqn = Drule.standard (sym RSN (2,trans) RS sym);
+val split_eqn = Drule.export_without_context (sym RSN (2,trans) RS sym);
 
 
 (*    [| a=b; b=c |] ==> reduce(a,c)  *)
-val red_trans = Drule.standard (trans RS red_if_equal);
+val red_trans = Drule.export_without_context (trans RS red_if_equal);
 
 (*For REWRITE rule: Make a reduction rule for simplification, e.g.
   [| a: C(0); ... ; a=c: C(0) |] ==> rec(0,a,b) = c: C(0) *)
--- a/src/Pure/Concurrent/future.ML	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/Pure/Concurrent/future.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -88,22 +88,24 @@
 
 (* datatype future *)
 
+type 'a result = 'a Exn.result Single_Assignment.var;
+
 datatype 'a future = Future of
  {promised: bool,
   task: task,
   group: group,
-  result: 'a Exn.result option Synchronized.var};
+  result: 'a result};
 
 fun task_of (Future {task, ...}) = task;
 fun group_of (Future {group, ...}) = group;
 fun result_of (Future {result, ...}) = result;
 
-fun peek x = Synchronized.value (result_of x);
+fun peek x = Single_Assignment.peek (result_of x);
 fun is_finished x = is_some (peek x);
 
 fun assign_result group result res =
   let
-    val _ = Synchronized.assign result (K (SOME res));
+    val _ = Single_Assignment.assign result res;
     val ok =
       (case res of
         Exn.Exn exn => (Task_Queue.cancel_group group exn; false)
@@ -167,7 +169,7 @@
 
 fun future_job group (e: unit -> 'a) =
   let
-    val result = Synchronized.var "future" (NONE: 'a Exn.result option);
+    val result = Single_Assignment.var "future" : 'a result;
     fun job ok =
       let
         val res =
@@ -409,9 +411,6 @@
       Exn.Exn (Exn.EXCEPTIONS (Exn.flatten_list (Task_Queue.group_status (group_of x))))
   | SOME res => res);
 
-fun passive_wait x =
-  Synchronized.readonly_access (result_of x) (fn NONE => NONE | SOME _ => SOME ());
-
 fun join_next deps = (*requires SYNCHRONIZED*)
   if null deps then NONE
   else
@@ -438,7 +437,7 @@
   else
     (case worker_task () of
       SOME task => join_depend task (map task_of xs)
-    | NONE => List.app passive_wait xs;
+    | NONE => List.app (ignore o Single_Assignment.await o result_of) xs;
     map get_result xs);
 
 end;
@@ -452,7 +451,7 @@
 fun value (x: 'a) =
   let
     val group = Task_Queue.new_group NONE;
-    val result = Synchronized.var "value" NONE : 'a Exn.result option Synchronized.var;
+    val result = Single_Assignment.var "value" : 'a result;
     val _ = assign_result group result (Exn.Result x);
   in Future {promised = false, task = Task_Queue.dummy_task, group = group, result = result} end;
 
@@ -476,7 +475,7 @@
 
 fun promise_group group : 'a future =
   let
-    val result = Synchronized.var "promise" (NONE: 'a Exn.result option);
+    val result = Single_Assignment.var "promise" : 'a result;
     val task = SYNCHRONIZED "enqueue" (fn () =>
       Unsynchronized.change_result queue (Task_Queue.enqueue_passive group));
   in Future {promised = true, task = task, group = group, result = result} end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/Concurrent/single_assignment.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -0,0 +1,57 @@
+(*  Title:      Pure/Concurrent/single_assignment.ML
+    Author:     Makarius
+
+Single-assignment variables with locking/signalling.
+*)
+
+signature SINGLE_ASSIGNMENT =
+sig
+  type 'a var
+  val var: string -> 'a var
+  val peek: 'a var -> 'a option
+  val await: 'a var -> 'a
+  val assign: 'a var -> 'a -> unit
+end;
+
+structure Single_Assignment: SINGLE_ASSIGNMENT =
+struct
+
+abstype 'a var = Var of
+ {name: string,
+  lock: Mutex.mutex,
+  cond: ConditionVar.conditionVar,
+  var: 'a SingleAssignment.saref}
+with
+
+fun var name = Var
+ {name = name,
+  lock = Mutex.mutex (),
+  cond = ConditionVar.conditionVar (),
+  var = SingleAssignment.saref ()};
+
+fun peek (Var {var, ...}) = SingleAssignment.savalue var;
+
+fun await (v as Var {name, lock, cond, var}) =
+  SimpleThread.synchronized name lock (fn () =>
+    let
+      fun wait () =
+        (case peek v of
+          NONE =>
+            (case Multithreading.sync_wait NONE NONE cond lock of
+              Exn.Result _ => wait ()
+            | Exn.Exn exn => reraise exn)
+        | SOME x => x);
+    in wait () end);
+
+fun assign (v as Var {name, lock, cond, var}) x =
+  SimpleThread.synchronized name lock (fn () =>
+    (case peek v of
+      SOME _ => raise Fail ("Duplicate assignment to variable " ^ quote name)
+    | NONE =>
+        uninterruptible (fn _ => fn () =>
+         (SingleAssignment.saset (var, x); ConditionVar.broadcast cond)) ()));
+
+end;
+
+end;
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/Concurrent/single_assignment_sequential.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -0,0 +1,30 @@
+(*  Title:      Pure/Concurrent/single_assignment_sequential.ML
+    Author:     Makarius
+
+Single-assignment variables (sequential version).
+*)
+
+structure Single_Assignment: SINGLE_ASSIGNMENT =
+struct
+
+abstype 'a var = Var of 'a SingleAssignment.saref
+with
+
+fun var _ = Var (SingleAssignment.saref ());
+
+fun peek (Var var) = SingleAssignment.savalue var;
+
+fun await v =
+  (case peek v of
+    SOME x => x
+  | NONE => Thread.unavailable ());
+
+fun assign (v as Var var) x =
+  (case peek v of
+    SOME _ => raise Fail "Duplicate assignment to variable"
+  | NONE => SingleAssignment.saset (var, x));
+
+end;
+
+end;
+
--- a/src/Pure/Concurrent/synchronized.ML	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/Pure/Concurrent/synchronized.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -11,10 +11,8 @@
   val value: 'a var -> 'a
   val timed_access: 'a var -> ('a -> Time.time option) -> ('a -> ('b * 'a) option) -> 'b option
   val guarded_access: 'a var -> ('a -> ('b * 'a) option) -> 'b
-  val readonly_access: 'a var -> ('a -> 'b option) -> 'b
   val change_result: 'a var -> ('a -> 'b * 'a) -> 'b
   val change: 'a var -> ('a -> 'a) -> unit
-  val assign: 'a var -> ('a -> 'a) -> unit
 end;
 
 structure Synchronized: SYNCHRONIZED =
@@ -22,11 +20,12 @@
 
 (* state variables *)
 
-datatype 'a var = Var of
+abstype 'a var = Var of
  {name: string,
   lock: Mutex.mutex,
   cond: ConditionVar.conditionVar,
-  var: 'a Unsynchronized.ref};
+  var: 'a Unsynchronized.ref}
+with
 
 fun var name x = Var
  {name = name,
@@ -39,7 +38,7 @@
 
 (* synchronized access *)
 
-fun access {time_limit, readonly, finish} (Var {name, lock, cond, var}) f =
+fun timed_access (Var {name, lock, cond, var}) time_limit f =
   SimpleThread.synchronized name lock (fn () =>
     let
       fun try_change () =
@@ -51,36 +50,19 @@
               | Exn.Result false => NONE
               | Exn.Exn exn => reraise exn)
           | SOME (y, x') =>
-              if readonly then SOME y
-              else
-                let
-                  val _ = magic_immutability_test var
-                    andalso raise Fail ("Attempt to change finished variable " ^ quote name);
-                  val _ = var := x';
-                  val _ = if finish then magic_immutability_mark var else ();
-                in SOME y end)
+              uninterruptible (fn _ => fn () =>
+                (var := x'; ConditionVar.broadcast cond; SOME y)) ())
         end;
-      val res = try_change ();
-      val _ = ConditionVar.broadcast cond;
-    in res end);
-
-fun timed_access var time_limit f =
-  access {time_limit = time_limit, readonly = false, finish = false} var f;
+    in try_change () end);
 
 fun guarded_access var f = the (timed_access var (K NONE) f);
 
-fun readonly_access var f =
-  the (access {time_limit = K NONE, readonly = true, finish = false} var
-    (fn x => (case f x of NONE => NONE | SOME y => SOME (y, x))));
-
 
 (* unconditional change *)
 
 fun change_result var f = guarded_access var (SOME o f);
 fun change var f = change_result var (fn x => ((), f x));
 
-fun assign var f =
-  the (access {time_limit = K NONE, readonly = false, finish = true} var
-    (fn x => SOME ((), f x)));
+end;
 
 end;
--- a/src/Pure/Concurrent/synchronized_sequential.ML	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/Pure/Concurrent/synchronized_sequential.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -20,13 +20,8 @@
 
 fun guarded_access var f = the (timed_access var (K NONE) f);
 
-fun readonly_access var f =
-  guarded_access var (fn x => (case f x of NONE => NONE | SOME y => SOME (y, x)));
-
 fun change_result var f = guarded_access var (SOME o f);
 fun change var f = change_result var (fn x => ((), f x));
 
-val assign = change;
-
 end;
 end;
--- a/src/Pure/Concurrent/task_queue.ML	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/Pure/Concurrent/task_queue.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -127,7 +127,7 @@
 val empty = make_queue Inttab.empty Task_Graph.empty;
 
 fun all_passive (Queue {jobs, ...}) =
-  Task_Graph.get_first NONE
+  Task_Graph.get_first
     ((fn Job _ => SOME () | Running _ => SOME () | Passive => NONE) o #2 o #1 o #2) jobs |> is_none;
 
 
@@ -204,7 +204,7 @@
           if is_ready deps group then SOME (task, group, rev list) else NONE
       | ready _ = NONE;
   in
-    (case Task_Graph.get_first NONE ready jobs of
+    (case Task_Graph.get_first ready jobs of
       NONE => (NONE, queue)
     | SOME (result as (task, _, _)) =>
         let val jobs' = set_job task (Running thread) jobs
--- a/src/Pure/General/file.ML	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/Pure/General/file.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -14,7 +14,6 @@
   val full_path: Path.T -> Path.T
   val tmp_path: Path.T -> Path.T
   val isabelle_tool: string -> string -> int
-  val system_command: string -> unit
   eqtype ident
   val rep_ident: ident -> string
   val ident: Path.T -> ident option
@@ -75,11 +74,11 @@
         then SOME path
         else NONE
       end handle OS.SysErr _ => NONE) of
-    SOME path => system (shell_quote path ^ " " ^ args)
+    SOME path => bash (shell_quote path ^ " " ^ args)
   | NONE => (writeln ("Unknown Isabelle tool: " ^ name); 2));
 
 fun system_command cmd =
-  if system cmd <> 0 then error ("System command failed: " ^ cmd)
+  if bash cmd <> 0 then error ("System command failed: " ^ cmd)
   else ();
 
 
@@ -116,7 +115,7 @@
               SOME id => id
             | NONE =>
                 let val (id, rc) =  (*potentially slow*)
-                  system_out ("\"$ISABELLE_HOME/lib/scripts/fileident\" " ^ shell_quote physical_path)
+                  bash_output ("\"$ISABELLE_HOME/lib/scripts/fileident\" " ^ shell_quote physical_path)
                 in
                   if id <> "" andalso rc = 0 then (update_cache (physical_path, (mod_time, id)); id)
                   else error ("Failed to identify file " ^ quote (Path.implode path) ^ " by " ^ cmd)
--- a/src/Pure/General/graph.ML	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/Pure/General/graph.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -15,8 +15,7 @@
   val is_empty: 'a T -> bool
   val keys: 'a T -> key list
   val dest: 'a T -> (key * key list) list
-  val get_first: key option -> (key * ('a * (key list * key list)) -> 'b option) ->
-    'a T -> 'b option
+  val get_first: (key * ('a * (key list * key list)) -> 'b option) -> 'a T -> 'b option
   val fold: (key * ('a * (key list * key list)) -> 'b -> 'b) -> 'a T -> 'b -> 'b
   val minimals: 'a T -> key list
   val maximals: 'a T -> key list
@@ -89,7 +88,7 @@
 fun keys (Graph tab) = Table.keys tab;
 fun dest (Graph tab) = map (fn (x, (_, (_, succs))) => (x, succs)) (Table.dest tab);
 
-fun get_first b f (Graph tab) = Table.get_first b f tab;
+fun get_first f (Graph tab) = Table.get_first f tab;
 fun fold_graph f (Graph tab) = Table.fold f tab;
 
 fun minimals G = fold_graph (fn (m, (_, ([], _))) => cons m | _ => I) G [];
--- a/src/Pure/General/secure.ML	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/Pure/General/secure.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -15,8 +15,8 @@
   val toplevel_pp: string list -> string -> unit
   val open_unsynchronized: unit -> unit
   val commit: unit -> unit
-  val system_out: string -> string * int
-  val system: string -> int
+  val bash_output: string -> string * int
+  val bash: string -> int
 end;
 
 structure Secure: SECURE =
@@ -61,12 +61,12 @@
 
 fun secure_shell () = deny_secure "Cannot execute shell commands in secure mode";
 
-val orig_system_out = system_out;
+val orig_bash_output = bash_output;
 
-fun system_out s = (secure_shell (); orig_system_out s);
+fun bash_output s = (secure_shell (); orig_bash_output s);
 
-fun system s =
-  (case system_out s of
+fun bash s =
+  (case bash_output s of
     ("", rc) => rc
   | (out, rc) => (writeln (perhaps (try (unsuffix "\n")) out); rc));
 
@@ -78,5 +78,5 @@
 fun use s = Secure.use_file ML_Parse.global_context true s
   handle ERROR msg => (writeln msg; error "ML error");
 val toplevel_pp = Secure.toplevel_pp;
-val system_out = Secure.system_out;
-val system = Secure.system;
+val bash_output = Secure.bash_output;
+val bash = Secure.bash;
--- a/src/Pure/General/table.ML	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/Pure/General/table.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -28,7 +28,7 @@
   val keys: 'a table -> key list
   val min_key: 'a table -> key option
   val max_key: 'a table -> key option
-  val get_first: key option -> (key * 'a -> 'b option) -> 'a table -> 'b option
+  val get_first: (key * 'a -> 'b option) -> 'a table -> 'b option
   val exists: (key * 'a -> bool) -> 'a table -> bool
   val forall: (key * 'a -> bool) -> 'a table -> bool
   val lookup: 'a table -> key -> 'a option
@@ -131,40 +131,32 @@
 
 (* get_first *)
 
-fun get_first boundary f tab =
+fun get_first f =
   let
-    val check =
-      (case boundary of
-        NONE => K true
-      | SOME b => (fn k => Key.ord (b, k) = LESS));
-    fun apply (k, x) = if check k then f (k, x) else NONE;
-    fun get_bounded tb k = if check k then get tb else NONE
-    and get Empty = NONE
+    fun get Empty = NONE
       | get (Branch2 (left, (k, x), right)) =
-          (case get_bounded left k of
+          (case get left of
             NONE =>
-              (case apply (k, x) of
+              (case f (k, x) of
                 NONE => get right
               | some => some)
           | some => some)
       | get (Branch3 (left, (k1, x1), mid, (k2, x2), right)) =
-          (case get_bounded left k1 of
+          (case get left of
             NONE =>
-              (case apply (k1, x1) of
+              (case f (k1, x1) of
                 NONE =>
-                  (case get_bounded mid k2 of
+                  (case get mid of
                     NONE =>
-                      (case apply (k2, x2) of
+                      (case f (k2, x2) of
                         NONE => get right
                       | some => some)
                   | some => some)
               | some => some)
           | some => some);
-  in get tab end;
+  in get end;
 
-fun exists pred =
-  is_some o get_first NONE (fn entry => if pred entry then SOME () else NONE);
-
+fun exists pred = is_some o get_first (fn entry => if pred entry then SOME () else NONE);
 fun forall pred = not o exists (not o pred);
 
 
--- a/src/Pure/IsaMakefile	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/Pure/IsaMakefile	Mon Feb 08 14:08:32 2010 +0100
@@ -21,19 +21,19 @@
 
 ## Pure
 
-BOOTSTRAP_FILES = ML-Systems/compiler_polyml-5.0.ML			\
+BOOTSTRAP_FILES = ML-Systems/bash.ML ML-Systems/compiler_polyml-5.0.ML	\
   ML-Systems/compiler_polyml-5.2.ML ML-Systems/compiler_polyml-5.3.ML	\
-  ML-Systems/ml_name_space.ML				\
-  ML-Systems/ml_pretty.ML ML-Systems/mosml.ML				\
+  ML-Systems/ml_name_space.ML ML-Systems/ml_pretty.ML			\
   ML-Systems/multithreading.ML ML-Systems/multithreading_polyml.ML	\
   ML-Systems/overloading_smlnj.ML ML-Systems/polyml-5.0.ML		\
   ML-Systems/polyml-5.1.ML ML-Systems/polyml-5.2.ML			\
   ML-Systems/polyml-5.2.1.ML ML-Systems/polyml.ML			\
   ML-Systems/polyml_common.ML ML-Systems/pp_polyml.ML			\
-  ML-Systems/proper_int.ML ML-Systems/smlnj.ML				\
-  ML-Systems/system_shell.ML ML-Systems/thread_dummy.ML			\
-  ML-Systems/timing.ML ML-Systems/time_limit.ML				\
-  ML-Systems/universal.ML ML-Systems/unsynchronized.ML
+  ML-Systems/proper_int.ML ML-Systems/single_assignment.ML		\
+  ML-Systems/single_assignment_polyml.ML ML-Systems/smlnj.ML		\
+  ML-Systems/thread_dummy.ML ML-Systems/timing.ML			\
+  ML-Systems/time_limit.ML ML-Systems/universal.ML			\
+  ML-Systems/unsynchronized.ML
 
 RAW: $(OUT)/RAW
 
@@ -47,17 +47,18 @@
   Concurrent/future.ML Concurrent/lazy.ML				\
   Concurrent/lazy_sequential.ML Concurrent/mailbox.ML			\
   Concurrent/par_list.ML Concurrent/par_list_sequential.ML		\
-  Concurrent/simple_thread.ML Concurrent/synchronized.ML		\
-  Concurrent/synchronized_sequential.ML Concurrent/task_queue.ML	\
-  General/alist.ML General/antiquote.ML General/balanced_tree.ML	\
-  General/basics.ML General/binding.ML General/buffer.ML		\
-  General/exn.ML General/file.ML General/graph.ML General/heap.ML	\
-  General/integer.ML General/long_name.ML General/markup.ML		\
-  General/name_space.ML General/ord_list.ML General/output.ML		\
-  General/path.ML General/position.ML General/pretty.ML			\
-  General/print_mode.ML General/properties.ML General/queue.ML		\
-  General/same.ML General/scan.ML General/secure.ML General/seq.ML	\
-  General/source.ML General/stack.ML General/symbol.ML			\
+  Concurrent/simple_thread.ML Concurrent/single_assignment.ML		\
+  Concurrent/single_assignment_sequential.ML				\
+  Concurrent/synchronized.ML Concurrent/synchronized_sequential.ML	\
+  Concurrent/task_queue.ML General/alist.ML General/antiquote.ML	\
+  General/balanced_tree.ML General/basics.ML General/binding.ML		\
+  General/buffer.ML General/exn.ML General/file.ML General/graph.ML	\
+  General/heap.ML General/integer.ML General/long_name.ML		\
+  General/markup.ML General/name_space.ML General/ord_list.ML		\
+  General/output.ML General/path.ML General/position.ML			\
+  General/pretty.ML General/print_mode.ML General/properties.ML		\
+  General/queue.ML General/same.ML General/scan.ML General/secure.ML	\
+  General/seq.ML General/source.ML General/stack.ML General/symbol.ML	\
   General/symbol_pos.ML General/table.ML General/url.ML General/xml.ML	\
   General/yxml.ML Isar/args.ML Isar/attrib.ML Isar/auto_bind.ML		\
   Isar/calculation.ML Isar/class.ML Isar/class_target.ML Isar/code.ML	\
--- a/src/Pure/Isar/args.ML	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/Pure/Isar/args.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -1,7 +1,7 @@
 (*  Title:      Pure/Isar/args.ML
     Author:     Markus Wenzel, TU Muenchen
 
-Parsing with implicit value assigment.  Concrete argument syntax of
+Parsing with implicit value assignment.  Concrete argument syntax of
 attributes, methods etc.
 *)
 
--- a/src/Pure/Isar/attrib.ML	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/Pure/Isar/attrib.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -298,7 +298,7 @@
   setup (Binding.name "params")
     (Scan.lift (P.and_list1 (Scan.repeat Args.name)) >> Rule_Cases.params)
     "named rule parameters" #>
-  setup (Binding.name "standard") (Scan.succeed (Thm.rule_attribute (K Drule.standard)))
+  setup (Binding.name "standard") (Scan.succeed (Thm.rule_attribute (K Drule.export_without_context)))
     "result put into standard form (legacy)" #>
   setup (Binding.name "rule_format") rule_format "result put into canonical rule format" #>
   setup (Binding.name "elim_format") (Scan.succeed elim_format)
--- a/src/Pure/Isar/class.ML	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/Pure/Isar/class.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -86,7 +86,7 @@
       | SOME prop => Logic.mk_implies (Morphism.term const_morph
           ((map_types o map_atyps) (K aT) prop), of_class_prop_concl);
     val sup_of_classes = map (snd o rules thy) sups;
-    val loc_axiom_intros = map Drule.standard' (Locale.axioms_of thy class);
+    val loc_axiom_intros = map Drule.export_without_context_open (Locale.axioms_of thy class);
     val axclass_intro = #intro (AxClass.get_info thy class);
     val base_sort_trivs = Thm.of_sort (Thm.ctyp_of thy aT, base_sort);
     val tac = REPEAT (SOMEGOAL
--- a/src/Pure/Isar/class_target.ML	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/Pure/Isar/class_target.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -233,7 +233,7 @@
 fun register_subclass (sub, sup) some_dep_morph some_wit export thy =
   let
     val intros = (snd o rules thy) sup :: map_filter I
-      [Option.map (Drule.standard' o Element.conclude_witness) some_wit,
+      [Option.map (Drule.export_without_context_open o Element.conclude_witness) some_wit,
         (fst o rules thy) sub];
     val tac = EVERY (map (TRYALL o Tactic.rtac) intros);
     val classrel = Skip_Proof.prove_global thy [] [] (Logic.mk_classrel (sub, sup))
--- a/src/Pure/Isar/expression.ML	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/Pure/Isar/expression.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -699,7 +699,7 @@
             |> PureThy.note_thmss ""
                  [((Binding.conceal (Binding.name introN), []), [([intro], [Locale.intro_add])]),
                   ((Binding.conceal (Binding.name axiomsN), []),
-                    [(map (Drule.standard o Element.conclude_witness) axioms, [])])]
+                    [(map (Drule.export_without_context o Element.conclude_witness) axioms, [])])]
             ||> Sign.restore_naming thy''';
         in (SOME statement, SOME intro, axioms, thy'''') end;
   in ((a_pred, a_intro, a_axioms), (b_pred, b_intro, b_axioms), thy'''') end;
--- a/src/Pure/Isar/skip_proof.ML	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/Pure/Isar/skip_proof.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -39,6 +39,6 @@
       else tac args st);
 
 fun prove_global thy xs asms prop tac =
-  Drule.standard (prove (ProofContext.init thy) xs asms prop tac);
+  Drule.export_without_context (prove (ProofContext.init thy) xs asms prop tac);
 
 end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/ML-Systems/bash.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -0,0 +1,43 @@
+(*  Title:      Pure/ML-Systems/bash.ML
+    Author:     Makarius
+
+Generic GNU bash processes (no provisions to propagate interrupts, but
+could work via the controlling tty).
+*)
+
+local
+
+fun read_file name =
+  let val is = TextIO.openIn name
+  in Exn.release (Exn.capture TextIO.inputAll is before TextIO.closeIn is) end;
+
+fun write_file name txt =
+  let val os = TextIO.openOut name
+  in Exn.release (Exn.capture TextIO.output (os, txt) before TextIO.closeOut os) end;
+
+in
+
+fun bash_output script =
+  let
+    val script_name = OS.FileSys.tmpName ();
+    val _ = write_file script_name script;
+
+    val output_name = OS.FileSys.tmpName ();
+
+    val status =
+      OS.Process.system ("exec \"$ISABELLE_HOME/lib/scripts/bash\" nogroup " ^
+        script_name ^ " /dev/null " ^ output_name);
+    val rc =
+      (case Posix.Process.fromStatus status of
+        Posix.Process.W_EXITED => 0
+      | Posix.Process.W_EXITSTATUS w => Word8.toInt w
+      | Posix.Process.W_SIGNALED s => 256 + LargeWord.toInt (Posix.Signal.toWord s)
+      | Posix.Process.W_STOPPED s => 512 + LargeWord.toInt (Posix.Signal.toWord s));
+
+    val output = read_file output_name handle IO.Io _ => "";
+    val _ = OS.FileSys.remove script_name handle OS.SysErr _ => ();
+    val _ = OS.FileSys.remove output_name handle OS.SysErr _ => ();
+  in (output, rc) end;
+
+end;
+
--- a/src/Pure/ML-Systems/mosml.ML	Mon Feb 08 14:06:58 2010 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,259 +0,0 @@
-(*  Title:      Pure/ML-Systems/mosml.ML
-    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
-    Author:     Makarius
-
-Compatibility file for Moscow ML 2.01
-
-NOTE: saving images does not work; may run it interactively as follows:
-
-$ cd Isabelle/src/Pure
-$ isabelle-process RAW_ML_SYSTEM
-> val ml_system = "mosml";
-> use "ML-Systems/mosml.ML";
-> use "ROOT.ML";
-> Session.finish ();
-> cd "../FOL";
-> use "ROOT.ML";
-> Session.finish ();
-> cd "../ZF";
-> use "ROOT.ML";
-*)
-
-(** ML system related **)
-
-val ml_system_fix_ints = false;
-
-fun forget_structure _ = ();
-
-load "Obj";
-load "Word";
-load "Bool";
-load "Int";
-load "Real";
-load "ListPair";
-load "OS";
-load "Process";
-load "FileSys";
-load "IO";
-load "CharVector";
-
-exception Interrupt;
-fun reraise exn = raise exn;
-
-use "ML-Systems/unsynchronized.ML";
-use "General/exn.ML";
-use "ML-Systems/universal.ML";
-use "ML-Systems/thread_dummy.ML";
-use "ML-Systems/multithreading.ML";
-use "ML-Systems/time_limit.ML";
-use "ML-Systems/ml_name_space.ML";
-use "ML-Systems/ml_pretty.ML";
-use "ML-Systems/use_context.ML";
-
-
-(*low-level pointer equality*)
-local val cast : 'a -> int = Obj.magic
-in fun pointer_eq (x:'a, y:'a) = (cast x = cast y) end;
-
-(*proper implementation available?*)
-structure IntInf =
-struct
-  fun divMod (x, y) = (x div y, x mod y);
-
-  local
-    fun log 1 a = a
-      | log n a = log (n div 2) (a + 1);
-  in
-    fun log2 n = if n > 0 then log n 0 else raise Domain;
-  end;
-
-  open Int;
-end;
-
-structure Substring =
-struct
-  open Substring;
-  val full = all;
-end;
-
-structure Real =
-struct
-  open Real;
-  val realFloor = real o floor;
-end;
-
-structure String =
-struct
-  fun isSuffix s1 s2 =
-    let val n1 = size s1 and n2 = size s2
-    in if n1 = n2 then s1 = s2 else n1 <= n2 andalso String.substring (s2, n2 - n1, n1) = s1 end;
-  open String;
-end;
-
-structure Time =
-struct
-  open Time;
-  fun toString t = Time.toString t
-    handle Overflow => Real.toString (Time.toReal t);   (*workaround Y2004 bug*)
-end;
-
-structure OS =
-struct
-  open OS
-  structure Process =
-  struct
-    open Process
-    fun sleep _ = raise Fail "OS.Process.sleep undefined"
-  end;
-  structure FileSys = FileSys
-end;
-
-exception Option = Option.Option;
-
-
-(*limit the printing depth*)
-local
-  val depth = ref 10;
-in
-  fun get_print_depth () = ! depth;
-  fun print_depth n =
-   (depth := n;
-    Meta.printDepth := n div 2;
-    Meta.printLength := n);
-end;
-
-(*dummy implementation*)
-fun toplevel_pp _ _ _ = ();
-
-(*dummy implementation*)
-fun ml_prompts p1 p2 = ();
-
-(*dummy implementation*)
-fun profile (n: int) f x = f x;
-
-(*dummy implementation*)
-fun exception_trace f = f ();
-
-
-
-(** Compiler-independent timing functions **)
-
-load "Timer";
-
-fun start_timing () =
-  let
-    val timer = Timer.startCPUTimer ();
-    val time = Timer.checkCPUTimer timer;
-  in (timer, time) end;
-
-fun end_timing (timer, {gc, sys, usr}) =
-  let
-    open Time;  (*...for Time.toString, Time.+ and Time.- *)
-    val {gc = gc2, sys = sys2, usr = usr2} = Timer.checkCPUTimer timer;
-    val user = usr2 - usr + gc2 - gc;
-    val all = user + sys2 - sys;
-    val message = "User " ^ toString user ^ "  All "^ toString all ^ " secs" handle Time => "";
-  in {message = message, user = user, all = all} end;
-
-
-(* SML basis library fixes *)
-
-structure TextIO =
-struct
-  fun canInput _ = raise IO.Io {cause = Fail "undefined", function = "canInput", name = ""};
-  open TextIO;
-  fun inputLine is =
-    let val s = TextIO.inputLine is
-    in if s = "" then NONE else SOME s end;
-  fun getOutstream _ = ();
-  structure StreamIO =
-  struct
-    fun setBufferMode _ = ();
-  end
-end;
-
-structure IO =
-struct
-  open IO;
-  val BLOCK_BUF = ();
-end;
-
-
-(* ML command execution *)
-
-(*Can one redirect 'use' directly to an instream?*)
-fun use_text ({tune_source, ...}: use_context) _ _ txt =
-  let
-    val tmp_name = FileSys.tmpName ();
-    val tmp_file = TextIO.openOut tmp_name;
-  in
-    TextIO.output (tmp_file, tune_source txt);
-    TextIO.closeOut tmp_file;
-    use tmp_name;
-    FileSys.remove tmp_name
-  end;
-
-fun use_file _ _ name = use name;
-
-
-
-(** interrupts **)      (*Note: may get into race conditions*)
-
-(* FIXME proper implementation available? *)
-
-fun interruptible f x = f x;
-fun uninterruptible f x = f (fn (g: 'c -> 'd) => g) x;
-
-
-
-(** OS related **)
-
-(*dummy implementation*)
-structure Posix =
-struct
-  structure ProcEnv =
-  struct
-    fun getpid () = 0;
-  end;  
-end;
-
-local
-
-fun read_file name =
-  let val is = TextIO.openIn name
-  in Exn.release (Exn.capture TextIO.inputAll is before TextIO.closeIn is) end;
-
-fun write_file name txt =
-  let val os = TextIO.openOut name
-  in Exn.release (Exn.capture TextIO.output (os, txt) before TextIO.closeOut os) end;
-
-in
-
-fun system_out script =
-  let
-    val script_name = OS.FileSys.tmpName ();
-    val _ = write_file script_name script;
-
-    val output_name = OS.FileSys.tmpName ();
-
-    val status =
-      OS.Process.system ("perl -w \"$ISABELLE_HOME/lib/scripts/system.pl\" nogroup " ^
-        script_name ^ " /dev/null " ^ output_name);
-    val rc = if status = OS.Process.success then 0 else 1;
-
-    val output = read_file output_name handle IO.Io _ => "";
-    val _ = OS.FileSys.remove script_name handle OS.SysErr _ => ();
-    val _ = OS.FileSys.remove output_name handle OS.SysErr _ => ();
-  in (output, rc) end;
-
-end;
-
-val cd = OS.FileSys.chDir;
-val pwd = OS.FileSys.getDir;
-
-val process_id = Int.toString o Posix.ProcEnv.getpid;
-
-fun getenv var =
-  (case Process.getEnv var of
-    NONE => ""
-  | SOME txt => txt);
--- a/src/Pure/ML-Systems/multithreading_polyml.ML	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/Pure/ML-Systems/multithreading_polyml.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -8,7 +8,7 @@
 sig
   val interruptible: ('a -> 'b) -> 'a -> 'b
   val uninterruptible: ((('c -> 'd) -> 'c -> 'd) -> 'a -> 'b) -> 'a -> 'b
-  val system_out: string -> string * int
+  val bash_output: string -> string * int
   structure TimeLimit: TIME_LIMIT
 end;
 
@@ -156,9 +156,9 @@
 end;
 
 
-(* system shell processes, with propagation of interrupts *)
+(* GNU bash processes, with propagation of interrupts *)
 
-fun system_out script = with_attributes no_interrupts (fn orig_atts =>
+fun bash_output script = with_attributes no_interrupts (fn orig_atts =>
   let
     val script_name = OS.FileSys.tmpName ();
     val _ = write_file script_name script;
@@ -180,7 +180,7 @@
     val system_thread = Thread.fork (fn () =>
       let
         val status =
-          OS.Process.system ("perl -w \"$ISABELLE_HOME/lib/scripts/system.pl\" group " ^
+          OS.Process.system ("exec \"$ISABELLE_HOME/lib/scripts/bash\" group " ^
             script_name ^ " " ^ pid_name ^ " " ^ output_name);
         val res =
           (case Posix.Process.fromStatus status of
--- a/src/Pure/ML-Systems/polyml_common.ML	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/Pure/ML-Systems/polyml_common.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -6,10 +6,11 @@
 exception Interrupt = SML90.Interrupt;
 
 use "General/exn.ML";
+use "ML-Systems/single_assignment_polyml.ML";
 use "ML-Systems/multithreading.ML";
 use "ML-Systems/time_limit.ML";
 use "ML-Systems/timing.ML";
-use "ML-Systems/system_shell.ML";
+use "ML-Systems/bash.ML";
 use "ML-Systems/ml_pretty.ML";
 use "ML-Systems/use_context.ML";
 
@@ -131,12 +132,3 @@
         val _ = RunCall.run_call1 RuntimeCalls.POLY_SYS_profiler 0;
       in Exn.release res end;
 
-
-(* magic immutability -- for internal use only! *)
-
-fun magic_immutability_mark (r: 'a Unsynchronized.ref) =
-  ignore (RunCall.run_call1 RuntimeCalls.POLY_SYS_lockseg r);
-
-fun magic_immutability_test (r: 'a Unsynchronized.ref) =
-  Word8.andb (0wx40, RunCall.run_call1 RuntimeCalls.POLY_SYS_get_flags r) = 0w0;
-
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/ML-Systems/single_assignment.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -0,0 +1,33 @@
+(*  Title:      Pure/ML-Systems/single_assignment.ML
+    Author:     Makarius
+
+References with single assignment.  Unsynchronized!
+*)
+
+signature SINGLE_ASSIGNMENT =
+sig
+  type 'a saref
+  exception Locked
+  val saref: unit -> 'a saref
+  val savalue: 'a saref -> 'a option
+  val saset: 'a saref * 'a -> unit
+end;
+
+structure SingleAssignment: SINGLE_ASSIGNMENT =
+struct
+
+exception Locked;
+
+abstype 'a saref = SARef of 'a option ref
+with
+
+fun saref () = SARef (ref NONE);
+
+fun savalue (SARef r) = ! r;
+
+fun saset (SARef (r as ref NONE), x) = r := SOME x
+  | saset _ = raise Locked;
+
+end;
+
+end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/ML-Systems/single_assignment_polyml.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -0,0 +1,35 @@
+(*  Title:      Pure/ML-Systems/single_assignment_polyml.ML
+    Author:     Makarius
+
+References with single assignment.  Unsynchronized!  Emulates
+structure SingleAssignment from Poly/ML 5.4.
+*)
+
+signature SINGLE_ASSIGNMENT =
+sig
+  type 'a saref
+  exception Locked
+  val saref: unit -> 'a saref
+  val savalue: 'a saref -> 'a option
+  val saset: 'a saref * 'a -> unit
+end;
+
+structure SingleAssignment: SINGLE_ASSIGNMENT =
+struct
+
+exception Locked;
+
+abstype 'a saref = SARef of 'a option ref
+with
+
+fun saref () = SARef (ref NONE);
+
+fun savalue (SARef r) = ! r;
+
+fun saset (SARef (r as ref NONE), x) =
+      (r := SOME x; RunCall.run_call1 RuntimeCalls.POLY_SYS_lockseg r)
+  | saset _ = raise Locked;
+
+end;
+
+end;
--- a/src/Pure/ML-Systems/smlnj.ML	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/Pure/ML-Systems/smlnj.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -10,11 +10,12 @@
 use "ML-Systems/unsynchronized.ML";
 use "ML-Systems/overloading_smlnj.ML";
 use "General/exn.ML";
+use "ML-Systems/single_assignment.ML";
 use "ML-Systems/universal.ML";
 use "ML-Systems/thread_dummy.ML";
 use "ML-Systems/multithreading.ML";
 use "ML-Systems/timing.ML";
-use "ML-Systems/system_shell.ML";
+use "ML-Systems/bash.ML";
 use "ML-Systems/ml_name_space.ML";
 use "ML-Systems/ml_pretty.ML";
 use "ML-Systems/use_context.ML";
@@ -66,10 +67,6 @@
   (Control.primaryPrompt := p1; Control.secondaryPrompt := p2);
 
 (*dummy implementation*)
-fun magic_immutability_test _ = false;
-fun magic_immutability_mark _ = ();
-
-(*dummy implementation*)
 fun profile (n: int) f x = f x;
 
 (*dummy implementation*)
@@ -179,7 +176,7 @@
 
 (* system command execution *)
 
-val system_out = (fn (output, rc) => (output, mk_int rc)) o system_out;
+val bash_output = (fn (output, rc) => (output, mk_int rc)) o bash_output;
 
 fun process_id pid =
   Word.fmt StringCvt.DEC (Word.fromLargeWord (Posix.Process.pidToWord (Posix.ProcEnv.getpid ())));
--- a/src/Pure/ML-Systems/system_shell.ML	Mon Feb 08 14:06:58 2010 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,43 +0,0 @@
-(*  Title:      Pure/ML-Systems/system_shell.ML
-    Author:     Makarius
-
-Generic system shell processes (no provisions to propagate interrupts;
-might still work via the controlling tty).
-*)
-
-local
-
-fun read_file name =
-  let val is = TextIO.openIn name
-  in Exn.release (Exn.capture TextIO.inputAll is before TextIO.closeIn is) end;
-
-fun write_file name txt =
-  let val os = TextIO.openOut name
-  in Exn.release (Exn.capture TextIO.output (os, txt) before TextIO.closeOut os) end;
-
-in
-
-fun system_out script =
-  let
-    val script_name = OS.FileSys.tmpName ();
-    val _ = write_file script_name script;
-
-    val output_name = OS.FileSys.tmpName ();
-
-    val status =
-      OS.Process.system ("perl -w \"$ISABELLE_HOME/lib/scripts/system.pl\" nogroup " ^
-        script_name ^ " /dev/null " ^ output_name);
-    val rc =
-      (case Posix.Process.fromStatus status of
-        Posix.Process.W_EXITED => 0
-      | Posix.Process.W_EXITSTATUS w => Word8.toInt w
-      | Posix.Process.W_SIGNALED s => 256 + LargeWord.toInt (Posix.Signal.toWord s)
-      | Posix.Process.W_STOPPED s => 512 + LargeWord.toInt (Posix.Signal.toWord s));
-
-    val output = read_file output_name handle IO.Io _ => "";
-    val _ = OS.FileSys.remove script_name handle OS.SysErr _ => ();
-    val _ = OS.FileSys.remove output_name handle OS.SysErr _ => ();
-  in (output, rc) end;
-
-end;
-
--- a/src/Pure/ML/ml_antiquote.ML	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/Pure/ML/ml_antiquote.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -38,17 +38,17 @@
 
 fun macro name scan = ML_Context.add_antiq name
   (fn _ => scan :|-- (fn ctxt => Scan.depend (fn _ => Scan.succeed
-    (Context.Proof ctxt, fn {background, ...} => (K ("", ""), background)))));
+    (Context.Proof ctxt, fn background => (K ("", ""), background)))));
 
 fun inline name scan = ML_Context.add_antiq name
-  (fn _ => scan >> (fn s => fn {background, ...} => (K ("", s), background)));
+  (fn _ => scan >> (fn s => fn background => (K ("", s), background)));
 
 fun declaration kind name scan = ML_Context.add_antiq name
-  (fn _ => scan >> (fn s => fn {struct_name, background} =>
+  (fn _ => scan >> (fn s => fn background =>
     let
       val (a, background') = variant name background;
       val env = kind ^ " " ^ a ^ " = " ^ s ^ ";\n";
-      val body = struct_name ^ "." ^ a;
+      val body = "Isabelle." ^ a;
     in (K (env, body), background') end));
 
 val value = declaration "val";
--- a/src/Pure/ML/ml_context.ML	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/Pure/ML/ml_context.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -22,9 +22,7 @@
   val stored_thms: thm list Unsynchronized.ref
   val ml_store_thm: string * thm -> unit
   val ml_store_thms: string * thm list -> unit
-  type antiq =
-    {struct_name: string, background: Proof.context} ->
-      (Proof.context -> string * string) * Proof.context
+  type antiq = Proof.context -> (Proof.context -> string * string) * Proof.context
   val add_antiq: string -> (Position.T -> antiq context_parser) -> unit
   val trace: bool Unsynchronized.ref
   val eval_antiquotes: ML_Lex.token Antiquote.antiquote list * Position.T ->
@@ -72,8 +70,8 @@
 val ml_store_thms = ml_store "";
 fun ml_store_thm (name, th) = ml_store "hd" (name, [th]);
 
-fun bind_thm (name, thm) = ml_store_thm (name, Drule.standard thm);
-fun bind_thms (name, thms) = ml_store_thms (name, map Drule.standard thms);
+fun bind_thm (name, thm) = ml_store_thm (name, Drule.export_without_context thm);
+fun bind_thms (name, thms) = ml_store_thms (name, map Drule.export_without_context thms);
 
 fun thm name = ProofContext.get_thm (the_local_context ()) name;
 fun thms name = ProofContext.get_thms (the_local_context ()) name;
@@ -84,9 +82,7 @@
 
 (* antiquotation commands *)
 
-type antiq =
-  {struct_name: string, background: Proof.context} ->
-    (Proof.context -> string * string) * Proof.context;
+type antiq = Proof.context -> (Proof.context -> string * string) * Proof.context;
 
 local
 
@@ -123,8 +119,7 @@
   P.!!! (P.position P.xname -- Args.parse --| Scan.ahead P.eof)
   >> (fn ((x, pos), y) => Args.src ((x, y), pos));
 
-val struct_name = "Isabelle";
-val begin_env = ML_Lex.tokenize ("structure " ^ struct_name ^ " =\nstruct\n");
+val begin_env = ML_Lex.tokenize "structure Isabelle =\nstruct\n";
 val end_env = ML_Lex.tokenize "end;";
 val reset_env = ML_Lex.tokenize "structure Isabelle = struct end";
 
@@ -151,7 +146,7 @@
                 let
                   val context = Stack.top scope;
                   val (f, context') = antiquotation (T.read_antiq lex antiq (ss, #1 range)) context;
-                  val (decl, background') = f {background = background, struct_name = struct_name};
+                  val (decl, background') = f background;
                   val decl' = decl #> pairself (ML_Lex.tokenize #> map (ML_Lex.set_range range));
                 in (decl', (Stack.map_top (K context') scope, background')) end
             | expand (Antiquote.Open _) (scope, background) =
--- a/src/Pure/ML/ml_thms.ML	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/Pure/ML/ml_thms.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -34,12 +34,12 @@
 (* fact references *)
 
 fun thm_antiq kind scan = ML_Context.add_antiq kind
-  (fn _ => scan >> (fn ths => fn {struct_name, background} =>
+  (fn _ => scan >> (fn ths => fn background =>
     let
       val i = serial ();
       val (a, background') = background
         |> ML_Antiquote.variant kind ||> put_thms (i, ths);
-      val ml = (thm_bind kind a i, struct_name ^ "." ^ a);
+      val ml = (thm_bind kind a i, "Isabelle." ^ a);
     in (K ml, background') end));
 
 val _ = thm_antiq "thm" (Attrib.thm >> single);
@@ -56,7 +56,7 @@
   (fn _ => Args.context -- Args.mode "open" --
       Scan.lift (OuterParse.and_list1 (Scan.repeat1 goal) --
         (by |-- Method.parse -- Scan.option Method.parse)) >>
-    (fn ((ctxt, is_open), (raw_propss, methods)) => fn {struct_name, background} =>
+    (fn ((ctxt, is_open), (raw_propss, methods)) => fn background =>
       let
         val propss = burrow (map (rpair []) o Syntax.read_props ctxt) raw_propss;
         val i = serial ();
@@ -69,8 +69,7 @@
         val (a, background') = background
           |> ML_Antiquote.variant "lemma" ||> put_thms (i, the_thms ctxt' i);
         val ml =
-         (thm_bind (if length (flat propss) = 1 then "thm" else "thms") a i,
-          struct_name ^ "." ^ a);
+          (thm_bind (if length (flat propss) = 1 then "thm" else "thms") a i, "Isabelle." ^ a);
       in (K ml, background') end));
 
 end;
--- a/src/Pure/ROOT.ML	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/Pure/ROOT.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -57,6 +57,10 @@
 
 use "Concurrent/simple_thread.ML";
 
+use "Concurrent/single_assignment.ML";
+if Multithreading.available then ()
+else use "Concurrent/single_assignment_sequential.ML";
+
 use "Concurrent/synchronized.ML";
 if Multithreading.available then ()
 else use "Concurrent/synchronized_sequential.ML";
--- a/src/Pure/System/isabelle_system.scala	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/Pure/System/isabelle_system.scala	Mon Feb 08 14:08:32 2010 +0100
@@ -162,7 +162,7 @@
 
   /** system tools **/
 
-  def system_out(script: String): (String, Int) =
+  def bash_output(script: String): (String, Int) =
   {
     Standard_System.with_tmp_file("isabelle_script") { script_file =>
       Standard_System.with_tmp_file("isabelle_pid") { pid_file =>
@@ -170,8 +170,7 @@
 
           Standard_System.write_file(script_file, script)
 
-          val proc = execute(true, "perl", "-w",
-            expand_path("$ISABELLE_HOME/lib/scripts/system.pl"), "group",
+          val proc = execute(true, expand_path("$ISABELLE_HOME/lib/scripts/bash"), "group",
             script_file.getPath, pid_file.getPath, output_file.getPath)
 
           def kill(strict: Boolean) =
@@ -310,20 +309,29 @@
 
   val font_family = "IsabelleText"
 
-  private def check_font(): Boolean =
-    new Font(font_family, Font.PLAIN, 1).getFamily == font_family
-
-  private def create_font(name: String) =
-    Font.createFont(Font.TRUETYPE_FONT, platform_file(name))
+  def get_font(bold: Boolean): Font =
+    new Font(font_family, if (bold) Font.BOLD else Font.PLAIN, 1)
 
   def install_fonts()
   {
+    def create_font(bold: Boolean): Font =
+    {
+      val name =
+        if (bold) "$ISABELLE_HOME/lib/fonts/IsabelleTextBold.ttf"
+        else "$ISABELLE_HOME/lib/fonts/IsabelleText.ttf"
+      Font.createFont(Font.TRUETYPE_FONT, platform_file(name))
+    }
+    def check_font() = get_font(false).getFamily == font_family
+
     if (!check_font()) {
+      val font = create_font(false)
+      val bold_font = create_font(true)
+
       val ge = GraphicsEnvironment.getLocalGraphicsEnvironment()
-      ge.registerFont(create_font("$ISABELLE_HOME/lib/fonts/IsabelleText.ttf"))
-      ge.registerFont(create_font("$ISABELLE_HOME/lib/fonts/IsabelleTextBold.ttf"))
-      if (!check_font())
-        error("Failed to install IsabelleText fonts")
+      ge.registerFont(font)
+      // workaround strange problem with Apple's Java 1.6 font manager
+      if (bold_font.getFamily == font_family) ge.registerFont(bold_font)
+      if (!check_font()) error("Failed to install IsabelleText fonts")
     }
   }
 }
--- a/src/Pure/System/platform.scala	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/Pure/System/platform.scala	Mon Feb 08 14:08:32 2010 +0100
@@ -55,15 +55,15 @@
 
   /* Swing look-and-feel */
 
+  private def find_laf(name: String): Option[String] =
+    UIManager.getInstalledLookAndFeels().find(_.getName == name).map(_.getClassName)
+
   def look_and_feel(): String =
   {
     if (is_windows || is_macos) UIManager.getSystemLookAndFeelClassName()
-    else {
-      UIManager.getInstalledLookAndFeels().find(laf => laf.getName == "Nimbus") match {
-        case None => UIManager.getCrossPlatformLookAndFeelClassName()
-        case Some(laf) => laf.getClassName
-      }
-    }
+    else
+      find_laf("Nimbus") orElse find_laf("GTK+") getOrElse
+      UIManager.getCrossPlatformLookAndFeelClassName()
   }
 }
 
--- a/src/Pure/System/session.scala	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/Pure/System/session.scala	Mon Feb 08 14:08:32 2010 +0100
@@ -122,14 +122,13 @@
         // global status message
         result.body match {
 
-          // document state assigment
+          // document state assignment
           case List(XML.Elem(Markup.ASSIGN, _, edits)) if target_id.isDefined =>
             documents.get(target_id.get) match {
               case Some(doc) =>
                 val states =
                   for {
-                    XML.Elem(Markup.EDIT, (Markup.ID, cmd_id) :: (Markup.STATE, state_id) :: _, _)
-                      <- edits
+                    XML.Elem(Markup.EDIT, (Markup.ID, cmd_id) :: (Markup.STATE, state_id) :: _, _) <- edits
                     cmd <- lookup_command(cmd_id)
                   } yield {
                     val st = cmd.assign_state(state_id)
--- a/src/Pure/Thy/completion.scala	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/Pure/Thy/completion.scala	Mon Feb 08 14:08:32 2010 +0100
@@ -88,12 +88,11 @@
     abbrevs_lex.parse(abbrevs_lex.keyword, new Library.Reverse(line)) match {
       case abbrevs_lex.Success(rev_a, _) =>
         val (word, c) = abbrevs_map(rev_a)
-        if (word == c) None
-        else Some(word, List(c))
+        Some(word, List(c))
       case _ =>
         Completion.Parse.read(line) match {
           case Some(word) =>
-            words_lex.completions(word).map(words_map(_)).filter(_ != word) match {
+            words_lex.completions(word).map(words_map(_)) match {
               case Nil => None
               case cs => Some(word, cs.sort(_ < _))
             }
--- a/src/Pure/Thy/present.ML	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/Pure/Thy/present.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -328,7 +328,7 @@
       \-n '" ^ name ^ "' -t '" ^ tags ^ "' " ^ File.shell_path path ^ " 2>&1";
     val doc_path = Path.append result_path (Path.ext format (Path.basic name));
     val _ = if verbose then writeln s else ();
-    val (out, rc) = system_out s;
+    val (out, rc) = bash_output s;
     val _ =
       if not (File.exists doc_path) orelse rc <> 0 then
         cat_error out ("Failed to build document " ^ quote (show_path doc_path))
--- a/src/Pure/axclass.ML	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/Pure/axclass.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -484,10 +484,10 @@
       def_thy
       |> Sign.mandatory_path (Binding.name_of bconst)
       |> PureThy.note_thmss ""
-        [((Binding.name introN, []), [([Drule.standard raw_intro], [])]),
-         ((Binding.name superN, []), [(map Drule.standard raw_classrel, [])]),
+        [((Binding.name introN, []), [([Drule.export_without_context raw_intro], [])]),
+         ((Binding.name superN, []), [(map Drule.export_without_context raw_classrel, [])]),
          ((Binding.name axiomsN, []),
-           [(map (fn th => Drule.standard (class_triv RS th)) raw_axioms, [])])]
+           [(map (fn th => Drule.export_without_context (class_triv RS th)) raw_axioms, [])])]
       ||> Sign.restore_naming def_thy;
 
 
--- a/src/Pure/drule.ML	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/Pure/drule.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -75,8 +75,8 @@
   val beta_conv: cterm -> cterm -> cterm
   val types_sorts: thm -> (indexname-> typ option) * (indexname-> sort option)
   val flexflex_unique: thm -> thm
-  val standard: thm -> thm
-  val standard': thm -> thm
+  val export_without_context: thm -> thm
+  val export_without_context_open: thm -> thm
   val get_def: theory -> xstring -> thm
   val store_thm: binding -> thm -> thm
   val store_standard_thm: binding -> thm -> thm
@@ -303,9 +303,9 @@
     |  _   => raise THM("flexflex_unique: multiple unifiers", 0, [th]);
 
 
-(* legacy standard operations *)
+(* old-style export without context *)
 
-val standard' =
+val export_without_context_open =
   implies_intr_hyps
   #> forall_intr_frees
   #> `Thm.maxidx_of
@@ -315,9 +315,9 @@
     #> zero_var_indexes
     #> Thm.varifyT);
 
-val standard =
+val export_without_context =
   flexflex_unique
-  #> standard'
+  #> export_without_context_open
   #> Thm.close_derivation;
 
 
@@ -459,8 +459,8 @@
 fun store_thm_open name th =
   Context.>>> (Context.map_theory_result (PureThy.store_thm_open (name, th)));
 
-fun store_standard_thm name th = store_thm name (standard th);
-fun store_standard_thm_open name thm = store_thm_open name (standard' thm);
+fun store_standard_thm name th = store_thm name (export_without_context th);
+fun store_standard_thm_open name thm = store_thm_open name (export_without_context_open thm);
 
 val reflexive_thm =
   let val cx = certify (Var(("x",0),TVar(("'a",0),[])))
@@ -708,12 +708,12 @@
 val protect = Thm.capply (certify Logic.protectC);
 
 val protectI =
-  store_thm (Binding.conceal (Binding.name "protectI"))
-    (standard (Thm.equal_elim (Thm.symmetric prop_def) (Thm.assume A)));
+  store_standard_thm (Binding.conceal (Binding.name "protectI"))
+    (Thm.equal_elim (Thm.symmetric prop_def) (Thm.assume A));
 
 val protectD =
-  store_thm (Binding.conceal (Binding.name "protectD"))
-    (standard (Thm.equal_elim prop_def (Thm.assume (protect A))));
+  store_standard_thm (Binding.conceal (Binding.name "protectD"))
+    (Thm.equal_elim prop_def (Thm.assume (protect A)));
 
 val protect_cong =
   store_standard_thm_open (Binding.name "protect_cong") (Thm.reflexive (protect A));
@@ -730,8 +730,8 @@
 (* term *)
 
 val termI =
-  store_thm (Binding.conceal (Binding.name "termI"))
-    (standard (Thm.equal_elim (Thm.symmetric term_def) (Thm.forall_intr A (Thm.trivial A))));
+  store_standard_thm (Binding.conceal (Binding.name "termI"))
+    (Thm.equal_elim (Thm.symmetric term_def) (Thm.forall_intr A (Thm.trivial A)));
 
 fun mk_term ct =
   let
@@ -759,15 +759,14 @@
 (* sort_constraint *)
 
 val sort_constraintI =
-  store_thm (Binding.conceal (Binding.name "sort_constraintI"))
-    (standard (Thm.equal_elim (Thm.symmetric sort_constraint_def) (mk_term T)));
+  store_standard_thm (Binding.conceal (Binding.name "sort_constraintI"))
+    (Thm.equal_elim (Thm.symmetric sort_constraint_def) (mk_term T));
 
 val sort_constraint_eq =
-  store_thm (Binding.conceal (Binding.name "sort_constraint_eq"))
-    (standard
-      (Thm.equal_intr
-        (Thm.implies_intr CA (Thm.implies_elim (Thm.assume CA) (Thm.unvarify sort_constraintI)))
-        (implies_intr_list [A, C] (Thm.assume A))));
+  store_standard_thm (Binding.conceal (Binding.name "sort_constraint_eq"))
+    (Thm.equal_intr
+      (Thm.implies_intr CA (Thm.implies_elim (Thm.assume CA) (Thm.unvarify sort_constraintI)))
+      (implies_intr_list [A, C] (Thm.assume A)));
 
 end;
 
@@ -983,5 +982,5 @@
 
 end;
 
-structure BasicDrule: BASIC_DRULE = Drule;
-open BasicDrule;
+structure Basic_Drule: BASIC_DRULE = Drule;
+open Basic_Drule;
--- a/src/Pure/goal.ML	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/Pure/goal.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -210,7 +210,7 @@
 fun prove ctxt xs asms prop tac = hd (prove_common true ctxt xs asms [prop] tac);
 
 fun prove_global thy xs asms prop tac =
-  Drule.standard (prove (ProofContext.init thy) xs asms prop tac);
+  Drule.export_without_context (prove (ProofContext.init thy) xs asms prop tac);
 
 
 
--- a/src/Pure/old_goals.ML	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/Pure/old_goals.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -305,7 +305,7 @@
             val th = Goal.conclude ath
             val thy' = Thm.theory_of_thm th
             val {hyps, prop, ...} = Thm.rep_thm th
-            val final_th = Drule.standard th
+            val final_th = Drule.export_without_context th
         in  if not check then final_th
             else if not (Theory.eq_thy(thy,thy')) then !result_error_fn state
                 ("Theory of proof state has changed!" ^
@@ -512,7 +512,7 @@
             0 => thm
           | i => !result_error_fn thm (string_of_int i ^ " unsolved goals!"))
   in
-    Drule.standard (implies_intr_list As
+    Drule.export_without_context (implies_intr_list As
       (check (Seq.pull (EVERY (f asms) (trivial B)))))
   end;
 
--- a/src/Pure/term.ML	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/Pure/term.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -12,10 +12,10 @@
 
 signature BASIC_TERM =
 sig
-  eqtype indexname
-  eqtype class
-  eqtype sort
-  eqtype arity
+  type indexname = string * int
+  type class = string
+  type sort = class list
+  type arity = string * sort list * sort
   datatype typ =
     Type  of string * typ list |
     TFree of string * sort |
--- a/src/Pure/type_infer.ML	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/Pure/type_infer.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -284,7 +284,7 @@
       | meets _ tye_idx = tye_idx;
 
 
-    (* occurs check and assigment *)
+    (* occurs check and assignment *)
 
     fun occurs_check tye i (Param (i', S)) =
           if i = i' then raise NO_UNIFIER ("Occurs check!", tye)
--- a/src/Sequents/simpdata.ML	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/Sequents/simpdata.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -49,9 +49,9 @@
 
 (*Congruence rules for = or <-> (instead of ==)*)
 fun mk_meta_cong rl =
-  Drule.standard(mk_meta_eq (mk_meta_prems rl))
-  handle THM _ =>
-  error("Premises and conclusion of congruence rules must use =-equality or <->");
+  Drule.export_without_context(mk_meta_eq (mk_meta_prems rl))
+    handle THM _ =>
+      error("Premises and conclusion of congruence rules must use =-equality or <->");
 
 
 (*** Standard simpsets ***)
--- a/src/Tools/Code/code_eval.ML	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/Tools/Code/code_eval.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -107,25 +107,25 @@
     val _ = map2 check_base constrs constrs'';
   in "datatype " ^ tyco'' ^ " = datatype " ^ Long_Name.append all_struct_name tyco'' end;
 
-fun print_code struct_name is_first print_it ctxt =
+fun print_code is_first print_it ctxt =
   let
     val (_, (_, (struct_code_name, acc_code))) = CodeAntiqData.get ctxt;
     val (ml_code, (tycos_map, consts_map)) = Lazy.force acc_code;
     val ml_code = if is_first then ml_code
       else "";
-    val all_struct_name = Long_Name.append struct_name struct_code_name;
+    val all_struct_name = "Isabelle." ^ struct_code_name;
   in (ml_code, print_it all_struct_name tycos_map consts_map) end;
 
 in
 
-fun ml_code_antiq raw_const {struct_name, background} =
+fun ml_code_antiq raw_const background =
   let
     val const = Code.check_const (ProofContext.theory_of background) raw_const;
     val is_first = is_first_occ background;
     val background' = register_const const background;
-  in (print_code struct_name is_first (print_const const), background') end;
+  in (print_code is_first (print_const const), background') end;
 
-fun ml_code_datatype_antiq (raw_tyco, raw_constrs) {struct_name, background} =
+fun ml_code_datatype_antiq (raw_tyco, raw_constrs) background =
   let
     val thy = ProofContext.theory_of background;
     val tyco = Sign.intern_type thy raw_tyco;
@@ -135,7 +135,7 @@
       else error ("Type " ^ quote tyco ^ ": given constructors diverge from real constructors")
     val is_first = is_first_occ background;
     val background' = register_datatype tyco constrs background;
-  in (print_code struct_name is_first (print_datatype tyco constrs), background') end;
+  in (print_code is_first (print_datatype tyco constrs), background') end;
 
 end; (*local*)
 
--- a/src/Tools/Compute_Oracle/am_ghc.ML	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/Tools/Compute_Oracle/am_ghc.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -228,7 +228,7 @@
         val module_file = tmp_file (module^".hs")
         val object_file = tmp_file (module^".o")
         val _ = writeTextFile module_file source
-        val _ = system ((!ghc)^" -c "^module_file)
+        val _ = bash ((!ghc)^" -c "^module_file)
         val _ = if not (fileExists object_file) then raise Compile ("Failure compiling haskell code (GHC_PATH = '"^(!ghc)^"')") else ()
     in
         (guid, module_file, arity)      
@@ -309,7 +309,7 @@
         val term = print_term arity_of 0 t
         val call_source = "module "^call^" where\n\nimport "^module^"\n\ncall = "^module^".calc \""^result_file^"\" ("^term^")"
         val _ = writeTextFile call_file call_source
-        val _ = system ((!ghc)^" -e \""^call^".call\" "^module_file^" "^call_file)
+        val _ = bash ((!ghc)^" -e \""^call^".call\" "^module_file^" "^call_file)
         val result = readResultFile result_file handle IO.Io _ => raise Run ("Failure running haskell compiler (GHC_PATH = '"^(!ghc)^"')")
         val t' = parse_result arity_of result
         val _ = OS.FileSys.remove call_file
--- a/src/Tools/jEdit/README_BUILD	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/Tools/jEdit/README_BUILD	Mon Feb 08 14:08:32 2010 +0100
@@ -2,10 +2,10 @@
 Requirements to build from sources
 ==================================
 
-* Proper Java JRE/JDK from Sun, e.g. 1.6.0_17
+* Proper Java JRE/JDK from Sun, e.g. 1.6.0_18
   http://java.sun.com/javase/downloads/index.jsp
 
-* Netbeans 6.7
+* Netbeans 6.7.1
   http://www.netbeans.org/downloads/index.html
 
 * Scala for Netbeans: version 6.7v1 for NB 6.7
@@ -13,7 +13,7 @@
   http://blogtrader.net/dcaoyuan/category/NetBeans
   http://wiki.netbeans.org/Scala
 
-* jEdit 4.3 (final)
+* jEdit 4.3.1 (final)
   http://www.jedit.org/
 
   Netbeans Project "jEdit": install official sources as ./contrib/jEdit/.
--- a/src/Tools/jEdit/src/jedit/isabelle_sidekick.scala	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/Tools/jEdit/src/jedit/isabelle_sidekick.scala	Mon Feb 08 14:08:32 2010 +0100
@@ -88,11 +88,11 @@
       case None => null
       case Some((word, cs)) =>
         val ds =
-          if (Isabelle_Encoding.is_active(buffer))
+          (if (Isabelle_Encoding.is_active(buffer))
             cs.map(Isabelle.system.symbols.decode(_)).sort(_ < _)
-          else cs
-        new SideKickCompletion(pane.getView, word, ds.toArray.asInstanceOf[Array[Object]]) { }
+           else cs).filter(_ != word)
+        if (ds.isEmpty) null
+        else new SideKickCompletion(pane.getView, word, ds.toArray.asInstanceOf[Array[Object]]) { }
     }
   }
-
 }
--- a/src/ZF/Tools/datatype_package.ML	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/ZF/Tools/datatype_package.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -292,7 +292,7 @@
          rtac case_trans 1,
          REPEAT (resolve_tac [refl, split_trans, Su.case_inl RS trans, Su.case_inr RS trans] 1)]);
 
-  val free_iffs = map Drule.standard (con_defs RL [@{thm def_swap_iff}]);
+  val free_iffs = map Drule.export_without_context (con_defs RL [@{thm def_swap_iff}]);
 
   val case_eqns = map prove_case_eqn (flat con_ty_lists ~~ case_args ~~ tl con_defs);
 
@@ -338,7 +338,7 @@
   val constructors =
       map (head_of o #1 o Logic.dest_equals o #prop o rep_thm) (tl con_defs);
 
-  val free_SEs = map Drule.standard (Ind_Syntax.mk_free_SEs free_iffs);
+  val free_SEs = map Drule.export_without_context (Ind_Syntax.mk_free_SEs free_iffs);
 
   val {intrs, elim, induct, mutual_induct, ...} = ind_result
 
--- a/src/ZF/Tools/inductive_package.ML	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/ZF/Tools/inductive_package.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -193,9 +193,9 @@
         [rtac (@{thm Collect_subset} RS @{thm bnd_monoI}) 1,
          REPEAT (ares_tac (@{thms basic_monos} @ monos) 1)]);
 
-  val dom_subset = Drule.standard (big_rec_def RS Fp.subs);
+  val dom_subset = Drule.export_without_context (big_rec_def RS Fp.subs);
 
-  val unfold = Drule.standard ([big_rec_def, bnd_mono] MRS Fp.Tarski);
+  val unfold = Drule.export_without_context ([big_rec_def, bnd_mono] MRS Fp.Tarski);
 
   (********)
   val dummy = writeln "  Proving the introduction rules...";
@@ -205,7 +205,7 @@
   val Part_trans =
       case rec_names of
            [_] => asm_rl
-         | _   => Drule.standard (@{thm Part_subset} RS @{thm subset_trans});
+         | _   => Drule.export_without_context (@{thm Part_subset} RS @{thm subset_trans});
 
   (*To type-check recursive occurrences of the inductive sets, possibly
     enclosed in some monotonic operator M.*)
@@ -272,7 +272,7 @@
     rule_by_tactic
       (basic_elim_tac THEN ALLGOALS (asm_full_simp_tac ss) THEN basic_elim_tac)
       (Thm.assume A RS elim)
-      |> Drule.standard';
+      |> Drule.export_without_context_open;
 
   fun induction_rules raw_induct thy =
    let
@@ -503,7 +503,7 @@
      val Const (@{const_name Trueprop}, _) $ (pred_var $ _) = concl_of induct0
 
      val induct = CP.split_rule_var(pred_var, elem_type-->FOLogic.oT, induct0)
-                  |> Drule.standard
+                  |> Drule.export_without_context
      and mutual_induct = CP.remove_split mutual_induct_fsplit
 
      val ([induct', mutual_induct'], thy') =
@@ -514,7 +514,7 @@
     in ((thy', induct'), mutual_induct')
     end;  (*of induction_rules*)
 
-  val raw_induct = Drule.standard ([big_rec_def, bnd_mono] MRS Fp.induct)
+  val raw_induct = Drule.export_without_context ([big_rec_def, bnd_mono] MRS Fp.induct)
 
   val ((thy2, induct), mutual_induct) =
     if not coind then induction_rules raw_induct thy1
--- a/src/ZF/ind_syntax.ML	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/ZF/ind_syntax.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -114,7 +114,7 @@
   | tryres (th, []) = raise THM("tryres", 0, [th]);
 
 fun gen_make_elim elim_rls rl =
-      Drule.standard (tryres (rl, elim_rls @ [revcut_rl]));
+  Drule.export_without_context (tryres (rl, elim_rls @ [revcut_rl]));
 
 (*Turns iff rules into safe elimination rules*)
 fun mk_free_SEs iffs = map (gen_make_elim [@{thm conjE}, @{thm FalseE}]) (iffs RL [@{thm iffD1}]);
--- a/src/ZF/int_arith.ML	Mon Feb 08 14:06:58 2010 +0100
+++ b/src/ZF/int_arith.ML	Mon Feb 08 14:08:32 2010 +0100
@@ -110,9 +110,8 @@
 (*To let us treat subtraction as addition*)
 val diff_simps = [@{thm zdiff_def}, @{thm zminus_zadd_distrib}, @{thm zminus_zminus}];
 
-(*push the unary minus down: - x * y = x * - y *)
-val int_minus_mult_eq_1_to_2 =
-    [@{thm zmult_zminus}, @{thm zmult_zminus_right} RS sym] MRS trans |> Drule.standard;
+(*push the unary minus down*)
+val int_minus_mult_eq_1_to_2 = @{lemma "$- w $* z = w $* $- z" by simp};
 
 (*to extract again any uncancelled minuses*)
 val int_minus_from_mult_simps =