47269
|
1 |
(*<*)
|
|
2 |
theory Bool_nat_list
|
|
3 |
imports Main
|
|
4 |
begin
|
|
5 |
(*>*)
|
|
6 |
|
|
7 |
text{*
|
|
8 |
\vspace{-4ex}
|
|
9 |
\section{\texorpdfstring{Types @{typ bool}, @{typ nat} and @{text list}}{Types bool, nat and list}}
|
|
10 |
|
|
11 |
These are the most important predefined types. We go through them one by one.
|
|
12 |
Based on examples we learn how to define (possibly recursive) functions and
|
|
13 |
prove theorems about them by induction and simplification.
|
|
14 |
|
|
15 |
\subsection{Type @{typ bool}}
|
|
16 |
|
|
17 |
The type of boolean values is a predefined datatype
|
|
18 |
@{datatype[display] bool}
|
|
19 |
with the two values @{const True} and @{const False} and
|
|
20 |
with many predefined functions: @{text "\<not>"}, @{text "\<and>"}, @{text "\<or>"}, @{text
|
|
21 |
"\<longrightarrow>"} etc. Here is how conjunction could be defined by pattern matching:
|
|
22 |
*}
|
|
23 |
|
|
24 |
fun conj :: "bool \<Rightarrow> bool \<Rightarrow> bool" where
|
|
25 |
"conj True True = True" |
|
|
26 |
"conj _ _ = False"
|
|
27 |
|
|
28 |
text{* Both the datatype and function definitions roughly follow the syntax
|
|
29 |
of functional programming languages.
|
|
30 |
|
|
31 |
\subsection{Type @{typ nat}}
|
|
32 |
|
|
33 |
Natural numbers are another predefined datatype:
|
|
34 |
@{datatype[display] nat}
|
|
35 |
All values of type @{typ nat} are generated by the constructors
|
|
36 |
@{text 0} and @{const Suc}. Thus the values of type @{typ nat} are
|
|
37 |
@{text 0}, @{term"Suc 0"}, @{term"Suc(Suc 0)"} etc.
|
|
38 |
There are many predefined functions: @{text "+"}, @{text "*"}, @{text
|
|
39 |
"\<le>"}, etc. Here is how you could define your own addition:
|
|
40 |
*}
|
|
41 |
|
|
42 |
fun add :: "nat \<Rightarrow> nat \<Rightarrow> nat" where
|
|
43 |
"add 0 n = n" |
|
|
44 |
"add (Suc m) n = Suc(add m n)"
|
|
45 |
|
|
46 |
text{* And here is a proof of the fact that @{prop"add m 0 = m"}: *}
|
|
47 |
|
|
48 |
lemma add_02: "add m 0 = m"
|
|
49 |
apply(induction m)
|
|
50 |
apply(auto)
|
|
51 |
done
|
|
52 |
(*<*)
|
|
53 |
lemma "add m 0 = m"
|
|
54 |
apply(induction m)
|
|
55 |
(*>*)
|
|
56 |
txt{* The \isacom{lemma} command starts the proof and gives the lemma
|
|
57 |
a name, @{text add_02}. Properties of recursively defined functions
|
|
58 |
need to be established by induction in most cases.
|
|
59 |
Command \isacom{apply}@{text"(induction m)"} instructs Isabelle to
|
|
60 |
start a proof by induction on @{text m}. In response, it will show the
|
|
61 |
following proof state:
|
|
62 |
@{subgoals[display,indent=0]}
|
|
63 |
The numbered lines are known as \emph{subgoals}.
|
|
64 |
The first subgoal is the base case, the second one the induction step.
|
|
65 |
The prefix @{text"\<And>m."} is Isabelle's way of saying ``for an arbitrary but fixed @{text m}''. The @{text"\<Longrightarrow>"} separates assumptions from the conclusion.
|
|
66 |
The command \isacom{apply}@{text"(auto)"} instructs Isabelle to try
|
|
67 |
and prove all subgoals automatically, essentially by simplifying them.
|
|
68 |
Because both subgoals are easy, Isabelle can do it.
|
|
69 |
The base case @{prop"add 0 0 = 0"} holds by definition of @{const add},
|
|
70 |
and the induction step is almost as simple:
|
|
71 |
@{text"add\<^raw:~>(Suc m) 0 = Suc(add m 0) = Suc m"}
|
|
72 |
using first the definition of @{const add} and then the induction hypothesis.
|
|
73 |
In summary, both subproofs rely on simplification with function definitions and
|
|
74 |
the induction hypothesis.
|
|
75 |
As a result of that final \isacom{done}, Isabelle associates the lemma
|
|
76 |
just proved with its name. You can now inspect the lemma with the command
|
|
77 |
*}
|
|
78 |
|
|
79 |
thm add_02
|
|
80 |
|
|
81 |
txt{* which displays @{thm[show_question_marks,display] add_02} The free
|
|
82 |
variable @{text m} has been replaced by the \concept{unknown}
|
|
83 |
@{text"?m"}. There is no logical difference between the two but an
|
|
84 |
operational one: unknowns can be instantiated, which is what you want after
|
|
85 |
some lemma has been proved.
|
|
86 |
|
|
87 |
Note that there is also a proof method @{text induct}, which behaves almost
|
|
88 |
like @{text induction}; the difference is explained in \autoref{ch:Isar}.
|
|
89 |
|
|
90 |
\begin{warn}
|
|
91 |
Terminology: We use \concept{lemma}, \concept{theorem} and \concept{rule}
|
|
92 |
interchangeably for propositions that have been proved.
|
|
93 |
\end{warn}
|
|
94 |
\begin{warn}
|
|
95 |
Numerals (@{text 0}, @{text 1}, @{text 2}, \dots) and most of the standard
|
|
96 |
arithmetic operations (@{text "+"}, @{text "-"}, @{text "*"}, @{text"\<le>"},
|
|
97 |
@{text"<"} etc) are overloaded: they are available
|
|
98 |
not just for natural numbers but for other types as well.
|
|
99 |
For example, given the goal @{text"x + 0 = x"}, there is nothing to indicate
|
|
100 |
that you are talking about natural numbers. Hence Isabelle can only infer
|
|
101 |
that @{term x} is of some arbitrary type where @{text 0} and @{text"+"}
|
|
102 |
exist. As a consequence, you will be unable to prove the
|
|
103 |
goal. To alert you to such pitfalls, Isabelle flags numerals without a
|
|
104 |
fixed type in its output: @{prop"x+0 = x"}. In this particular example,
|
|
105 |
you need to include
|
|
106 |
an explicit type constraint, for example @{text"x+0 = (x::nat)"}. If there
|
|
107 |
is enough contextual information this may not be necessary: @{prop"Suc x =
|
|
108 |
x"} automatically implies @{text"x::nat"} because @{term Suc} is not
|
|
109 |
overloaded.
|
|
110 |
\end{warn}
|
|
111 |
|
|
112 |
\subsubsection{An informal proof}
|
|
113 |
|
|
114 |
Above we gave some terse informal explanation of the proof of
|
|
115 |
@{prop"add m 0 = m"}. A more detailed informal exposition of the lemma
|
|
116 |
might look like this:
|
|
117 |
\bigskip
|
|
118 |
|
|
119 |
\noindent
|
|
120 |
\textbf{Lemma} @{prop"add m 0 = m"}
|
|
121 |
|
|
122 |
\noindent
|
|
123 |
\textbf{Proof} by induction on @{text m}.
|
|
124 |
\begin{itemize}
|
|
125 |
\item Case @{text 0} (the base case): @{prop"add 0 0 = 0"}
|
|
126 |
holds by definition of @{const add}.
|
|
127 |
\item Case @{term"Suc m"} (the induction step):
|
|
128 |
We assume @{prop"add m 0 = m"}, the induction hypothesis (IH),
|
|
129 |
and we need to show @{text"add (Suc m) 0 = Suc m"}.
|
|
130 |
The proof is as follows:\smallskip
|
|
131 |
|
|
132 |
\begin{tabular}{@ {}rcl@ {\quad}l@ {}}
|
|
133 |
@{term "add (Suc m) 0"} &@{text"="}& @{term"Suc(add m 0)"}
|
|
134 |
& by definition of @{text add}\\
|
|
135 |
&@{text"="}& @{term "Suc m"} & by IH
|
|
136 |
\end{tabular}
|
|
137 |
\end{itemize}
|
|
138 |
Throughout this book, \concept{IH} will stand for ``induction hypothesis''.
|
|
139 |
|
|
140 |
We have now seen three proofs of @{prop"add m 0 = 0"}: the Isabelle one, the
|
|
141 |
terse 4 lines explaining the base case and the induction step, and just now a
|
|
142 |
model of a traditional inductive proof. The three proofs differ in the level
|
|
143 |
of detail given and the intended reader: the Isabelle proof is for the
|
|
144 |
machine, the informal proofs are for humans. Although this book concentrates
|
|
145 |
of Isabelle proofs, it is important to be able to rephrase those proofs
|
|
146 |
as informal text comprehensible to a reader familiar with traditional
|
|
147 |
mathematical proofs. Later on we will introduce an Isabelle proof language
|
|
148 |
that is closer to traditional informal mathematical language and is often
|
|
149 |
directly readable.
|
|
150 |
|
|
151 |
\subsection{Type @{text list}}
|
|
152 |
|
|
153 |
Although lists are already predefined, we define our own copy just for
|
|
154 |
demonstration purposes:
|
|
155 |
*}
|
|
156 |
(*<*)
|
|
157 |
apply(auto)
|
|
158 |
done
|
|
159 |
declare [[names_short]]
|
|
160 |
(*>*)
|
47302
|
161 |
datatype 'a list = Nil | Cons 'a "'a list"
|
47269
|
162 |
|
47302
|
163 |
text{*
|
|
164 |
\begin{itemize}
|
|
165 |
\item Type @{typ "'a list"} is the type of list over elements of type @{typ 'a}. Because @{typ 'a} is a type variable, lists are in fact \concept{polymorphic}: the elements of a list can be of arbitrary type (but must all be of the same type).
|
|
166 |
\item Lists have two constructors: @{const Nil}, the empty list, and @{const Cons}, which puts an element (of type @{typ 'a}) in front of a list (of type @{typ "'a list"}).
|
|
167 |
Hence all lists are of the form @{const Nil}, or @{term"Cons x Nil"},
|
47269
|
168 |
or @{term"Cons x (Cons y Nil)"} etc.
|
47302
|
169 |
\item \isacom{datatype} requires no quotation marks on the
|
|
170 |
left-hand side, but on the right-hand side each of the argument
|
|
171 |
types of a constructor needs to be enclosed in quotation marks, unless
|
|
172 |
it is just an identifier (e.g.\ @{typ nat} or @{typ 'a}).
|
|
173 |
\end{itemize}
|
47269
|
174 |
We also define two standard functions, append and reverse: *}
|
|
175 |
|
|
176 |
fun app :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" where
|
|
177 |
"app Nil ys = ys" |
|
|
178 |
"app (Cons x xs) ys = Cons x (app xs ys)"
|
|
179 |
|
|
180 |
fun rev :: "'a list \<Rightarrow> 'a list" where
|
|
181 |
"rev Nil = Nil" |
|
|
182 |
"rev (Cons x xs) = app (rev xs) (Cons x Nil)"
|
|
183 |
|
|
184 |
text{* By default, variables @{text xs}, @{text ys} and @{text zs} are of
|
|
185 |
@{text list} type.
|
|
186 |
|
|
187 |
Command \isacom{value} evaluates a term. For example, *}
|
|
188 |
|
|
189 |
value "rev(Cons True (Cons False Nil))"
|
|
190 |
|
|
191 |
text{* yields the result @{value "rev(Cons True (Cons False Nil))"}. This works symbolically, too: *}
|
|
192 |
|
|
193 |
value "rev(Cons a (Cons b Nil))"
|
|
194 |
|
|
195 |
text{* yields @{value "rev(Cons a (Cons b Nil))"}.
|
|
196 |
\medskip
|
|
197 |
|
47302
|
198 |
Figure~\ref{fig:MyList} shows the theory created so far.
|
|
199 |
% Notice where the
|
|
200 |
%quotations marks are needed that we mostly sweep under the carpet. In
|
|
201 |
%particular, notice that \isacom{datatype} requires no quotation marks on the
|
|
202 |
%left-hand side, but that on the right-hand side each of the argument
|
|
203 |
%types of a constructor needs to be enclosed in quotation marks.
|
47269
|
204 |
|
|
205 |
\begin{figure}[htbp]
|
|
206 |
\begin{alltt}
|
|
207 |
\input{Thys/MyList.thy}\end{alltt}
|
|
208 |
\caption{A Theory of Lists}
|
|
209 |
\label{fig:MyList}
|
|
210 |
\end{figure}
|
|
211 |
|
|
212 |
\subsubsection{Structural Induction for Lists}
|
|
213 |
|
|
214 |
Just as for natural numbers, there is a proof principle of induction for
|
|
215 |
lists. Induction over a list is essentially induction over the length of
|
|
216 |
the list, although the length remains implicit. To prove that some property
|
|
217 |
@{text P} holds for all lists @{text xs}, i.e.\ \mbox{@{prop"P(xs)"}},
|
|
218 |
you need to prove
|
|
219 |
\begin{enumerate}
|
|
220 |
\item the base case @{prop"P(Nil)"} and
|
|
221 |
\item the inductive case @{prop"P(Cons x xs)"} under the assumption @{prop"P(xs)"}, for some arbitrary but fixed @{text xs}.
|
|
222 |
\end{enumerate}
|
|
223 |
This is often called \concept{structural induction}.
|
|
224 |
|
|
225 |
\subsection{The Proof Process}
|
|
226 |
|
|
227 |
We will now demonstrate the typical proof process, which involves
|
|
228 |
the formulation and proof of auxiliary lemmas.
|
|
229 |
Our goal is to show that reversing a list twice produces the original
|
|
230 |
list. *}
|
|
231 |
|
|
232 |
theorem rev_rev [simp]: "rev(rev xs) = xs"
|
|
233 |
|
|
234 |
txt{* Commands \isacom{theorem} and \isacom{lemma} are
|
|
235 |
interchangeable and merely indicate the importance we attach to a
|
|
236 |
proposition. Via the bracketed attribute @{text simp} we also tell Isabelle
|
|
237 |
to make the eventual theorem a \concept{simplification rule}: future proofs
|
|
238 |
involving simplification will replace occurrences of @{term"rev(rev xs)"} by
|
|
239 |
@{term"xs"}. The proof is by induction: *}
|
|
240 |
|
|
241 |
apply(induction xs)
|
|
242 |
|
|
243 |
txt{*
|
|
244 |
As explained above, we obtain two subgoals, namely the base case (@{const Nil}) and the induction step (@{const Cons}):
|
|
245 |
@{subgoals[display,indent=0,margin=65]}
|
|
246 |
Let us try to solve both goals automatically:
|
|
247 |
*}
|
|
248 |
|
|
249 |
apply(auto)
|
|
250 |
|
|
251 |
txt{*Subgoal~1 is proved, and disappears; the simplified version
|
|
252 |
of subgoal~2 becomes the new subgoal~1:
|
|
253 |
@{subgoals[display,indent=0,margin=70]}
|
|
254 |
In order to simplify this subgoal further, a lemma suggests itself.
|
|
255 |
|
|
256 |
\subsubsection{A First Lemma}
|
|
257 |
|
|
258 |
We insert the following lemma in front of the main theorem:
|
|
259 |
*}
|
|
260 |
(*<*)
|
|
261 |
oops
|
|
262 |
(*>*)
|
|
263 |
lemma rev_app [simp]: "rev(app xs ys) = app (rev ys) (rev xs)"
|
|
264 |
|
|
265 |
txt{* There are two variables that we could induct on: @{text xs} and
|
|
266 |
@{text ys}. Because @{const app} is defined by recursion on
|
|
267 |
the first argument, @{text xs} is the correct one:
|
|
268 |
*}
|
|
269 |
|
|
270 |
apply(induction xs)
|
|
271 |
|
|
272 |
txt{* This time not even the base case is solved automatically: *}
|
|
273 |
apply(auto)
|
|
274 |
txt{*
|
|
275 |
\vspace{-5ex}
|
|
276 |
@{subgoals[display,goals_limit=1]}
|
|
277 |
Again, we need to abandon this proof attempt and prove another simple lemma
|
|
278 |
first.
|
|
279 |
|
|
280 |
\subsubsection{A Second Lemma}
|
|
281 |
|
|
282 |
We again try the canonical proof procedure:
|
|
283 |
*}
|
|
284 |
(*<*)
|
|
285 |
oops
|
|
286 |
(*>*)
|
|
287 |
lemma app_Nil2 [simp]: "app xs Nil = xs"
|
|
288 |
apply(induction xs)
|
|
289 |
apply(auto)
|
|
290 |
done
|
|
291 |
|
|
292 |
text{*
|
|
293 |
Thankfully, this worked.
|
|
294 |
Now we can continue with our stuck proof attempt of the first lemma:
|
|
295 |
*}
|
|
296 |
|
|
297 |
lemma rev_app [simp]: "rev(app xs ys) = app (rev ys) (rev xs)"
|
|
298 |
apply(induction xs)
|
|
299 |
apply(auto)
|
|
300 |
|
|
301 |
txt{*
|
|
302 |
We find that this time @{text"auto"} solves the base case, but the
|
|
303 |
induction step merely simplifies to
|
|
304 |
@{subgoals[display,indent=0,goals_limit=1]}
|
|
305 |
The the missing lemma is associativity of @{const app},
|
|
306 |
which we insert in front of the failed lemma @{text rev_app}.
|
|
307 |
|
|
308 |
\subsubsection{Associativity of @{const app}}
|
|
309 |
|
|
310 |
The canonical proof procedure succeeds without further ado:
|
|
311 |
*}
|
|
312 |
(*<*)oops(*>*)
|
|
313 |
lemma app_assoc [simp]: "app (app xs ys) zs = app xs (app ys zs)"
|
|
314 |
apply(induction xs)
|
|
315 |
apply(auto)
|
|
316 |
done
|
|
317 |
(*<*)
|
|
318 |
lemma rev_app [simp]: "rev(app xs ys) = app (rev ys)(rev xs)"
|
|
319 |
apply(induction xs)
|
|
320 |
apply(auto)
|
|
321 |
done
|
|
322 |
|
|
323 |
theorem rev_rev [simp]: "rev(rev xs) = xs"
|
|
324 |
apply(induction xs)
|
|
325 |
apply(auto)
|
|
326 |
done
|
|
327 |
(*>*)
|
|
328 |
text{*
|
|
329 |
Finally the proofs of @{thm[source] rev_app} and @{thm[source] rev_rev}
|
|
330 |
succeed, too.
|
|
331 |
|
|
332 |
\subsubsection{Another informal proof}
|
|
333 |
|
|
334 |
Here is the informal proof of associativity of @{const app}
|
|
335 |
corresponding to the Isabelle proof above.
|
|
336 |
\bigskip
|
|
337 |
|
|
338 |
\noindent
|
|
339 |
\textbf{Lemma} @{prop"app (app xs ys) zs = app xs (app ys zs)"}
|
|
340 |
|
|
341 |
\noindent
|
|
342 |
\textbf{Proof} by induction on @{text xs}.
|
|
343 |
\begin{itemize}
|
|
344 |
\item Case @{text Nil}: \ @{prop"app (app Nil ys) zs = app ys zs"} @{text"="}
|
|
345 |
\mbox{@{term"app Nil (app ys zs)"}} \ holds by definition of @{text app}.
|
|
346 |
\item Case @{text"Cons x xs"}: We assume
|
|
347 |
\begin{center} \hfill @{term"app (app xs ys) zs"} @{text"="}
|
|
348 |
@{term"app xs (app ys zs)"} \hfill (IH) \end{center}
|
|
349 |
and we need to show
|
|
350 |
\begin{center} @{prop"app (app (Cons x xs) ys) zs = app (Cons x xs) (app ys zs)"}.\end{center}
|
|
351 |
The proof is as follows:\smallskip
|
|
352 |
|
|
353 |
\begin{tabular}{@ {}l@ {\quad}l@ {}}
|
|
354 |
@{term"app (app (Cons x xs) ys) zs"}\\
|
|
355 |
@{text"= app (Cons x (app xs ys)) zs"} & by definition of @{text app}\\
|
|
356 |
@{text"= Cons x (app (app xs ys) zs)"} & by definition of @{text app}\\
|
|
357 |
@{text"= Cons x (app xs (app ys zs))"} & by IH\\
|
|
358 |
@{text"= app (Cons x xs) (app ys zs)"} & by definition of @{text app}
|
|
359 |
\end{tabular}
|
|
360 |
\end{itemize}
|
|
361 |
\medskip
|
|
362 |
|
|
363 |
\noindent Didn't we say earlier that all proofs are by simplification? But
|
|
364 |
in both cases, going from left to right, the last equality step is not a
|
|
365 |
simplification at all! In the base case it is @{prop"app ys zs = app Nil (app
|
|
366 |
ys zs)"}. It appears almost mysterious because we suddenly complicate the
|
|
367 |
term by appending @{text Nil} on the left. What is really going on is this:
|
|
368 |
when proving some equality \mbox{@{prop"s = t"}}, both @{text s} and @{text t} are
|
|
369 |
simplified to some common term @{text u}. This heuristic for equality proofs
|
|
370 |
works well for a functional programming context like ours. In the base case
|
|
371 |
@{text s} is @{term"app (app Nil ys) zs"}, @{text t} is @{term"app Nil (app
|
|
372 |
ys zs)"}, and @{text u} is @{term"app ys zs"}.
|
|
373 |
|
|
374 |
\subsection{Predefined lists}
|
|
375 |
\label{sec:predeflists}
|
|
376 |
|
|
377 |
Isabelle's predefined lists are the same as the ones above, but with
|
|
378 |
more syntactic sugar:
|
|
379 |
\begin{itemize}
|
|
380 |
\item @{text "[]"} is @{const Nil},
|
|
381 |
\item @{term"x # xs"} is @{term"Cons x xs"},
|
|
382 |
\item @{text"[x\<^isub>1, \<dots>, x\<^isub>n]"} is @{text"x\<^isub>1 # \<dots> # x\<^isub>n # []"}, and
|
|
383 |
\item @{term "xs @ ys"} is @{term"app xs ys"}.
|
|
384 |
\end{itemize}
|
|
385 |
There is also a large library of predefined functions.
|
|
386 |
The most important ones are the length function
|
|
387 |
@{text"length :: 'a list \<Rightarrow> nat"} (with the obvious definition),
|
|
388 |
and the map function that applies a function to all elements of a list:
|
|
389 |
\begin{isabelle}
|
47306
|
390 |
\isacom{fun} @{const map} @{text"::"} @{typ[source] "('a \<Rightarrow> 'b) \<Rightarrow> 'a list \<Rightarrow> 'b list"}\\
|
|
391 |
@{text"\""}@{thm map.simps(1)}@{text"\" |"}\\
|
|
392 |
@{text"\""}@{thm map.simps(2)}@{text"\""}
|
47269
|
393 |
\end{isabelle}
|
47306
|
394 |
\sem
|
47269
|
395 |
Also useful are the \concept{head} of a list, its first element,
|
|
396 |
and the \concept{tail}, the rest of the list:
|
|
397 |
\begin{isabelle}
|
|
398 |
\isacom{fun} @{text"hd :: 'a list \<Rightarrow> 'a"}\\
|
|
399 |
@{prop"hd(x#xs) = x"}
|
|
400 |
\end{isabelle}
|
|
401 |
\begin{isabelle}
|
|
402 |
\isacom{fun} @{text"tl :: 'a list \<Rightarrow> 'a list"}\\
|
|
403 |
@{prop"tl [] = []"} @{text"|"}\\
|
|
404 |
@{prop"tl(x#xs) = xs"}
|
|
405 |
\end{isabelle}
|
|
406 |
Note that since HOL is a logic of total functions, @{term"hd []"} is defined,
|
|
407 |
but we do now know what the result is. That is, @{term"hd []"} is not undefined
|
|
408 |
but underdefined.
|
47306
|
409 |
\endsem
|
|
410 |
%
|
47269
|
411 |
*}
|
|
412 |
(*<*)
|
|
413 |
end
|
|
414 |
(*>*)
|