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