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