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