| author | wenzelm | 
| Thu, 07 Oct 2010 19:05:42 +0100 | |
| changeset 39822 | 0de42180febe | 
| parent 39821 | bf164c153d10 | 
| child 39825 | f9066b94bf07 | 
| permissions | -rw-r--r-- | 
| 29755 | 1  | 
theory Prelim  | 
2  | 
imports Base  | 
|
3  | 
begin  | 
|
| 18537 | 4  | 
|
5  | 
chapter {* Preliminaries *}
 | 
|
6  | 
||
| 20429 | 7  | 
section {* Contexts \label{sec:context} *}
 | 
| 18537 | 8  | 
|
| 20429 | 9  | 
text {*
 | 
| 20451 | 10  | 
A logical context represents the background that is required for  | 
11  | 
formulating statements and composing proofs. It acts as a medium to  | 
|
12  | 
produce formal content, depending on earlier material (declarations,  | 
|
13  | 
results etc.).  | 
|
| 18537 | 14  | 
|
| 20451 | 15  | 
For example, derivations within the Isabelle/Pure logic can be  | 
16  | 
  described as a judgment @{text "\<Gamma> \<turnstile>\<^sub>\<Theta> \<phi>"}, which means that a
 | 
|
| 20429 | 17  | 
  proposition @{text "\<phi>"} is derivable from hypotheses @{text "\<Gamma>"}
 | 
18  | 
  within the theory @{text "\<Theta>"}.  There are logical reasons for
 | 
|
| 20451 | 19  | 
  keeping @{text "\<Theta>"} and @{text "\<Gamma>"} separate: theories can be
 | 
20  | 
liberal about supporting type constructors and schematic  | 
|
21  | 
polymorphism of constants and axioms, while the inner calculus of  | 
|
22  | 
  @{text "\<Gamma> \<turnstile> \<phi>"} is strictly limited to Simple Type Theory (with
 | 
|
23  | 
fixed type variables in the assumptions).  | 
|
| 18537 | 24  | 
|
| 20429 | 25  | 
\medskip Contexts and derivations are linked by the following key  | 
26  | 
principles:  | 
|
27  | 
||
28  | 
  \begin{itemize}
 | 
|
29  | 
||
30  | 
\item Transfer: monotonicity of derivations admits results to be  | 
|
| 20451 | 31  | 
  transferred into a \emph{larger} context, i.e.\ @{text "\<Gamma> \<turnstile>\<^sub>\<Theta>
 | 
32  | 
  \<phi>"} implies @{text "\<Gamma>' \<turnstile>\<^sub>\<Theta>\<^sub>' \<phi>"} for contexts @{text "\<Theta>'
 | 
|
33  | 
  \<supseteq> \<Theta>"} and @{text "\<Gamma>' \<supseteq> \<Gamma>"}.
 | 
|
| 18537 | 34  | 
|
| 20429 | 35  | 
\item Export: discharge of hypotheses admits results to be exported  | 
| 20451 | 36  | 
  into a \emph{smaller} context, i.e.\ @{text "\<Gamma>' \<turnstile>\<^sub>\<Theta> \<phi>"}
 | 
37  | 
  implies @{text "\<Gamma> \<turnstile>\<^sub>\<Theta> \<Delta> \<Longrightarrow> \<phi>"} where @{text "\<Gamma>' \<supseteq> \<Gamma>"} and
 | 
|
38  | 
  @{text "\<Delta> = \<Gamma>' - \<Gamma>"}.  Note that @{text "\<Theta>"} remains unchanged here,
 | 
|
39  | 
  only the @{text "\<Gamma>"} part is affected.
 | 
|
| 20429 | 40  | 
|
41  | 
  \end{itemize}
 | 
|
| 18537 | 42  | 
|
| 20451 | 43  | 
\medskip By modeling the main characteristics of the primitive  | 
44  | 
  @{text "\<Theta>"} and @{text "\<Gamma>"} above, and abstracting over any
 | 
|
45  | 
particular logical content, we arrive at the fundamental notions of  | 
|
46  | 
  \emph{theory context} and \emph{proof context} in Isabelle/Isar.
 | 
|
47  | 
  These implement a certain policy to manage arbitrary \emph{context
 | 
|
48  | 
data}. There is a strongly-typed mechanism to declare new kinds of  | 
|
| 20429 | 49  | 
data at compile time.  | 
| 18537 | 50  | 
|
| 20451 | 51  | 
The internal bootstrap process of Isabelle/Pure eventually reaches a  | 
52  | 
  stage where certain data slots provide the logical content of @{text
 | 
|
53  | 
  "\<Theta>"} and @{text "\<Gamma>"} sketched above, but this does not stop there!
 | 
|
54  | 
Various additional data slots support all kinds of mechanisms that  | 
|
55  | 
are not necessarily part of the core logic.  | 
|
| 18537 | 56  | 
|
| 20429 | 57  | 
For example, there would be data for canonical introduction and  | 
58  | 
elimination rules for arbitrary operators (depending on the  | 
|
59  | 
object-logic and application), which enables users to perform  | 
|
| 20451 | 60  | 
  standard proof steps implicitly (cf.\ the @{text "rule"} method
 | 
61  | 
  \cite{isabelle-isar-ref}).
 | 
|
| 18537 | 62  | 
|
| 20451 | 63  | 
\medskip Thus Isabelle/Isar is able to bring forth more and more  | 
64  | 
concepts successively. In particular, an object-logic like  | 
|
65  | 
Isabelle/HOL continues the Isabelle/Pure setup by adding specific  | 
|
66  | 
components for automated reasoning (classical reasoner, tableau  | 
|
67  | 
prover, structured induction etc.) and derived specification  | 
|
68  | 
mechanisms (inductive predicates, recursive functions etc.). All of  | 
|
69  | 
this is ultimately based on the generic data management by theory  | 
|
70  | 
and proof contexts introduced here.  | 
|
| 18537 | 71  | 
*}  | 
72  | 
||
73  | 
||
74  | 
subsection {* Theory context \label{sec:context-theory} *}
 | 
|
75  | 
||
| 34921 | 76  | 
text {* A \emph{theory} is a data container with explicit name and
 | 
77  | 
unique identifier. Theories are related by a (nominal) sub-theory  | 
|
| 20451 | 78  | 
relation, which corresponds to the dependency graph of the original  | 
79  | 
construction; each theory is derived from a certain sub-graph of  | 
|
| 34921 | 80  | 
ancestor theories. To this end, the system maintains a set of  | 
81  | 
symbolic ``identification stamps'' within each theory.  | 
|
| 18537 | 82  | 
|
| 34921 | 83  | 
In order to avoid the full-scale overhead of explicit sub-theory  | 
84  | 
identification of arbitrary intermediate stages, a theory is  | 
|
85  | 
  switched into @{text "draft"} mode under certain circumstances.  A
 | 
|
86  | 
draft theory acts like a linear type, where updates invalidate  | 
|
87  | 
  earlier versions.  An invalidated draft is called \emph{stale}.
 | 
|
| 20429 | 88  | 
|
| 34921 | 89  | 
  The @{text "checkpoint"} operation produces a safe stepping stone
 | 
90  | 
that will survive the next update without becoming stale: both the  | 
|
91  | 
old and the new theory remain valid and are related by the  | 
|
92  | 
sub-theory relation. Checkpointing essentially recovers purely  | 
|
93  | 
functional theory values, at the expense of some extra internal  | 
|
94  | 
bookkeeping.  | 
|
| 20447 | 95  | 
|
96  | 
  The @{text "copy"} operation produces an auxiliary version that has
 | 
|
97  | 
the same data content, but is unrelated to the original: updates of  | 
|
98  | 
the copy do not affect the original, neither does the sub-theory  | 
|
99  | 
relation hold.  | 
|
| 20429 | 100  | 
|
| 34921 | 101  | 
  The @{text "merge"} operation produces the least upper bound of two
 | 
102  | 
theories, which actually degenerates into absorption of one theory  | 
|
103  | 
into the other (according to the nominal sub-theory relation).  | 
|
104  | 
||
105  | 
  The @{text "begin"} operation starts a new theory by importing
 | 
|
106  | 
several parent theories and entering a special mode of nameless  | 
|
107  | 
  incremental updates, until the final @{text "end"} operation is
 | 
|
108  | 
performed.  | 
|
109  | 
||
| 20447 | 110  | 
  \medskip The example in \figref{fig:ex-theory} below shows a theory
 | 
| 20451 | 111  | 
  graph derived from @{text "Pure"}, with theory @{text "Length"}
 | 
112  | 
  importing @{text "Nat"} and @{text "List"}.  The body of @{text
 | 
|
113  | 
"Length"} consists of a sequence of updates, working mostly on  | 
|
| 34921 | 114  | 
drafts internally, while transaction boundaries of Isar top-level  | 
115  | 
  commands (\secref{sec:isar-toplevel}) are guaranteed to be safe
 | 
|
116  | 
checkpoints.  | 
|
| 20447 | 117  | 
|
118  | 
  \begin{figure}[htb]
 | 
|
119  | 
  \begin{center}
 | 
|
| 20429 | 120  | 
  \begin{tabular}{rcccl}
 | 
| 20447 | 121  | 
        &            & @{text "Pure"} \\
 | 
122  | 
        &            & @{text "\<down>"} \\
 | 
|
123  | 
        &            & @{text "FOL"} \\
 | 
|
| 18537 | 124  | 
& $\swarrow$ & & $\searrow$ & \\  | 
| 21852 | 125  | 
  @{text "Nat"} &    &              &            & @{text "List"} \\
 | 
| 18537 | 126  | 
& $\searrow$ & & $\swarrow$ \\  | 
| 20447 | 127  | 
        &            & @{text "Length"} \\
 | 
| 26864 | 128  | 
        &            & \multicolumn{3}{l}{~~@{keyword "imports"}} \\
 | 
129  | 
        &            & \multicolumn{3}{l}{~~@{keyword "begin"}} \\
 | 
|
| 18537 | 130  | 
& & $\vdots$~~ \\  | 
| 20447 | 131  | 
        &            & @{text "\<bullet>"}~~ \\
 | 
132  | 
& & $\vdots$~~ \\  | 
|
133  | 
        &            & @{text "\<bullet>"}~~ \\
 | 
|
134  | 
& & $\vdots$~~ \\  | 
|
| 26864 | 135  | 
        &            & \multicolumn{3}{l}{~~@{command "end"}} \\
 | 
| 20429 | 136  | 
  \end{tabular}
 | 
| 20451 | 137  | 
  \caption{A theory definition depending on ancestors}\label{fig:ex-theory}
 | 
| 20447 | 138  | 
  \end{center}
 | 
139  | 
  \end{figure}
 | 
|
| 20451 | 140  | 
|
141  | 
  \medskip There is a separate notion of \emph{theory reference} for
 | 
|
142  | 
maintaining a live link to an evolving theory context: updates on  | 
|
| 39821 | 143  | 
drafts are propagated automatically. Dynamic updating stops when  | 
144  | 
  the next @{text "checkpoint"} is reached.
 | 
|
| 20451 | 145  | 
|
146  | 
Derived entities may store a theory reference in order to indicate  | 
|
| 39821 | 147  | 
the formal context from which they are derived. This implicitly  | 
148  | 
assumes monotonic reasoning, because the referenced context may  | 
|
149  | 
become larger without further notice.  | 
|
| 18537 | 150  | 
*}  | 
151  | 
||
| 20430 | 152  | 
text %mlref {*
 | 
| 20447 | 153  | 
  \begin{mldecls}
 | 
154  | 
  @{index_ML_type theory} \\
 | 
|
155  | 
  @{index_ML Theory.subthy: "theory * theory -> bool"} \\
 | 
|
156  | 
  @{index_ML Theory.checkpoint: "theory -> theory"} \\
 | 
|
| 20547 | 157  | 
  @{index_ML Theory.copy: "theory -> theory"} \\
 | 
| 34921 | 158  | 
  @{index_ML Theory.merge: "theory * theory -> theory"} \\
 | 
159  | 
  @{index_ML Theory.begin_theory: "string -> theory list -> theory"} \\
 | 
|
| 20547 | 160  | 
  \end{mldecls}
 | 
161  | 
  \begin{mldecls}
 | 
|
| 20447 | 162  | 
  @{index_ML_type theory_ref} \\
 | 
163  | 
  @{index_ML Theory.deref: "theory_ref -> theory"} \\
 | 
|
| 
24137
 
8d7896398147
replaced Theory.self_ref by Theory.check_thy, which now produces a checked ref;
 
wenzelm 
parents: 
22869 
diff
changeset
 | 
164  | 
  @{index_ML Theory.check_thy: "theory -> theory_ref"} \\
 | 
| 20447 | 165  | 
  \end{mldecls}
 | 
166  | 
||
167  | 
  \begin{description}
 | 
|
168  | 
||
| 20451 | 169  | 
  \item @{ML_type theory} represents theory contexts.  This is
 | 
| 39821 | 170  | 
essentially a linear type, with explicit runtime checking.  | 
171  | 
Primitive theory operations destroy the original version, which then  | 
|
172  | 
becomes ``stale''. This can be prevented by explicit checkpointing,  | 
|
173  | 
which the system does at least at the boundary of toplevel command  | 
|
174  | 
  transactions \secref{sec:isar-toplevel}.
 | 
|
| 20447 | 175  | 
|
| 34921 | 176  | 
  \item @{ML "Theory.subthy"}~@{text "(thy\<^sub>1, thy\<^sub>2)"} compares theories
 | 
177  | 
according to the intrinsic graph structure of the construction.  | 
|
178  | 
This sub-theory relation is a nominal approximation of inclusion  | 
|
179  | 
  (@{text "\<subseteq>"}) of the corresponding content (according to the
 | 
|
180  | 
semantics of the ML modules that implement the data).  | 
|
| 20447 | 181  | 
|
182  | 
  \item @{ML "Theory.checkpoint"}~@{text "thy"} produces a safe
 | 
|
| 34921 | 183  | 
  stepping stone in the linear development of @{text "thy"}.  This
 | 
184  | 
changes the old theory, but the next update will result in two  | 
|
185  | 
related, valid theories.  | 
|
| 20447 | 186  | 
|
187  | 
  \item @{ML "Theory.copy"}~@{text "thy"} produces a variant of @{text
 | 
|
| 34921 | 188  | 
"thy"} with the same data. The copy is not related to the original,  | 
189  | 
but the original is unchanged.  | 
|
190  | 
||
191  | 
  \item @{ML "Theory.merge"}~@{text "(thy\<^sub>1, thy\<^sub>2)"} absorbs one theory
 | 
|
192  | 
  into the other, without changing @{text "thy\<^sub>1"} or @{text "thy\<^sub>2"}.
 | 
|
193  | 
This version of ad-hoc theory merge fails for unrelated theories!  | 
|
194  | 
||
195  | 
  \item @{ML "Theory.begin_theory"}~@{text "name parents"} constructs
 | 
|
196  | 
  a new theory based on the given parents.  This {\ML} function is
 | 
|
197  | 
normally not invoked directly.  | 
|
| 20447 | 198  | 
|
| 20451 | 199  | 
  \item @{ML_type theory_ref} represents a sliding reference to an
 | 
200  | 
always valid theory; updates on the original are propagated  | 
|
| 20447 | 201  | 
automatically.  | 
202  | 
||
| 
24137
 
8d7896398147
replaced Theory.self_ref by Theory.check_thy, which now produces a checked ref;
 
wenzelm 
parents: 
22869 
diff
changeset
 | 
203  | 
  \item @{ML "Theory.deref"}~@{text "thy_ref"} turns a @{ML_type
 | 
| 
 
8d7896398147
replaced Theory.self_ref by Theory.check_thy, which now produces a checked ref;
 
wenzelm 
parents: 
22869 
diff
changeset
 | 
204  | 
  "theory_ref"} into an @{ML_type "theory"} value.  As the referenced
 | 
| 
 
8d7896398147
replaced Theory.self_ref by Theory.check_thy, which now produces a checked ref;
 
wenzelm 
parents: 
22869 
diff
changeset
 | 
205  | 
  theory evolves monotonically over time, later invocations of @{ML
 | 
| 20451 | 206  | 
"Theory.deref"} may refer to a larger context.  | 
| 20447 | 207  | 
|
| 
24137
 
8d7896398147
replaced Theory.self_ref by Theory.check_thy, which now produces a checked ref;
 
wenzelm 
parents: 
22869 
diff
changeset
 | 
208  | 
  \item @{ML "Theory.check_thy"}~@{text "thy"} produces a @{ML_type
 | 
| 
 
8d7896398147
replaced Theory.self_ref by Theory.check_thy, which now produces a checked ref;
 
wenzelm 
parents: 
22869 
diff
changeset
 | 
209  | 
  "theory_ref"} from a valid @{ML_type "theory"} value.
 | 
| 
 
8d7896398147
replaced Theory.self_ref by Theory.check_thy, which now produces a checked ref;
 
wenzelm 
parents: 
22869 
diff
changeset
 | 
210  | 
|
| 20447 | 211  | 
  \end{description}
 | 
| 20430 | 212  | 
*}  | 
213  | 
||
| 18537 | 214  | 
|
215  | 
subsection {* Proof context \label{sec:context-proof} *}
 | 
|
216  | 
||
| 34921 | 217  | 
text {* A proof context is a container for pure data with a
 | 
| 39821 | 218  | 
  back-reference to the theory from which it is derived.  The @{text
 | 
219  | 
"init"} operation creates a proof context from a given theory.  | 
|
| 34921 | 220  | 
Modifications to draft theories are propagated to the proof context  | 
221  | 
  as usual, but there is also an explicit @{text "transfer"} operation
 | 
|
222  | 
to force resynchronization with more substantial updates to the  | 
|
223  | 
underlying theory.  | 
|
| 20429 | 224  | 
|
| 34921 | 225  | 
Entities derived in a proof context need to record logical  | 
| 20447 | 226  | 
requirements explicitly, since there is no separate context  | 
| 34921 | 227  | 
identification or symbolic inclusion as for theories. For example,  | 
228  | 
  hypotheses used in primitive derivations (cf.\ \secref{sec:thms})
 | 
|
229  | 
  are recorded separately within the sequent @{text "\<Gamma> \<turnstile> \<phi>"}, just to
 | 
|
230  | 
make double sure. Results could still leak into an alien proof  | 
|
231  | 
context due to programming errors, but Isabelle/Isar includes some  | 
|
232  | 
extra validity checks in critical positions, notably at the end of a  | 
|
233  | 
sub-proof.  | 
|
| 20429 | 234  | 
|
| 20451 | 235  | 
Proof contexts may be manipulated arbitrarily, although the common  | 
236  | 
discipline is to follow block structure as a mental model: a given  | 
|
237  | 
context is extended consecutively, and results are exported back  | 
|
| 34921 | 238  | 
into the original context. Note that an Isar proof state models  | 
| 20451 | 239  | 
block-structured reasoning explicitly, using a stack of proof  | 
| 34921 | 240  | 
contexts internally. For various technical reasons, the background  | 
241  | 
theory of an Isar proof state must not be changed while the proof is  | 
|
242  | 
still under construction!  | 
|
| 18537 | 243  | 
*}  | 
244  | 
||
| 20449 | 245  | 
text %mlref {*
 | 
246  | 
  \begin{mldecls}
 | 
|
247  | 
  @{index_ML_type Proof.context} \\
 | 
|
| 36611 | 248  | 
  @{index_ML ProofContext.init_global: "theory -> Proof.context"} \\
 | 
| 20449 | 249  | 
  @{index_ML ProofContext.theory_of: "Proof.context -> theory"} \\
 | 
250  | 
  @{index_ML ProofContext.transfer: "theory -> Proof.context -> Proof.context"} \\
 | 
|
251  | 
  \end{mldecls}
 | 
|
252  | 
||
253  | 
  \begin{description}
 | 
|
254  | 
||
255  | 
  \item @{ML_type Proof.context} represents proof contexts.  Elements
 | 
|
256  | 
of this type are essentially pure values, with a sliding reference  | 
|
257  | 
to the background theory.  | 
|
258  | 
||
| 36611 | 259  | 
  \item @{ML ProofContext.init_global}~@{text "thy"} produces a proof context
 | 
| 20449 | 260  | 
  derived from @{text "thy"}, initializing all data.
 | 
261  | 
||
262  | 
  \item @{ML ProofContext.theory_of}~@{text "ctxt"} selects the
 | 
|
| 20451 | 263  | 
  background theory from @{text "ctxt"}, dereferencing its internal
 | 
264  | 
  @{ML_type theory_ref}.
 | 
|
| 20449 | 265  | 
|
266  | 
  \item @{ML ProofContext.transfer}~@{text "thy ctxt"} promotes the
 | 
|
267  | 
  background theory of @{text "ctxt"} to the super theory @{text
 | 
|
268  | 
"thy"}.  | 
|
269  | 
||
270  | 
  \end{description}
 | 
|
271  | 
*}  | 
|
272  | 
||
| 20430 | 273  | 
|
| 20451 | 274  | 
subsection {* Generic contexts \label{sec:generic-context} *}
 | 
| 20429 | 275  | 
|
| 20449 | 276  | 
text {*
 | 
277  | 
A generic context is the disjoint sum of either a theory or proof  | 
|
| 20451 | 278  | 
context. Occasionally, this enables uniform treatment of generic  | 
| 20450 | 279  | 
context data, typically extra-logical information. Operations on  | 
| 20449 | 280  | 
generic contexts include the usual injections, partial selections,  | 
281  | 
and combinators for lifting operations on either component of the  | 
|
282  | 
disjoint sum.  | 
|
283  | 
||
284  | 
  Moreover, there are total operations @{text "theory_of"} and @{text
 | 
|
285  | 
"proof_of"} to convert a generic context into either kind: a theory  | 
|
| 20451 | 286  | 
can always be selected from the sum, while a proof context might  | 
| 34921 | 287  | 
  have to be constructed by an ad-hoc @{text "init"} operation, which
 | 
288  | 
incurs a small runtime overhead.  | 
|
| 20449 | 289  | 
*}  | 
| 20430 | 290  | 
|
| 20449 | 291  | 
text %mlref {*
 | 
292  | 
  \begin{mldecls}
 | 
|
293  | 
  @{index_ML_type Context.generic} \\
 | 
|
294  | 
  @{index_ML Context.theory_of: "Context.generic -> theory"} \\
 | 
|
295  | 
  @{index_ML Context.proof_of: "Context.generic -> Proof.context"} \\
 | 
|
296  | 
  \end{mldecls}
 | 
|
297  | 
||
298  | 
  \begin{description}
 | 
|
| 20430 | 299  | 
|
| 20449 | 300  | 
  \item @{ML_type Context.generic} is the direct sum of @{ML_type
 | 
| 20451 | 301  | 
  "theory"} and @{ML_type "Proof.context"}, with the datatype
 | 
302  | 
  constructors @{ML "Context.Theory"} and @{ML "Context.Proof"}.
 | 
|
| 20449 | 303  | 
|
304  | 
  \item @{ML Context.theory_of}~@{text "context"} always produces a
 | 
|
305  | 
  theory from the generic @{text "context"}, using @{ML
 | 
|
306  | 
"ProofContext.theory_of"} as required.  | 
|
307  | 
||
308  | 
  \item @{ML Context.proof_of}~@{text "context"} always produces a
 | 
|
309  | 
  proof context from the generic @{text "context"}, using @{ML
 | 
|
| 36611 | 310  | 
"ProofContext.init_global"} as required (note that this re-initializes the  | 
| 20451 | 311  | 
context data with each invocation).  | 
| 20449 | 312  | 
|
313  | 
  \end{description}
 | 
|
314  | 
*}  | 
|
| 20437 | 315  | 
|
| 20476 | 316  | 
|
317  | 
subsection {* Context data \label{sec:context-data} *}
 | 
|
| 20447 | 318  | 
|
| 
33524
 
a08e6c1cbc04
updated functor Theory_Data, Proof_Data, Generic_Data;
 
wenzelm 
parents: 
33174 
diff
changeset
 | 
319  | 
text {* The main purpose of theory and proof contexts is to manage
 | 
| 
 
a08e6c1cbc04
updated functor Theory_Data, Proof_Data, Generic_Data;
 
wenzelm 
parents: 
33174 
diff
changeset
 | 
320  | 
arbitrary (pure) data. New data types can be declared incrementally  | 
| 
 
a08e6c1cbc04
updated functor Theory_Data, Proof_Data, Generic_Data;
 
wenzelm 
parents: 
33174 
diff
changeset
 | 
321  | 
at compile time. There are separate declaration mechanisms for any  | 
| 
 
a08e6c1cbc04
updated functor Theory_Data, Proof_Data, Generic_Data;
 
wenzelm 
parents: 
33174 
diff
changeset
 | 
322  | 
of the three kinds of contexts: theory, proof, generic.  | 
| 20449 | 323  | 
|
| 
33524
 
a08e6c1cbc04
updated functor Theory_Data, Proof_Data, Generic_Data;
 
wenzelm 
parents: 
33174 
diff
changeset
 | 
324  | 
  \paragraph{Theory data} declarations need to implement the following
 | 
| 
 
a08e6c1cbc04
updated functor Theory_Data, Proof_Data, Generic_Data;
 
wenzelm 
parents: 
33174 
diff
changeset
 | 
325  | 
SML signature:  | 
| 20449 | 326  | 
|
327  | 
\medskip  | 
|
328  | 
  \begin{tabular}{ll}
 | 
|
| 22869 | 329  | 
  @{text "\<type> T"} & representing type \\
 | 
330  | 
  @{text "\<val> empty: T"} & empty default value \\
 | 
|
331  | 
  @{text "\<val> extend: T \<rightarrow> T"} & re-initialize on import \\
 | 
|
332  | 
  @{text "\<val> merge: T \<times> T \<rightarrow> T"} & join on import \\
 | 
|
| 20449 | 333  | 
  \end{tabular}
 | 
334  | 
\medskip  | 
|
335  | 
||
| 22869 | 336  | 
  \noindent The @{text "empty"} value acts as initial default for
 | 
337  | 
  \emph{any} theory that does not declare actual data content; @{text
 | 
|
| 
33524
 
a08e6c1cbc04
updated functor Theory_Data, Proof_Data, Generic_Data;
 
wenzelm 
parents: 
33174 
diff
changeset
 | 
338  | 
  "extend"} is acts like a unitary version of @{text "merge"}.
 | 
| 20449 | 339  | 
|
| 34921 | 340  | 
  Implementing @{text "merge"} can be tricky.  The general idea is
 | 
341  | 
  that @{text "merge (data\<^sub>1, data\<^sub>2)"} inserts those parts of @{text
 | 
|
342  | 
  "data\<^sub>2"} into @{text "data\<^sub>1"} that are not yet present, while
 | 
|
343  | 
  keeping the general order of things.  The @{ML Library.merge}
 | 
|
344  | 
function on plain lists may serve as canonical template.  | 
|
345  | 
||
346  | 
Particularly note that shared parts of the data must not be  | 
|
347  | 
duplicated by naive concatenation, or a theory graph that is like a  | 
|
348  | 
chain of diamonds would cause an exponential blowup!  | 
|
349  | 
||
| 
33524
 
a08e6c1cbc04
updated functor Theory_Data, Proof_Data, Generic_Data;
 
wenzelm 
parents: 
33174 
diff
changeset
 | 
350  | 
  \paragraph{Proof context data} declarations need to implement the
 | 
| 
 
a08e6c1cbc04
updated functor Theory_Data, Proof_Data, Generic_Data;
 
wenzelm 
parents: 
33174 
diff
changeset
 | 
351  | 
following SML signature:  | 
| 20449 | 352  | 
|
353  | 
\medskip  | 
|
354  | 
  \begin{tabular}{ll}
 | 
|
| 22869 | 355  | 
  @{text "\<type> T"} & representing type \\
 | 
356  | 
  @{text "\<val> init: theory \<rightarrow> T"} & produce initial value \\
 | 
|
| 20449 | 357  | 
  \end{tabular}
 | 
358  | 
\medskip  | 
|
359  | 
||
360  | 
  \noindent The @{text "init"} operation is supposed to produce a pure
 | 
|
| 34921 | 361  | 
value from the given background theory and should be somehow  | 
362  | 
``immediate''. Whenever a proof context is initialized, which  | 
|
363  | 
  happens frequently, the the system invokes the @{text "init"}
 | 
|
| 39821 | 364  | 
  operation of \emph{all} theory data slots ever declared.  This also
 | 
365  | 
means that one needs to be economic about the total number of proof  | 
|
366  | 
data declarations in the system, i.e.\ each ML module should declare  | 
|
367  | 
at most one, sometimes two data slots for its internal use.  | 
|
368  | 
Repeated data declarations to simulate a record type should be  | 
|
369  | 
avoided!  | 
|
| 20449 | 370  | 
|
| 20451 | 371  | 
  \paragraph{Generic data} provides a hybrid interface for both theory
 | 
| 
33524
 
a08e6c1cbc04
updated functor Theory_Data, Proof_Data, Generic_Data;
 
wenzelm 
parents: 
33174 
diff
changeset
 | 
372  | 
  and proof data.  The @{text "init"} operation for proof contexts is
 | 
| 
 
a08e6c1cbc04
updated functor Theory_Data, Proof_Data, Generic_Data;
 
wenzelm 
parents: 
33174 
diff
changeset
 | 
373  | 
predefined to select the current data value from the background  | 
| 
 
a08e6c1cbc04
updated functor Theory_Data, Proof_Data, Generic_Data;
 
wenzelm 
parents: 
33174 
diff
changeset
 | 
374  | 
theory.  | 
| 20449 | 375  | 
|
| 39821 | 376  | 
  \bigskip Any of the above data declarations over type @{text "T"}
 | 
377  | 
result in an ML structure with the following signature:  | 
|
| 20449 | 378  | 
|
379  | 
\medskip  | 
|
380  | 
  \begin{tabular}{ll}
 | 
|
381  | 
  @{text "get: context \<rightarrow> T"} \\
 | 
|
382  | 
  @{text "put: T \<rightarrow> context \<rightarrow> context"} \\
 | 
|
383  | 
  @{text "map: (T \<rightarrow> T) \<rightarrow> context \<rightarrow> context"} \\
 | 
|
384  | 
  \end{tabular}
 | 
|
385  | 
\medskip  | 
|
386  | 
||
| 34921 | 387  | 
\noindent These other operations provide exclusive access for the  | 
388  | 
particular kind of context (theory, proof, or generic context).  | 
|
| 39821 | 389  | 
This interface observes the ML discipline for types and scopes:  | 
390  | 
there is no other way to access the corresponding data slot of a  | 
|
391  | 
context. By keeping these operations private, an Isabelle/ML module  | 
|
392  | 
may maintain abstract values authentically.  | 
|
| 20447 | 393  | 
*}  | 
394  | 
||
| 20450 | 395  | 
text %mlref {*
 | 
396  | 
  \begin{mldecls}
 | 
|
| 
33524
 
a08e6c1cbc04
updated functor Theory_Data, Proof_Data, Generic_Data;
 
wenzelm 
parents: 
33174 
diff
changeset
 | 
397  | 
  @{index_ML_functor Theory_Data} \\
 | 
| 
 
a08e6c1cbc04
updated functor Theory_Data, Proof_Data, Generic_Data;
 
wenzelm 
parents: 
33174 
diff
changeset
 | 
398  | 
  @{index_ML_functor Proof_Data} \\
 | 
| 
 
a08e6c1cbc04
updated functor Theory_Data, Proof_Data, Generic_Data;
 
wenzelm 
parents: 
33174 
diff
changeset
 | 
399  | 
  @{index_ML_functor Generic_Data} \\
 | 
| 20450 | 400  | 
  \end{mldecls}
 | 
401  | 
||
402  | 
  \begin{description}
 | 
|
403  | 
||
| 
33524
 
a08e6c1cbc04
updated functor Theory_Data, Proof_Data, Generic_Data;
 
wenzelm 
parents: 
33174 
diff
changeset
 | 
404  | 
  \item @{ML_functor Theory_Data}@{text "(spec)"} declares data for
 | 
| 20450 | 405  | 
  type @{ML_type theory} according to the specification provided as
 | 
| 20451 | 406  | 
argument structure. The resulting structure provides data init and  | 
407  | 
access operations as described above.  | 
|
| 20450 | 408  | 
|
| 
33524
 
a08e6c1cbc04
updated functor Theory_Data, Proof_Data, Generic_Data;
 
wenzelm 
parents: 
33174 
diff
changeset
 | 
409  | 
  \item @{ML_functor Proof_Data}@{text "(spec)"} is analogous to
 | 
| 
 
a08e6c1cbc04
updated functor Theory_Data, Proof_Data, Generic_Data;
 
wenzelm 
parents: 
33174 
diff
changeset
 | 
410  | 
  @{ML_functor Theory_Data} for type @{ML_type Proof.context}.
 | 
| 20450 | 411  | 
|
| 
33524
 
a08e6c1cbc04
updated functor Theory_Data, Proof_Data, Generic_Data;
 
wenzelm 
parents: 
33174 
diff
changeset
 | 
412  | 
  \item @{ML_functor Generic_Data}@{text "(spec)"} is analogous to
 | 
| 
 
a08e6c1cbc04
updated functor Theory_Data, Proof_Data, Generic_Data;
 
wenzelm 
parents: 
33174 
diff
changeset
 | 
413  | 
  @{ML_functor Theory_Data} for type @{ML_type Context.generic}.
 | 
| 20450 | 414  | 
|
415  | 
  \end{description}
 | 
|
416  | 
*}  | 
|
417  | 
||
| 34928 | 418  | 
text %mlex {*
 | 
419  | 
The following artificial example demonstrates theory  | 
|
420  | 
data: we maintain a set of terms that are supposed to be wellformed  | 
|
421  | 
wrt.\ the enclosing theory. The public interface is as follows:  | 
|
422  | 
*}  | 
|
423  | 
||
424  | 
ML {*
 | 
|
425  | 
signature WELLFORMED_TERMS =  | 
|
426  | 
sig  | 
|
427  | 
val get: theory -> term list  | 
|
428  | 
val add: term -> theory -> theory  | 
|
429  | 
end;  | 
|
430  | 
*}  | 
|
431  | 
||
432  | 
text {* \noindent The implementation uses private theory data
 | 
|
433  | 
internally, and only exposes an operation that involves explicit  | 
|
434  | 
argument checking wrt.\ the given theory. *}  | 
|
435  | 
||
436  | 
ML {*
 | 
|
437  | 
structure Wellformed_Terms: WELLFORMED_TERMS =  | 
|
438  | 
struct  | 
|
439  | 
||
440  | 
structure Terms = Theory_Data  | 
|
441  | 
(  | 
|
| 39687 | 442  | 
type T = term Ord_List.T;  | 
| 34928 | 443  | 
val empty = [];  | 
444  | 
val extend = I;  | 
|
445  | 
fun merge (ts1, ts2) =  | 
|
| 39687 | 446  | 
Ord_List.union Term_Ord.fast_term_ord ts1 ts2;  | 
| 34928 | 447  | 
)  | 
448  | 
||
449  | 
val get = Terms.get;  | 
|
450  | 
||
451  | 
fun add raw_t thy =  | 
|
| 39821 | 452  | 
let  | 
453  | 
val t = Sign.cert_term thy raw_t;  | 
|
454  | 
in  | 
|
455  | 
Terms.map (Ord_List.insert Term_Ord.fast_term_ord t) thy  | 
|
456  | 
end;  | 
|
| 34928 | 457  | 
|
458  | 
end;  | 
|
459  | 
*}  | 
|
460  | 
||
| 39821 | 461  | 
text {* \noindent We use @{ML_type "term Ord_List.T"} for reasonably
 | 
462  | 
efficient representation of a set of terms: all operations are  | 
|
463  | 
linear in the number of stored elements. Here we assume that users  | 
|
464  | 
of this module do not care about the declaration order, since that  | 
|
465  | 
data structure forces its own arrangement of elements.  | 
|
| 34928 | 466  | 
|
467  | 
  Observe how the @{verbatim merge} operation joins the data slots of
 | 
|
| 39687 | 468  | 
  the two constituents: @{ML Ord_List.union} prevents duplication of
 | 
| 34928 | 469  | 
common data from different branches, thus avoiding the danger of  | 
| 39821 | 470  | 
exponential blowup. Plain list append etc.\ must never be used for  | 
471  | 
theory data merges!  | 
|
| 34928 | 472  | 
|
473  | 
\medskip Our intended invariant is achieved as follows:  | 
|
474  | 
  \begin{enumerate}
 | 
|
475  | 
||
476  | 
  \item @{ML Wellformed_Terms.add} only admits terms that have passed
 | 
|
477  | 
  the @{ML Sign.cert_term} check of the given theory at that point.
 | 
|
478  | 
||
479  | 
  \item Wellformedness in the sense of @{ML Sign.cert_term} is
 | 
|
480  | 
monotonic wrt.\ the sub-theory relation. So our data can move  | 
|
481  | 
upwards in the hierarchy (via extension or merges), and maintain  | 
|
482  | 
wellformedness without further checks.  | 
|
483  | 
||
484  | 
  \end{enumerate}
 | 
|
485  | 
||
486  | 
Note that all basic operations of the inference kernel (which  | 
|
487  | 
  includes @{ML Sign.cert_term}) observe this monotonicity principle,
 | 
|
488  | 
but other user-space tools don't. For example, fully-featured  | 
|
489  | 
  type-inference via @{ML Syntax.check_term} (cf.\
 | 
|
490  | 
  \secref{sec:term-check}) is not necessarily monotonic wrt.\ the
 | 
|
491  | 
background theory, since constraints of term constants can be  | 
|
| 39821 | 492  | 
modified by later declarations, for example.  | 
| 34928 | 493  | 
|
494  | 
In most cases, user-space context data does not have to take such  | 
|
495  | 
invariants too seriously. The situation is different in the  | 
|
496  | 
implementation of the inference kernel itself, which uses the very  | 
|
497  | 
same data mechanisms for types, constants, axioms etc.  | 
|
498  | 
*}  | 
|
499  | 
||
| 20447 | 500  | 
|
| 26872 | 501  | 
section {* Names \label{sec:names} *}
 | 
| 20451 | 502  | 
|
| 34925 | 503  | 
text {* In principle, a name is just a string, but there are various
 | 
504  | 
conventions for representing additional structure. For example,  | 
|
| 
34927
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
505  | 
  ``@{text "Foo.bar.baz"}'' is considered as a long name consisting of
 | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
506  | 
  qualifier @{text "Foo.bar"} and base name @{text "baz"}.  The
 | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
507  | 
individual constituents of a name may have further substructure,  | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
508  | 
e.g.\ the string ``\verb,\,\verb,<alpha>,'' encodes as a single  | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
509  | 
symbol.  | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
510  | 
|
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
511  | 
\medskip Subsequently, we shall introduce specific categories of  | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
512  | 
names. Roughly speaking these correspond to logical entities as  | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
513  | 
follows:  | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
514  | 
  \begin{itemize}
 | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
515  | 
|
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
516  | 
  \item Basic names (\secref{sec:basic-name}): free and bound
 | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
517  | 
variables.  | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
518  | 
|
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
519  | 
  \item Indexed names (\secref{sec:indexname}): schematic variables.
 | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
520  | 
|
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
521  | 
  \item Long names (\secref{sec:long-name}): constants of any kind
 | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
522  | 
(type constructors, term constants, other concepts defined in user  | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
523  | 
space). Such entities are typically managed via name spaces  | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
524  | 
  (\secref{sec:name-space}).
 | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
525  | 
|
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
526  | 
  \end{itemize}
 | 
| 20451 | 527  | 
*}  | 
| 20437 | 528  | 
|
529  | 
||
530  | 
subsection {* Strings of symbols *}
 | 
|
531  | 
||
| 34925 | 532  | 
text {* A \emph{symbol} constitutes the smallest textual unit in
 | 
533  | 
Isabelle --- raw ML characters are normally not encountered at all!  | 
|
534  | 
Isabelle strings consist of a sequence of symbols, represented as a  | 
|
535  | 
packed string or an exploded list of strings. Each symbol is in  | 
|
536  | 
itself a small string, which has either one of the following forms:  | 
|
| 20437 | 537  | 
|
| 20451 | 538  | 
  \begin{enumerate}
 | 
| 20437 | 539  | 
|
| 
37533
 
d775bd70f571
explicit treatment of UTF8 character sequences as Isabelle symbols;
 
wenzelm 
parents: 
36611 
diff
changeset
 | 
540  | 
  \item a single ASCII character ``@{text "c"}'', for example
 | 
| 
 
d775bd70f571
explicit treatment of UTF8 character sequences as Isabelle symbols;
 
wenzelm 
parents: 
36611 
diff
changeset
 | 
541  | 
``\verb,a,'',  | 
| 
 
d775bd70f571
explicit treatment of UTF8 character sequences as Isabelle symbols;
 
wenzelm 
parents: 
36611 
diff
changeset
 | 
542  | 
|
| 
 
d775bd70f571
explicit treatment of UTF8 character sequences as Isabelle symbols;
 
wenzelm 
parents: 
36611 
diff
changeset
 | 
543  | 
\item a codepoint according to UTF8 (non-ASCII byte sequence),  | 
| 20437 | 544  | 
|
| 20488 | 545  | 
  \item a regular symbol ``\verb,\,\verb,<,@{text "ident"}\verb,>,'',
 | 
| 20476 | 546  | 
for example ``\verb,\,\verb,<alpha>,'',  | 
| 20437 | 547  | 
|
| 20488 | 548  | 
  \item a control symbol ``\verb,\,\verb,<^,@{text "ident"}\verb,>,'',
 | 
| 20476 | 549  | 
for example ``\verb,\,\verb,<^bold>,'',  | 
| 20437 | 550  | 
|
| 20488 | 551  | 
  \item a raw symbol ``\verb,\,\verb,<^raw:,@{text text}\verb,>,''
 | 
| 34925 | 552  | 
  where @{text text} consists of printable characters excluding
 | 
| 20476 | 553  | 
``\verb,.,'' and ``\verb,>,'', for example  | 
554  | 
  ``\verb,\,\verb,<^raw:$\sum_{i = 1}^n$>,'',
 | 
|
| 20437 | 555  | 
|
| 20488 | 556  | 
  \item a numbered raw control symbol ``\verb,\,\verb,<^raw,@{text
 | 
| 20476 | 557  | 
  n}\verb,>, where @{text n} consists of digits, for example
 | 
| 20451 | 558  | 
``\verb,\,\verb,<^raw42>,''.  | 
| 20437 | 559  | 
|
| 20451 | 560  | 
  \end{enumerate}
 | 
| 20437 | 561  | 
|
| 20476 | 562  | 
  \noindent The @{text "ident"} syntax for symbol names is @{text
 | 
563  | 
  "letter (letter | digit)\<^sup>*"}, where @{text "letter =
 | 
|
564  | 
  A..Za..z"} and @{text "digit = 0..9"}.  There are infinitely many
 | 
|
565  | 
regular symbols and control symbols, but a fixed collection of  | 
|
566  | 
standard symbols is treated specifically. For example,  | 
|
| 20488 | 567  | 
``\verb,\,\verb,<alpha>,'' is classified as a letter, which means it  | 
568  | 
may occur within regular Isabelle identifiers.  | 
|
| 20437 | 569  | 
|
| 
37533
 
d775bd70f571
explicit treatment of UTF8 character sequences as Isabelle symbols;
 
wenzelm 
parents: 
36611 
diff
changeset
 | 
570  | 
The character set underlying Isabelle symbols is 7-bit ASCII, but  | 
| 
 
d775bd70f571
explicit treatment of UTF8 character sequences as Isabelle symbols;
 
wenzelm 
parents: 
36611 
diff
changeset
 | 
571  | 
8-bit character sequences are passed-through unchanged. Unicode/UCS  | 
| 
 
d775bd70f571
explicit treatment of UTF8 character sequences as Isabelle symbols;
 
wenzelm 
parents: 
36611 
diff
changeset
 | 
572  | 
data in UTF-8 encoding is processed in a non-strict fashion, such  | 
| 
 
d775bd70f571
explicit treatment of UTF8 character sequences as Isabelle symbols;
 
wenzelm 
parents: 
36611 
diff
changeset
 | 
573  | 
that well-formed code sequences are recognized  | 
| 
 
d775bd70f571
explicit treatment of UTF8 character sequences as Isabelle symbols;
 
wenzelm 
parents: 
36611 
diff
changeset
 | 
574  | 
  accordingly.\footnote{Note that ISO-Latin-1 differs from UTF-8 only
 | 
| 
 
d775bd70f571
explicit treatment of UTF8 character sequences as Isabelle symbols;
 
wenzelm 
parents: 
36611 
diff
changeset
 | 
575  | 
in some special punctuation characters that even have replacements  | 
| 
 
d775bd70f571
explicit treatment of UTF8 character sequences as Isabelle symbols;
 
wenzelm 
parents: 
36611 
diff
changeset
 | 
576  | 
within the standard collection of Isabelle symbols. Text consisting  | 
| 
 
d775bd70f571
explicit treatment of UTF8 character sequences as Isabelle symbols;
 
wenzelm 
parents: 
36611 
diff
changeset
 | 
577  | 
of ASCII plus accented letters can be processed in either encoding.}  | 
| 
 
d775bd70f571
explicit treatment of UTF8 character sequences as Isabelle symbols;
 
wenzelm 
parents: 
36611 
diff
changeset
 | 
578  | 
Unicode provides its own collection of mathematical symbols, but  | 
| 
 
d775bd70f571
explicit treatment of UTF8 character sequences as Isabelle symbols;
 
wenzelm 
parents: 
36611 
diff
changeset
 | 
579  | 
within the core Isabelle/ML world there is no link to the standard  | 
| 
 
d775bd70f571
explicit treatment of UTF8 character sequences as Isabelle symbols;
 
wenzelm 
parents: 
36611 
diff
changeset
 | 
580  | 
collection of Isabelle regular symbols.  | 
| 20476 | 581  | 
|
582  | 
\medskip Output of Isabelle symbols depends on the print mode  | 
|
| 29758 | 583  | 
  (\secref{print-mode}).  For example, the standard {\LaTeX} setup of
 | 
584  | 
the Isabelle document preparation system would present  | 
|
| 20451 | 585  | 
  ``\verb,\,\verb,<alpha>,'' as @{text "\<alpha>"}, and
 | 
586  | 
  ``\verb,\,\verb,<^bold>,\verb,\,\verb,<alpha>,'' as @{text
 | 
|
| 34925 | 587  | 
"\<^bold>\<alpha>"}. On-screen rendering usually works by mapping a finite  | 
588  | 
subset of Isabelle symbols to suitable Unicode characters.  | 
|
| 20451 | 589  | 
*}  | 
| 20437 | 590  | 
|
591  | 
text %mlref {*
 | 
|
592  | 
  \begin{mldecls}
 | 
|
| 34921 | 593  | 
  @{index_ML_type "Symbol.symbol": string} \\
 | 
| 20437 | 594  | 
  @{index_ML Symbol.explode: "string -> Symbol.symbol list"} \\
 | 
595  | 
  @{index_ML Symbol.is_letter: "Symbol.symbol -> bool"} \\
 | 
|
596  | 
  @{index_ML Symbol.is_digit: "Symbol.symbol -> bool"} \\
 | 
|
597  | 
  @{index_ML Symbol.is_quasi: "Symbol.symbol -> bool"} \\
 | 
|
| 20547 | 598  | 
  @{index_ML Symbol.is_blank: "Symbol.symbol -> bool"} \\
 | 
599  | 
  \end{mldecls}
 | 
|
600  | 
  \begin{mldecls}
 | 
|
| 20437 | 601  | 
  @{index_ML_type "Symbol.sym"} \\
 | 
602  | 
  @{index_ML Symbol.decode: "Symbol.symbol -> Symbol.sym"} \\
 | 
|
603  | 
  \end{mldecls}
 | 
|
604  | 
||
605  | 
  \begin{description}
 | 
|
606  | 
||
| 20488 | 607  | 
  \item @{ML_type "Symbol.symbol"} represents individual Isabelle
 | 
| 34921 | 608  | 
symbols.  | 
| 20437 | 609  | 
|
| 20476 | 610  | 
  \item @{ML "Symbol.explode"}~@{text "str"} produces a symbol list
 | 
| 39821 | 611  | 
  from the packed form.  This function supersedes @{ML
 | 
| 20476 | 612  | 
"String.explode"} for virtually all purposes of manipulating text in  | 
| 34925 | 613  | 
  Isabelle!\footnote{The runtime overhead for exploded strings is
 | 
614  | 
mainly that of the list structure: individual symbols that happen to  | 
|
| 39821 | 615  | 
be a singleton string do not require extra memory in Poly/ML.}  | 
| 20437 | 616  | 
|
617  | 
  \item @{ML "Symbol.is_letter"}, @{ML "Symbol.is_digit"}, @{ML
 | 
|
| 20476 | 618  | 
  "Symbol.is_quasi"}, @{ML "Symbol.is_blank"} classify standard
 | 
619  | 
symbols according to fixed syntactic conventions of Isabelle, cf.\  | 
|
620  | 
  \cite{isabelle-isar-ref}.
 | 
|
| 20437 | 621  | 
|
622  | 
  \item @{ML_type "Symbol.sym"} is a concrete datatype that represents
 | 
|
| 20488 | 623  | 
  the different kinds of symbols explicitly, with constructors @{ML
 | 
| 
37533
 
d775bd70f571
explicit treatment of UTF8 character sequences as Isabelle symbols;
 
wenzelm 
parents: 
36611 
diff
changeset
 | 
624  | 
  "Symbol.Char"}, @{ML "Symbol.Sym"}, @{ML "Symbol.UTF8"}, @{ML
 | 
| 
 
d775bd70f571
explicit treatment of UTF8 character sequences as Isabelle symbols;
 
wenzelm 
parents: 
36611 
diff
changeset
 | 
625  | 
  "Symbol.Ctrl"}, @{ML "Symbol.Raw"}.
 | 
| 20437 | 626  | 
|
627  | 
  \item @{ML "Symbol.decode"} converts the string representation of a
 | 
|
| 20451 | 628  | 
symbol into the datatype version.  | 
| 20437 | 629  | 
|
630  | 
  \end{description}
 | 
|
| 34925 | 631  | 
|
632  | 
  \paragraph{Historical note.} In the original SML90 standard the
 | 
|
633  | 
  primitive ML type @{ML_type char} did not exists, and the basic @{ML
 | 
|
634  | 
"explode: string -> string list"} operation would produce a list of  | 
|
635  | 
singleton strings as in Isabelle/ML today. When SML97 came out,  | 
|
| 
34927
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
636  | 
Isabelle did not adopt its slightly anachronistic 8-bit characters,  | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
637  | 
but the idea of exploding a string into a list of small strings was  | 
| 34925 | 638  | 
extended to ``symbols'' as explained above. Thus Isabelle sources  | 
639  | 
can refer to an infinite store of user-defined symbols, without  | 
|
640  | 
having to worry about the multitude of Unicode encodings.  | 
|
| 20437 | 641  | 
*}  | 
642  | 
||
643  | 
||
| 
34927
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
644  | 
subsection {* Basic names \label{sec:basic-name} *}
 | 
| 20476 | 645  | 
|
646  | 
text {*
 | 
|
647  | 
  A \emph{basic name} essentially consists of a single Isabelle
 | 
|
648  | 
identifier. There are conventions to mark separate classes of basic  | 
|
| 29761 | 649  | 
names, by attaching a suffix of underscores: one underscore means  | 
650  | 
  \emph{internal name}, two underscores means \emph{Skolem name},
 | 
|
651  | 
  three underscores means \emph{internal Skolem name}.
 | 
|
| 20476 | 652  | 
|
653  | 
  For example, the basic name @{text "foo"} has the internal version
 | 
|
654  | 
  @{text "foo_"}, with Skolem versions @{text "foo__"} and @{text
 | 
|
655  | 
"foo___"}, respectively.  | 
|
656  | 
||
| 20488 | 657  | 
These special versions provide copies of the basic name space, apart  | 
658  | 
from anything that normally appears in the user text. For example,  | 
|
659  | 
system generated variables in Isar proof contexts are usually marked  | 
|
| 34926 | 660  | 
  as internal, which prevents mysterious names like @{text "xaa"} to
 | 
661  | 
appear in human-readable text.  | 
|
| 20476 | 662  | 
|
| 20488 | 663  | 
\medskip Manipulating binding scopes often requires on-the-fly  | 
664  | 
  renamings.  A \emph{name context} contains a collection of already
 | 
|
665  | 
  used names.  The @{text "declare"} operation adds names to the
 | 
|
666  | 
context.  | 
|
| 20476 | 667  | 
|
| 20488 | 668  | 
  The @{text "invents"} operation derives a number of fresh names from
 | 
669  | 
a given starting point. For example, the first three names derived  | 
|
670  | 
  from @{text "a"} are @{text "a"}, @{text "b"}, @{text "c"}.
 | 
|
| 20476 | 671  | 
|
672  | 
  The @{text "variants"} operation produces fresh names by
 | 
|
| 20488 | 673  | 
  incrementing tentative names as base-26 numbers (with digits @{text
 | 
674  | 
  "a..z"}) until all clashes are resolved.  For example, name @{text
 | 
|
675  | 
  "foo"} results in variants @{text "fooa"}, @{text "foob"}, @{text
 | 
|
676  | 
  "fooc"}, \dots, @{text "fooaa"}, @{text "fooab"} etc.; each renaming
 | 
|
677  | 
step picks the next unused variant from this sequence.  | 
|
| 20476 | 678  | 
*}  | 
679  | 
||
680  | 
text %mlref {*
 | 
|
681  | 
  \begin{mldecls}
 | 
|
682  | 
  @{index_ML Name.internal: "string -> string"} \\
 | 
|
| 20547 | 683  | 
  @{index_ML Name.skolem: "string -> string"} \\
 | 
684  | 
  \end{mldecls}
 | 
|
685  | 
  \begin{mldecls}
 | 
|
| 20476 | 686  | 
  @{index_ML_type Name.context} \\
 | 
687  | 
  @{index_ML Name.context: Name.context} \\
 | 
|
688  | 
  @{index_ML Name.declare: "string -> Name.context -> Name.context"} \\
 | 
|
689  | 
  @{index_ML Name.invents: "Name.context -> string -> int -> string list"} \\
 | 
|
690  | 
  @{index_ML Name.variants: "string list -> Name.context -> string list * Name.context"} \\
 | 
|
691  | 
  \end{mldecls}
 | 
|
| 34926 | 692  | 
  \begin{mldecls}
 | 
693  | 
  @{index_ML Variable.names_of: "Proof.context -> Name.context"} \\
 | 
|
694  | 
  \end{mldecls}
 | 
|
| 20476 | 695  | 
|
696  | 
  \begin{description}
 | 
|
697  | 
||
698  | 
  \item @{ML Name.internal}~@{text "name"} produces an internal name
 | 
|
699  | 
by adding one underscore.  | 
|
700  | 
||
701  | 
  \item @{ML Name.skolem}~@{text "name"} produces a Skolem name by
 | 
|
702  | 
adding two underscores.  | 
|
703  | 
||
704  | 
  \item @{ML_type Name.context} represents the context of already used
 | 
|
705  | 
  names; the initial value is @{ML "Name.context"}.
 | 
|
706  | 
||
| 20488 | 707  | 
  \item @{ML Name.declare}~@{text "name"} enters a used name into the
 | 
708  | 
context.  | 
|
| 20437 | 709  | 
|
| 20488 | 710  | 
  \item @{ML Name.invents}~@{text "context name n"} produces @{text
 | 
711  | 
  "n"} fresh names derived from @{text "name"}.
 | 
|
712  | 
||
713  | 
  \item @{ML Name.variants}~@{text "names context"} produces fresh
 | 
|
| 29761 | 714  | 
  variants of @{text "names"}; the result is entered into the context.
 | 
| 20476 | 715  | 
|
| 34926 | 716  | 
  \item @{ML Variable.names_of}~@{text "ctxt"} retrieves the context
 | 
717  | 
of declared type and term variable names. Projecting a proof  | 
|
718  | 
context down to a primitive name context is occasionally useful when  | 
|
719  | 
invoking lower-level operations. Regular management of ``fresh  | 
|
720  | 
  variables'' is done by suitable operations of structure @{ML_struct
 | 
|
721  | 
Variable}, which is also able to provide an official status of  | 
|
722  | 
``locally fixed variable'' within the logical environment (cf.\  | 
|
723  | 
  \secref{sec:variables}).
 | 
|
724  | 
||
| 20476 | 725  | 
  \end{description}
 | 
726  | 
*}  | 
|
727  | 
||
728  | 
||
| 
34927
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
729  | 
subsection {* Indexed names \label{sec:indexname} *}
 | 
| 20476 | 730  | 
|
731  | 
text {*
 | 
|
732  | 
  An \emph{indexed name} (or @{text "indexname"}) is a pair of a basic
 | 
|
| 20488 | 733  | 
name and a natural number. This representation allows efficient  | 
734  | 
renaming by incrementing the second component only. The canonical  | 
|
735  | 
way to rename two collections of indexnames apart from each other is  | 
|
736  | 
  this: determine the maximum index @{text "maxidx"} of the first
 | 
|
737  | 
collection, then increment all indexes of the second collection by  | 
|
738  | 
  @{text "maxidx + 1"}; the maximum index of an empty collection is
 | 
|
739  | 
  @{text "-1"}.
 | 
|
| 20476 | 740  | 
|
| 
34927
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
741  | 
Occasionally, basic names are injected into the same pair type of  | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
742  | 
  indexed names: then @{text "(x, -1)"} is used to encode the basic
 | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
743  | 
  name @{text "x"}.
 | 
| 20488 | 744  | 
|
745  | 
\medskip Isabelle syntax observes the following rules for  | 
|
746  | 
  representing an indexname @{text "(x, i)"} as a packed string:
 | 
|
| 20476 | 747  | 
|
748  | 
  \begin{itemize}
 | 
|
749  | 
||
| 20479 | 750  | 
  \item @{text "?x"} if @{text "x"} does not end with a digit and @{text "i = 0"},
 | 
| 20476 | 751  | 
|
752  | 
  \item @{text "?xi"} if @{text "x"} does not end with a digit,
 | 
|
753  | 
||
| 20488 | 754  | 
  \item @{text "?x.i"} otherwise.
 | 
| 20476 | 755  | 
|
756  | 
  \end{itemize}
 | 
|
| 20470 | 757  | 
|
| 
34927
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
758  | 
Indexnames may acquire large index numbers after several maxidx  | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
759  | 
shifts have been applied. Results are usually normalized towards  | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
760  | 
  @{text "0"} at certain checkpoints, notably at the end of a proof.
 | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
761  | 
This works by producing variants of the corresponding basic name  | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
762  | 
  components.  For example, the collection @{text "?x1, ?x7, ?x42"}
 | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
763  | 
  becomes @{text "?x, ?xa, ?xb"}.
 | 
| 20476 | 764  | 
*}  | 
765  | 
||
766  | 
text %mlref {*
 | 
|
767  | 
  \begin{mldecls}
 | 
|
768  | 
  @{index_ML_type indexname} \\
 | 
|
769  | 
  \end{mldecls}
 | 
|
770  | 
||
771  | 
  \begin{description}
 | 
|
772  | 
||
773  | 
  \item @{ML_type indexname} represents indexed names.  This is an
 | 
|
774  | 
  abbreviation for @{ML_type "string * int"}.  The second component is
 | 
|
775  | 
  usually non-negative, except for situations where @{text "(x, -1)"}
 | 
|
| 34926 | 776  | 
is used to inject basic names into this type. Other negative  | 
777  | 
indexes should not be used.  | 
|
| 20476 | 778  | 
|
779  | 
  \end{description}
 | 
|
780  | 
*}  | 
|
781  | 
||
782  | 
||
| 
34927
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
783  | 
subsection {* Long names \label{sec:long-name} *}
 | 
| 20476 | 784  | 
|
| 
34927
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
785  | 
text {* A \emph{long name} consists of a sequence of non-empty name
 | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
786  | 
components. The packed representation uses a dot as separator, as  | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
787  | 
  in ``@{text "A.b.c"}''.  The last component is called \emph{base
 | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
788  | 
  name}, the remaining prefix is called \emph{qualifier} (which may be
 | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
789  | 
empty). The qualifier can be understood as the access path to the  | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
790  | 
named entity while passing through some nested block-structure,  | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
791  | 
although our free-form long names do not really enforce any strict  | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
792  | 
discipline.  | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
793  | 
|
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
794  | 
  For example, an item named ``@{text "A.b.c"}'' may be understood as
 | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
795  | 
  a local entity @{text "c"}, within a local structure @{text "b"},
 | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
796  | 
  within a global structure @{text "A"}.  In practice, long names
 | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
797  | 
usually represent 1--3 levels of qualification. User ML code should  | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
798  | 
not make any assumptions about the particular structure of long  | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
799  | 
names!  | 
| 20437 | 800  | 
|
| 20476 | 801  | 
The empty name is commonly used as an indication of unnamed  | 
| 
34927
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
802  | 
entities, or entities that are not entered into the corresponding  | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
803  | 
name space, whenever this makes any sense. The basic operations on  | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
804  | 
long names map empty names again to empty names.  | 
| 20437 | 805  | 
*}  | 
806  | 
||
| 20476 | 807  | 
text %mlref {*
 | 
808  | 
  \begin{mldecls}
 | 
|
| 30365 | 809  | 
  @{index_ML Long_Name.base_name: "string -> string"} \\
 | 
810  | 
  @{index_ML Long_Name.qualifier: "string -> string"} \\
 | 
|
811  | 
  @{index_ML Long_Name.append: "string -> string -> string"} \\
 | 
|
812  | 
  @{index_ML Long_Name.implode: "string list -> string"} \\
 | 
|
813  | 
  @{index_ML Long_Name.explode: "string -> string list"} \\
 | 
|
| 20547 | 814  | 
  \end{mldecls}
 | 
| 
34927
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
815  | 
|
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
816  | 
  \begin{description}
 | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
817  | 
|
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
818  | 
  \item @{ML Long_Name.base_name}~@{text "name"} returns the base name
 | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
819  | 
of a long name.  | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
820  | 
|
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
821  | 
  \item @{ML Long_Name.qualifier}~@{text "name"} returns the qualifier
 | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
822  | 
of a long name.  | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
823  | 
|
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
824  | 
  \item @{ML Long_Name.append}~@{text "name\<^isub>1 name\<^isub>2"} appends two long
 | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
825  | 
names.  | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
826  | 
|
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
827  | 
  \item @{ML Long_Name.implode}~@{text "names"} and @{ML
 | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
828  | 
  Long_Name.explode}~@{text "name"} convert between the packed string
 | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
829  | 
representation and the explicit list form of long names.  | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
830  | 
|
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
831  | 
  \end{description}
 | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
832  | 
*}  | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
833  | 
|
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
834  | 
|
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
835  | 
subsection {* Name spaces \label{sec:name-space} *}
 | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
836  | 
|
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
837  | 
text {* A @{text "name space"} manages a collection of long names,
 | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
838  | 
together with a mapping between partially qualified external names  | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
839  | 
and fully qualified internal names (in both directions). Note that  | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
840  | 
  the corresponding @{text "intern"} and @{text "extern"} operations
 | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
841  | 
  are mostly used for parsing and printing only!  The @{text
 | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
842  | 
"declare"} operation augments a name space according to the accesses  | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
843  | 
determined by a given binding, and a naming policy from the context.  | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
844  | 
|
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
845  | 
  \medskip A @{text "binding"} specifies details about the prospective
 | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
846  | 
long name of a newly introduced formal entity. It consists of a  | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
847  | 
base name, prefixes for qualification (separate ones for system  | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
848  | 
infrastructure and user-space mechanisms), a slot for the original  | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
849  | 
source position, and some additional flags.  | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
850  | 
|
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
851  | 
  \medskip A @{text "naming"} provides some additional details for
 | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
852  | 
producing a long name from a binding. Normally, the naming is  | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
853  | 
  implicit in the theory or proof context.  The @{text "full"}
 | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
854  | 
operation (and its variants for different context types) produces a  | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
855  | 
fully qualified internal name to be entered into a name space. The  | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
856  | 
main equation of this ``chemical reaction'' when binding new  | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
857  | 
entities in a context is as follows:  | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
858  | 
|
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
859  | 
\smallskip  | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
860  | 
  \begin{tabular}{l}
 | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
861  | 
  @{text "binding + naming \<longrightarrow> long name + name space accesses"}
 | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
862  | 
  \end{tabular}
 | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
863  | 
\smallskip  | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
864  | 
|
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
865  | 
\medskip As a general principle, there is a separate name space for  | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
866  | 
each kind of formal entity, e.g.\ fact, logical constant, type  | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
867  | 
constructor, type class. It is usually clear from the occurrence in  | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
868  | 
concrete syntax (or from the scope) which kind of entity a name  | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
869  | 
  refers to.  For example, the very same name @{text "c"} may be used
 | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
870  | 
uniformly for a constant, type constructor, and type class.  | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
871  | 
|
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
872  | 
There are common schemes to name derived entities systematically  | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
873  | 
according to the name of the main logical entity involved, e.g.\  | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
874  | 
  fact @{text "c.intro"} for a canonical introduction rule related to
 | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
875  | 
  constant @{text "c"}.  This technique of mapping names from one
 | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
876  | 
space into another requires some care in order to avoid conflicts.  | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
877  | 
In particular, theorem names derived from a type constructor or type  | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
878  | 
class are better suffixed in addition to the usual qualification,  | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
879  | 
  e.g.\ @{text "c_type.intro"} and @{text "c_class.intro"} for
 | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
880  | 
  theorems related to type @{text "c"} and class @{text "c"},
 | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
881  | 
respectively.  | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
882  | 
*}  | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
883  | 
|
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
884  | 
text %mlref {*
 | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
885  | 
  \begin{mldecls}
 | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
886  | 
  @{index_ML_type binding} \\
 | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
887  | 
  @{index_ML Binding.empty: binding} \\
 | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
888  | 
  @{index_ML Binding.name: "string -> binding"} \\
 | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
889  | 
  @{index_ML Binding.qualify: "bool -> string -> binding -> binding"} \\
 | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
890  | 
  @{index_ML Binding.prefix: "bool -> string -> binding -> binding"} \\
 | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
891  | 
  @{index_ML Binding.conceal: "binding -> binding"} \\
 | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
892  | 
  @{index_ML Binding.str_of: "binding -> string"} \\
 | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
893  | 
  \end{mldecls}
 | 
| 20547 | 894  | 
  \begin{mldecls}
 | 
| 33174 | 895  | 
  @{index_ML_type Name_Space.naming} \\
 | 
896  | 
  @{index_ML Name_Space.default_naming: Name_Space.naming} \\
 | 
|
897  | 
  @{index_ML Name_Space.add_path: "string -> Name_Space.naming -> Name_Space.naming"} \\
 | 
|
898  | 
  @{index_ML Name_Space.full_name: "Name_Space.naming -> binding -> string"} \\
 | 
|
| 20547 | 899  | 
  \end{mldecls}
 | 
900  | 
  \begin{mldecls}
 | 
|
| 33174 | 901  | 
  @{index_ML_type Name_Space.T} \\
 | 
902  | 
  @{index_ML Name_Space.empty: "string -> Name_Space.T"} \\
 | 
|
903  | 
  @{index_ML Name_Space.merge: "Name_Space.T * Name_Space.T -> Name_Space.T"} \\
 | 
|
904  | 
  @{index_ML Name_Space.declare: "bool -> Name_Space.naming -> binding -> Name_Space.T ->
 | 
|
905  | 
string * Name_Space.T"} \\  | 
|
906  | 
  @{index_ML Name_Space.intern: "Name_Space.T -> string -> string"} \\
 | 
|
907  | 
  @{index_ML Name_Space.extern: "Name_Space.T -> string -> string"} \\
 | 
|
| 
34927
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
908  | 
  @{index_ML Name_Space.is_concealed: "Name_Space.T -> string -> bool"}
 | 
| 20476 | 909  | 
  \end{mldecls}
 | 
| 20437 | 910  | 
|
| 20476 | 911  | 
  \begin{description}
 | 
912  | 
||
| 
34927
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
913  | 
  \item @{ML_type binding} represents the abstract concept of name
 | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
914  | 
bindings.  | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
915  | 
|
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
916  | 
  \item @{ML Binding.empty} is the empty binding.
 | 
| 20476 | 917  | 
|
| 
34927
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
918  | 
  \item @{ML Binding.name}~@{text "name"} produces a binding with base
 | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
919  | 
  name @{text "name"}.
 | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
920  | 
|
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
921  | 
  \item @{ML Binding.qualify}~@{text "mandatory name binding"}
 | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
922  | 
  prefixes qualifier @{text "name"} to @{text "binding"}.  The @{text
 | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
923  | 
"mandatory"} flag tells if this name component always needs to be  | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
924  | 
  given in name space accesses --- this is mostly @{text "false"} in
 | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
925  | 
practice. Note that this part of qualification is typically used in  | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
926  | 
derived specification mechanisms.  | 
| 20437 | 927  | 
|
| 
34927
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
928  | 
  \item @{ML Binding.prefix} is similar to @{ML Binding.qualify}, but
 | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
929  | 
affects the system prefix. This part of extra qualification is  | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
930  | 
typically used in the infrastructure for modular specifications,  | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
931  | 
  notably ``local theory targets'' (see also \chref{ch:local-theory}).
 | 
| 20437 | 932  | 
|
| 
34927
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
933  | 
  \item @{ML Binding.conceal}~@{text "binding"} indicates that the
 | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
934  | 
binding shall refer to an entity that serves foundational purposes  | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
935  | 
only. This flag helps to mark implementation details of  | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
936  | 
specification mechanism etc. Other tools should not depend on the  | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
937  | 
  particulars of concealed entities (cf.\ @{ML
 | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
938  | 
Name_Space.is_concealed}).  | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
939  | 
|
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
940  | 
  \item @{ML Binding.str_of}~@{text "binding"} produces a string
 | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
941  | 
representation for human-readable output, together with some formal  | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
942  | 
markup that might get used in GUI front-ends, for example.  | 
| 20476 | 943  | 
|
| 33174 | 944  | 
  \item @{ML_type Name_Space.naming} represents the abstract concept of
 | 
| 20476 | 945  | 
a naming policy.  | 
| 20437 | 946  | 
|
| 33174 | 947  | 
  \item @{ML Name_Space.default_naming} is the default naming policy.
 | 
| 20476 | 948  | 
In a theory context, this is usually augmented by a path prefix  | 
949  | 
consisting of the theory name.  | 
|
950  | 
||
| 33174 | 951  | 
  \item @{ML Name_Space.add_path}~@{text "path naming"} augments the
 | 
| 20488 | 952  | 
naming policy by extending its path component.  | 
| 20437 | 953  | 
|
| 33174 | 954  | 
  \item @{ML Name_Space.full_name}~@{text "naming binding"} turns a
 | 
| 
30281
 
9ad15d8ed311
renamed NameSpace.base to NameSpace.base_name (in accordance with "full_name");
 
wenzelm 
parents: 
30272 
diff
changeset
 | 
955  | 
name binding (usually a basic name) into the fully qualified  | 
| 29008 | 956  | 
internal name, according to the given naming policy.  | 
| 20476 | 957  | 
|
| 33174 | 958  | 
  \item @{ML_type Name_Space.T} represents name spaces.
 | 
| 20476 | 959  | 
|
| 33174 | 960  | 
  \item @{ML Name_Space.empty}~@{text "kind"} and @{ML Name_Space.merge}~@{text
 | 
| 20488 | 961  | 
"(space\<^isub>1, space\<^isub>2)"} are the canonical operations for  | 
962  | 
maintaining name spaces according to theory data management  | 
|
| 33174 | 963  | 
  (\secref{sec:context-data}); @{text "kind"} is a formal comment
 | 
964  | 
to characterize the purpose of a name space.  | 
|
| 20437 | 965  | 
|
| 33174 | 966  | 
  \item @{ML Name_Space.declare}~@{text "strict naming bindings
 | 
967  | 
space"} enters a name binding as fully qualified internal name into  | 
|
968  | 
the name space, with external accesses determined by the naming  | 
|
969  | 
policy.  | 
|
| 20476 | 970  | 
|
| 33174 | 971  | 
  \item @{ML Name_Space.intern}~@{text "space name"} internalizes a
 | 
| 20476 | 972  | 
(partially qualified) external name.  | 
| 20437 | 973  | 
|
| 20488 | 974  | 
This operation is mostly for parsing! Note that fully qualified  | 
| 20476 | 975  | 
  names stemming from declarations are produced via @{ML
 | 
| 33174 | 976  | 
  "Name_Space.full_name"} and @{ML "Name_Space.declare"}
 | 
| 29008 | 977  | 
  (or their derivatives for @{ML_type theory} and
 | 
| 20488 | 978  | 
  @{ML_type Proof.context}).
 | 
| 20437 | 979  | 
|
| 33174 | 980  | 
  \item @{ML Name_Space.extern}~@{text "space name"} externalizes a
 | 
| 20476 | 981  | 
(fully qualified) internal name.  | 
982  | 
||
| 
30281
 
9ad15d8ed311
renamed NameSpace.base to NameSpace.base_name (in accordance with "full_name");
 
wenzelm 
parents: 
30272 
diff
changeset
 | 
983  | 
This operation is mostly for printing! User code should not rely on  | 
| 
 
9ad15d8ed311
renamed NameSpace.base to NameSpace.base_name (in accordance with "full_name");
 
wenzelm 
parents: 
30272 
diff
changeset
 | 
984  | 
the precise result too much.  | 
| 20476 | 985  | 
|
| 
34927
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
986  | 
  \item @{ML Name_Space.is_concealed}~@{text "space name"} indicates
 | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
987  | 
  whether @{text "name"} refers to a strictly private entity that
 | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
988  | 
other tools are supposed to ignore!  | 
| 
 
c4c02ac736a6
more details on long names, binding/naming, name space;
 
wenzelm 
parents: 
34926 
diff
changeset
 | 
989  | 
|
| 20476 | 990  | 
  \end{description}
 | 
991  | 
*}  | 
|
| 30272 | 992  | 
|
| 18537 | 993  | 
end  |