author | wenzelm |
Tue, 29 Aug 2006 18:56:11 +0200 | |
changeset 20430 | fd646e926983 |
parent 20429 | 116255c9209b |
child 20437 | 0eb5e30fd620 |
permissions | -rw-r--r-- |
18537 | 1 |
|
2 |
(* $Id$ *) |
|
3 |
||
4 |
theory prelim imports base begin |
|
5 |
||
6 |
chapter {* Preliminaries *} |
|
7 |
||
8 |
section {* Named entities *} |
|
9 |
||
10 |
text {* Named entities of different kinds (logical constant, type, |
|
11 |
type class, theorem, method etc.) live in separate name spaces. It is |
|
12 |
usually clear from the occurrence of a name which kind of entity it |
|
13 |
refers to. For example, proof method @{text "foo"} vs.\ theorem |
|
14 |
@{text "foo"} vs.\ logical constant @{text "foo"} are easily |
|
15 |
distinguished by means of the syntactic context. A notable exception |
|
16 |
are logical identifiers within a term (\secref{sec:terms}): constants, |
|
17 |
fixed variables, and bound variables all share the same identifier |
|
18 |
syntax, but are distinguished by their scope. |
|
19 |
||
20 |
Each name space is organized as a collection of \emph{qualified |
|
21 |
names}, which consist of a sequence of basic name components separated |
|
22 |
by dots: @{text "Bar.bar.foo"}, @{text "Bar.foo"}, and @{text "foo"} |
|
23 |
are examples for valid qualified names. Name components are |
|
24 |
subdivided into \emph{symbols}, which constitute the smallest textual |
|
25 |
unit in Isabelle --- raw characters are normally not encountered |
|
26 |
directly. *} |
|
27 |
||
28 |
||
29 |
subsection {* Strings of symbols *} |
|
30 |
||
31 |
text {* Isabelle strings consist of a sequence of |
|
32 |
symbols\glossary{Symbol}{The smalles unit of text in Isabelle, |
|
33 |
subsumes plain ASCII characters as well as an infinite collection of |
|
34 |
named symbols (for greek, math etc.).}, which are either packed as an |
|
35 |
actual @{text "string"}, or represented as a list. Each symbol is in |
|
36 |
itself a small string of the following form: |
|
37 |
||
38 |
\begin{enumerate} |
|
39 |
||
40 |
\item either a singleton ASCII character ``@{text "c"}'' (with |
|
41 |
character code 0--127), for example ``\verb,a,'', |
|
42 |
||
43 |
\item or a regular symbol ``\verb,\,\verb,<,@{text "ident"}\verb,>,'', |
|
44 |
for example ``\verb,\,\verb,<alpha>,'', |
|
45 |
||
46 |
\item or a control symbol ``\verb,\,\verb,<^,@{text |
|
47 |
"ident"}\verb,>,'', for example ``\verb,\,\verb,<^bold>,'', |
|
48 |
||
49 |
\item or a raw control symbol ``\verb,\,\verb,<^raw:,@{text |
|
20205
7b2958d3d575
raw symbols: disallow dot to avoid confusion in NameSpace.unpack;
wenzelm
parents:
18554
diff
changeset
|
50 |
"\<dots>"}\verb,>,'' where ``@{text "\<dots>"}'' refers to any |
7b2958d3d575
raw symbols: disallow dot to avoid confusion in NameSpace.unpack;
wenzelm
parents:
18554
diff
changeset
|
51 |
printable ASCII character (excluding ``\verb,.,'' and ``\verb,>,'') or |
20214 | 52 |
non-ASCII character, for example ``\verb,\,\verb,<^raw:$\sum_{i = 1}^n$>,'', |
18537 | 53 |
|
54 |
\item or a numbered raw control symbol ``\verb,\,\verb,<^raw,@{text |
|
55 |
"nnn"}\verb,>, where @{text "nnn"} are digits, for example |
|
56 |
``\verb,\,\verb,<^raw42>,''. |
|
57 |
||
58 |
\end{enumerate} |
|
59 |
||
60 |
The @{text "ident"} syntax for symbol names is @{text "letter (letter |
|
61 |
| digit)\<^sup>*"}, where @{text "letter = A..Za..Z"} and @{text |
|
62 |
"digit = 0..9"}. There are infinitely many regular symbols and |
|
63 |
control symbols available, but a certain collection of standard |
|
64 |
symbols is treated specifically. For example, |
|
65 |
``\verb,\,\verb,<alpha>,'' is classified as a (non-ASCII) letter, |
|
66 |
which means it may occur within regular Isabelle identifier syntax. |
|
67 |
||
68 |
Output of symbols depends on the print mode (\secref{sec:print-mode}). |
|
69 |
For example, the standard {\LaTeX} setup of the Isabelle document |
|
70 |
preparation system would present ``\verb,\,\verb,<alpha>,'' as @{text |
|
71 |
"\<alpha>"}, and ``\verb,\,\verb,<^bold>,\verb,\,\verb,<alpha>,'' as @{text |
|
72 |
"\<^bold>\<alpha>"}. |
|
73 |
||
74 |
\medskip It is important to note that the character set underlying |
|
75 |
Isabelle symbols is plain 7-bit ASCII. Since 8-bit characters are |
|
76 |
passed through transparently, Isabelle may easily process actual |
|
77 |
Unicode/UCS data (using the well-known UTF-8 encoding, for example). |
|
78 |
Unicode provides its own collection of mathematical symbols, but there |
|
79 |
is presently no link to Isabelle's named ones; both kinds of symbols |
|
80 |
coexist independently. *} |
|
81 |
||
82 |
text %mlref {* |
|
83 |
\begin{mldecls} |
|
84 |
@{index_ML_type "Symbol.symbol"} \\ |
|
85 |
@{index_ML Symbol.explode: "string -> Symbol.symbol list"} \\ |
|
86 |
@{index_ML Symbol.is_letter: "Symbol.symbol -> bool"} \\ |
|
87 |
@{index_ML Symbol.is_digit: "Symbol.symbol -> bool"} \\ |
|
88 |
@{index_ML Symbol.is_quasi: "Symbol.symbol -> bool"} \\ |
|
89 |
@{index_ML Symbol.is_blank: "Symbol.symbol -> bool"} \\ |
|
90 |
@{index_ML_type "Symbol.sym"} \\ |
|
91 |
@{index_ML Symbol.decode: "Symbol.symbol -> Symbol.sym"} \\ |
|
92 |
\end{mldecls} |
|
93 |
||
94 |
\begin{description} |
|
95 |
||
96 |
\item @{ML_type "Symbol.symbol"} represents Isabelle symbols; this type |
|
97 |
is merely an alias for @{ML_type "string"}, but emphasizes the |
|
98 |
specific format encountered here. |
|
99 |
||
100 |
\item @{ML "Symbol.explode"}~@{text "s"} produces an actual symbol |
|
101 |
list from the packed form usually encountered as user input. This |
|
102 |
function replaces @{ML "String.explode"} for virtually all purposes |
|
103 |
of manipulating text in Isabelle! Plain @{text "implode"} may be |
|
104 |
used for the reverse operation. |
|
105 |
||
106 |
\item @{ML "Symbol.is_letter"}, @{ML "Symbol.is_digit"}, @{ML |
|
107 |
"Symbol.is_quasi"}, @{ML "Symbol.is_blank"} classify certain symbols |
|
108 |
(both ASCII and several named ones) according to fixed syntactic |
|
109 |
convections of Isabelle, e.g.\ see \cite{isabelle-isar-ref}. |
|
110 |
||
111 |
\item @{ML_type "Symbol.sym"} is a concrete datatype that represents |
|
112 |
the different kinds of symbols explicitly as @{ML "Symbol.Char"}, |
|
113 |
@{ML "Symbol.Sym"}, @{ML "Symbol.Ctrl"}, or @{ML "Symbol.Raw"}. |
|
114 |
||
115 |
\item @{ML "Symbol.decode"} converts the string representation of a |
|
116 |
symbol into the explicit datatype version. |
|
117 |
||
118 |
\end{description} |
|
119 |
*} |
|
120 |
||
121 |
||
20429 | 122 |
subsection {* Simple names *} |
123 |
||
124 |
text FIXME |
|
125 |
||
126 |
||
18537 | 127 |
subsection {* Qualified names and name spaces *} |
128 |
||
129 |
text %FIXME {* Qualified names are constructed according to implicit naming |
|
130 |
principles of the present context. |
|
131 |
||
132 |
||
133 |
The last component is called \emph{base name}; the remaining prefix of |
|
134 |
qualification may be empty. |
|
135 |
||
136 |
Some practical conventions help to organize named entities more |
|
137 |
systematically: |
|
138 |
||
139 |
\begin{itemize} |
|
140 |
||
141 |
\item Names are qualified first by the theory name, second by an |
|
142 |
optional ``structure''. For example, a constant @{text "c"} declared |
|
143 |
as part of a certain structure @{text "b"} (say a type definition) in |
|
144 |
theory @{text "A"} will be named @{text "A.b.c"} internally. |
|
145 |
||
146 |
\item |
|
147 |
||
148 |
\item |
|
149 |
||
150 |
\item |
|
151 |
||
152 |
\item |
|
153 |
||
154 |
\end{itemize} |
|
155 |
||
156 |
Names of different kinds of entities are basically independent, but |
|
157 |
some practical naming conventions relate them to each other. For |
|
158 |
example, a constant @{text "foo"} may be accompanied with theorems |
|
159 |
@{text "foo.intro"}, @{text "foo.elim"}, @{text "foo.simps"} etc. The |
|
160 |
same may happen for a type @{text "foo"}, which is then apt to cause |
|
161 |
clashes in the theorem name space! To avoid this, we occasionally |
|
162 |
follow an additional convention of suffixes that determine the |
|
163 |
original kind of entity that a name has been derived. For example, |
|
164 |
constant @{text "foo"} is associated with theorem @{text "foo.intro"}, |
|
165 |
type @{text "foo"} with theorem @{text "foo_type.intro"}, and type |
|
166 |
class @{text "foo"} with @{text "foo_class.intro"}. |
|
167 |
||
168 |
*} |
|
169 |
||
170 |
||
171 |
section {* Structured output *} |
|
172 |
||
173 |
subsection {* Pretty printing *} |
|
174 |
||
175 |
text FIXME |
|
176 |
||
177 |
subsection {* Output channels *} |
|
178 |
||
179 |
text FIXME |
|
180 |
||
181 |
subsection {* Print modes *} |
|
182 |
||
183 |
text FIXME |
|
184 |
||
185 |
||
20429 | 186 |
section {* Contexts \label{sec:context} *} |
18537 | 187 |
|
20429 | 188 |
text {* |
189 |
A logical context represents the background that is taken for |
|
190 |
granted when formulating statements and composing proofs. It acts |
|
191 |
as a medium to produce formal content, depending on earlier material |
|
192 |
(declarations, results etc.). |
|
18537 | 193 |
|
20429 | 194 |
In particular, derivations within the primitive Pure logic can be |
195 |
described as a judgment @{text "\<Gamma> \<turnstile>\<^sub>\<Theta> \<phi>"}, meaning that a |
|
196 |
proposition @{text "\<phi>"} is derivable from hypotheses @{text "\<Gamma>"} |
|
197 |
within the theory @{text "\<Theta>"}. There are logical reasons for |
|
198 |
keeping @{text "\<Theta>"} and @{text "\<Gamma>"} separate: theories support type |
|
199 |
constructors and schematic polymorphism of constants and axioms, |
|
200 |
while the inner calculus of @{text "\<Gamma> \<turnstile> \<phi>"} is limited to Simple |
|
201 |
Type Theory (with fixed type variables in the assumptions). |
|
18537 | 202 |
|
20429 | 203 |
\medskip Contexts and derivations are linked by the following key |
204 |
principles: |
|
205 |
||
206 |
\begin{itemize} |
|
207 |
||
208 |
\item Transfer: monotonicity of derivations admits results to be |
|
209 |
transferred into a larger context, i.e.\ @{text "\<Gamma> \<turnstile>\<^sub>\<Theta> \<phi>"} |
|
210 |
implies @{text "\<Gamma>' \<turnstile>\<^sub>\<Theta>\<^sub>' \<phi>"} for contexts @{text "\<Theta>' \<supseteq> |
|
211 |
\<Theta>"} and @{text "\<Gamma>' \<supseteq> \<Gamma>"}. |
|
18537 | 212 |
|
20429 | 213 |
\item Export: discharge of hypotheses admits results to be exported |
214 |
into a smaller context, i.e.\ @{text "\<Gamma>' \<turnstile>\<^sub>\<Theta> \<phi>"} implies |
|
215 |
@{text "\<Gamma> \<turnstile>\<^sub>\<Theta> \<Delta> \<Longrightarrow> \<phi>"} where @{text "\<Gamma>' \<supseteq> \<Gamma>"} and @{text "\<Delta> = |
|
216 |
\<Gamma>' - \<Gamma>"}. Note that @{text "\<Theta>"} remains unchanged here, only the |
|
217 |
@{text "\<Gamma>"} part is affected. |
|
218 |
||
219 |
\end{itemize} |
|
18537 | 220 |
|
20429 | 221 |
\medskip Isabelle/Isar provides two different notions of abstract |
222 |
containers called \emph{theory context} and \emph{proof context}, |
|
223 |
respectively. These model the main characteristics of the primitive |
|
224 |
@{text "\<Theta>"} and @{text "\<Gamma>"} above, without subscribing to any |
|
225 |
particular kind of content yet. Instead, contexts merely impose a |
|
226 |
certain policy of managing arbitrary \emph{context data}. The |
|
227 |
system provides strongly typed mechanisms to declare new kinds of |
|
228 |
data at compile time. |
|
18537 | 229 |
|
20429 | 230 |
Thus the internal bootstrap process of Isabelle/Pure eventually |
231 |
reaches a stage where certain data slots provide the logical content |
|
232 |
of @{text "\<Theta>"} and @{text "\<Gamma>"} sketched above, but this does not |
|
233 |
stop there! Various additional data slots support all kinds of |
|
234 |
mechanisms that are not necessarily part of the core logic. |
|
18537 | 235 |
|
20429 | 236 |
For example, there would be data for canonical introduction and |
237 |
elimination rules for arbitrary operators (depending on the |
|
238 |
object-logic and application), which enables users to perform |
|
239 |
standard proof steps implicitly (cf.\ the @{text "rule"} method). |
|
18537 | 240 |
|
20429 | 241 |
Isabelle is able to bring forth more and more concepts successively. |
242 |
In particular, an object-logic like Isabelle/HOL continues the |
|
243 |
Isabelle/Pure setup by adding specific components for automated |
|
244 |
reasoning (classical reasoner, tableau prover, structured induction |
|
245 |
etc.) and derived specification mechanisms (inductive predicates, |
|
246 |
recursive functions etc.). All of this is based on the generic data |
|
247 |
management by theory and proof contexts. |
|
18537 | 248 |
*} |
249 |
||
250 |
||
251 |
subsection {* Theory context \label{sec:context-theory} *} |
|
252 |
||
20429 | 253 |
text {* |
254 |
Each theory is explicitly named and holds a unique identifier. |
|
255 |
There is a separate \emph{theory reference} for pointing backwards |
|
256 |
to the enclosing theory context of derived entities. Theories are |
|
257 |
related by a (nominal) sub-theory relation, which corresponds to the |
|
258 |
canonical dependency graph: each theory is derived from a certain |
|
259 |
sub-graph of ancestor theories. The @{text "merge"} of two theories |
|
260 |
refers to the least upper bound, which actually degenerates into |
|
261 |
absorption of one theory into the other, due to the nominal |
|
262 |
sub-theory relation this. |
|
18537 | 263 |
|
20429 | 264 |
The @{text "begin"} operation starts a new theory by importing |
265 |
several parent theories and entering a special @{text "draft"} mode, |
|
266 |
which is sustained until the final @{text "end"} operation. A draft |
|
267 |
mode theory acts like a linear type, where updates invalidate |
|
268 |
earlier drafts, but theory reference values will be propagated |
|
269 |
automatically. Thus derived entities that ``belong'' to a draft |
|
270 |
might be transferred spontaneously to a larger context. An |
|
271 |
invalidated draft is called ``stale''. The @{text "copy"} operation |
|
272 |
produces an auxiliary version with the same data content, but is |
|
273 |
unrelated to the original: updates of the copy do not affect the |
|
274 |
original, neither does the sub-theory relation hold. |
|
275 |
||
276 |
The example below shows a theory graph derived from @{text "Pure"}. |
|
277 |
Theory @{text "Length"} imports @{text "Nat"} and @{text "List"}. |
|
278 |
The linear draft mode is enabled during the ``@{text "\<dots>"}'' stage of |
|
279 |
the theory body. |
|
280 |
||
281 |
\bigskip |
|
282 |
\begin{tabular}{rcccl} |
|
18537 | 283 |
& & $Pure$ \\ |
284 |
& & $\downarrow$ \\ |
|
285 |
& & $FOL$ \\ |
|
286 |
& $\swarrow$ & & $\searrow$ & \\ |
|
287 |
$Nat$ & & & & $List$ \\ |
|
288 |
& $\searrow$ & & $\swarrow$ \\ |
|
289 |
& & $Length$ \\ |
|
290 |
& & \multicolumn{3}{l}{~~$\isarkeyword{imports}$} \\ |
|
291 |
& & \multicolumn{3}{l}{~~$\isarkeyword{begin}$} \\ |
|
292 |
& & $\vdots$~~ \\ |
|
293 |
& & \multicolumn{3}{l}{~~$\isarkeyword{end}$} \\ |
|
20429 | 294 |
\end{tabular} |
295 |
||
296 |
\medskip In practice, derived theory operations mostly operate |
|
297 |
drafts, namely the body of the current portion of theory that the |
|
298 |
user happens to be composing. |
|
18537 | 299 |
|
20429 | 300 |
\medskip There is also a builtin theory history mechanism that amends for |
301 |
the destructive behaviour of theory drafts. The @{text |
|
302 |
"checkpoint"} operation produces an intermediate stepping stone that |
|
303 |
survives the next update unscathed: both the original and the |
|
304 |
changed theory remain valid and are related by the sub-theory |
|
305 |
relation. This recovering of pure theory values comes at the cost |
|
306 |
of extra internal bookeeping. The cumulative effect of |
|
307 |
checkpointing is purged by the @{text "finish"} operation. |
|
18537 | 308 |
|
20429 | 309 |
History operations are usually managed by the system, e.g.\ notably |
310 |
in the Isar transaction loop. |
|
18537 | 311 |
|
20429 | 312 |
\medskip |
313 |
FIXME theory data |
|
18537 | 314 |
*} |
315 |
||
20430 | 316 |
text %mlref {* |
317 |
*} |
|
318 |
||
18537 | 319 |
|
320 |
subsection {* Proof context \label{sec:context-proof} *} |
|
321 |
||
322 |
text {* |
|
20429 | 323 |
A proof context is an arbitrary container that is initialized from a |
324 |
given theory. The result contains a back-reference to the theory it |
|
325 |
belongs to, together with pure data. No further bookkeeping is |
|
326 |
required here, thanks to the lack of destructive features. |
|
327 |
||
328 |
There is no restriction on producing proof contexts, although the |
|
329 |
usual discipline is to follow block structure as a mental model: a |
|
330 |
given context is extended consecutively, results are exported back |
|
331 |
into the original context. In particular, the concept of Isar proof |
|
332 |
state models block-structured reasoning explicitly, using a stack of |
|
333 |
proof contexts. |
|
334 |
||
335 |
Due to the lack of identification and back-referencing, entities |
|
336 |
derived in a proof context need to record inherent logical |
|
337 |
requirements explicitly. For example, hypotheses used in a |
|
338 |
derivation will be recorded separately within the sequent @{text "\<Gamma> |
|
339 |
\<turnstile> \<phi>"}, just to make double sure. Results could leak into an alien |
|
340 |
proof context do to programming errors, but Isabelle/Isar |
|
341 |
occasionally includes extra validity checks at the end of a |
|
342 |
sub-proof. |
|
343 |
||
344 |
\medskip |
|
345 |
FIXME proof data |
|
18537 | 346 |
|
347 |
\glossary{Proof context}{The static context of a structured proof, |
|
348 |
acts like a local ``theory'' of the current portion of Isar proof |
|
349 |
text, generalizes the idea of local hypotheses @{text "\<Gamma>"} in |
|
350 |
judgments @{text "\<Gamma> \<turnstile> \<phi>"} of natural deduction calculi. There is a |
|
351 |
generic notion of introducing and discharging hypotheses. Arbritrary |
|
352 |
auxiliary context data may be adjoined.} |
|
353 |
||
354 |
*} |
|
355 |
||
20430 | 356 |
text %mlref {* FIXME *} |
357 |
||
20429 | 358 |
|
359 |
subsection {* Generic contexts *} |
|
360 |
||
20430 | 361 |
text FIXME |
362 |
||
363 |
text %mlref {* FIXME *} |
|
364 |
||
18537 | 365 |
end |