author | wenzelm |
Sun, 13 Apr 1997 19:10:27 +0200 | |
changeset 2941 | b228efa26ea3 |
parent 1904 | 4f77c2fd8f3d |
child 3103 | 98af809fee46 |
permissions | -rw-r--r-- |
105 | 1 |
%% $Id$ |
284 | 2 |
\part{Advanced Methods} |
105 | 3 |
Before continuing, it might be wise to try some of your own examples in |
4 |
Isabelle, reinforcing your knowledge of the basic functions. |
|
5 |
||
6 |
Look through {\em Isabelle's Object-Logics\/} and try proving some simple |
|
7 |
theorems. You probably should begin with first-order logic ({\tt FOL} |
|
8 |
or~{\tt LK}). Try working some of the examples provided, and others from |
|
9 |
the literature. Set theory~({\tt ZF}) and Constructive Type Theory~({\tt |
|
10 |
CTT}) form a richer world for mathematical reasoning and, again, many |
|
11 |
examples are in the literature. Higher-order logic~({\tt HOL}) is |
|
331 | 12 |
Isabelle's most sophisticated logic because its types and functions are |
13 |
identified with those of the meta-logic. |
|
105 | 14 |
|
15 |
Choose a logic that you already understand. Isabelle is a proof |
|
16 |
tool, not a teaching tool; if you do not know how to do a particular proof |
|
17 |
on paper, then you certainly will not be able to do it on the machine. |
|
18 |
Even experienced users plan large proofs on paper. |
|
19 |
||
20 |
We have covered only the bare essentials of Isabelle, but enough to perform |
|
21 |
substantial proofs. By occasionally dipping into the {\em Reference |
|
22 |
Manual}, you can learn additional tactics, subgoal commands and tacticals. |
|
23 |
||
24 |
||
25 |
\section{Deriving rules in Isabelle} |
|
26 |
\index{rules!derived} |
|
27 |
A mathematical development goes through a progression of stages. Each |
|
28 |
stage defines some concepts and derives rules about them. We shall see how |
|
29 |
to derive rules, perhaps involving definitions, using Isabelle. The |
|
310 | 30 |
following section will explain how to declare types, constants, rules and |
105 | 31 |
definitions. |
32 |
||
33 |
||
296 | 34 |
\subsection{Deriving a rule using tactics and meta-level assumptions} |
35 |
\label{deriving-example} |
|
310 | 36 |
\index{examples!of deriving rules}\index{assumptions!of main goal} |
296 | 37 |
|
307 | 38 |
The subgoal module supports the derivation of rules, as discussed in |
39 |
\S\ref{deriving}. The \ttindex{goal} command, when supplied a goal of the |
|
40 |
form $\List{\theta@1; \ldots; \theta@k} \Imp \phi$, creates $\phi\Imp\phi$ |
|
41 |
as the initial proof state and returns a list consisting of the theorems |
|
42 |
${\theta@i\;[\theta@i]}$, for $i=1$, \ldots,~$k$. These meta-assumptions |
|
310 | 43 |
are also recorded internally, allowing {\tt result} to discharge them |
307 | 44 |
in the original order. |
105 | 45 |
|
307 | 46 |
Let us derive $\conj$ elimination using Isabelle. |
310 | 47 |
Until now, calling {\tt goal} has returned an empty list, which we have |
105 | 48 |
thrown away. In this example, the list contains the two premises of the |
49 |
rule. We bind them to the \ML\ identifiers {\tt major} and {\tt |
|
50 |
minor}:\footnote{Some ML compilers will print a message such as {\em |
|
51 |
binding not exhaustive}. This warns that {\tt goal} must return a |
|
52 |
2-element list. Otherwise, the pattern-match will fail; ML will |
|
310 | 53 |
raise exception \xdx{Match}.} |
105 | 54 |
\begin{ttbox} |
55 |
val [major,minor] = goal FOL.thy |
|
56 |
"[| P&Q; [| P; Q |] ==> R |] ==> R"; |
|
57 |
{\out Level 0} |
|
58 |
{\out R} |
|
59 |
{\out 1. R} |
|
60 |
{\out val major = "P & Q [P & Q]" : thm} |
|
61 |
{\out val minor = "[| P; Q |] ==> R [[| P; Q |] ==> R]" : thm} |
|
62 |
\end{ttbox} |
|
63 |
Look at the minor premise, recalling that meta-level assumptions are |
|
64 |
shown in brackets. Using {\tt minor}, we reduce $R$ to the subgoals |
|
65 |
$P$ and~$Q$: |
|
66 |
\begin{ttbox} |
|
67 |
by (resolve_tac [minor] 1); |
|
68 |
{\out Level 1} |
|
69 |
{\out R} |
|
70 |
{\out 1. P} |
|
71 |
{\out 2. Q} |
|
72 |
\end{ttbox} |
|
73 |
Deviating from~\S\ref{deriving}, we apply $({\conj}E1)$ forwards from the |
|
74 |
assumption $P\conj Q$ to obtain the theorem~$P\;[P\conj Q]$. |
|
75 |
\begin{ttbox} |
|
76 |
major RS conjunct1; |
|
77 |
{\out val it = "P [P & Q]" : thm} |
|
78 |
\ttbreak |
|
79 |
by (resolve_tac [major RS conjunct1] 1); |
|
80 |
{\out Level 2} |
|
81 |
{\out R} |
|
82 |
{\out 1. Q} |
|
83 |
\end{ttbox} |
|
84 |
Similarly, we solve the subgoal involving~$Q$. |
|
85 |
\begin{ttbox} |
|
86 |
major RS conjunct2; |
|
87 |
{\out val it = "Q [P & Q]" : thm} |
|
88 |
by (resolve_tac [major RS conjunct2] 1); |
|
89 |
{\out Level 3} |
|
90 |
{\out R} |
|
91 |
{\out No subgoals!} |
|
92 |
\end{ttbox} |
|
93 |
Calling \ttindex{topthm} returns the current proof state as a theorem. |
|
94 |
Note that it contains assumptions. Calling \ttindex{result} discharges the |
|
95 |
assumptions --- both occurrences of $P\conj Q$ are discharged as one --- |
|
96 |
and makes the variables schematic. |
|
97 |
\begin{ttbox} |
|
98 |
topthm(); |
|
99 |
{\out val it = "R [P & Q, P & Q, [| P; Q |] ==> R]" : thm} |
|
100 |
val conjE = result(); |
|
101 |
{\out val conjE = "[| ?P & ?Q; [| ?P; ?Q |] ==> ?R |] ==> ?R" : thm} |
|
102 |
\end{ttbox} |
|
103 |
||
104 |
||
105 |
\subsection{Definitions and derived rules} \label{definitions} |
|
310 | 106 |
\index{rules!derived}\index{definitions!and derived rules|(} |
107 |
||
105 | 108 |
Definitions are expressed as meta-level equalities. Let us define negation |
109 |
and the if-and-only-if connective: |
|
110 |
\begin{eqnarray*} |
|
111 |
\neg \Var{P} & \equiv & \Var{P}\imp\bot \\ |
|
112 |
\Var{P}\bimp \Var{Q} & \equiv & |
|
113 |
(\Var{P}\imp \Var{Q}) \conj (\Var{Q}\imp \Var{P}) |
|
114 |
\end{eqnarray*} |
|
331 | 115 |
\index{meta-rewriting}% |
105 | 116 |
Isabelle permits {\bf meta-level rewriting} using definitions such as |
117 |
these. {\bf Unfolding} replaces every instance |
|
331 | 118 |
of $\neg \Var{P}$ by the corresponding instance of ${\Var{P}\imp\bot}$. For |
105 | 119 |
example, $\forall x.\neg (P(x)\conj \neg R(x,0))$ unfolds to |
120 |
\[ \forall x.(P(x)\conj R(x,0)\imp\bot)\imp\bot. \] |
|
121 |
{\bf Folding} a definition replaces occurrences of the right-hand side by |
|
122 |
the left. The occurrences need not be free in the entire formula. |
|
123 |
||
124 |
When you define new concepts, you should derive rules asserting their |
|
125 |
abstract properties, and then forget their definitions. This supports |
|
331 | 126 |
modularity: if you later change the definitions without affecting their |
105 | 127 |
abstract properties, then most of your proofs will carry through without |
128 |
change. Indiscriminate unfolding makes a subgoal grow exponentially, |
|
129 |
becoming unreadable. |
|
130 |
||
131 |
Taking this point of view, Isabelle does not unfold definitions |
|
132 |
automatically during proofs. Rewriting must be explicit and selective. |
|
133 |
Isabelle provides tactics and meta-rules for rewriting, and a version of |
|
134 |
the {\tt goal} command that unfolds the conclusion and premises of the rule |
|
135 |
being derived. |
|
136 |
||
137 |
For example, the intuitionistic definition of negation given above may seem |
|
138 |
peculiar. Using Isabelle, we shall derive pleasanter negation rules: |
|
139 |
\[ \infer[({\neg}I)]{\neg P}{\infer*{\bot}{[P]}} \qquad |
|
140 |
\infer[({\neg}E)]{Q}{\neg P & P} \] |
|
296 | 141 |
This requires proving the following meta-formulae: |
105 | 142 |
$$ (P\Imp\bot) \Imp \neg P \eqno(\neg I)$$ |
143 |
$$ \List{\neg P; P} \Imp Q. \eqno(\neg E)$$ |
|
144 |
||
145 |
||
296 | 146 |
\subsection{Deriving the $\neg$ introduction rule} |
310 | 147 |
To derive $(\neg I)$, we may call {\tt goal} with the appropriate |
105 | 148 |
formula. Again, {\tt goal} returns a list consisting of the rule's |
296 | 149 |
premises. We bind this one-element list to the \ML\ identifier {\tt |
150 |
prems}. |
|
105 | 151 |
\begin{ttbox} |
152 |
val prems = goal FOL.thy "(P ==> False) ==> ~P"; |
|
153 |
{\out Level 0} |
|
154 |
{\out ~P} |
|
155 |
{\out 1. ~P} |
|
156 |
{\out val prems = ["P ==> False [P ==> False]"] : thm list} |
|
157 |
\end{ttbox} |
|
310 | 158 |
Calling \ttindex{rewrite_goals_tac} with \tdx{not_def}, which is the |
105 | 159 |
definition of negation, unfolds that definition in the subgoals. It leaves |
160 |
the main goal alone. |
|
161 |
\begin{ttbox} |
|
162 |
not_def; |
|
163 |
{\out val it = "~?P == ?P --> False" : thm} |
|
164 |
by (rewrite_goals_tac [not_def]); |
|
165 |
{\out Level 1} |
|
166 |
{\out ~P} |
|
167 |
{\out 1. P --> False} |
|
168 |
\end{ttbox} |
|
310 | 169 |
Using \tdx{impI} and the premise, we reduce subgoal~1 to a triviality: |
105 | 170 |
\begin{ttbox} |
171 |
by (resolve_tac [impI] 1); |
|
172 |
{\out Level 2} |
|
173 |
{\out ~P} |
|
174 |
{\out 1. P ==> False} |
|
175 |
\ttbreak |
|
176 |
by (resolve_tac prems 1); |
|
177 |
{\out Level 3} |
|
178 |
{\out ~P} |
|
179 |
{\out 1. P ==> P} |
|
180 |
\end{ttbox} |
|
296 | 181 |
The rest of the proof is routine. Note the form of the final result. |
105 | 182 |
\begin{ttbox} |
183 |
by (assume_tac 1); |
|
184 |
{\out Level 4} |
|
185 |
{\out ~P} |
|
186 |
{\out No subgoals!} |
|
296 | 187 |
\ttbreak |
105 | 188 |
val notI = result(); |
189 |
{\out val notI = "(?P ==> False) ==> ~?P" : thm} |
|
190 |
\end{ttbox} |
|
310 | 191 |
\indexbold{*notI theorem} |
105 | 192 |
|
193 |
There is a simpler way of conducting this proof. The \ttindex{goalw} |
|
310 | 194 |
command starts a backward proof, as does {\tt goal}, but it also |
296 | 195 |
unfolds definitions. Thus there is no need to call |
196 |
\ttindex{rewrite_goals_tac}: |
|
105 | 197 |
\begin{ttbox} |
198 |
val prems = goalw FOL.thy [not_def] |
|
199 |
"(P ==> False) ==> ~P"; |
|
200 |
{\out Level 0} |
|
201 |
{\out ~P} |
|
202 |
{\out 1. P --> False} |
|
203 |
{\out val prems = ["P ==> False [P ==> False]"] : thm list} |
|
204 |
\end{ttbox} |
|
205 |
||
206 |
||
296 | 207 |
\subsection{Deriving the $\neg$ elimination rule} |
284 | 208 |
Let us derive the rule $(\neg E)$. The proof follows that of~{\tt conjE} |
296 | 209 |
above, with an additional step to unfold negation in the major premise. |
210 |
Although the {\tt goalw} command is best for this, let us |
|
310 | 211 |
try~{\tt goal} to see another way of unfolding definitions. After |
212 |
binding the premises to \ML\ identifiers, we apply \tdx{FalseE}: |
|
105 | 213 |
\begin{ttbox} |
214 |
val [major,minor] = goal FOL.thy "[| ~P; P |] ==> R"; |
|
215 |
{\out Level 0} |
|
216 |
{\out R} |
|
217 |
{\out 1. R} |
|
218 |
{\out val major = "~ P [~ P]" : thm} |
|
219 |
{\out val minor = "P [P]" : thm} |
|
220 |
\ttbreak |
|
221 |
by (resolve_tac [FalseE] 1); |
|
222 |
{\out Level 1} |
|
223 |
{\out R} |
|
224 |
{\out 1. False} |
|
284 | 225 |
\end{ttbox} |
226 |
Everything follows from falsity. And we can prove falsity using the |
|
227 |
premises and Modus Ponens: |
|
228 |
\begin{ttbox} |
|
105 | 229 |
by (resolve_tac [mp] 1); |
230 |
{\out Level 2} |
|
231 |
{\out R} |
|
232 |
{\out 1. ?P1 --> False} |
|
233 |
{\out 2. ?P1} |
|
234 |
\end{ttbox} |
|
235 |
For subgoal~1, we transform the major premise from~$\neg P$ |
|
236 |
to~${P\imp\bot}$. The function \ttindex{rewrite_rule}, given a list of |
|
296 | 237 |
definitions, unfolds them in a theorem. Rewriting does not |
105 | 238 |
affect the theorem's hypothesis, which remains~$\neg P$: |
239 |
\begin{ttbox} |
|
240 |
rewrite_rule [not_def] major; |
|
241 |
{\out val it = "P --> False [~P]" : thm} |
|
242 |
by (resolve_tac [it] 1); |
|
243 |
{\out Level 3} |
|
244 |
{\out R} |
|
245 |
{\out 1. P} |
|
246 |
\end{ttbox} |
|
296 | 247 |
The subgoal {\tt?P1} has been instantiated to~{\tt P}, which we can prove |
284 | 248 |
using the minor premise: |
105 | 249 |
\begin{ttbox} |
250 |
by (resolve_tac [minor] 1); |
|
251 |
{\out Level 4} |
|
252 |
{\out R} |
|
253 |
{\out No subgoals!} |
|
254 |
val notE = result(); |
|
255 |
{\out val notE = "[| ~?P; ?P |] ==> ?R" : thm} |
|
256 |
\end{ttbox} |
|
310 | 257 |
\indexbold{*notE theorem} |
105 | 258 |
|
259 |
\medskip |
|
284 | 260 |
Again, there is a simpler way of conducting this proof. Recall that |
261 |
the \ttindex{goalw} command unfolds definitions the conclusion; it also |
|
262 |
unfolds definitions in the premises: |
|
105 | 263 |
\begin{ttbox} |
264 |
val [major,minor] = goalw FOL.thy [not_def] |
|
265 |
"[| ~P; P |] ==> R"; |
|
266 |
{\out val major = "P --> False [~ P]" : thm} |
|
267 |
{\out val minor = "P [P]" : thm} |
|
268 |
\end{ttbox} |
|
296 | 269 |
Observe the difference in {\tt major}; the premises are unfolded without |
270 |
calling~\ttindex{rewrite_rule}. Incidentally, the four calls to |
|
271 |
\ttindex{resolve_tac} above can be collapsed to one, with the help |
|
284 | 272 |
of~\ttindex{RS}; this is a typical example of forward reasoning from a |
273 |
complex premise. |
|
105 | 274 |
\begin{ttbox} |
275 |
minor RS (major RS mp RS FalseE); |
|
276 |
{\out val it = "?P [P, ~P]" : thm} |
|
277 |
by (resolve_tac [it] 1); |
|
278 |
{\out Level 1} |
|
279 |
{\out R} |
|
280 |
{\out No subgoals!} |
|
281 |
\end{ttbox} |
|
310 | 282 |
\index{definitions!and derived rules|)} |
105 | 283 |
|
310 | 284 |
\goodbreak\medskip\index{*"!"! symbol!in main goal} |
284 | 285 |
Finally, here is a trick that is sometimes useful. If the goal |
105 | 286 |
has an outermost meta-quantifier, then \ttindex{goal} and \ttindex{goalw} |
284 | 287 |
do not return the rule's premises in the list of theorems; instead, the |
288 |
premises become assumptions in subgoal~1. |
|
289 |
%%%It does not matter which variables are quantified over. |
|
105 | 290 |
\begin{ttbox} |
291 |
goalw FOL.thy [not_def] "!!P R. [| ~P; P |] ==> R"; |
|
292 |
{\out Level 0} |
|
293 |
{\out !!P R. [| ~ P; P |] ==> R} |
|
294 |
{\out 1. !!P R. [| P --> False; P |] ==> R} |
|
295 |
val it = [] : thm list |
|
296 |
\end{ttbox} |
|
297 |
The proof continues as before. But instead of referring to \ML\ |
|
310 | 298 |
identifiers, we refer to assumptions using {\tt eresolve_tac} or |
299 |
{\tt assume_tac}: |
|
105 | 300 |
\begin{ttbox} |
301 |
by (resolve_tac [FalseE] 1); |
|
302 |
{\out Level 1} |
|
303 |
{\out !!P R. [| ~ P; P |] ==> R} |
|
304 |
{\out 1. !!P R. [| P --> False; P |] ==> False} |
|
305 |
\ttbreak |
|
306 |
by (eresolve_tac [mp] 1); |
|
307 |
{\out Level 2} |
|
308 |
{\out !!P R. [| ~ P; P |] ==> R} |
|
309 |
{\out 1. !!P R. P ==> P} |
|
310 |
\ttbreak |
|
311 |
by (assume_tac 1); |
|
312 |
{\out Level 3} |
|
313 |
{\out !!P R. [| ~ P; P |] ==> R} |
|
314 |
{\out No subgoals!} |
|
315 |
\end{ttbox} |
|
316 |
Calling \ttindex{result} strips the meta-quantifiers, so the resulting |
|
317 |
theorem is the same as before. |
|
318 |
\begin{ttbox} |
|
319 |
val notE = result(); |
|
320 |
{\out val notE = "[| ~?P; ?P |] ==> ?R" : thm} |
|
321 |
\end{ttbox} |
|
322 |
Do not use the {\tt!!}\ trick if the premises contain meta-level |
|
323 |
connectives, because \ttindex{eresolve_tac} and \ttindex{assume_tac} would |
|
324 |
not be able to handle the resulting assumptions. The trick is not suitable |
|
325 |
for deriving the introduction rule~$(\neg I)$. |
|
326 |
||
327 |
||
284 | 328 |
\section{Defining theories}\label{sec:defining-theories} |
105 | 329 |
\index{theories!defining|(} |
310 | 330 |
|
105 | 331 |
Isabelle makes no distinction between simple extensions of a logic --- like |
332 |
defining a type~$bool$ with constants~$true$ and~$false$ --- and defining |
|
333 |
an entire logic. A theory definition has the form |
|
334 |
\begin{ttbox} |
|
335 |
\(T\) = \(S@1\) + \(\cdots\) + \(S@n\) + |
|
336 |
classes {\it class declarations} |
|
337 |
default {\it sort} |
|
331 | 338 |
types {\it type declarations and synonyms} |
105 | 339 |
arities {\it arity declarations} |
340 |
consts {\it constant declarations} |
|
841 | 341 |
translations {\it translation declarations} |
1185 | 342 |
defs {\it definitions} |
105 | 343 |
rules {\it rule declarations} |
344 |
end |
|
345 |
ML {\it ML code} |
|
346 |
\end{ttbox} |
|
347 |
This declares the theory $T$ to extend the existing theories |
|
348 |
$S@1$,~\ldots,~$S@n$. It may declare new classes, types, arities |
|
349 |
(overloadings of existing types), constants and rules; it can specify the |
|
350 |
default sort for type variables. A constant declaration can specify an |
|
351 |
associated concrete syntax. The translations section specifies rewrite |
|
352 |
rules on abstract syntax trees, for defining notations and abbreviations. |
|
310 | 353 |
\index{*ML section} |
354 |
The {\tt ML} section contains code to perform arbitrary syntactic |
|
284 | 355 |
transformations. The main declaration forms are discussed below. |
303 | 356 |
The full syntax can be found in \iflabelundefined{app:TheorySyntax}{the |
307 | 357 |
appendix of the {\it Reference Manual}}{App.\ts\ref{app:TheorySyntax}}. |
105 | 358 |
|
1084 | 359 |
All the declaration parts can be omitted or repeated and may appear in any |
360 |
order, except that the {\ML} section must be last. In the simplest case, $T$ |
|
361 |
is just the union of $S@1$,~\ldots,~$S@n$. New theories always extend one or |
|
362 |
more other theories, inheriting their types, constants, syntax, etc. The |
|
363 |
theory \thydx{Pure} contains nothing but Isabelle's meta-logic. |
|
105 | 364 |
|
331 | 365 |
Each theory definition must reside in a separate file, whose name is the |
366 |
theory's with {\tt.thy} appended. For example, theory {\tt ListFn} resides |
|
367 |
on a file named {\tt ListFn.thy}. Isabelle uses this convention to locate the |
|
368 |
file containing a given theory; \ttindexbold{use_thy} automatically loads a |
|
369 |
theory's parents before loading the theory itself. |
|
105 | 370 |
|
331 | 371 |
Calling \ttindexbold{use_thy}~{\tt"{\it T\/}"} reads a theory from the |
372 |
file {\it T}{\tt.thy}, writes the corresponding {\ML} code to the file |
|
373 |
{\tt.{\it T}.thy.ML}, reads the latter file, and deletes it if no errors |
|
374 |
occurred. This declares the {\ML} structure~$T$, which contains a component |
|
109 | 375 |
{\tt thy} denoting the new theory, a component for each rule, and everything |
376 |
declared in {\it ML code}. |
|
105 | 377 |
|
378 |
Errors may arise during the translation to {\ML} (say, a misspelled keyword) |
|
379 |
or during creation of the new theory (say, a type error in a rule). But if |
|
331 | 380 |
all goes well, {\tt use_thy} will finally read the file {\it T}{\tt.ML}, if |
105 | 381 |
it exists. This file typically begins with the {\ML} declaration {\tt |
382 |
open}~$T$ and contains proofs that refer to the components of~$T$. |
|
383 |
||
296 | 384 |
When a theory file is modified, many theories may have to be reloaded. |
385 |
Isabelle records the modification times and dependencies of theory files. |
|
331 | 386 |
See |
387 |
\iflabelundefined{sec:reloading-theories}{the {\em Reference Manual\/}}% |
|
388 |
{\S\ref{sec:reloading-theories}} |
|
296 | 389 |
for more details. |
390 |
||
105 | 391 |
|
1084 | 392 |
\subsection{Declaring constants, definitions and rules} |
310 | 393 |
\indexbold{constants!declaring}\index{rules!declaring} |
394 |
||
1084 | 395 |
Most theories simply declare constants, definitions and rules. The {\bf |
396 |
constant declaration part} has the form |
|
105 | 397 |
\begin{ttbox} |
1387 | 398 |
consts \(c@1\) :: \(\tau@1\) |
105 | 399 |
\vdots |
1387 | 400 |
\(c@n\) :: \(\tau@n\) |
105 | 401 |
\end{ttbox} |
402 |
where $c@1$, \ldots, $c@n$ are constants and $\tau@1$, \ldots, $\tau@n$ are |
|
1387 | 403 |
types. The types must be enclosed in quotation marks if they contain |
404 |
user-declared infix type constructors like {\tt *}. Each |
|
105 | 405 |
constant must be enclosed in quotation marks unless it is a valid |
406 |
identifier. To declare $c@1$, \ldots, $c@n$ as constants of type $\tau$, |
|
407 |
the $n$ declarations may be abbreviated to a single line: |
|
408 |
\begin{ttbox} |
|
1387 | 409 |
\(c@1\), \ldots, \(c@n\) :: \(\tau\) |
105 | 410 |
\end{ttbox} |
411 |
The {\bf rule declaration part} has the form |
|
412 |
\begin{ttbox} |
|
413 |
rules \(id@1\) "\(rule@1\)" |
|
414 |
\vdots |
|
415 |
\(id@n\) "\(rule@n\)" |
|
416 |
\end{ttbox} |
|
417 |
where $id@1$, \ldots, $id@n$ are \ML{} identifiers and $rule@1$, \ldots, |
|
284 | 418 |
$rule@n$ are expressions of type~$prop$. Each rule {\em must\/} be |
419 |
enclosed in quotation marks. |
|
420 |
||
1084 | 421 |
\indexbold{definitions} The {\bf definition part} is similar, but with the |
422 |
keyword {\tt defs} instead of {\tt rules}. {\bf Definitions} are rules of the |
|
1649
c4901f7161c5
Added 'constdefs' and extended the section on 'defs'
nipkow
parents:
1468
diff
changeset
|
423 |
form $s \equiv t$, and should serve only as abbreviations. The simplest form |
c4901f7161c5
Added 'constdefs' and extended the section on 'defs'
nipkow
parents:
1468
diff
changeset
|
424 |
of a definition is $f \equiv t$, where $f$ is a constant. Isabelle also |
1904 | 425 |
allows a derived form where the arguments of~$f$ appear on the left: |
1649
c4901f7161c5
Added 'constdefs' and extended the section on 'defs'
nipkow
parents:
1468
diff
changeset
|
426 |
\begin{itemize} |
c4901f7161c5
Added 'constdefs' and extended the section on 'defs'
nipkow
parents:
1468
diff
changeset
|
427 |
\item In uncurried syntax: $f(x@1,\dots,x@n) \equiv u$ (\FOL, \ZF, \dots) |
c4901f7161c5
Added 'constdefs' and extended the section on 'defs'
nipkow
parents:
1468
diff
changeset
|
428 |
\item In curried syntax: $f~x@1~\dots~x@n \equiv u$ (\HOL, \HOLCF) |
c4901f7161c5
Added 'constdefs' and extended the section on 'defs'
nipkow
parents:
1468
diff
changeset
|
429 |
\end{itemize} |
c4901f7161c5
Added 'constdefs' and extended the section on 'defs'
nipkow
parents:
1468
diff
changeset
|
430 |
Isabelle checks for common errors in definitions, such as extra variables on |
c4901f7161c5
Added 'constdefs' and extended the section on 'defs'
nipkow
parents:
1468
diff
changeset
|
431 |
the right-hand side. Determined users can write non-conservative |
c4901f7161c5
Added 'constdefs' and extended the section on 'defs'
nipkow
parents:
1468
diff
changeset
|
432 |
`definitions' by using mutual recursion, for example; the consequences of |
c4901f7161c5
Added 'constdefs' and extended the section on 'defs'
nipkow
parents:
1468
diff
changeset
|
433 |
such actions are their responsibility. |
1084 | 434 |
|
435 |
\index{examples!of theories} |
|
436 |
This theory extends first-order logic by declaring and defining two constants, |
|
437 |
{\em nand} and {\em xor}: |
|
284 | 438 |
\begin{ttbox} |
105 | 439 |
Gate = FOL + |
1387 | 440 |
consts nand,xor :: [o,o] => o |
1084 | 441 |
defs nand_def "nand(P,Q) == ~(P & Q)" |
105 | 442 |
xor_def "xor(P,Q) == P & ~Q | ~P & Q" |
443 |
end |
|
444 |
\end{ttbox} |
|
445 |
||
1649
c4901f7161c5
Added 'constdefs' and extended the section on 'defs'
nipkow
parents:
1468
diff
changeset
|
446 |
Declaring and defining constants can be combined: |
c4901f7161c5
Added 'constdefs' and extended the section on 'defs'
nipkow
parents:
1468
diff
changeset
|
447 |
\begin{ttbox} |
c4901f7161c5
Added 'constdefs' and extended the section on 'defs'
nipkow
parents:
1468
diff
changeset
|
448 |
Gate = FOL + |
c4901f7161c5
Added 'constdefs' and extended the section on 'defs'
nipkow
parents:
1468
diff
changeset
|
449 |
constdefs nand :: [o,o] => o |
c4901f7161c5
Added 'constdefs' and extended the section on 'defs'
nipkow
parents:
1468
diff
changeset
|
450 |
"nand(P,Q) == ~(P & Q)" |
c4901f7161c5
Added 'constdefs' and extended the section on 'defs'
nipkow
parents:
1468
diff
changeset
|
451 |
xor :: [o,o] => o |
c4901f7161c5
Added 'constdefs' and extended the section on 'defs'
nipkow
parents:
1468
diff
changeset
|
452 |
"xor(P,Q) == P & ~Q | ~P & Q" |
c4901f7161c5
Added 'constdefs' and extended the section on 'defs'
nipkow
parents:
1468
diff
changeset
|
453 |
end |
c4901f7161c5
Added 'constdefs' and extended the section on 'defs'
nipkow
parents:
1468
diff
changeset
|
454 |
\end{ttbox} |
c4901f7161c5
Added 'constdefs' and extended the section on 'defs'
nipkow
parents:
1468
diff
changeset
|
455 |
{\tt constdefs} generates the names {\tt nand_def} and {\tt xor_def} |
c4901f7161c5
Added 'constdefs' and extended the section on 'defs'
nipkow
parents:
1468
diff
changeset
|
456 |
automatically, which is why it is restricted to alphanumeric identifiers. In |
c4901f7161c5
Added 'constdefs' and extended the section on 'defs'
nipkow
parents:
1468
diff
changeset
|
457 |
general it has the form |
c4901f7161c5
Added 'constdefs' and extended the section on 'defs'
nipkow
parents:
1468
diff
changeset
|
458 |
\begin{ttbox} |
c4901f7161c5
Added 'constdefs' and extended the section on 'defs'
nipkow
parents:
1468
diff
changeset
|
459 |
constdefs \(id@1\) :: \(\tau@1\) |
c4901f7161c5
Added 'constdefs' and extended the section on 'defs'
nipkow
parents:
1468
diff
changeset
|
460 |
"\(id@1 \equiv \dots\)" |
c4901f7161c5
Added 'constdefs' and extended the section on 'defs'
nipkow
parents:
1468
diff
changeset
|
461 |
\vdots |
c4901f7161c5
Added 'constdefs' and extended the section on 'defs'
nipkow
parents:
1468
diff
changeset
|
462 |
\(id@n\) :: \(\tau@n\) |
c4901f7161c5
Added 'constdefs' and extended the section on 'defs'
nipkow
parents:
1468
diff
changeset
|
463 |
"\(id@n \equiv \dots\)" |
c4901f7161c5
Added 'constdefs' and extended the section on 'defs'
nipkow
parents:
1468
diff
changeset
|
464 |
\end{ttbox} |
c4901f7161c5
Added 'constdefs' and extended the section on 'defs'
nipkow
parents:
1468
diff
changeset
|
465 |
|
c4901f7161c5
Added 'constdefs' and extended the section on 'defs'
nipkow
parents:
1468
diff
changeset
|
466 |
|
1366
3f3c25d3ec04
Inserted warning about defs with extra vars on rhs.
nipkow
parents:
1185
diff
changeset
|
467 |
\begin{warn} |
3f3c25d3ec04
Inserted warning about defs with extra vars on rhs.
nipkow
parents:
1185
diff
changeset
|
468 |
A common mistake when writing definitions is to introduce extra free variables |
1468 | 469 |
on the right-hand side as in the following fictitious definition: |
1366
3f3c25d3ec04
Inserted warning about defs with extra vars on rhs.
nipkow
parents:
1185
diff
changeset
|
470 |
\begin{ttbox} |
3f3c25d3ec04
Inserted warning about defs with extra vars on rhs.
nipkow
parents:
1185
diff
changeset
|
471 |
defs prime_def "prime(p) == (m divides p) --> (m=1 | m=p)" |
3f3c25d3ec04
Inserted warning about defs with extra vars on rhs.
nipkow
parents:
1185
diff
changeset
|
472 |
\end{ttbox} |
3f3c25d3ec04
Inserted warning about defs with extra vars on rhs.
nipkow
parents:
1185
diff
changeset
|
473 |
Isabelle rejects this ``definition'' because of the extra {\tt m} on the |
1467 | 474 |
right-hand side, which would introduce an inconsistency. What you should have |
1366
3f3c25d3ec04
Inserted warning about defs with extra vars on rhs.
nipkow
parents:
1185
diff
changeset
|
475 |
written is |
3f3c25d3ec04
Inserted warning about defs with extra vars on rhs.
nipkow
parents:
1185
diff
changeset
|
476 |
\begin{ttbox} |
3f3c25d3ec04
Inserted warning about defs with extra vars on rhs.
nipkow
parents:
1185
diff
changeset
|
477 |
defs prime_def "prime(p) == ALL m. (m divides p) --> (m=1 | m=p)" |
3f3c25d3ec04
Inserted warning about defs with extra vars on rhs.
nipkow
parents:
1185
diff
changeset
|
478 |
\end{ttbox} |
3f3c25d3ec04
Inserted warning about defs with extra vars on rhs.
nipkow
parents:
1185
diff
changeset
|
479 |
\end{warn} |
105 | 480 |
|
481 |
\subsection{Declaring type constructors} |
|
303 | 482 |
\indexbold{types!declaring}\indexbold{arities!declaring} |
284 | 483 |
% |
105 | 484 |
Types are composed of type variables and {\bf type constructors}. Each |
284 | 485 |
type constructor takes a fixed number of arguments. They are declared |
486 |
with an \ML-like syntax. If $list$ takes one type argument, $tree$ takes |
|
487 |
two arguments and $nat$ takes no arguments, then these type constructors |
|
488 |
can be declared by |
|
105 | 489 |
\begin{ttbox} |
284 | 490 |
types 'a list |
491 |
('a,'b) tree |
|
492 |
nat |
|
105 | 493 |
\end{ttbox} |
284 | 494 |
|
495 |
The {\bf type declaration part} has the general form |
|
496 |
\begin{ttbox} |
|
497 |
types \(tids@1\) \(id@1\) |
|
498 |
\vdots |
|
841 | 499 |
\(tids@n\) \(id@n\) |
284 | 500 |
\end{ttbox} |
501 |
where $id@1$, \ldots, $id@n$ are identifiers and $tids@1$, \ldots, $tids@n$ |
|
502 |
are type argument lists as shown in the example above. It declares each |
|
503 |
$id@i$ as a type constructor with the specified number of argument places. |
|
105 | 504 |
|
505 |
The {\bf arity declaration part} has the form |
|
506 |
\begin{ttbox} |
|
507 |
arities \(tycon@1\) :: \(arity@1\) |
|
508 |
\vdots |
|
509 |
\(tycon@n\) :: \(arity@n\) |
|
510 |
\end{ttbox} |
|
511 |
where $tycon@1$, \ldots, $tycon@n$ are identifiers and $arity@1$, \ldots, |
|
512 |
$arity@n$ are arities. Arity declarations add arities to existing |
|
296 | 513 |
types; they do not declare the types themselves. |
105 | 514 |
In the simplest case, for an 0-place type constructor, an arity is simply |
515 |
the type's class. Let us declare a type~$bool$ of class $term$, with |
|
284 | 516 |
constants $tt$ and~$ff$. (In first-order logic, booleans are |
517 |
distinct from formulae, which have type $o::logic$.) |
|
105 | 518 |
\index{examples!of theories} |
284 | 519 |
\begin{ttbox} |
105 | 520 |
Bool = FOL + |
284 | 521 |
types bool |
105 | 522 |
arities bool :: term |
1387 | 523 |
consts tt,ff :: bool |
105 | 524 |
end |
525 |
\end{ttbox} |
|
296 | 526 |
A $k$-place type constructor may have arities of the form |
527 |
$(s@1,\ldots,s@k)c$, where $s@1,\ldots,s@n$ are sorts and $c$ is a class. |
|
528 |
Each sort specifies a type argument; it has the form $\{c@1,\ldots,c@m\}$, |
|
529 |
where $c@1$, \dots,~$c@m$ are classes. Mostly we deal with singleton |
|
530 |
sorts, and may abbreviate them by dropping the braces. The arity |
|
531 |
$(term)term$ is short for $(\{term\})term$. Recall the discussion in |
|
532 |
\S\ref{polymorphic}. |
|
105 | 533 |
|
534 |
A type constructor may be overloaded (subject to certain conditions) by |
|
296 | 535 |
appearing in several arity declarations. For instance, the function type |
331 | 536 |
constructor~$fun$ has the arity $(logic,logic)logic$; in higher-order |
105 | 537 |
logic, it is declared also to have arity $(term,term)term$. |
538 |
||
539 |
Theory {\tt List} declares the 1-place type constructor $list$, gives |
|
284 | 540 |
it arity $(term)term$, and declares constants $Nil$ and $Cons$ with |
296 | 541 |
polymorphic types:% |
542 |
\footnote{In the {\tt consts} part, type variable {\tt'a} has the default |
|
543 |
sort, which is {\tt term}. See the {\em Reference Manual\/} |
|
544 |
\iflabelundefined{sec:ref-defining-theories}{}% |
|
545 |
{(\S\ref{sec:ref-defining-theories})} for more information.} |
|
105 | 546 |
\index{examples!of theories} |
284 | 547 |
\begin{ttbox} |
105 | 548 |
List = FOL + |
284 | 549 |
types 'a list |
105 | 550 |
arities list :: (term)term |
1387 | 551 |
consts Nil :: 'a list |
552 |
Cons :: ['a, 'a list] => 'a list |
|
105 | 553 |
end |
554 |
\end{ttbox} |
|
284 | 555 |
Multiple arity declarations may be abbreviated to a single line: |
105 | 556 |
\begin{ttbox} |
557 |
arities \(tycon@1\), \ldots, \(tycon@n\) :: \(arity\) |
|
558 |
\end{ttbox} |
|
559 |
||
560 |
\begin{warn} |
|
561 |
Arity declarations resemble constant declarations, but there are {\it no\/} |
|
562 |
quotation marks! Types and rules must be quoted because the theory |
|
563 |
translator passes them verbatim to the {\ML} output file. |
|
564 |
\end{warn} |
|
565 |
||
331 | 566 |
\subsection{Type synonyms}\indexbold{type synonyms} |
303 | 567 |
Isabelle supports {\bf type synonyms} ({\bf abbreviations}) which are similar |
307 | 568 |
to those found in \ML. Such synonyms are defined in the type declaration part |
303 | 569 |
and are fairly self explanatory: |
570 |
\begin{ttbox} |
|
1387 | 571 |
types gate = [o,o] => o |
572 |
'a pred = 'a => o |
|
573 |
('a,'b)nuf = 'b => 'a |
|
303 | 574 |
\end{ttbox} |
575 |
Type declarations and synonyms can be mixed arbitrarily: |
|
576 |
\begin{ttbox} |
|
577 |
types nat |
|
1387 | 578 |
'a stream = nat => 'a |
579 |
signal = nat stream |
|
303 | 580 |
'a list |
581 |
\end{ttbox} |
|
307 | 582 |
A synonym is merely an abbreviation for some existing type expression. Hence |
583 |
synonyms may not be recursive! Internally all synonyms are fully expanded. As |
|
584 |
a consequence Isabelle output never contains synonyms. Their main purpose is |
|
585 |
to improve the readability of theories. Synonyms can be used just like any |
|
303 | 586 |
other type: |
587 |
\begin{ttbox} |
|
1387 | 588 |
consts and,or :: gate |
589 |
negate :: signal => signal |
|
303 | 590 |
\end{ttbox} |
591 |
||
348 | 592 |
\subsection{Infix and mixfix operators} |
310 | 593 |
\index{infixes}\index{examples!of theories} |
594 |
||
595 |
Infix or mixfix syntax may be attached to constants. Consider the |
|
596 |
following theory: |
|
284 | 597 |
\begin{ttbox} |
105 | 598 |
Gate2 = FOL + |
1387 | 599 |
consts "~&" :: [o,o] => o (infixl 35) |
600 |
"#" :: [o,o] => o (infixl 30) |
|
1084 | 601 |
defs nand_def "P ~& Q == ~(P & Q)" |
105 | 602 |
xor_def "P # Q == P & ~Q | ~P & Q" |
603 |
end |
|
604 |
\end{ttbox} |
|
310 | 605 |
The constant declaration part declares two left-associating infix operators |
606 |
with their priorities, or precedences; they are $\nand$ of priority~35 and |
|
607 |
$\xor$ of priority~30. Hence $P \xor Q \xor R$ is parsed as $(P\xor Q) |
|
608 |
\xor R$ and $P \xor Q \nand R$ as $P \xor (Q \nand R)$. Note the quotation |
|
609 |
marks in \verb|"~&"| and \verb|"#"|. |
|
105 | 610 |
|
611 |
The constants \hbox{\verb|op ~&|} and \hbox{\verb|op #|} are declared |
|
612 |
automatically, just as in \ML. Hence you may write propositions like |
|
613 |
\verb|op #(True) == op ~&(True)|, which asserts that the functions $\lambda |
|
614 |
Q.True \xor Q$ and $\lambda Q.True \nand Q$ are identical. |
|
615 |
||
310 | 616 |
\bigskip\index{mixfix declarations} |
617 |
{\bf Mixfix} operators may have arbitrary context-free syntaxes. Let us |
|
618 |
add a line to the constant declaration part: |
|
284 | 619 |
\begin{ttbox} |
1387 | 620 |
If :: [o,o,o] => o ("if _ then _ else _") |
105 | 621 |
\end{ttbox} |
310 | 622 |
This declares a constant $If$ of type $[o,o,o] \To o$ with concrete syntax {\tt |
296 | 623 |
if~$P$ then~$Q$ else~$R$} as well as {\tt If($P$,$Q$,$R$)}. Underscores |
310 | 624 |
denote argument positions. |
105 | 625 |
|
310 | 626 |
The declaration above does not allow the {\tt if}-{\tt then}-{\tt else} |
627 |
construct to be split across several lines, even if it is too long to fit |
|
628 |
on one line. Pretty-printing information can be added to specify the |
|
629 |
layout of mixfix operators. For details, see |
|
630 |
\iflabelundefined{Defining-Logics}% |
|
631 |
{the {\it Reference Manual}, chapter `Defining Logics'}% |
|
632 |
{Chap.\ts\ref{Defining-Logics}}. |
|
633 |
||
634 |
Mixfix declarations can be annotated with priorities, just like |
|
105 | 635 |
infixes. The example above is just a shorthand for |
284 | 636 |
\begin{ttbox} |
1387 | 637 |
If :: [o,o,o] => o ("if _ then _ else _" [0,0,0] 1000) |
105 | 638 |
\end{ttbox} |
310 | 639 |
The numeric components determine priorities. The list of integers |
640 |
defines, for each argument position, the minimal priority an expression |
|
641 |
at that position must have. The final integer is the priority of the |
|
105 | 642 |
construct itself. In the example above, any argument expression is |
310 | 643 |
acceptable because priorities are non-negative, and conditionals may |
644 |
appear everywhere because 1000 is the highest priority. On the other |
|
645 |
hand, the declaration |
|
284 | 646 |
\begin{ttbox} |
1387 | 647 |
If :: [o,o,o] => o ("if _ then _ else _" [100,0,0] 99) |
105 | 648 |
\end{ttbox} |
284 | 649 |
defines concrete syntax for a conditional whose first argument cannot have |
310 | 650 |
the form {\tt if~$P$ then~$Q$ else~$R$} because it must have a priority |
651 |
of at least~100. We may of course write |
|
284 | 652 |
\begin{quote}\tt |
653 |
if (if $P$ then $Q$ else $R$) then $S$ else $T$ |
|
156 | 654 |
\end{quote} |
310 | 655 |
because expressions in parentheses have maximal priority. |
105 | 656 |
|
657 |
Binary type constructors, like products and sums, may also be declared as |
|
658 |
infixes. The type declaration below introduces a type constructor~$*$ with |
|
659 |
infix notation $\alpha*\beta$, together with the mixfix notation |
|
1084 | 660 |
${<}\_,\_{>}$ for pairs. We also see a rule declaration part. |
310 | 661 |
\index{examples!of theories}\index{mixfix declarations} |
105 | 662 |
\begin{ttbox} |
663 |
Prod = FOL + |
|
284 | 664 |
types ('a,'b) "*" (infixl 20) |
105 | 665 |
arities "*" :: (term,term)term |
666 |
consts fst :: "'a * 'b => 'a" |
|
667 |
snd :: "'a * 'b => 'b" |
|
668 |
Pair :: "['a,'b] => 'a * 'b" ("(1<_,/_>)") |
|
669 |
rules fst "fst(<a,b>) = a" |
|
670 |
snd "snd(<a,b>) = b" |
|
671 |
end |
|
672 |
\end{ttbox} |
|
673 |
||
674 |
\begin{warn} |
|
675 |
The name of the type constructor is~{\tt *} and not {\tt op~*}, as it would |
|
676 |
be in the case of an infix constant. Only infix type constructors can have |
|
677 |
symbolic names like~{\tt *}. There is no general mixfix syntax for types. |
|
678 |
\end{warn} |
|
679 |
||
680 |
||
681 |
\subsection{Overloading} |
|
682 |
\index{overloading}\index{examples!of theories} |
|
683 |
The {\bf class declaration part} has the form |
|
684 |
\begin{ttbox} |
|
685 |
classes \(id@1\) < \(c@1\) |
|
686 |
\vdots |
|
687 |
\(id@n\) < \(c@n\) |
|
688 |
\end{ttbox} |
|
689 |
where $id@1$, \ldots, $id@n$ are identifiers and $c@1$, \ldots, $c@n$ are |
|
690 |
existing classes. It declares each $id@i$ as a new class, a subclass |
|
691 |
of~$c@i$. In the general case, an identifier may be declared to be a |
|
692 |
subclass of $k$ existing classes: |
|
693 |
\begin{ttbox} |
|
694 |
\(id\) < \(c@1\), \ldots, \(c@k\) |
|
695 |
\end{ttbox} |
|
296 | 696 |
Type classes allow constants to be overloaded. As suggested in |
307 | 697 |
\S\ref{polymorphic}, let us define the class $arith$ of arithmetic |
296 | 698 |
types with the constants ${+} :: [\alpha,\alpha]\To \alpha$ and $0,1 {::} |
699 |
\alpha$, for $\alpha{::}arith$. We introduce $arith$ as a subclass of |
|
700 |
$term$ and add the three polymorphic constants of this class. |
|
310 | 701 |
\index{examples!of theories}\index{constants!overloaded} |
105 | 702 |
\begin{ttbox} |
703 |
Arith = FOL + |
|
704 |
classes arith < term |
|
1387 | 705 |
consts "0" :: 'a::arith ("0") |
706 |
"1" :: 'a::arith ("1") |
|
707 |
"+" :: ['a::arith,'a] => 'a (infixl 60) |
|
105 | 708 |
end |
709 |
\end{ttbox} |
|
710 |
No rules are declared for these constants: we merely introduce their |
|
711 |
names without specifying properties. On the other hand, classes |
|
712 |
with rules make it possible to prove {\bf generic} theorems. Such |
|
713 |
theorems hold for all instances, all types in that class. |
|
714 |
||
715 |
We can now obtain distinct versions of the constants of $arith$ by |
|
716 |
declaring certain types to be of class $arith$. For example, let us |
|
717 |
declare the 0-place type constructors $bool$ and $nat$: |
|
718 |
\index{examples!of theories} |
|
719 |
\begin{ttbox} |
|
720 |
BoolNat = Arith + |
|
348 | 721 |
types bool nat |
722 |
arities bool, nat :: arith |
|
1387 | 723 |
consts Suc :: nat=>nat |
284 | 724 |
\ttbreak |
105 | 725 |
rules add0 "0 + n = n::nat" |
726 |
addS "Suc(m)+n = Suc(m+n)" |
|
727 |
nat1 "1 = Suc(0)" |
|
728 |
or0l "0 + x = x::bool" |
|
729 |
or0r "x + 0 = x::bool" |
|
730 |
or1l "1 + x = 1::bool" |
|
731 |
or1r "x + 1 = 1::bool" |
|
732 |
end |
|
733 |
\end{ttbox} |
|
734 |
Because $nat$ and $bool$ have class $arith$, we can use $0$, $1$ and $+$ at |
|
735 |
either type. The type constraints in the axioms are vital. Without |
|
736 |
constraints, the $x$ in $1+x = x$ would have type $\alpha{::}arith$ |
|
737 |
and the axiom would hold for any type of class $arith$. This would |
|
284 | 738 |
collapse $nat$ to a trivial type: |
105 | 739 |
\[ Suc(1) = Suc(0+1) = Suc(0)+1 = 1+1 = 1! \] |
296 | 740 |
|
105 | 741 |
|
296 | 742 |
\section{Theory example: the natural numbers} |
743 |
||
744 |
We shall now work through a small example of formalized mathematics |
|
105 | 745 |
demonstrating many of the theory extension features. |
746 |
||
747 |
||
748 |
\subsection{Extending first-order logic with the natural numbers} |
|
749 |
\index{examples!of theories} |
|
750 |
||
284 | 751 |
Section\ts\ref{sec:logical-syntax} has formalized a first-order logic, |
752 |
including a type~$nat$ and the constants $0::nat$ and $Suc::nat\To nat$. |
|
753 |
Let us introduce the Peano axioms for mathematical induction and the |
|
310 | 754 |
freeness of $0$ and~$Suc$:\index{axioms!Peano} |
307 | 755 |
\[ \vcenter{\infer[(induct)]{P[n/x]}{P[0/x] & \infer*{P[Suc(x)/x]}{[P]}}} |
105 | 756 |
\qquad \parbox{4.5cm}{provided $x$ is not free in any assumption except~$P$} |
757 |
\] |
|
758 |
\[ \infer[(Suc\_inject)]{m=n}{Suc(m)=Suc(n)} \qquad |
|
759 |
\infer[(Suc\_neq\_0)]{R}{Suc(m)=0} |
|
760 |
\] |
|
761 |
Mathematical induction asserts that $P(n)$ is true, for any $n::nat$, |
|
762 |
provided $P(0)$ holds and that $P(x)$ implies $P(Suc(x))$ for all~$x$. |
|
763 |
Some authors express the induction step as $\forall x. P(x)\imp P(Suc(x))$. |
|
764 |
To avoid making induction require the presence of other connectives, we |
|
765 |
formalize mathematical induction as |
|
766 |
$$ \List{P(0); \Forall x. P(x)\Imp P(Suc(x))} \Imp P(n). \eqno(induct) $$ |
|
767 |
||
768 |
\noindent |
|
769 |
Similarly, to avoid expressing the other rules using~$\forall$, $\imp$ |
|
770 |
and~$\neg$, we take advantage of the meta-logic;\footnote |
|
771 |
{On the other hand, the axioms $Suc(m)=Suc(n) \bimp m=n$ |
|
772 |
and $\neg(Suc(m)=0)$ are logically equivalent to those given, and work |
|
773 |
better with Isabelle's simplifier.} |
|
774 |
$(Suc\_neq\_0)$ is |
|
775 |
an elimination rule for $Suc(m)=0$: |
|
776 |
$$ Suc(m)=Suc(n) \Imp m=n \eqno(Suc\_inject) $$ |
|
777 |
$$ Suc(m)=0 \Imp R \eqno(Suc\_neq\_0) $$ |
|
778 |
||
779 |
\noindent |
|
780 |
We shall also define a primitive recursion operator, $rec$. Traditionally, |
|
781 |
primitive recursion takes a natural number~$a$ and a 2-place function~$f$, |
|
782 |
and obeys the equations |
|
783 |
\begin{eqnarray*} |
|
784 |
rec(0,a,f) & = & a \\ |
|
785 |
rec(Suc(m),a,f) & = & f(m, rec(m,a,f)) |
|
786 |
\end{eqnarray*} |
|
787 |
Addition, defined by $m+n \equiv rec(m,n,\lambda x\,y.Suc(y))$, |
|
788 |
should satisfy |
|
789 |
\begin{eqnarray*} |
|
790 |
0+n & = & n \\ |
|
791 |
Suc(m)+n & = & Suc(m+n) |
|
792 |
\end{eqnarray*} |
|
296 | 793 |
Primitive recursion appears to pose difficulties: first-order logic has no |
794 |
function-valued expressions. We again take advantage of the meta-logic, |
|
795 |
which does have functions. We also generalise primitive recursion to be |
|
105 | 796 |
polymorphic over any type of class~$term$, and declare the addition |
797 |
function: |
|
798 |
\begin{eqnarray*} |
|
799 |
rec & :: & [nat, \alpha{::}term, [nat,\alpha]\To\alpha] \To\alpha \\ |
|
800 |
+ & :: & [nat,nat]\To nat |
|
801 |
\end{eqnarray*} |
|
802 |
||
803 |
||
804 |
\subsection{Declaring the theory to Isabelle} |
|
805 |
\index{examples!of theories} |
|
310 | 806 |
Let us create the theory \thydx{Nat} starting from theory~\verb$FOL$, |
105 | 807 |
which contains only classical logic with no natural numbers. We declare |
307 | 808 |
the 0-place type constructor $nat$ and the associated constants. Note that |
809 |
the constant~0 requires a mixfix annotation because~0 is not a legal |
|
810 |
identifier, and could not otherwise be written in terms: |
|
310 | 811 |
\begin{ttbox}\index{mixfix declarations} |
105 | 812 |
Nat = FOL + |
284 | 813 |
types nat |
105 | 814 |
arities nat :: term |
1387 | 815 |
consts "0" :: nat ("0") |
816 |
Suc :: nat=>nat |
|
817 |
rec :: [nat, 'a, [nat,'a]=>'a] => 'a |
|
818 |
"+" :: [nat, nat] => nat (infixl 60) |
|
296 | 819 |
rules Suc_inject "Suc(m)=Suc(n) ==> m=n" |
105 | 820 |
Suc_neq_0 "Suc(m)=0 ==> R" |
296 | 821 |
induct "[| P(0); !!x. P(x) ==> P(Suc(x)) |] ==> P(n)" |
105 | 822 |
rec_0 "rec(0,a,f) = a" |
823 |
rec_Suc "rec(Suc(m), a, f) = f(m, rec(m,a,f))" |
|
296 | 824 |
add_def "m+n == rec(m, n, \%x y. Suc(y))" |
105 | 825 |
end |
826 |
\end{ttbox} |
|
827 |
In axiom {\tt add_def}, recall that \verb|%| stands for~$\lambda$. |
|
296 | 828 |
Loading this theory file creates the \ML\ structure {\tt Nat}, which |
829 |
contains the theory and axioms. Opening structure {\tt Nat} lets us write |
|
830 |
{\tt induct} instead of {\tt Nat.induct}, and so forth. |
|
105 | 831 |
\begin{ttbox} |
832 |
open Nat; |
|
833 |
\end{ttbox} |
|
296 | 834 |
|
835 |
\subsection{Proving some recursion equations} |
|
331 | 836 |
File {\tt FOL/ex/Nat.ML} contains proofs involving this theory of the |
105 | 837 |
natural numbers. As a trivial example, let us derive recursion equations |
838 |
for \verb$+$. Here is the zero case: |
|
284 | 839 |
\begin{ttbox} |
105 | 840 |
goalw Nat.thy [add_def] "0+n = n"; |
841 |
{\out Level 0} |
|
842 |
{\out 0 + n = n} |
|
284 | 843 |
{\out 1. rec(0,n,\%x y. Suc(y)) = n} |
105 | 844 |
\ttbreak |
845 |
by (resolve_tac [rec_0] 1); |
|
846 |
{\out Level 1} |
|
847 |
{\out 0 + n = n} |
|
848 |
{\out No subgoals!} |
|
849 |
val add_0 = result(); |
|
284 | 850 |
\end{ttbox} |
105 | 851 |
And here is the successor case: |
284 | 852 |
\begin{ttbox} |
105 | 853 |
goalw Nat.thy [add_def] "Suc(m)+n = Suc(m+n)"; |
854 |
{\out Level 0} |
|
855 |
{\out Suc(m) + n = Suc(m + n)} |
|
284 | 856 |
{\out 1. rec(Suc(m),n,\%x y. Suc(y)) = Suc(rec(m,n,\%x y. Suc(y)))} |
105 | 857 |
\ttbreak |
858 |
by (resolve_tac [rec_Suc] 1); |
|
859 |
{\out Level 1} |
|
860 |
{\out Suc(m) + n = Suc(m + n)} |
|
861 |
{\out No subgoals!} |
|
862 |
val add_Suc = result(); |
|
284 | 863 |
\end{ttbox} |
105 | 864 |
The induction rule raises some complications, which are discussed next. |
865 |
\index{theories!defining|)} |
|
866 |
||
867 |
||
868 |
\section{Refinement with explicit instantiation} |
|
310 | 869 |
\index{resolution!with instantiation} |
870 |
\index{instantiation|(} |
|
871 |
||
105 | 872 |
In order to employ mathematical induction, we need to refine a subgoal by |
873 |
the rule~$(induct)$. The conclusion of this rule is $\Var{P}(\Var{n})$, |
|
874 |
which is highly ambiguous in higher-order unification. It matches every |
|
875 |
way that a formula can be regarded as depending on a subterm of type~$nat$. |
|
876 |
To get round this problem, we could make the induction rule conclude |
|
877 |
$\forall n.\Var{P}(n)$ --- but putting a subgoal into this form requires |
|
878 |
refinement by~$(\forall E)$, which is equally hard! |
|
879 |
||
880 |
The tactic {\tt res_inst_tac}, like {\tt resolve_tac}, refines a subgoal by |
|
881 |
a rule. But it also accepts explicit instantiations for the rule's |
|
882 |
schematic variables. |
|
883 |
\begin{description} |
|
310 | 884 |
\item[\ttindex{res_inst_tac} {\it insts} {\it thm} {\it i}] |
105 | 885 |
instantiates the rule {\it thm} with the instantiations {\it insts}, and |
886 |
then performs resolution on subgoal~$i$. |
|
887 |
||
310 | 888 |
\item[\ttindex{eres_inst_tac}] |
889 |
and \ttindex{dres_inst_tac} are similar, but perform elim-resolution |
|
105 | 890 |
and destruct-resolution, respectively. |
891 |
\end{description} |
|
892 |
The list {\it insts} consists of pairs $[(v@1,e@1), \ldots, (v@n,e@n)]$, |
|
893 |
where $v@1$, \ldots, $v@n$ are names of schematic variables in the rule --- |
|
307 | 894 |
with no leading question marks! --- and $e@1$, \ldots, $e@n$ are |
105 | 895 |
expressions giving their instantiations. The expressions are type-checked |
896 |
in the context of a particular subgoal: free variables receive the same |
|
897 |
types as they have in the subgoal, and parameters may appear. Type |
|
898 |
variable instantiations may appear in~{\it insts}, but they are seldom |
|
899 |
required: {\tt res_inst_tac} instantiates type variables automatically |
|
900 |
whenever the type of~$e@i$ is an instance of the type of~$\Var{v@i}$. |
|
901 |
||
902 |
\subsection{A simple proof by induction} |
|
310 | 903 |
\index{examples!of induction} |
105 | 904 |
Let us prove that no natural number~$k$ equals its own successor. To |
905 |
use~$(induct)$, we instantiate~$\Var{n}$ to~$k$; Isabelle finds a good |
|
906 |
instantiation for~$\Var{P}$. |
|
284 | 907 |
\begin{ttbox} |
105 | 908 |
goal Nat.thy "~ (Suc(k) = k)"; |
909 |
{\out Level 0} |
|
459 | 910 |
{\out Suc(k) ~= k} |
911 |
{\out 1. Suc(k) ~= k} |
|
105 | 912 |
\ttbreak |
913 |
by (res_inst_tac [("n","k")] induct 1); |
|
914 |
{\out Level 1} |
|
459 | 915 |
{\out Suc(k) ~= k} |
916 |
{\out 1. Suc(0) ~= 0} |
|
917 |
{\out 2. !!x. Suc(x) ~= x ==> Suc(Suc(x)) ~= Suc(x)} |
|
284 | 918 |
\end{ttbox} |
105 | 919 |
We should check that Isabelle has correctly applied induction. Subgoal~1 |
920 |
is the base case, with $k$ replaced by~0. Subgoal~2 is the inductive step, |
|
921 |
with $k$ replaced by~$Suc(x)$ and with an induction hypothesis for~$x$. |
|
310 | 922 |
The rest of the proof demonstrates~\tdx{notI}, \tdx{notE} and the |
923 |
other rules of theory {\tt Nat}. The base case holds by~\ttindex{Suc_neq_0}: |
|
284 | 924 |
\begin{ttbox} |
105 | 925 |
by (resolve_tac [notI] 1); |
926 |
{\out Level 2} |
|
459 | 927 |
{\out Suc(k) ~= k} |
105 | 928 |
{\out 1. Suc(0) = 0 ==> False} |
459 | 929 |
{\out 2. !!x. Suc(x) ~= x ==> Suc(Suc(x)) ~= Suc(x)} |
105 | 930 |
\ttbreak |
931 |
by (eresolve_tac [Suc_neq_0] 1); |
|
932 |
{\out Level 3} |
|
459 | 933 |
{\out Suc(k) ~= k} |
934 |
{\out 1. !!x. Suc(x) ~= x ==> Suc(Suc(x)) ~= Suc(x)} |
|
284 | 935 |
\end{ttbox} |
105 | 936 |
The inductive step holds by the contrapositive of~\ttindex{Suc_inject}. |
284 | 937 |
Negation rules transform the subgoal into that of proving $Suc(x)=x$ from |
938 |
$Suc(Suc(x)) = Suc(x)$: |
|
939 |
\begin{ttbox} |
|
105 | 940 |
by (resolve_tac [notI] 1); |
941 |
{\out Level 4} |
|
459 | 942 |
{\out Suc(k) ~= k} |
943 |
{\out 1. !!x. [| Suc(x) ~= x; Suc(Suc(x)) = Suc(x) |] ==> False} |
|
105 | 944 |
\ttbreak |
945 |
by (eresolve_tac [notE] 1); |
|
946 |
{\out Level 5} |
|
459 | 947 |
{\out Suc(k) ~= k} |
105 | 948 |
{\out 1. !!x. Suc(Suc(x)) = Suc(x) ==> Suc(x) = x} |
949 |
\ttbreak |
|
950 |
by (eresolve_tac [Suc_inject] 1); |
|
951 |
{\out Level 6} |
|
459 | 952 |
{\out Suc(k) ~= k} |
105 | 953 |
{\out No subgoals!} |
284 | 954 |
\end{ttbox} |
105 | 955 |
|
956 |
||
957 |
\subsection{An example of ambiguity in {\tt resolve_tac}} |
|
958 |
\index{examples!of induction}\index{unification!higher-order} |
|
959 |
If you try the example above, you may observe that {\tt res_inst_tac} is |
|
960 |
not actually needed. Almost by chance, \ttindex{resolve_tac} finds the right |
|
961 |
instantiation for~$(induct)$ to yield the desired next state. With more |
|
962 |
complex formulae, our luck fails. |
|
284 | 963 |
\begin{ttbox} |
105 | 964 |
goal Nat.thy "(k+m)+n = k+(m+n)"; |
965 |
{\out Level 0} |
|
966 |
{\out k + m + n = k + (m + n)} |
|
967 |
{\out 1. k + m + n = k + (m + n)} |
|
968 |
\ttbreak |
|
969 |
by (resolve_tac [induct] 1); |
|
970 |
{\out Level 1} |
|
971 |
{\out k + m + n = k + (m + n)} |
|
972 |
{\out 1. k + m + n = 0} |
|
973 |
{\out 2. !!x. k + m + n = x ==> k + m + n = Suc(x)} |
|
284 | 974 |
\end{ttbox} |
975 |
This proof requires induction on~$k$. The occurrence of~0 in subgoal~1 |
|
976 |
indicates that induction has been applied to the term~$k+(m+n)$; this |
|
977 |
application is sound but will not lead to a proof here. Fortunately, |
|
978 |
Isabelle can (lazily!) generate all the valid applications of induction. |
|
979 |
The \ttindex{back} command causes backtracking to an alternative outcome of |
|
980 |
the tactic. |
|
981 |
\begin{ttbox} |
|
105 | 982 |
back(); |
983 |
{\out Level 1} |
|
984 |
{\out k + m + n = k + (m + n)} |
|
985 |
{\out 1. k + m + n = k + 0} |
|
986 |
{\out 2. !!x. k + m + n = k + x ==> k + m + n = k + Suc(x)} |
|
284 | 987 |
\end{ttbox} |
988 |
Now induction has been applied to~$m+n$. This is equally useless. Let us |
|
989 |
call \ttindex{back} again. |
|
990 |
\begin{ttbox} |
|
105 | 991 |
back(); |
992 |
{\out Level 1} |
|
993 |
{\out k + m + n = k + (m + n)} |
|
994 |
{\out 1. k + m + 0 = k + (m + 0)} |
|
284 | 995 |
{\out 2. !!x. k + m + x = k + (m + x) ==>} |
996 |
{\out k + m + Suc(x) = k + (m + Suc(x))} |
|
997 |
\end{ttbox} |
|
105 | 998 |
Now induction has been applied to~$n$. What is the next alternative? |
284 | 999 |
\begin{ttbox} |
105 | 1000 |
back(); |
1001 |
{\out Level 1} |
|
1002 |
{\out k + m + n = k + (m + n)} |
|
1003 |
{\out 1. k + m + n = k + (m + 0)} |
|
1004 |
{\out 2. !!x. k + m + n = k + (m + x) ==> k + m + n = k + (m + Suc(x))} |
|
284 | 1005 |
\end{ttbox} |
105 | 1006 |
Inspecting subgoal~1 reveals that induction has been applied to just the |
1007 |
second occurrence of~$n$. This perfectly legitimate induction is useless |
|
310 | 1008 |
here. |
1009 |
||
1010 |
The main goal admits fourteen different applications of induction. The |
|
1011 |
number is exponential in the size of the formula. |
|
105 | 1012 |
|
1013 |
\subsection{Proving that addition is associative} |
|
331 | 1014 |
Let us invoke the induction rule properly, using~{\tt |
310 | 1015 |
res_inst_tac}. At the same time, we shall have a glimpse at Isabelle's |
1016 |
simplification tactics, which are described in |
|
1017 |
\iflabelundefined{simp-chap}% |
|
1018 |
{the {\em Reference Manual}}{Chap.\ts\ref{simp-chap}}. |
|
284 | 1019 |
|
310 | 1020 |
\index{simplification}\index{examples!of simplification} |
1021 |
||
1022 |
Isabelle's simplification tactics repeatedly apply equations to a subgoal, |
|
1023 |
perhaps proving it. For efficiency, the rewrite rules must be |
|
1024 |
packaged into a {\bf simplification set},\index{simplification sets} |
|
1025 |
or {\bf simpset}. We take the standard simpset for first-order logic and |
|
331 | 1026 |
insert the equations proved in the previous section, namely |
310 | 1027 |
$0+n=n$ and ${\tt Suc}(m)+n={\tt Suc}(m+n)$: |
284 | 1028 |
\begin{ttbox} |
459 | 1029 |
val add_ss = FOL_ss addsimps [add_0, add_Suc]; |
284 | 1030 |
\end{ttbox} |
105 | 1031 |
We state the goal for associativity of addition, and |
1032 |
use \ttindex{res_inst_tac} to invoke induction on~$k$: |
|
284 | 1033 |
\begin{ttbox} |
105 | 1034 |
goal Nat.thy "(k+m)+n = k+(m+n)"; |
1035 |
{\out Level 0} |
|
1036 |
{\out k + m + n = k + (m + n)} |
|
1037 |
{\out 1. k + m + n = k + (m + n)} |
|
1038 |
\ttbreak |
|
1039 |
by (res_inst_tac [("n","k")] induct 1); |
|
1040 |
{\out Level 1} |
|
1041 |
{\out k + m + n = k + (m + n)} |
|
1042 |
{\out 1. 0 + m + n = 0 + (m + n)} |
|
284 | 1043 |
{\out 2. !!x. x + m + n = x + (m + n) ==>} |
1044 |
{\out Suc(x) + m + n = Suc(x) + (m + n)} |
|
1045 |
\end{ttbox} |
|
105 | 1046 |
The base case holds easily; both sides reduce to $m+n$. The |
1047 |
tactic~\ttindex{simp_tac} rewrites with respect to the given simplification |
|
331 | 1048 |
set, applying the rewrite rules for addition: |
284 | 1049 |
\begin{ttbox} |
105 | 1050 |
by (simp_tac add_ss 1); |
1051 |
{\out Level 2} |
|
1052 |
{\out k + m + n = k + (m + n)} |
|
284 | 1053 |
{\out 1. !!x. x + m + n = x + (m + n) ==>} |
1054 |
{\out Suc(x) + m + n = Suc(x) + (m + n)} |
|
1055 |
\end{ttbox} |
|
331 | 1056 |
The inductive step requires rewriting by the equations for addition |
105 | 1057 |
together the induction hypothesis, which is also an equation. The |
1058 |
tactic~\ttindex{asm_simp_tac} rewrites using a simplification set and any |
|
1059 |
useful assumptions: |
|
284 | 1060 |
\begin{ttbox} |
105 | 1061 |
by (asm_simp_tac add_ss 1); |
1062 |
{\out Level 3} |
|
1063 |
{\out k + m + n = k + (m + n)} |
|
1064 |
{\out No subgoals!} |
|
284 | 1065 |
\end{ttbox} |
310 | 1066 |
\index{instantiation|)} |
105 | 1067 |
|
1068 |
||
284 | 1069 |
\section{A Prolog interpreter} |
105 | 1070 |
\index{Prolog interpreter|bold} |
284 | 1071 |
To demonstrate the power of tacticals, let us construct a Prolog |
105 | 1072 |
interpreter and execute programs involving lists.\footnote{To run these |
331 | 1073 |
examples, see the file {\tt FOL/ex/Prolog.ML}.} The Prolog program |
105 | 1074 |
consists of a theory. We declare a type constructor for lists, with an |
1075 |
arity declaration to say that $(\tau)list$ is of class~$term$ |
|
1076 |
provided~$\tau$ is: |
|
1077 |
\begin{eqnarray*} |
|
1078 |
list & :: & (term)term |
|
1079 |
\end{eqnarray*} |
|
1080 |
We declare four constants: the empty list~$Nil$; the infix list |
|
1081 |
constructor~{:}; the list concatenation predicate~$app$; the list reverse |
|
284 | 1082 |
predicate~$rev$. (In Prolog, functions on lists are expressed as |
105 | 1083 |
predicates.) |
1084 |
\begin{eqnarray*} |
|
1085 |
Nil & :: & \alpha list \\ |
|
1086 |
{:} & :: & [\alpha,\alpha list] \To \alpha list \\ |
|
1087 |
app & :: & [\alpha list,\alpha list,\alpha list] \To o \\ |
|
1088 |
rev & :: & [\alpha list,\alpha list] \To o |
|
1089 |
\end{eqnarray*} |
|
284 | 1090 |
The predicate $app$ should satisfy the Prolog-style rules |
105 | 1091 |
\[ {app(Nil,ys,ys)} \qquad |
1092 |
{app(xs,ys,zs) \over app(x:xs, ys, x:zs)} \] |
|
1093 |
We define the naive version of $rev$, which calls~$app$: |
|
1094 |
\[ {rev(Nil,Nil)} \qquad |
|
1095 |
{rev(xs,ys)\quad app(ys, x:Nil, zs) \over |
|
1096 |
rev(x:xs, zs)} |
|
1097 |
\] |
|
1098 |
||
1099 |
\index{examples!of theories} |
|
310 | 1100 |
Theory \thydx{Prolog} extends first-order logic in order to make use |
105 | 1101 |
of the class~$term$ and the type~$o$. The interpreter does not use the |
310 | 1102 |
rules of~{\tt FOL}. |
105 | 1103 |
\begin{ttbox} |
1104 |
Prolog = FOL + |
|
296 | 1105 |
types 'a list |
105 | 1106 |
arities list :: (term)term |
1387 | 1107 |
consts Nil :: 'a list |
1108 |
":" :: ['a, 'a list]=> 'a list (infixr 60) |
|
1109 |
app :: ['a list, 'a list, 'a list] => o |
|
1110 |
rev :: ['a list, 'a list] => o |
|
105 | 1111 |
rules appNil "app(Nil,ys,ys)" |
1112 |
appCons "app(xs,ys,zs) ==> app(x:xs, ys, x:zs)" |
|
1113 |
revNil "rev(Nil,Nil)" |
|
1114 |
revCons "[| rev(xs,ys); app(ys,x:Nil,zs) |] ==> rev(x:xs,zs)" |
|
1115 |
end |
|
1116 |
\end{ttbox} |
|
1117 |
\subsection{Simple executions} |
|
284 | 1118 |
Repeated application of the rules solves Prolog goals. Let us |
105 | 1119 |
append the lists $[a,b,c]$ and~$[d,e]$. As the rules are applied, the |
1120 |
answer builds up in~{\tt ?x}. |
|
1121 |
\begin{ttbox} |
|
1122 |
goal Prolog.thy "app(a:b:c:Nil, d:e:Nil, ?x)"; |
|
1123 |
{\out Level 0} |
|
1124 |
{\out app(a : b : c : Nil, d : e : Nil, ?x)} |
|
1125 |
{\out 1. app(a : b : c : Nil, d : e : Nil, ?x)} |
|
1126 |
\ttbreak |
|
1127 |
by (resolve_tac [appNil,appCons] 1); |
|
1128 |
{\out Level 1} |
|
1129 |
{\out app(a : b : c : Nil, d : e : Nil, a : ?zs1)} |
|
1130 |
{\out 1. app(b : c : Nil, d : e : Nil, ?zs1)} |
|
1131 |
\ttbreak |
|
1132 |
by (resolve_tac [appNil,appCons] 1); |
|
1133 |
{\out Level 2} |
|
1134 |
{\out app(a : b : c : Nil, d : e : Nil, a : b : ?zs2)} |
|
1135 |
{\out 1. app(c : Nil, d : e : Nil, ?zs2)} |
|
1136 |
\end{ttbox} |
|
1137 |
At this point, the first two elements of the result are~$a$ and~$b$. |
|
1138 |
\begin{ttbox} |
|
1139 |
by (resolve_tac [appNil,appCons] 1); |
|
1140 |
{\out Level 3} |
|
1141 |
{\out app(a : b : c : Nil, d : e : Nil, a : b : c : ?zs3)} |
|
1142 |
{\out 1. app(Nil, d : e : Nil, ?zs3)} |
|
1143 |
\ttbreak |
|
1144 |
by (resolve_tac [appNil,appCons] 1); |
|
1145 |
{\out Level 4} |
|
1146 |
{\out app(a : b : c : Nil, d : e : Nil, a : b : c : d : e : Nil)} |
|
1147 |
{\out No subgoals!} |
|
1148 |
\end{ttbox} |
|
1149 |
||
284 | 1150 |
Prolog can run functions backwards. Which list can be appended |
105 | 1151 |
with $[c,d]$ to produce $[a,b,c,d]$? |
1152 |
Using \ttindex{REPEAT}, we find the answer at once, $[a,b]$: |
|
1153 |
\begin{ttbox} |
|
1154 |
goal Prolog.thy "app(?x, c:d:Nil, a:b:c:d:Nil)"; |
|
1155 |
{\out Level 0} |
|
1156 |
{\out app(?x, c : d : Nil, a : b : c : d : Nil)} |
|
1157 |
{\out 1. app(?x, c : d : Nil, a : b : c : d : Nil)} |
|
1158 |
\ttbreak |
|
1159 |
by (REPEAT (resolve_tac [appNil,appCons] 1)); |
|
1160 |
{\out Level 1} |
|
1161 |
{\out app(a : b : Nil, c : d : Nil, a : b : c : d : Nil)} |
|
1162 |
{\out No subgoals!} |
|
1163 |
\end{ttbox} |
|
1164 |
||
1165 |
||
310 | 1166 |
\subsection{Backtracking}\index{backtracking!Prolog style} |
296 | 1167 |
Prolog backtracking can answer questions that have multiple solutions. |
1168 |
Which lists $x$ and $y$ can be appended to form the list $[a,b,c,d]$? This |
|
1169 |
question has five solutions. Using \ttindex{REPEAT} to apply the rules, we |
|
1170 |
quickly find the first solution, namely $x=[]$ and $y=[a,b,c,d]$: |
|
105 | 1171 |
\begin{ttbox} |
1172 |
goal Prolog.thy "app(?x, ?y, a:b:c:d:Nil)"; |
|
1173 |
{\out Level 0} |
|
1174 |
{\out app(?x, ?y, a : b : c : d : Nil)} |
|
1175 |
{\out 1. app(?x, ?y, a : b : c : d : Nil)} |
|
1176 |
\ttbreak |
|
1177 |
by (REPEAT (resolve_tac [appNil,appCons] 1)); |
|
1178 |
{\out Level 1} |
|
1179 |
{\out app(Nil, a : b : c : d : Nil, a : b : c : d : Nil)} |
|
1180 |
{\out No subgoals!} |
|
1181 |
\end{ttbox} |
|
284 | 1182 |
Isabelle can lazily generate all the possibilities. The \ttindex{back} |
1183 |
command returns the tactic's next outcome, namely $x=[a]$ and $y=[b,c,d]$: |
|
105 | 1184 |
\begin{ttbox} |
1185 |
back(); |
|
1186 |
{\out Level 1} |
|
1187 |
{\out app(a : Nil, b : c : d : Nil, a : b : c : d : Nil)} |
|
1188 |
{\out No subgoals!} |
|
1189 |
\end{ttbox} |
|
1190 |
The other solutions are generated similarly. |
|
1191 |
\begin{ttbox} |
|
1192 |
back(); |
|
1193 |
{\out Level 1} |
|
1194 |
{\out app(a : b : Nil, c : d : Nil, a : b : c : d : Nil)} |
|
1195 |
{\out No subgoals!} |
|
1196 |
\ttbreak |
|
1197 |
back(); |
|
1198 |
{\out Level 1} |
|
1199 |
{\out app(a : b : c : Nil, d : Nil, a : b : c : d : Nil)} |
|
1200 |
{\out No subgoals!} |
|
1201 |
\ttbreak |
|
1202 |
back(); |
|
1203 |
{\out Level 1} |
|
1204 |
{\out app(a : b : c : d : Nil, Nil, a : b : c : d : Nil)} |
|
1205 |
{\out No subgoals!} |
|
1206 |
\end{ttbox} |
|
1207 |
||
1208 |
||
1209 |
\subsection{Depth-first search} |
|
1210 |
\index{search!depth-first} |
|
1211 |
Now let us try $rev$, reversing a list. |
|
1212 |
Bundle the rules together as the \ML{} identifier {\tt rules}. Naive |
|
1213 |
reverse requires 120 inferences for this 14-element list, but the tactic |
|
1214 |
terminates in a few seconds. |
|
1215 |
\begin{ttbox} |
|
1216 |
goal Prolog.thy "rev(a:b:c:d:e:f:g:h:i:j:k:l:m:n:Nil, ?w)"; |
|
1217 |
{\out Level 0} |
|
1218 |
{\out rev(a : b : c : d : e : f : g : h : i : j : k : l : m : n : Nil, ?w)} |
|
284 | 1219 |
{\out 1. rev(a : b : c : d : e : f : g : h : i : j : k : l : m : n : Nil,} |
1220 |
{\out ?w)} |
|
1221 |
\ttbreak |
|
105 | 1222 |
val rules = [appNil,appCons,revNil,revCons]; |
1223 |
\ttbreak |
|
1224 |
by (REPEAT (resolve_tac rules 1)); |
|
1225 |
{\out Level 1} |
|
1226 |
{\out rev(a : b : c : d : e : f : g : h : i : j : k : l : m : n : Nil,} |
|
1227 |
{\out n : m : l : k : j : i : h : g : f : e : d : c : b : a : Nil)} |
|
1228 |
{\out No subgoals!} |
|
1229 |
\end{ttbox} |
|
1230 |
We may execute $rev$ backwards. This, too, should reverse a list. What |
|
1231 |
is the reverse of $[a,b,c]$? |
|
1232 |
\begin{ttbox} |
|
1233 |
goal Prolog.thy "rev(?x, a:b:c:Nil)"; |
|
1234 |
{\out Level 0} |
|
1235 |
{\out rev(?x, a : b : c : Nil)} |
|
1236 |
{\out 1. rev(?x, a : b : c : Nil)} |
|
1237 |
\ttbreak |
|
1238 |
by (REPEAT (resolve_tac rules 1)); |
|
1239 |
{\out Level 1} |
|
1240 |
{\out rev(?x1 : Nil, a : b : c : Nil)} |
|
1241 |
{\out 1. app(Nil, ?x1 : Nil, a : b : c : Nil)} |
|
1242 |
\end{ttbox} |
|
1243 |
The tactic has failed to find a solution! It reached a dead end at |
|
331 | 1244 |
subgoal~1: there is no~$\Var{x@1}$ such that [] appended with~$[\Var{x@1}]$ |
105 | 1245 |
equals~$[a,b,c]$. Backtracking explores other outcomes. |
1246 |
\begin{ttbox} |
|
1247 |
back(); |
|
1248 |
{\out Level 1} |
|
1249 |
{\out rev(?x1 : a : Nil, a : b : c : Nil)} |
|
1250 |
{\out 1. app(Nil, ?x1 : Nil, b : c : Nil)} |
|
1251 |
\end{ttbox} |
|
1252 |
This too is a dead end, but the next outcome is successful. |
|
1253 |
\begin{ttbox} |
|
1254 |
back(); |
|
1255 |
{\out Level 1} |
|
1256 |
{\out rev(c : b : a : Nil, a : b : c : Nil)} |
|
1257 |
{\out No subgoals!} |
|
1258 |
\end{ttbox} |
|
310 | 1259 |
\ttindex{REPEAT} goes wrong because it is only a repetition tactical, not a |
1260 |
search tactical. {\tt REPEAT} stops when it cannot continue, regardless of |
|
1261 |
which state is reached. The tactical \ttindex{DEPTH_FIRST} searches for a |
|
1262 |
satisfactory state, as specified by an \ML{} predicate. Below, |
|
105 | 1263 |
\ttindex{has_fewer_prems} specifies that the proof state should have no |
310 | 1264 |
subgoals. |
105 | 1265 |
\begin{ttbox} |
1266 |
val prolog_tac = DEPTH_FIRST (has_fewer_prems 1) |
|
1267 |
(resolve_tac rules 1); |
|
1268 |
\end{ttbox} |
|
284 | 1269 |
Since Prolog uses depth-first search, this tactic is a (slow!) |
296 | 1270 |
Prolog interpreter. We return to the start of the proof using |
1271 |
\ttindex{choplev}, and apply {\tt prolog_tac}: |
|
105 | 1272 |
\begin{ttbox} |
1273 |
choplev 0; |
|
1274 |
{\out Level 0} |
|
1275 |
{\out rev(?x, a : b : c : Nil)} |
|
1276 |
{\out 1. rev(?x, a : b : c : Nil)} |
|
1277 |
\ttbreak |
|
1278 |
by (DEPTH_FIRST (has_fewer_prems 1) (resolve_tac rules 1)); |
|
1279 |
{\out Level 1} |
|
1280 |
{\out rev(c : b : a : Nil, a : b : c : Nil)} |
|
1281 |
{\out No subgoals!} |
|
1282 |
\end{ttbox} |
|
1283 |
Let us try {\tt prolog_tac} on one more example, containing four unknowns: |
|
1284 |
\begin{ttbox} |
|
1285 |
goal Prolog.thy "rev(a:?x:c:?y:Nil, d:?z:b:?u)"; |
|
1286 |
{\out Level 0} |
|
1287 |
{\out rev(a : ?x : c : ?y : Nil, d : ?z : b : ?u)} |
|
1288 |
{\out 1. rev(a : ?x : c : ?y : Nil, d : ?z : b : ?u)} |
|
1289 |
\ttbreak |
|
1290 |
by prolog_tac; |
|
1291 |
{\out Level 1} |
|
1292 |
{\out rev(a : b : c : d : Nil, d : c : b : a : Nil)} |
|
1293 |
{\out No subgoals!} |
|
1294 |
\end{ttbox} |
|
284 | 1295 |
Although Isabelle is much slower than a Prolog system, Isabelle |
156 | 1296 |
tactics can exploit logic programming techniques. |
1297 |