author | blanchet |
Mon, 30 Apr 2012 21:59:10 +0200 | |
changeset 47845 | 2a2bc13669bd |
parent 38432 | 439f50a241c1 |
child 48966 | 6e15de7dd871 |
permissions | -rw-r--r-- |
15136 | 1 |
theory ToyList |
26729 | 2 |
imports Datatype |
15136 | 3 |
begin |
8745 | 4 |
|
5 |
text{*\noindent |
|
26729 | 6 |
HOL already has a predefined theory of lists called @{text List} --- |
7 |
@{text ToyList} is merely a small fragment of it chosen as an example. In |
|
8745 | 8 |
contrast to what is recommended in \S\ref{sec:Basic:Theories}, |
26729 | 9 |
@{text ToyList} is not based on @{text Main} but on @{text Datatype}, a |
8745 | 10 |
theory that contains pretty much everything but lists, thus avoiding |
11 |
ambiguities caused by defining lists twice. |
|
12 |
*} |
|
13 |
||
14 |
datatype 'a list = Nil ("[]") |
|
15 |
| Cons 'a "'a list" (infixr "#" 65); |
|
16 |
||
17 |
text{*\noindent |
|
12327 | 18 |
The datatype\index{datatype@\isacommand {datatype} (command)} |
19 |
\tydx{list} introduces two |
|
11428 | 20 |
constructors \cdx{Nil} and \cdx{Cons}, the |
9541 | 21 |
empty~list and the operator that adds an element to the front of a list. For |
9792 | 22 |
example, the term \isa{Cons True (Cons False Nil)} is a value of |
23 |
type @{typ"bool list"}, namely the list with the elements @{term"True"} and |
|
11450 | 24 |
@{term"False"}. Because this notation quickly becomes unwieldy, the |
8745 | 25 |
datatype declaration is annotated with an alternative syntax: instead of |
9792 | 26 |
@{term[source]Nil} and \isa{Cons x xs} we can write |
15364 | 27 |
@{term"[]"}\index{$HOL2list@\isa{[]}|bold} and |
28 |
@{term"x # xs"}\index{$HOL2list@\isa{\#}|bold}. In fact, this |
|
11450 | 29 |
alternative syntax is the familiar one. Thus the list \isa{Cons True |
9541 | 30 |
(Cons False Nil)} becomes @{term"True # False # []"}. The annotation |
11428 | 31 |
\isacommand{infixr}\index{infixr@\isacommand{infixr} (annotation)} |
32 |
means that @{text"#"} associates to |
|
11450 | 33 |
the right: the term @{term"x # y # z"} is read as @{text"x # (y # z)"} |
9792 | 34 |
and not as @{text"(x # y) # z"}. |
10971 | 35 |
The @{text 65} is the priority of the infix @{text"#"}. |
8745 | 36 |
|
37 |
\begin{warn} |
|
13191 | 38 |
Syntax annotations can be powerful, but they are difficult to master and |
11456 | 39 |
are never necessary. You |
9792 | 40 |
could drop them from theory @{text"ToyList"} and go back to the identifiers |
27015 | 41 |
@{term[source]Nil} and @{term[source]Cons}. Novices should avoid using |
10795 | 42 |
syntax annotations in their own theories. |
8745 | 43 |
\end{warn} |
27015 | 44 |
Next, two functions @{text"app"} and \cdx{rev} are defined recursively, |
45 |
in this order, because Isabelle insists on definition before use: |
|
8745 | 46 |
*} |
47 |
||
27015 | 48 |
primrec app :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" (infixr "@" 65) where |
49 |
"[] @ ys = ys" | |
|
50 |
"(x # xs) @ ys = x # (xs @ ys)" |
|
51 |
||
52 |
primrec rev :: "'a list \<Rightarrow> 'a list" where |
|
53 |
"rev [] = []" | |
|
54 |
"rev (x # xs) = (rev xs) @ (x # [])" |
|
8745 | 55 |
|
27015 | 56 |
text{*\noindent |
57 |
Each function definition is of the form |
|
58 |
\begin{center} |
|
59 |
\isacommand{primrec} \textit{name} @{text"::"} \textit{type} \textit{(optional syntax)} \isakeyword{where} \textit{equations} |
|
60 |
\end{center} |
|
61 |
The equations must be separated by @{text"|"}. |
|
62 |
% |
|
63 |
Function @{text"app"} is annotated with concrete syntax. Instead of the |
|
10790 | 64 |
prefix syntax @{text"app xs ys"} the infix |
15364 | 65 |
@{term"xs @ ys"}\index{$HOL2list@\isa{\at}|bold} becomes the preferred |
27015 | 66 |
form. |
8745 | 67 |
|
27015 | 68 |
\index{*rev (constant)|(}\index{append function|(} |
10790 | 69 |
The equations for @{text"app"} and @{term"rev"} hardly need comments: |
70 |
@{text"app"} appends two lists and @{term"rev"} reverses a list. The |
|
11428 | 71 |
keyword \commdx{primrec} indicates that the recursion is |
10790 | 72 |
of a particularly primitive kind where each recursive call peels off a datatype |
8771 | 73 |
constructor from one of the arguments. Thus the |
10654 | 74 |
recursion always terminates, i.e.\ the function is \textbf{total}. |
11428 | 75 |
\index{functions!total} |
8745 | 76 |
|
77 |
The termination requirement is absolutely essential in HOL, a logic of total |
|
78 |
functions. If we were to drop it, inconsistencies would quickly arise: the |
|
79 |
``definition'' $f(n) = f(n)+1$ immediately leads to $0 = 1$ by subtracting |
|
80 |
$f(n)$ on both sides. |
|
81 |
% However, this is a subtle issue that we cannot discuss here further. |
|
82 |
||
83 |
\begin{warn} |
|
11456 | 84 |
As we have indicated, the requirement for total functions is an essential characteristic of HOL\@. It is only |
8745 | 85 |
because of totality that reasoning in HOL is comparatively easy. More |
11456 | 86 |
generally, the philosophy in HOL is to refrain from asserting arbitrary axioms (such as |
8745 | 87 |
function definitions whose totality has not been proved) because they |
88 |
quickly lead to inconsistencies. Instead, fixed constructs for introducing |
|
89 |
types and functions are offered (such as \isacommand{datatype} and |
|
90 |
\isacommand{primrec}) which are guaranteed to preserve consistency. |
|
91 |
\end{warn} |
|
92 |
||
11456 | 93 |
\index{syntax}% |
8745 | 94 |
A remark about syntax. The textual definition of a theory follows a fixed |
10971 | 95 |
syntax with keywords like \isacommand{datatype} and \isacommand{end}. |
96 |
% (see Fig.~\ref{fig:keywords} in Appendix~\ref{sec:Appendix} for a full list). |
|
8745 | 97 |
Embedded in this syntax are the types and formulae of HOL, whose syntax is |
12631 | 98 |
extensible (see \S\ref{sec:concrete-syntax}), e.g.\ by new user-defined infix operators. |
10971 | 99 |
To distinguish the two levels, everything |
8745 | 100 |
HOL-specific (terms and types) should be enclosed in |
101 |
\texttt{"}\dots\texttt{"}. |
|
102 |
To lessen this burden, quotation marks around a single identifier can be |
|
27015 | 103 |
dropped, unless the identifier happens to be a keyword, for example |
104 |
\isa{"end"}. |
|
8745 | 105 |
When Isabelle prints a syntax error message, it refers to the HOL syntax as |
11456 | 106 |
the \textbf{inner syntax} and the enclosing theory language as the \textbf{outer syntax}. |
8745 | 107 |
|
38430
254a021ed66e
tuned text about "value" and added note on comments.
nipkow
parents:
27015
diff
changeset
|
108 |
Comments\index{comment} must be in enclosed in \texttt{(* }and\texttt{ *)}. |
254a021ed66e
tuned text about "value" and added note on comments.
nipkow
parents:
27015
diff
changeset
|
109 |
|
25342 | 110 |
\section{Evaluation} |
111 |
\index{evaluation} |
|
112 |
||
113 |
Assuming you have processed the declarations and definitions of |
|
114 |
\texttt{ToyList} presented so far, you may want to test your |
|
115 |
functions by running them. For example, what is the value of |
|
116 |
@{term"rev(True#False#[])"}? Command |
|
117 |
*} |
|
118 |
||
119 |
value "rev (True # False # [])" |
|
120 |
||
121 |
text{* \noindent yields the correct result @{term"False # True # []"}. |
|
122 |
But we can go beyond mere functional programming and evaluate terms with |
|
123 |
variables in them, executing functions symbolically: *} |
|
124 |
||
38430
254a021ed66e
tuned text about "value" and added note on comments.
nipkow
parents:
27015
diff
changeset
|
125 |
value "rev (a # b # c # [])" |
25342 | 126 |
|
38432
439f50a241c1
Using type real does not require a separate logic now.
nipkow
parents:
38430
diff
changeset
|
127 |
text{*\noindent yields @{term"c # b # a # []"}. |
439f50a241c1
Using type real does not require a separate logic now.
nipkow
parents:
38430
diff
changeset
|
128 |
|
10885 | 129 |
\section{An Introductory Proof} |
8745 | 130 |
\label{sec:intro-proof} |
131 |
||
25342 | 132 |
Having convinced ourselves (as well as one can by testing) that our |
133 |
definitions capture our intentions, we are ready to prove a few simple |
|
16360 | 134 |
theorems. This will illustrate not just the basic proof commands but |
135 |
also the typical proof process. |
|
8745 | 136 |
|
11457 | 137 |
\subsubsection*{Main Goal.} |
8745 | 138 |
|
139 |
Our goal is to show that reversing a list twice produces the original |
|
11456 | 140 |
list. |
8745 | 141 |
*} |
142 |
||
143 |
theorem rev_rev [simp]: "rev(rev xs) = xs"; |
|
144 |
||
11428 | 145 |
txt{*\index{theorem@\isacommand {theorem} (command)|bold}% |
10795 | 146 |
\noindent |
11456 | 147 |
This \isacommand{theorem} command does several things: |
8745 | 148 |
\begin{itemize} |
149 |
\item |
|
11456 | 150 |
It establishes a new theorem to be proved, namely @{prop"rev(rev xs) = xs"}. |
8745 | 151 |
\item |
11456 | 152 |
It gives that theorem the name @{text"rev_rev"}, for later reference. |
8745 | 153 |
\item |
11456 | 154 |
It tells Isabelle (via the bracketed attribute \attrdx{simp}) to take the eventual theorem as a simplification rule: future proofs involving |
9792 | 155 |
simplification will replace occurrences of @{term"rev(rev xs)"} by |
156 |
@{term"xs"}. |
|
11457 | 157 |
\end{itemize} |
8745 | 158 |
The name and the simplification attribute are optional. |
12332 | 159 |
Isabelle's response is to print the initial proof state consisting |
160 |
of some header information (like how many subgoals there are) followed by |
|
13868 | 161 |
@{subgoals[display,indent=0]} |
12332 | 162 |
For compactness reasons we omit the header in this tutorial. |
163 |
Until we have finished a proof, the \rmindex{proof state} proper |
|
164 |
always looks like this: |
|
9723 | 165 |
\begin{isabelle} |
8745 | 166 |
~1.~$G\sb{1}$\isanewline |
167 |
~~\vdots~~\isanewline |
|
168 |
~$n$.~$G\sb{n}$ |
|
9723 | 169 |
\end{isabelle} |
13868 | 170 |
The numbered lines contain the subgoals $G\sb{1}$, \dots, $G\sb{n}$ |
171 |
that we need to prove to establish the main goal.\index{subgoals} |
|
172 |
Initially there is only one subgoal, which is identical with the |
|
173 |
main goal. (If you always want to see the main goal as well, |
|
174 |
set the flag \isa{Proof.show_main_goal}\index{*show_main_goal (flag)} |
|
175 |
--- this flag used to be set by default.) |
|
8745 | 176 |
|
9792 | 177 |
Let us now get back to @{prop"rev(rev xs) = xs"}. Properties of recursively |
8745 | 178 |
defined functions are best established by induction. In this case there is |
11428 | 179 |
nothing obvious except induction on @{term"xs"}: |
8745 | 180 |
*} |
181 |
||
182 |
apply(induct_tac xs); |
|
183 |
||
11428 | 184 |
txt{*\noindent\index{*induct_tac (method)}% |
9792 | 185 |
This tells Isabelle to perform induction on variable @{term"xs"}. The suffix |
11428 | 186 |
@{term"tac"} stands for \textbf{tactic},\index{tactics} |
187 |
a synonym for ``theorem proving function''. |
|
8745 | 188 |
By default, induction acts on the first subgoal. The new proof state contains |
9792 | 189 |
two subgoals, namely the base case (@{term[source]Nil}) and the induction step |
190 |
(@{term[source]Cons}): |
|
10971 | 191 |
@{subgoals[display,indent=0,margin=65]} |
8745 | 192 |
|
11456 | 193 |
The induction step is an example of the general format of a subgoal:\index{subgoals} |
9723 | 194 |
\begin{isabelle} |
12327 | 195 |
~$i$.~{\isasymAnd}$x\sb{1}$~\dots$x\sb{n}$.~{\it assumptions}~{\isasymLongrightarrow}~{\it conclusion} |
10328 | 196 |
\end{isabelle}\index{$IsaAnd@\isasymAnd|bold} |
8745 | 197 |
The prefix of bound variables \isasymAnd$x\sb{1}$~\dots~$x\sb{n}$ can be |
198 |
ignored most of the time, or simply treated as a list of variables local to |
|
10302 | 199 |
this subgoal. Their deeper significance is explained in Chapter~\ref{chap:rules}. |
11456 | 200 |
The {\it assumptions}\index{assumptions!of subgoal} |
201 |
are the local assumptions for this subgoal and {\it |
|
202 |
conclusion}\index{conclusion!of subgoal} is the actual proposition to be proved. |
|
203 |
Typical proof steps |
|
204 |
that add new assumptions are induction and case distinction. In our example |
|
9541 | 205 |
the only assumption is the induction hypothesis @{term"rev (rev list) = |
9792 | 206 |
list"}, where @{term"list"} is a variable name chosen by Isabelle. If there |
8745 | 207 |
are multiple assumptions, they are enclosed in the bracket pair |
208 |
\indexboldpos{\isasymlbrakk}{$Isabrl} and |
|
209 |
\indexboldpos{\isasymrbrakk}{$Isabrr} and separated by semicolons. |
|
210 |
||
211 |
Let us try to solve both goals automatically: |
|
212 |
*} |
|
213 |
||
214 |
apply(auto); |
|
215 |
||
216 |
txt{*\noindent |
|
217 |
This command tells Isabelle to apply a proof strategy called |
|
9792 | 218 |
@{text"auto"} to all subgoals. Essentially, @{text"auto"} tries to |
10978 | 219 |
simplify the subgoals. In our case, subgoal~1 is solved completely (thanks |
9792 | 220 |
to the equation @{prop"rev [] = []"}) and disappears; the simplified version |
8745 | 221 |
of subgoal~2 becomes the new subgoal~1: |
10971 | 222 |
@{subgoals[display,indent=0,margin=70]} |
8745 | 223 |
In order to simplify this subgoal further, a lemma suggests itself. |
224 |
*} |
|
225 |
(*<*) |
|
226 |
oops |
|
227 |
(*>*) |
|
228 |
||
11428 | 229 |
subsubsection{*First Lemma*} |
9723 | 230 |
|
8745 | 231 |
text{* |
11428 | 232 |
\indexbold{abandoning a proof}\indexbold{proofs!abandoning} |
233 |
After abandoning the above proof attempt (at the shell level type |
|
234 |
\commdx{oops}) we start a new proof: |
|
8745 | 235 |
*} |
236 |
||
237 |
lemma rev_app [simp]: "rev(xs @ ys) = (rev ys) @ (rev xs)"; |
|
238 |
||
11428 | 239 |
txt{*\noindent The keywords \commdx{theorem} and |
240 |
\commdx{lemma} are interchangeable and merely indicate |
|
10971 | 241 |
the importance we attach to a proposition. Therefore we use the words |
11428 | 242 |
\emph{theorem} and \emph{lemma} pretty much interchangeably, too. |
8745 | 243 |
|
9792 | 244 |
There are two variables that we could induct on: @{term"xs"} and |
245 |
@{term"ys"}. Because @{text"@"} is defined by recursion on |
|
246 |
the first argument, @{term"xs"} is the correct one: |
|
8745 | 247 |
*} |
248 |
||
249 |
apply(induct_tac xs); |
|
250 |
||
251 |
txt{*\noindent |
|
252 |
This time not even the base case is solved automatically: |
|
253 |
*} |
|
254 |
||
255 |
apply(auto); |
|
256 |
||
257 |
txt{* |
|
10362 | 258 |
@{subgoals[display,indent=0,goals_limit=1]} |
259 |
Again, we need to abandon this proof attempt and prove another simple lemma |
|
260 |
first. In the future the step of abandoning an incomplete proof before |
|
261 |
embarking on the proof of a lemma usually remains implicit. |
|
8745 | 262 |
*} |
263 |
(*<*) |
|
264 |
oops |
|
265 |
(*>*) |
|
266 |
||
11428 | 267 |
subsubsection{*Second Lemma*} |
9723 | 268 |
|
8745 | 269 |
text{* |
11456 | 270 |
We again try the canonical proof procedure: |
8745 | 271 |
*} |
272 |
||
273 |
lemma app_Nil2 [simp]: "xs @ [] = xs"; |
|
274 |
apply(induct_tac xs); |
|
275 |
apply(auto); |
|
276 |
||
277 |
txt{* |
|
278 |
\noindent |
|
11456 | 279 |
It works, yielding the desired message @{text"No subgoals!"}: |
10362 | 280 |
@{goals[display,indent=0]} |
8745 | 281 |
We still need to confirm that the proof is now finished: |
282 |
*} |
|
283 |
||
10171 | 284 |
done |
8745 | 285 |
|
11428 | 286 |
text{*\noindent |
287 |
As a result of that final \commdx{done}, Isabelle associates the lemma just proved |
|
10171 | 288 |
with its name. In this tutorial, we sometimes omit to show that final \isacommand{done} |
289 |
if it is obvious from the context that the proof is finished. |
|
290 |
||
291 |
% Instead of \isacommand{apply} followed by a dot, you can simply write |
|
292 |
% \isacommand{by}\indexbold{by}, which we do most of the time. |
|
10971 | 293 |
Notice that in lemma @{thm[source]app_Nil2}, |
294 |
as printed out after the final \isacommand{done}, the free variable @{term"xs"} has been |
|
9792 | 295 |
replaced by the unknown @{text"?xs"}, just as explained in |
296 |
\S\ref{sec:variables}. |
|
8745 | 297 |
|
298 |
Going back to the proof of the first lemma |
|
299 |
*} |
|
300 |
||
301 |
lemma rev_app [simp]: "rev(xs @ ys) = (rev ys) @ (rev xs)"; |
|
302 |
apply(induct_tac xs); |
|
303 |
apply(auto); |
|
304 |
||
305 |
txt{* |
|
306 |
\noindent |
|
9792 | 307 |
we find that this time @{text"auto"} solves the base case, but the |
8745 | 308 |
induction step merely simplifies to |
10362 | 309 |
@{subgoals[display,indent=0,goals_limit=1]} |
9792 | 310 |
Now we need to remember that @{text"@"} associates to the right, and that |
311 |
@{text"#"} and @{text"@"} have the same priority (namely the @{text"65"} |
|
8745 | 312 |
in their \isacommand{infixr} annotation). Thus the conclusion really is |
9723 | 313 |
\begin{isabelle} |
9792 | 314 |
~~~~~(rev~ys~@~rev~list)~@~(a~\#~[])~=~rev~ys~@~(rev~list~@~(a~\#~[])) |
9723 | 315 |
\end{isabelle} |
9792 | 316 |
and the missing lemma is associativity of @{text"@"}. |
9723 | 317 |
*} |
318 |
(*<*)oops(*>*) |
|
8745 | 319 |
|
11456 | 320 |
subsubsection{*Third Lemma*} |
8745 | 321 |
|
9723 | 322 |
text{* |
11456 | 323 |
Abandoning the previous attempt, the canonical proof procedure |
324 |
succeeds without further ado. |
|
8745 | 325 |
*} |
326 |
||
327 |
lemma app_assoc [simp]: "(xs @ ys) @ zs = xs @ (ys @ zs)"; |
|
328 |
apply(induct_tac xs); |
|
10171 | 329 |
apply(auto); |
330 |
done |
|
8745 | 331 |
|
332 |
text{* |
|
333 |
\noindent |
|
11456 | 334 |
Now we can prove the first lemma: |
8745 | 335 |
*} |
336 |
||
337 |
lemma rev_app [simp]: "rev(xs @ ys) = (rev ys) @ (rev xs)"; |
|
338 |
apply(induct_tac xs); |
|
10171 | 339 |
apply(auto); |
340 |
done |
|
8745 | 341 |
|
342 |
text{*\noindent |
|
11456 | 343 |
Finally, we prove our main theorem: |
8745 | 344 |
*} |
345 |
||
346 |
theorem rev_rev [simp]: "rev(rev xs) = xs"; |
|
347 |
apply(induct_tac xs); |
|
10171 | 348 |
apply(auto); |
349 |
done |
|
8745 | 350 |
|
351 |
text{*\noindent |
|
11456 | 352 |
The final \commdx{end} tells Isabelle to close the current theory because |
353 |
we are finished with its development:% |
|
354 |
\index{*rev (constant)|)}\index{append function|)} |
|
8745 | 355 |
*} |
356 |
||
357 |
end |