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