doc-src/IsarImplementation/Thy/Prelim.thy
author wenzelm
Sun, 08 Mar 2009 17:37:18 +0100
changeset 30365 790129514c2d
parent 30281 9ad15d8ed311
child 33174 1f2051f41335
permissions -rw-r--r--
adapted to structure Long_Name;
wenzelm@30081
     1
theory Prelim
wenzelm@30081
     2
imports Base
wenzelm@30081
     3
begin
wenzelm@18537
     4
wenzelm@18537
     5
chapter {* Preliminaries *}
wenzelm@18537
     6
wenzelm@20437
     7
section {* Contexts \label{sec:context} *}
wenzelm@20437
     8
wenzelm@20437
     9
text {*
wenzelm@20451
    10
  A logical context represents the background that is required for
wenzelm@20451
    11
  formulating statements and composing proofs.  It acts as a medium to
wenzelm@20451
    12
  produce formal content, depending on earlier material (declarations,
wenzelm@20451
    13
  results etc.).
wenzelm@20437
    14
wenzelm@20451
    15
  For example, derivations within the Isabelle/Pure logic can be
wenzelm@20451
    16
  described as a judgment @{text "\<Gamma> \<turnstile>\<^sub>\<Theta> \<phi>"}, which means that a
wenzelm@20437
    17
  proposition @{text "\<phi>"} is derivable from hypotheses @{text "\<Gamma>"}
wenzelm@20437
    18
  within the theory @{text "\<Theta>"}.  There are logical reasons for
wenzelm@20451
    19
  keeping @{text "\<Theta>"} and @{text "\<Gamma>"} separate: theories can be
wenzelm@20451
    20
  liberal about supporting type constructors and schematic
wenzelm@20451
    21
  polymorphism of constants and axioms, while the inner calculus of
wenzelm@20451
    22
  @{text "\<Gamma> \<turnstile> \<phi>"} is strictly limited to Simple Type Theory (with
wenzelm@20451
    23
  fixed type variables in the assumptions).
wenzelm@20437
    24
wenzelm@20437
    25
  \medskip Contexts and derivations are linked by the following key
wenzelm@20437
    26
  principles:
wenzelm@20437
    27
wenzelm@20437
    28
  \begin{itemize}
wenzelm@20437
    29
wenzelm@20437
    30
  \item Transfer: monotonicity of derivations admits results to be
wenzelm@20451
    31
  transferred into a \emph{larger} context, i.e.\ @{text "\<Gamma> \<turnstile>\<^sub>\<Theta>
wenzelm@20451
    32
  \<phi>"} implies @{text "\<Gamma>' \<turnstile>\<^sub>\<Theta>\<^sub>' \<phi>"} for contexts @{text "\<Theta>'
wenzelm@20451
    33
  \<supseteq> \<Theta>"} and @{text "\<Gamma>' \<supseteq> \<Gamma>"}.
wenzelm@20437
    34
wenzelm@20437
    35
  \item Export: discharge of hypotheses admits results to be exported
wenzelm@20451
    36
  into a \emph{smaller} context, i.e.\ @{text "\<Gamma>' \<turnstile>\<^sub>\<Theta> \<phi>"}
wenzelm@20451
    37
  implies @{text "\<Gamma> \<turnstile>\<^sub>\<Theta> \<Delta> \<Longrightarrow> \<phi>"} where @{text "\<Gamma>' \<supseteq> \<Gamma>"} and
wenzelm@20451
    38
  @{text "\<Delta> = \<Gamma>' - \<Gamma>"}.  Note that @{text "\<Theta>"} remains unchanged here,
wenzelm@20451
    39
  only the @{text "\<Gamma>"} part is affected.
wenzelm@20437
    40
wenzelm@20437
    41
  \end{itemize}
wenzelm@20437
    42
wenzelm@20451
    43
  \medskip By modeling the main characteristics of the primitive
wenzelm@20451
    44
  @{text "\<Theta>"} and @{text "\<Gamma>"} above, and abstracting over any
wenzelm@20451
    45
  particular logical content, we arrive at the fundamental notions of
wenzelm@20451
    46
  \emph{theory context} and \emph{proof context} in Isabelle/Isar.
wenzelm@20451
    47
  These implement a certain policy to manage arbitrary \emph{context
wenzelm@20451
    48
  data}.  There is a strongly-typed mechanism to declare new kinds of
wenzelm@20437
    49
  data at compile time.
wenzelm@20437
    50
wenzelm@20451
    51
  The internal bootstrap process of Isabelle/Pure eventually reaches a
wenzelm@20451
    52
  stage where certain data slots provide the logical content of @{text
wenzelm@20451
    53
  "\<Theta>"} and @{text "\<Gamma>"} sketched above, but this does not stop there!
wenzelm@20451
    54
  Various additional data slots support all kinds of mechanisms that
wenzelm@20451
    55
  are not necessarily part of the core logic.
wenzelm@20437
    56
wenzelm@20437
    57
  For example, there would be data for canonical introduction and
wenzelm@20437
    58
  elimination rules for arbitrary operators (depending on the
wenzelm@20437
    59
  object-logic and application), which enables users to perform
wenzelm@20451
    60
  standard proof steps implicitly (cf.\ the @{text "rule"} method
wenzelm@20451
    61
  \cite{isabelle-isar-ref}).
wenzelm@20437
    62
wenzelm@20451
    63
  \medskip Thus Isabelle/Isar is able to bring forth more and more
wenzelm@20451
    64
  concepts successively.  In particular, an object-logic like
wenzelm@20451
    65
  Isabelle/HOL continues the Isabelle/Pure setup by adding specific
wenzelm@20451
    66
  components for automated reasoning (classical reasoner, tableau
wenzelm@20451
    67
  prover, structured induction etc.) and derived specification
wenzelm@20451
    68
  mechanisms (inductive predicates, recursive functions etc.).  All of
wenzelm@20451
    69
  this is ultimately based on the generic data management by theory
wenzelm@20451
    70
  and proof contexts introduced here.
wenzelm@20437
    71
*}
wenzelm@20437
    72
wenzelm@20437
    73
wenzelm@20437
    74
subsection {* Theory context \label{sec:context-theory} *}
wenzelm@20437
    75
wenzelm@20437
    76
text {*
wenzelm@30087
    77
  A \emph{theory} is a data container with explicit name and unique
wenzelm@20451
    78
  identifier.  Theories are related by a (nominal) sub-theory
wenzelm@20451
    79
  relation, which corresponds to the dependency graph of the original
wenzelm@20451
    80
  construction; each theory is derived from a certain sub-graph of
wenzelm@20451
    81
  ancestor theories.
wenzelm@20451
    82
wenzelm@20451
    83
  The @{text "merge"} operation produces the least upper bound of two
wenzelm@20451
    84
  theories, which actually degenerates into absorption of one theory
wenzelm@20451
    85
  into the other (due to the nominal sub-theory relation).
wenzelm@20437
    86
wenzelm@20437
    87
  The @{text "begin"} operation starts a new theory by importing
wenzelm@20437
    88
  several parent theories and entering a special @{text "draft"} mode,
wenzelm@20437
    89
  which is sustained until the final @{text "end"} operation.  A draft
wenzelm@20451
    90
  theory acts like a linear type, where updates invalidate earlier
wenzelm@20451
    91
  versions.  An invalidated draft is called ``stale''.
wenzelm@20437
    92
wenzelm@20447
    93
  The @{text "checkpoint"} operation produces an intermediate stepping
wenzelm@20451
    94
  stone that will survive the next update: both the original and the
wenzelm@20451
    95
  changed theory remain valid and are related by the sub-theory
wenzelm@20451
    96
  relation.  Checkpointing essentially recovers purely functional
wenzelm@20451
    97
  theory values, at the expense of some extra internal bookkeeping.
wenzelm@20437
    98
wenzelm@20447
    99
  The @{text "copy"} operation produces an auxiliary version that has
wenzelm@20447
   100
  the same data content, but is unrelated to the original: updates of
wenzelm@20447
   101
  the copy do not affect the original, neither does the sub-theory
wenzelm@20447
   102
  relation hold.
wenzelm@20447
   103
wenzelm@20447
   104
  \medskip The example in \figref{fig:ex-theory} below shows a theory
wenzelm@20451
   105
  graph derived from @{text "Pure"}, with theory @{text "Length"}
wenzelm@20451
   106
  importing @{text "Nat"} and @{text "List"}.  The body of @{text
wenzelm@20451
   107
  "Length"} consists of a sequence of updates, working mostly on
wenzelm@20451
   108
  drafts.  Intermediate checkpoints may occur as well, due to the
wenzelm@20451
   109
  history mechanism provided by the Isar top-level, cf.\
wenzelm@20451
   110
  \secref{sec:isar-toplevel}.
wenzelm@20447
   111
wenzelm@20447
   112
  \begin{figure}[htb]
wenzelm@20447
   113
  \begin{center}
wenzelm@20437
   114
  \begin{tabular}{rcccl}
wenzelm@20447
   115
        &            & @{text "Pure"} \\
wenzelm@20447
   116
        &            & @{text "\<down>"} \\
wenzelm@20447
   117
        &            & @{text "FOL"} \\
wenzelm@20437
   118
        & $\swarrow$ &              & $\searrow$ & \\
wenzelm@21852
   119
  @{text "Nat"} &    &              &            & @{text "List"} \\
wenzelm@20437
   120
        & $\searrow$ &              & $\swarrow$ \\
wenzelm@20447
   121
        &            & @{text "Length"} \\
wenzelm@26864
   122
        &            & \multicolumn{3}{l}{~~@{keyword "imports"}} \\
wenzelm@26864
   123
        &            & \multicolumn{3}{l}{~~@{keyword "begin"}} \\
wenzelm@20437
   124
        &            & $\vdots$~~ \\
wenzelm@20447
   125
        &            & @{text "\<bullet>"}~~ \\
wenzelm@20447
   126
        &            & $\vdots$~~ \\
wenzelm@20447
   127
        &            & @{text "\<bullet>"}~~ \\
wenzelm@20447
   128
        &            & $\vdots$~~ \\
wenzelm@26864
   129
        &            & \multicolumn{3}{l}{~~@{command "end"}} \\
wenzelm@20437
   130
  \end{tabular}
wenzelm@20451
   131
  \caption{A theory definition depending on ancestors}\label{fig:ex-theory}
wenzelm@20447
   132
  \end{center}
wenzelm@20447
   133
  \end{figure}
wenzelm@20451
   134
wenzelm@20451
   135
  \medskip There is a separate notion of \emph{theory reference} for
wenzelm@20451
   136
  maintaining a live link to an evolving theory context: updates on
wenzelm@20488
   137
  drafts are propagated automatically.  Dynamic updating stops after
wenzelm@20488
   138
  an explicit @{text "end"} only.
wenzelm@20451
   139
wenzelm@20451
   140
  Derived entities may store a theory reference in order to indicate
wenzelm@20451
   141
  the context they belong to.  This implicitly assumes monotonic
wenzelm@20451
   142
  reasoning, because the referenced context may become larger without
wenzelm@20451
   143
  further notice.
wenzelm@20437
   144
*}
wenzelm@20437
   145
wenzelm@20437
   146
text %mlref {*
wenzelm@20447
   147
  \begin{mldecls}
wenzelm@20447
   148
  @{index_ML_type theory} \\
wenzelm@20447
   149
  @{index_ML Theory.subthy: "theory * theory -> bool"} \\
wenzelm@20447
   150
  @{index_ML Theory.merge: "theory * theory -> theory"} \\
wenzelm@20447
   151
  @{index_ML Theory.checkpoint: "theory -> theory"} \\
wenzelm@20547
   152
  @{index_ML Theory.copy: "theory -> theory"} \\
wenzelm@20547
   153
  \end{mldecls}
wenzelm@20547
   154
  \begin{mldecls}
wenzelm@20447
   155
  @{index_ML_type theory_ref} \\
wenzelm@20447
   156
  @{index_ML Theory.deref: "theory_ref -> theory"} \\
wenzelm@24137
   157
  @{index_ML Theory.check_thy: "theory -> theory_ref"} \\
wenzelm@20447
   158
  \end{mldecls}
wenzelm@20447
   159
wenzelm@20447
   160
  \begin{description}
wenzelm@20447
   161
wenzelm@20451
   162
  \item @{ML_type theory} represents theory contexts.  This is
wenzelm@20451
   163
  essentially a linear type!  Most operations destroy the original
wenzelm@20451
   164
  version, which then becomes ``stale''.
wenzelm@20447
   165
wenzelm@20447
   166
  \item @{ML "Theory.subthy"}~@{text "(thy\<^sub>1, thy\<^sub>2)"}
wenzelm@20447
   167
  compares theories according to the inherent graph structure of the
wenzelm@20447
   168
  construction.  This sub-theory relation is a nominal approximation
wenzelm@20447
   169
  of inclusion (@{text "\<subseteq>"}) of the corresponding content.
wenzelm@20447
   170
wenzelm@20447
   171
  \item @{ML "Theory.merge"}~@{text "(thy\<^sub>1, thy\<^sub>2)"}
wenzelm@20447
   172
  absorbs one theory into the other.  This fails for unrelated
wenzelm@20447
   173
  theories!
wenzelm@20447
   174
wenzelm@20447
   175
  \item @{ML "Theory.checkpoint"}~@{text "thy"} produces a safe
wenzelm@20447
   176
  stepping stone in the linear development of @{text "thy"}.  The next
wenzelm@20447
   177
  update will result in two related, valid theories.
wenzelm@20447
   178
wenzelm@20447
   179
  \item @{ML "Theory.copy"}~@{text "thy"} produces a variant of @{text
wenzelm@20451
   180
  "thy"} that holds a copy of the same data.  The result is not
wenzelm@30087
   181
  related to the original; the original is unchanged.
wenzelm@20447
   182
wenzelm@20451
   183
  \item @{ML_type theory_ref} represents a sliding reference to an
wenzelm@20451
   184
  always valid theory; updates on the original are propagated
wenzelm@20447
   185
  automatically.
wenzelm@20447
   186
wenzelm@24137
   187
  \item @{ML "Theory.deref"}~@{text "thy_ref"} turns a @{ML_type
wenzelm@24137
   188
  "theory_ref"} into an @{ML_type "theory"} value.  As the referenced
wenzelm@24137
   189
  theory evolves monotonically over time, later invocations of @{ML
wenzelm@20451
   190
  "Theory.deref"} may refer to a larger context.
wenzelm@20447
   191
wenzelm@24137
   192
  \item @{ML "Theory.check_thy"}~@{text "thy"} produces a @{ML_type
wenzelm@24137
   193
  "theory_ref"} from a valid @{ML_type "theory"} value.
wenzelm@24137
   194
wenzelm@20447
   195
  \end{description}
wenzelm@20437
   196
*}
wenzelm@20437
   197
wenzelm@20437
   198
wenzelm@20437
   199
subsection {* Proof context \label{sec:context-proof} *}
wenzelm@20437
   200
wenzelm@20437
   201
text {*
wenzelm@20447
   202
  A proof context is a container for pure data with a back-reference
wenzelm@20449
   203
  to the theory it belongs to.  The @{text "init"} operation creates a
wenzelm@20451
   204
  proof context from a given theory.  Modifications to draft theories
wenzelm@20451
   205
  are propagated to the proof context as usual, but there is also an
wenzelm@20451
   206
  explicit @{text "transfer"} operation to force resynchronization
wenzelm@20451
   207
  with more substantial updates to the underlying theory.  The actual
wenzelm@20451
   208
  context data does not require any special bookkeeping, thanks to the
wenzelm@20451
   209
  lack of destructive features.
wenzelm@20437
   210
wenzelm@20447
   211
  Entities derived in a proof context need to record inherent logical
wenzelm@20447
   212
  requirements explicitly, since there is no separate context
wenzelm@20447
   213
  identification as for theories.  For example, hypotheses used in
wenzelm@20451
   214
  primitive derivations (cf.\ \secref{sec:thms}) are recorded
wenzelm@20447
   215
  separately within the sequent @{text "\<Gamma> \<turnstile> \<phi>"}, just to make double
wenzelm@30087
   216
  sure.  Results could still leak into an alien proof context due to
wenzelm@20447
   217
  programming errors, but Isabelle/Isar includes some extra validity
wenzelm@22438
   218
  checks in critical positions, notably at the end of a sub-proof.
wenzelm@20437
   219
wenzelm@20451
   220
  Proof contexts may be manipulated arbitrarily, although the common
wenzelm@20451
   221
  discipline is to follow block structure as a mental model: a given
wenzelm@20451
   222
  context is extended consecutively, and results are exported back
wenzelm@20451
   223
  into the original context.  Note that the Isar proof states model
wenzelm@20451
   224
  block-structured reasoning explicitly, using a stack of proof
wenzelm@30084
   225
  contexts internally.
wenzelm@20437
   226
*}
wenzelm@20437
   227
wenzelm@20449
   228
text %mlref {*
wenzelm@20449
   229
  \begin{mldecls}
wenzelm@20449
   230
  @{index_ML_type Proof.context} \\
wenzelm@20449
   231
  @{index_ML ProofContext.init: "theory -> Proof.context"} \\
wenzelm@20449
   232
  @{index_ML ProofContext.theory_of: "Proof.context -> theory"} \\
wenzelm@20449
   233
  @{index_ML ProofContext.transfer: "theory -> Proof.context -> Proof.context"} \\
wenzelm@20449
   234
  \end{mldecls}
wenzelm@20449
   235
wenzelm@20449
   236
  \begin{description}
wenzelm@20449
   237
wenzelm@20449
   238
  \item @{ML_type Proof.context} represents proof contexts.  Elements
wenzelm@20449
   239
  of this type are essentially pure values, with a sliding reference
wenzelm@20449
   240
  to the background theory.
wenzelm@20449
   241
wenzelm@20449
   242
  \item @{ML ProofContext.init}~@{text "thy"} produces a proof context
wenzelm@20449
   243
  derived from @{text "thy"}, initializing all data.
wenzelm@20449
   244
wenzelm@20449
   245
  \item @{ML ProofContext.theory_of}~@{text "ctxt"} selects the
wenzelm@20451
   246
  background theory from @{text "ctxt"}, dereferencing its internal
wenzelm@20451
   247
  @{ML_type theory_ref}.
wenzelm@20449
   248
wenzelm@20449
   249
  \item @{ML ProofContext.transfer}~@{text "thy ctxt"} promotes the
wenzelm@20449
   250
  background theory of @{text "ctxt"} to the super theory @{text
wenzelm@20449
   251
  "thy"}.
wenzelm@20449
   252
wenzelm@20449
   253
  \end{description}
wenzelm@20449
   254
*}
wenzelm@20449
   255
wenzelm@20437
   256
wenzelm@20451
   257
subsection {* Generic contexts \label{sec:generic-context} *}
wenzelm@20437
   258
wenzelm@20449
   259
text {*
wenzelm@20449
   260
  A generic context is the disjoint sum of either a theory or proof
wenzelm@20451
   261
  context.  Occasionally, this enables uniform treatment of generic
wenzelm@20450
   262
  context data, typically extra-logical information.  Operations on
wenzelm@20449
   263
  generic contexts include the usual injections, partial selections,
wenzelm@20449
   264
  and combinators for lifting operations on either component of the
wenzelm@20449
   265
  disjoint sum.
wenzelm@20437
   266
wenzelm@20449
   267
  Moreover, there are total operations @{text "theory_of"} and @{text
wenzelm@20449
   268
  "proof_of"} to convert a generic context into either kind: a theory
wenzelm@20451
   269
  can always be selected from the sum, while a proof context might
wenzelm@20451
   270
  have to be constructed by an ad-hoc @{text "init"} operation.
wenzelm@20449
   271
*}
wenzelm@20437
   272
wenzelm@20449
   273
text %mlref {*
wenzelm@20449
   274
  \begin{mldecls}
wenzelm@20449
   275
  @{index_ML_type Context.generic} \\
wenzelm@20449
   276
  @{index_ML Context.theory_of: "Context.generic -> theory"} \\
wenzelm@20449
   277
  @{index_ML Context.proof_of: "Context.generic -> Proof.context"} \\
wenzelm@20449
   278
  \end{mldecls}
wenzelm@20449
   279
wenzelm@20449
   280
  \begin{description}
wenzelm@20449
   281
wenzelm@20449
   282
  \item @{ML_type Context.generic} is the direct sum of @{ML_type
wenzelm@20451
   283
  "theory"} and @{ML_type "Proof.context"}, with the datatype
wenzelm@20451
   284
  constructors @{ML "Context.Theory"} and @{ML "Context.Proof"}.
wenzelm@20449
   285
wenzelm@20449
   286
  \item @{ML Context.theory_of}~@{text "context"} always produces a
wenzelm@20449
   287
  theory from the generic @{text "context"}, using @{ML
wenzelm@20449
   288
  "ProofContext.theory_of"} as required.
wenzelm@20449
   289
wenzelm@20449
   290
  \item @{ML Context.proof_of}~@{text "context"} always produces a
wenzelm@20449
   291
  proof context from the generic @{text "context"}, using @{ML
wenzelm@20451
   292
  "ProofContext.init"} as required (note that this re-initializes the
wenzelm@20451
   293
  context data with each invocation).
wenzelm@20449
   294
wenzelm@20449
   295
  \end{description}
wenzelm@20449
   296
*}
wenzelm@20437
   297
wenzelm@20476
   298
wenzelm@20476
   299
subsection {* Context data \label{sec:context-data} *}
wenzelm@20447
   300
wenzelm@20447
   301
text {*
wenzelm@20451
   302
  The main purpose of theory and proof contexts is to manage arbitrary
wenzelm@20451
   303
  data.  New data types can be declared incrementally at compile time.
wenzelm@20451
   304
  There are separate declaration mechanisms for any of the three kinds
wenzelm@20451
   305
  of contexts: theory, proof, generic.
wenzelm@20449
   306
wenzelm@20449
   307
  \paragraph{Theory data} may refer to destructive entities, which are
wenzelm@20451
   308
  maintained in direct correspondence to the linear evolution of
wenzelm@20451
   309
  theory values, including explicit copies.\footnote{Most existing
wenzelm@20451
   310
  instances of destructive theory data are merely historical relics
wenzelm@20451
   311
  (e.g.\ the destructive theorem storage, and destructive hints for
wenzelm@20451
   312
  the Simplifier and Classical rules).}  A theory data declaration
wenzelm@22869
   313
  needs to implement the following SML signature:
wenzelm@20449
   314
wenzelm@20449
   315
  \medskip
wenzelm@20449
   316
  \begin{tabular}{ll}
wenzelm@22869
   317
  @{text "\<type> T"} & representing type \\
wenzelm@22869
   318
  @{text "\<val> empty: T"} & empty default value \\
wenzelm@22869
   319
  @{text "\<val> copy: T \<rightarrow> T"} & refresh impure data \\
wenzelm@22869
   320
  @{text "\<val> extend: T \<rightarrow> T"} & re-initialize on import \\
wenzelm@22869
   321
  @{text "\<val> merge: T \<times> T \<rightarrow> T"} & join on import \\
wenzelm@20449
   322
  \end{tabular}
wenzelm@20449
   323
  \medskip
wenzelm@20449
   324
wenzelm@22869
   325
  \noindent The @{text "empty"} value acts as initial default for
wenzelm@22869
   326
  \emph{any} theory that does not declare actual data content; @{text
wenzelm@22869
   327
  "copy"} maintains persistent integrity for impure data, it is just
wenzelm@22869
   328
  the identity for pure values; @{text "extend"} is acts like a
wenzelm@22869
   329
  unitary version of @{text "merge"}, both operations should also
wenzelm@22869
   330
  include the functionality of @{text "copy"} for impure data.
wenzelm@20449
   331
wenzelm@20451
   332
  \paragraph{Proof context data} is purely functional.  A declaration
wenzelm@22869
   333
  needs to implement the following SML signature:
wenzelm@20449
   334
wenzelm@20449
   335
  \medskip
wenzelm@20449
   336
  \begin{tabular}{ll}
wenzelm@22869
   337
  @{text "\<type> T"} & representing type \\
wenzelm@22869
   338
  @{text "\<val> init: theory \<rightarrow> T"} & produce initial value \\
wenzelm@20449
   339
  \end{tabular}
wenzelm@20449
   340
  \medskip
wenzelm@20449
   341
wenzelm@20449
   342
  \noindent The @{text "init"} operation is supposed to produce a pure
wenzelm@22869
   343
  value from the given background theory.
wenzelm@20449
   344
wenzelm@20451
   345
  \paragraph{Generic data} provides a hybrid interface for both theory
wenzelm@20451
   346
  and proof data.  The declaration is essentially the same as for
wenzelm@22869
   347
  (pure) theory data, without @{text "copy"}.  The @{text "init"}
wenzelm@22869
   348
  operation for proof contexts merely selects the current data value
wenzelm@22869
   349
  from the background theory.
wenzelm@20449
   350
wenzelm@22869
   351
  \bigskip A data declaration of type @{text "T"} results in the
wenzelm@22869
   352
  following interface:
wenzelm@20449
   353
wenzelm@20449
   354
  \medskip
wenzelm@20449
   355
  \begin{tabular}{ll}
wenzelm@30087
   356
  @{text "init: theory \<rightarrow> T"} \\
wenzelm@20449
   357
  @{text "get: context \<rightarrow> T"} \\
wenzelm@20449
   358
  @{text "put: T \<rightarrow> context \<rightarrow> context"} \\
wenzelm@20449
   359
  @{text "map: (T \<rightarrow> T) \<rightarrow> context \<rightarrow> context"} \\
wenzelm@20449
   360
  \end{tabular}
wenzelm@20449
   361
  \medskip
wenzelm@20449
   362
wenzelm@22869
   363
  \noindent Here @{text "init"} is only applicable to impure theory
wenzelm@22869
   364
  data to install a fresh copy persistently (destructive update on
wenzelm@22869
   365
  uninitialized has no permanent effect).  The other operations provide
wenzelm@22869
   366
  access for the particular kind of context (theory, proof, or generic
wenzelm@22869
   367
  context).  Note that this is a safe interface: there is no other way
wenzelm@22869
   368
  to access the corresponding data slot of a context.  By keeping
wenzelm@22869
   369
  these operations private, a component may maintain abstract values
wenzelm@22869
   370
  authentically, without other components interfering.
wenzelm@20447
   371
*}
wenzelm@20447
   372
wenzelm@20450
   373
text %mlref {*
wenzelm@20450
   374
  \begin{mldecls}
wenzelm@20450
   375
  @{index_ML_functor TheoryDataFun} \\
wenzelm@20450
   376
  @{index_ML_functor ProofDataFun} \\
wenzelm@20450
   377
  @{index_ML_functor GenericDataFun} \\
wenzelm@20450
   378
  \end{mldecls}
wenzelm@20450
   379
wenzelm@20450
   380
  \begin{description}
wenzelm@20450
   381
wenzelm@20450
   382
  \item @{ML_functor TheoryDataFun}@{text "(spec)"} declares data for
wenzelm@20450
   383
  type @{ML_type theory} according to the specification provided as
wenzelm@20451
   384
  argument structure.  The resulting structure provides data init and
wenzelm@20451
   385
  access operations as described above.
wenzelm@20450
   386
wenzelm@20470
   387
  \item @{ML_functor ProofDataFun}@{text "(spec)"} is analogous to
wenzelm@20470
   388
  @{ML_functor TheoryDataFun} for type @{ML_type Proof.context}.
wenzelm@20450
   389
wenzelm@20470
   390
  \item @{ML_functor GenericDataFun}@{text "(spec)"} is analogous to
wenzelm@20470
   391
  @{ML_functor TheoryDataFun} for type @{ML_type Context.generic}.
wenzelm@20450
   392
wenzelm@20450
   393
  \end{description}
wenzelm@20450
   394
*}
wenzelm@20450
   395
wenzelm@20447
   396
wenzelm@26872
   397
section {* Names \label{sec:names} *}
wenzelm@18537
   398
wenzelm@20476
   399
text {*
wenzelm@20476
   400
  In principle, a name is just a string, but there are various
wenzelm@20488
   401
  convention for encoding additional structure.  For example, ``@{text
wenzelm@20488
   402
  "Foo.bar.baz"}'' is considered as a qualified name consisting of
wenzelm@20488
   403
  three basic name components.  The individual constituents of a name
wenzelm@20488
   404
  may have further substructure, e.g.\ the string
wenzelm@20488
   405
  ``\verb,\,\verb,<alpha>,'' encodes as a single symbol.
wenzelm@20451
   406
*}
wenzelm@18537
   407
wenzelm@18537
   408
wenzelm@18537
   409
subsection {* Strings of symbols *}
wenzelm@18537
   410
wenzelm@20476
   411
text {*
wenzelm@20476
   412
  A \emph{symbol} constitutes the smallest textual unit in Isabelle
wenzelm@20488
   413
  --- raw characters are normally not encountered at all.  Isabelle
wenzelm@20488
   414
  strings consist of a sequence of symbols, represented as a packed
wenzelm@20488
   415
  string or a list of strings.  Each symbol is in itself a small
wenzelm@20488
   416
  string, which has either one of the following forms:
wenzelm@18537
   417
wenzelm@20451
   418
  \begin{enumerate}
wenzelm@18537
   419
wenzelm@20488
   420
  \item a single ASCII character ``@{text "c"}'', for example
wenzelm@20488
   421
  ``\verb,a,'',
wenzelm@18537
   422
wenzelm@20488
   423
  \item a regular symbol ``\verb,\,\verb,<,@{text "ident"}\verb,>,'',
wenzelm@20476
   424
  for example ``\verb,\,\verb,<alpha>,'',
wenzelm@18537
   425
wenzelm@20488
   426
  \item a control symbol ``\verb,\,\verb,<^,@{text "ident"}\verb,>,'',
wenzelm@20476
   427
  for example ``\verb,\,\verb,<^bold>,'',
wenzelm@18537
   428
wenzelm@20488
   429
  \item a raw symbol ``\verb,\,\verb,<^raw:,@{text text}\verb,>,''
wenzelm@20488
   430
  where @{text text} constists of printable characters excluding
wenzelm@20476
   431
  ``\verb,.,'' and ``\verb,>,'', for example
wenzelm@20476
   432
  ``\verb,\,\verb,<^raw:$\sum_{i = 1}^n$>,'',
wenzelm@18537
   433
wenzelm@20488
   434
  \item a numbered raw control symbol ``\verb,\,\verb,<^raw,@{text
wenzelm@20476
   435
  n}\verb,>, where @{text n} consists of digits, for example
wenzelm@20451
   436
  ``\verb,\,\verb,<^raw42>,''.
wenzelm@18537
   437
wenzelm@20451
   438
  \end{enumerate}
wenzelm@18537
   439
wenzelm@20476
   440
  \noindent The @{text "ident"} syntax for symbol names is @{text
wenzelm@20476
   441
  "letter (letter | digit)\<^sup>*"}, where @{text "letter =
wenzelm@20476
   442
  A..Za..z"} and @{text "digit = 0..9"}.  There are infinitely many
wenzelm@20476
   443
  regular symbols and control symbols, but a fixed collection of
wenzelm@20476
   444
  standard symbols is treated specifically.  For example,
wenzelm@20488
   445
  ``\verb,\,\verb,<alpha>,'' is classified as a letter, which means it
wenzelm@20488
   446
  may occur within regular Isabelle identifiers.
wenzelm@18537
   447
wenzelm@20488
   448
  Since the character set underlying Isabelle symbols is 7-bit ASCII
wenzelm@20488
   449
  and 8-bit characters are passed through transparently, Isabelle may
wenzelm@20488
   450
  also process Unicode/UCS data in UTF-8 encoding.  Unicode provides
wenzelm@20488
   451
  its own collection of mathematical symbols, but there is no built-in
wenzelm@20488
   452
  link to the standard collection of Isabelle.
wenzelm@20476
   453
wenzelm@20476
   454
  \medskip Output of Isabelle symbols depends on the print mode
wenzelm@30084
   455
  (\secref{print-mode}).  For example, the standard {\LaTeX} setup of
wenzelm@30084
   456
  the Isabelle document preparation system would present
wenzelm@20451
   457
  ``\verb,\,\verb,<alpha>,'' as @{text "\<alpha>"}, and
wenzelm@20451
   458
  ``\verb,\,\verb,<^bold>,\verb,\,\verb,<alpha>,'' as @{text
wenzelm@20451
   459
  "\<^bold>\<alpha>"}.
wenzelm@20451
   460
*}
wenzelm@18537
   461
wenzelm@18537
   462
text %mlref {*
wenzelm@18537
   463
  \begin{mldecls}
wenzelm@18537
   464
  @{index_ML_type "Symbol.symbol"} \\
wenzelm@18537
   465
  @{index_ML Symbol.explode: "string -> Symbol.symbol list"} \\
wenzelm@18537
   466
  @{index_ML Symbol.is_letter: "Symbol.symbol -> bool"} \\
wenzelm@18537
   467
  @{index_ML Symbol.is_digit: "Symbol.symbol -> bool"} \\
wenzelm@18537
   468
  @{index_ML Symbol.is_quasi: "Symbol.symbol -> bool"} \\
wenzelm@20547
   469
  @{index_ML Symbol.is_blank: "Symbol.symbol -> bool"} \\
wenzelm@20547
   470
  \end{mldecls}
wenzelm@20547
   471
  \begin{mldecls}
wenzelm@18537
   472
  @{index_ML_type "Symbol.sym"} \\
wenzelm@18537
   473
  @{index_ML Symbol.decode: "Symbol.symbol -> Symbol.sym"} \\
wenzelm@18537
   474
  \end{mldecls}
wenzelm@18537
   475
wenzelm@18537
   476
  \begin{description}
wenzelm@18537
   477
wenzelm@20488
   478
  \item @{ML_type "Symbol.symbol"} represents individual Isabelle
wenzelm@20488
   479
  symbols; this is an alias for @{ML_type "string"}.
wenzelm@18537
   480
wenzelm@20476
   481
  \item @{ML "Symbol.explode"}~@{text "str"} produces a symbol list
wenzelm@20488
   482
  from the packed form.  This function supercedes @{ML
wenzelm@20476
   483
  "String.explode"} for virtually all purposes of manipulating text in
wenzelm@20476
   484
  Isabelle!
wenzelm@18537
   485
wenzelm@18537
   486
  \item @{ML "Symbol.is_letter"}, @{ML "Symbol.is_digit"}, @{ML
wenzelm@20476
   487
  "Symbol.is_quasi"}, @{ML "Symbol.is_blank"} classify standard
wenzelm@20476
   488
  symbols according to fixed syntactic conventions of Isabelle, cf.\
wenzelm@20476
   489
  \cite{isabelle-isar-ref}.
wenzelm@18537
   490
wenzelm@18537
   491
  \item @{ML_type "Symbol.sym"} is a concrete datatype that represents
wenzelm@20488
   492
  the different kinds of symbols explicitly, with constructors @{ML
wenzelm@20488
   493
  "Symbol.Char"}, @{ML "Symbol.Sym"}, @{ML "Symbol.Ctrl"}, @{ML
wenzelm@20451
   494
  "Symbol.Raw"}.
wenzelm@18537
   495
wenzelm@18537
   496
  \item @{ML "Symbol.decode"} converts the string representation of a
wenzelm@20451
   497
  symbol into the datatype version.
wenzelm@18537
   498
wenzelm@18537
   499
  \end{description}
wenzelm@18537
   500
*}
wenzelm@18537
   501
wenzelm@18537
   502
wenzelm@20476
   503
subsection {* Basic names \label{sec:basic-names} *}
wenzelm@18537
   504
wenzelm@20476
   505
text {*
wenzelm@20476
   506
  A \emph{basic name} essentially consists of a single Isabelle
wenzelm@20476
   507
  identifier.  There are conventions to mark separate classes of basic
wenzelm@30087
   508
  names, by attaching a suffix of underscores: one underscore means
wenzelm@30087
   509
  \emph{internal name}, two underscores means \emph{Skolem name},
wenzelm@30087
   510
  three underscores means \emph{internal Skolem name}.
wenzelm@20470
   511
wenzelm@20476
   512
  For example, the basic name @{text "foo"} has the internal version
wenzelm@20476
   513
  @{text "foo_"}, with Skolem versions @{text "foo__"} and @{text
wenzelm@20476
   514
  "foo___"}, respectively.
wenzelm@20476
   515
wenzelm@20488
   516
  These special versions provide copies of the basic name space, apart
wenzelm@20488
   517
  from anything that normally appears in the user text.  For example,
wenzelm@20488
   518
  system generated variables in Isar proof contexts are usually marked
wenzelm@20488
   519
  as internal, which prevents mysterious name references like @{text
wenzelm@20488
   520
  "xaa"} to appear in the text.
wenzelm@20476
   521
wenzelm@20488
   522
  \medskip Manipulating binding scopes often requires on-the-fly
wenzelm@20488
   523
  renamings.  A \emph{name context} contains a collection of already
wenzelm@20488
   524
  used names.  The @{text "declare"} operation adds names to the
wenzelm@20488
   525
  context.
wenzelm@20476
   526
wenzelm@20488
   527
  The @{text "invents"} operation derives a number of fresh names from
wenzelm@20488
   528
  a given starting point.  For example, the first three names derived
wenzelm@20488
   529
  from @{text "a"} are @{text "a"}, @{text "b"}, @{text "c"}.
wenzelm@20476
   530
wenzelm@20476
   531
  The @{text "variants"} operation produces fresh names by
wenzelm@20488
   532
  incrementing tentative names as base-26 numbers (with digits @{text
wenzelm@20488
   533
  "a..z"}) until all clashes are resolved.  For example, name @{text
wenzelm@20488
   534
  "foo"} results in variants @{text "fooa"}, @{text "foob"}, @{text
wenzelm@20488
   535
  "fooc"}, \dots, @{text "fooaa"}, @{text "fooab"} etc.; each renaming
wenzelm@20488
   536
  step picks the next unused variant from this sequence.
wenzelm@20476
   537
*}
wenzelm@20476
   538
wenzelm@20476
   539
text %mlref {*
wenzelm@20476
   540
  \begin{mldecls}
wenzelm@20476
   541
  @{index_ML Name.internal: "string -> string"} \\
wenzelm@20547
   542
  @{index_ML Name.skolem: "string -> string"} \\
wenzelm@20547
   543
  \end{mldecls}
wenzelm@20547
   544
  \begin{mldecls}
wenzelm@20476
   545
  @{index_ML_type Name.context} \\
wenzelm@20476
   546
  @{index_ML Name.context: Name.context} \\
wenzelm@20476
   547
  @{index_ML Name.declare: "string -> Name.context -> Name.context"} \\
wenzelm@20476
   548
  @{index_ML Name.invents: "Name.context -> string -> int -> string list"} \\
wenzelm@20476
   549
  @{index_ML Name.variants: "string list -> Name.context -> string list * Name.context"} \\
wenzelm@20476
   550
  \end{mldecls}
wenzelm@20476
   551
wenzelm@20476
   552
  \begin{description}
wenzelm@20476
   553
wenzelm@20476
   554
  \item @{ML Name.internal}~@{text "name"} produces an internal name
wenzelm@20476
   555
  by adding one underscore.
wenzelm@20476
   556
wenzelm@20476
   557
  \item @{ML Name.skolem}~@{text "name"} produces a Skolem name by
wenzelm@20476
   558
  adding two underscores.
wenzelm@20476
   559
wenzelm@20476
   560
  \item @{ML_type Name.context} represents the context of already used
wenzelm@20476
   561
  names; the initial value is @{ML "Name.context"}.
wenzelm@20476
   562
wenzelm@20488
   563
  \item @{ML Name.declare}~@{text "name"} enters a used name into the
wenzelm@20488
   564
  context.
wenzelm@20476
   565
wenzelm@20488
   566
  \item @{ML Name.invents}~@{text "context name n"} produces @{text
wenzelm@20488
   567
  "n"} fresh names derived from @{text "name"}.
wenzelm@20488
   568
wenzelm@20488
   569
  \item @{ML Name.variants}~@{text "names context"} produces fresh
wenzelm@30087
   570
  variants of @{text "names"}; the result is entered into the context.
wenzelm@20476
   571
wenzelm@20476
   572
  \end{description}
wenzelm@20476
   573
*}
wenzelm@20476
   574
wenzelm@20476
   575
wenzelm@20476
   576
subsection {* Indexed names *}
wenzelm@20476
   577
wenzelm@20476
   578
text {*
wenzelm@20476
   579
  An \emph{indexed name} (or @{text "indexname"}) is a pair of a basic
wenzelm@20488
   580
  name and a natural number.  This representation allows efficient
wenzelm@20488
   581
  renaming by incrementing the second component only.  The canonical
wenzelm@20488
   582
  way to rename two collections of indexnames apart from each other is
wenzelm@20488
   583
  this: determine the maximum index @{text "maxidx"} of the first
wenzelm@20488
   584
  collection, then increment all indexes of the second collection by
wenzelm@20488
   585
  @{text "maxidx + 1"}; the maximum index of an empty collection is
wenzelm@20488
   586
  @{text "-1"}.
wenzelm@20476
   587
wenzelm@20488
   588
  Occasionally, basic names and indexed names are injected into the
wenzelm@20488
   589
  same pair type: the (improper) indexname @{text "(x, -1)"} is used
wenzelm@20488
   590
  to encode basic names.
wenzelm@20488
   591
wenzelm@20488
   592
  \medskip Isabelle syntax observes the following rules for
wenzelm@20488
   593
  representing an indexname @{text "(x, i)"} as a packed string:
wenzelm@20476
   594
wenzelm@20476
   595
  \begin{itemize}
wenzelm@20476
   596
wenzelm@20479
   597
  \item @{text "?x"} if @{text "x"} does not end with a digit and @{text "i = 0"},
wenzelm@20476
   598
wenzelm@20476
   599
  \item @{text "?xi"} if @{text "x"} does not end with a digit,
wenzelm@20476
   600
wenzelm@20488
   601
  \item @{text "?x.i"} otherwise.
wenzelm@20476
   602
wenzelm@20476
   603
  \end{itemize}
wenzelm@20476
   604
wenzelm@20488
   605
  Indexnames may acquire large index numbers over time.  Results are
wenzelm@20488
   606
  normalized towards @{text "0"} at certain checkpoints, notably at
wenzelm@20488
   607
  the end of a proof.  This works by producing variants of the
wenzelm@20488
   608
  corresponding basic name components.  For example, the collection
wenzelm@20488
   609
  @{text "?x1, ?x7, ?x42"} becomes @{text "?x, ?xa, ?xb"}.
wenzelm@20476
   610
*}
wenzelm@20476
   611
wenzelm@20476
   612
text %mlref {*
wenzelm@20476
   613
  \begin{mldecls}
wenzelm@20476
   614
  @{index_ML_type indexname} \\
wenzelm@20476
   615
  \end{mldecls}
wenzelm@20476
   616
wenzelm@20476
   617
  \begin{description}
wenzelm@20476
   618
wenzelm@20476
   619
  \item @{ML_type indexname} represents indexed names.  This is an
wenzelm@20476
   620
  abbreviation for @{ML_type "string * int"}.  The second component is
wenzelm@20476
   621
  usually non-negative, except for situations where @{text "(x, -1)"}
wenzelm@20488
   622
  is used to embed basic names into this type.
wenzelm@20476
   623
wenzelm@20476
   624
  \end{description}
wenzelm@20476
   625
*}
wenzelm@20476
   626
wenzelm@20476
   627
wenzelm@20476
   628
subsection {* Qualified names and name spaces *}
wenzelm@20476
   629
wenzelm@20476
   630
text {*
wenzelm@20476
   631
  A \emph{qualified name} consists of a non-empty sequence of basic
wenzelm@20488
   632
  name components.  The packed representation uses a dot as separator,
wenzelm@20488
   633
  as in ``@{text "A.b.c"}''.  The last component is called \emph{base}
wenzelm@20488
   634
  name, the remaining prefix \emph{qualifier} (which may be empty).
wenzelm@20488
   635
  The idea of qualified names is to encode nested structures by
wenzelm@20488
   636
  recording the access paths as qualifiers.  For example, an item
wenzelm@20488
   637
  named ``@{text "A.b.c"}'' may be understood as a local entity @{text
wenzelm@20488
   638
  "c"}, within a local structure @{text "b"}, within a global
wenzelm@20488
   639
  structure @{text "A"}.  Typically, name space hierarchies consist of
wenzelm@20488
   640
  1--2 levels of qualification, but this need not be always so.
wenzelm@20451
   641
wenzelm@20476
   642
  The empty name is commonly used as an indication of unnamed
wenzelm@20488
   643
  entities, whenever this makes any sense.  The basic operations on
wenzelm@20488
   644
  qualified names are smart enough to pass through such improper names
wenzelm@20476
   645
  unchanged.
wenzelm@20451
   646
wenzelm@20476
   647
  \medskip A @{text "naming"} policy tells how to turn a name
wenzelm@20476
   648
  specification into a fully qualified internal name (by the @{text
wenzelm@20488
   649
  "full"} operation), and how fully qualified names may be accessed
wenzelm@20488
   650
  externally.  For example, the default naming policy is to prefix an
wenzelm@20488
   651
  implicit path: @{text "full x"} produces @{text "path.x"}, and the
wenzelm@20488
   652
  standard accesses for @{text "path.x"} include both @{text "x"} and
wenzelm@20488
   653
  @{text "path.x"}.  Normally, the naming is implicit in the theory or
wenzelm@20488
   654
  proof context; there are separate versions of the corresponding.
wenzelm@20476
   655
wenzelm@20476
   656
  \medskip A @{text "name space"} manages a collection of fully
wenzelm@20476
   657
  internalized names, together with a mapping between external names
wenzelm@20476
   658
  and internal names (in both directions).  The corresponding @{text
wenzelm@20476
   659
  "intern"} and @{text "extern"} operations are mostly used for
wenzelm@20476
   660
  parsing and printing only!  The @{text "declare"} operation augments
wenzelm@20488
   661
  a name space according to the accesses determined by the naming
wenzelm@20488
   662
  policy.
wenzelm@20476
   663
wenzelm@20488
   664
  \medskip As a general principle, there is a separate name space for
wenzelm@20488
   665
  each kind of formal entity, e.g.\ logical constant, type
wenzelm@20488
   666
  constructor, type class, theorem.  It is usually clear from the
wenzelm@20488
   667
  occurrence in concrete syntax (or from the scope) which kind of
wenzelm@20488
   668
  entity a name refers to.  For example, the very same name @{text
wenzelm@20488
   669
  "c"} may be used uniformly for a constant, type constructor, and
wenzelm@20488
   670
  type class.
wenzelm@20476
   671
wenzelm@20479
   672
  There are common schemes to name theorems systematically, according
wenzelm@20488
   673
  to the name of the main logical entity involved, e.g.\ @{text
wenzelm@20488
   674
  "c.intro"} for a canonical theorem related to constant @{text "c"}.
wenzelm@20488
   675
  This technique of mapping names from one space into another requires
wenzelm@20488
   676
  some care in order to avoid conflicts.  In particular, theorem names
wenzelm@20488
   677
  derived from a type constructor or type class are better suffixed in
wenzelm@20488
   678
  addition to the usual qualification, e.g.\ @{text "c_type.intro"}
wenzelm@20488
   679
  and @{text "c_class.intro"} for theorems related to type @{text "c"}
wenzelm@20488
   680
  and class @{text "c"}, respectively.
wenzelm@20451
   681
*}
wenzelm@18537
   682
wenzelm@20476
   683
text %mlref {*
wenzelm@20476
   684
  \begin{mldecls}
wenzelm@30365
   685
  @{index_ML Long_Name.base_name: "string -> string"} \\
wenzelm@30365
   686
  @{index_ML Long_Name.qualifier: "string -> string"} \\
wenzelm@30365
   687
  @{index_ML Long_Name.append: "string -> string -> string"} \\
wenzelm@30365
   688
  @{index_ML Long_Name.implode: "string list -> string"} \\
wenzelm@30365
   689
  @{index_ML Long_Name.explode: "string -> string list"} \\
wenzelm@20547
   690
  \end{mldecls}
wenzelm@20547
   691
  \begin{mldecls}
wenzelm@20476
   692
  @{index_ML_type NameSpace.naming} \\
wenzelm@20476
   693
  @{index_ML NameSpace.default_naming: NameSpace.naming} \\
wenzelm@20476
   694
  @{index_ML NameSpace.add_path: "string -> NameSpace.naming -> NameSpace.naming"} \\
haftmann@29581
   695
  @{index_ML NameSpace.full_name: "NameSpace.naming -> binding -> string"} \\
wenzelm@20547
   696
  \end{mldecls}
wenzelm@20547
   697
  \begin{mldecls}
wenzelm@20476
   698
  @{index_ML_type NameSpace.T} \\
wenzelm@20476
   699
  @{index_ML NameSpace.empty: NameSpace.T} \\
wenzelm@20476
   700
  @{index_ML NameSpace.merge: "NameSpace.T * NameSpace.T -> NameSpace.T"} \\
wenzelm@30281
   701
  @{index_ML NameSpace.declare: "NameSpace.naming -> binding -> NameSpace.T ->
wenzelm@30281
   702
  string * NameSpace.T"} \\
wenzelm@20476
   703
  @{index_ML NameSpace.intern: "NameSpace.T -> string -> string"} \\
wenzelm@20476
   704
  @{index_ML NameSpace.extern: "NameSpace.T -> string -> string"} \\
wenzelm@20476
   705
  \end{mldecls}
wenzelm@18537
   706
wenzelm@20476
   707
  \begin{description}
wenzelm@18537
   708
wenzelm@30365
   709
  \item @{ML Long_Name.base_name}~@{text "name"} returns the base name of a
wenzelm@20476
   710
  qualified name.
wenzelm@18537
   711
wenzelm@30365
   712
  \item @{ML Long_Name.qualifier}~@{text "name"} returns the qualifier
wenzelm@20476
   713
  of a qualified name.
wenzelm@18537
   714
wenzelm@30365
   715
  \item @{ML Long_Name.append}~@{text "name\<^isub>1 name\<^isub>2"}
wenzelm@20476
   716
  appends two qualified names.
wenzelm@18537
   717
wenzelm@30365
   718
  \item @{ML Long_Name.implode}~@{text "names"} and @{ML
wenzelm@30365
   719
  Long_Name.explode}~@{text "name"} convert between the packed string
wenzelm@20488
   720
  representation and the explicit list form of qualified names.
wenzelm@18537
   721
wenzelm@20476
   722
  \item @{ML_type NameSpace.naming} represents the abstract concept of
wenzelm@20476
   723
  a naming policy.
wenzelm@18537
   724
wenzelm@20476
   725
  \item @{ML NameSpace.default_naming} is the default naming policy.
wenzelm@20476
   726
  In a theory context, this is usually augmented by a path prefix
wenzelm@20476
   727
  consisting of the theory name.
wenzelm@18537
   728
wenzelm@20476
   729
  \item @{ML NameSpace.add_path}~@{text "path naming"} augments the
wenzelm@20488
   730
  naming policy by extending its path component.
wenzelm@18537
   731
wenzelm@30281
   732
  \item @{ML NameSpace.full_name}~@{text "naming binding"} turns a
wenzelm@30281
   733
  name binding (usually a basic name) into the fully qualified
haftmann@29001
   734
  internal name, according to the given naming policy.
wenzelm@20476
   735
wenzelm@20476
   736
  \item @{ML_type NameSpace.T} represents name spaces.
wenzelm@20476
   737
wenzelm@20476
   738
  \item @{ML NameSpace.empty} and @{ML NameSpace.merge}~@{text
wenzelm@20488
   739
  "(space\<^isub>1, space\<^isub>2)"} are the canonical operations for
wenzelm@20488
   740
  maintaining name spaces according to theory data management
wenzelm@20488
   741
  (\secref{sec:context-data}).
wenzelm@20476
   742
haftmann@29001
   743
  \item @{ML NameSpace.declare}~@{text "naming bindings space"} enters a
haftmann@29001
   744
  name binding as fully qualified internal name into the name space,
haftmann@29001
   745
  with external accesses determined by the naming policy.
wenzelm@20476
   746
wenzelm@20476
   747
  \item @{ML NameSpace.intern}~@{text "space name"} internalizes a
wenzelm@20476
   748
  (partially qualified) external name.
wenzelm@20476
   749
wenzelm@20488
   750
  This operation is mostly for parsing!  Note that fully qualified
wenzelm@20476
   751
  names stemming from declarations are produced via @{ML
haftmann@29001
   752
  "NameSpace.full_name"} and @{ML "NameSpace.declare"}
haftmann@29001
   753
  (or their derivatives for @{ML_type theory} and
wenzelm@20488
   754
  @{ML_type Proof.context}).
wenzelm@20476
   755
wenzelm@20476
   756
  \item @{ML NameSpace.extern}~@{text "space name"} externalizes a
wenzelm@20476
   757
  (fully qualified) internal name.
wenzelm@20476
   758
wenzelm@30281
   759
  This operation is mostly for printing!  User code should not rely on
wenzelm@30281
   760
  the precise result too much.
wenzelm@20476
   761
wenzelm@20476
   762
  \end{description}
wenzelm@20476
   763
*}
wenzelm@30273
   764
wenzelm@18537
   765
end