Theory Bool_nat_list
theory Bool_nat_list
imports Complex_Main
begin
text‹
\vspace{-4ex}
\section{\texorpdfstring{Types \<^typ>‹bool›, \<^typ>‹nat› and ‹list›}{Types bool, nat and list}}
These are the most important predefined types. We go through them one by one.
Based on examples we learn how to define (possibly recursive) functions and
prove theorems about them by induction and simplification.
\subsection{Type \indexed{\<^typ>‹bool›}{bool}}
The type of boolean values is a predefined datatype
@{datatype[display] bool}
with the two values \indexed{\<^const>‹True›}{True} and \indexed{\<^const>‹False›}{False} and
with many predefined functions: ‹¬›, ‹∧›, ‹∨›, ‹⟶›, etc. Here is how conjunction could be defined by pattern matching:
›
fun conj :: "bool ⇒ bool ⇒ bool" where
"conj True True = True" |
"conj _ _ = False"
text‹Both the datatype and function definitions roughly follow the syntax
of functional programming languages.
\subsection{Type \indexed{\<^typ>‹nat›}{nat}}
Natural numbers are another predefined datatype:
@{datatype[display] nat}\index{Suc@\<^const>‹Suc›}
All values of type \<^typ>‹nat› are generated by the constructors
‹0› and \<^const>‹Suc›. Thus the values of type \<^typ>‹nat› are
‹0›, \<^term>‹Suc 0›, \<^term>‹Suc(Suc 0)›, etc.
There are many predefined functions: ‹+›, ‹*›, ‹≤›, etc. Here is how you could define your own addition:
›
fun add :: "nat ⇒ nat ⇒ nat" where
"add 0 n = n" |
"add (Suc m) n = Suc(add m n)"
text‹And here is a proof of the fact that \<^prop>‹add m 0 = m›:›
lemma add_02: "add m 0 = m"
apply(induction m)
apply(auto)
done
lemma "add m 0 = m"
apply(induction m)
txt‹The \isacom{lemma} command starts the proof and gives the lemma
a name, ‹add_02›. Properties of recursively defined functions
need to be established by induction in most cases.
Command \isacom{apply}‹(induction m)› instructs Isabelle to
start a proof by induction on ‹m›. In response, it will show the
following proof state\ifsem\footnote{See page \pageref{proof-state} for how to
display the proof state.}\fi:
@{subgoals[display,indent=0]}
The numbered lines are known as \emph{subgoals}.
The first subgoal is the base case, the second one the induction step.
The prefix ‹⋀m.› is Isabelle's way of saying ``for an arbitrary but fixed ‹m›''. The ‹⟹› separates assumptions from the conclusion.
The command \isacom{apply}‹(auto)› instructs Isabelle to try
and prove all subgoals automatically, essentially by simplifying them.
Because both subgoals are easy, Isabelle can do it.
The base case \<^prop>‹add 0 0 = 0› holds by definition of \<^const>‹add›,
and the induction step is almost as simple:
‹add\<^latex>‹~›(Suc m) 0 = Suc(add m 0) = Suc m›
using first the definition of \<^const>‹add› and then the induction hypothesis.
In summary, both subproofs rely on simplification with function definitions and
the induction hypothesis.
As a result of that final \isacom{done}, Isabelle associates the lemma
just proved with its name. You can now inspect the lemma with the command
›
thm add_02
txt‹which displays @{thm[show_question_marks,display] add_02} The free
variable ‹m› has been replaced by the \concept{unknown}
‹?m›. There is no logical difference between the two but there is an
operational one: unknowns can be instantiated, which is what you want after
some lemma has been proved.
Note that there is also a proof method ‹induct›, which behaves almost
like ‹induction›; the difference is explained in \autoref{ch:Isar}.
\begin{warn}
Terminology: We use \concept{lemma}, \concept{theorem} and \concept{rule}
interchangeably for propositions that have been proved.
\end{warn}
\begin{warn}
Numerals (‹0›, ‹1›, ‹2›, \dots) and most of the standard
arithmetic operations (‹+›, ‹-›, ‹*›, ‹≤›,
‹<›, etc.) are overloaded: they are available
not just for natural numbers but for other types as well.
For example, given the goal ‹x + 0 = x›, there is nothing to indicate
that you are talking about natural numbers. Hence Isabelle can only infer
that \<^term>‹x› is of some arbitrary type where ‹0› and ‹+›
exist. As a consequence, you will be unable to prove the goal.
% To alert you to such pitfalls, Isabelle flags numerals without a
% fixed type in its output: @ {prop"x+0 = x"}.
In this particular example, you need to include
an explicit type constraint, for example ‹x+0 = (x::nat)›. If there
is enough contextual information this may not be necessary: \<^prop>‹Suc x =
x› automatically implies ‹x::nat› because \<^term>‹Suc› is not
overloaded.
\end{warn}
\subsubsection{An Informal Proof}
Above we gave some terse informal explanation of the proof of
\<^prop>‹add m 0 = m›. A more detailed informal exposition of the lemma
might look like this:
\bigskip
\noindent
\textbf{Lemma} \<^prop>‹add m 0 = m›
\noindent
\textbf{Proof} by induction on ‹m›.
\begin{itemize}
\item Case ‹0› (the base case): \<^prop>‹add 0 0 = 0›
holds by definition of \<^const>‹add›.
\item Case \<^term>‹Suc m› (the induction step):
We assume \<^prop>‹add m 0 = m›, the induction hypothesis (IH),
and we need to show ‹add (Suc m) 0 = Suc m›.
The proof is as follows:\smallskip
\begin{tabular}{@ {}rcl@ {\quad}l@ {}}
\<^term>‹add (Suc m) 0› &‹=›& \<^term>‹Suc(add m 0)›
& by definition of ‹add›\\
&‹=›& \<^term>‹Suc m› & by IH
\end{tabular}
\end{itemize}
Throughout this book, \concept{IH} will stand for ``induction hypothesis''.
We have now seen three proofs of \<^prop>‹add m 0 = 0›: the Isabelle one, the
terse four lines explaining the base case and the induction step, and just now a
model of a traditional inductive proof. The three proofs differ in the level
of detail given and the intended reader: the Isabelle proof is for the
machine, the informal proofs are for humans. Although this book concentrates
on Isabelle proofs, it is important to be able to rephrase those proofs
as informal text comprehensible to a reader familiar with traditional
mathematical proofs. Later on we will introduce an Isabelle proof language
that is closer to traditional informal mathematical language and is often
directly readable.
\subsection{Type \indexed{‹list›}{list}}
Although lists are already predefined, we define our own copy for
demonstration purposes:
›
apply(auto)
done
declare [[names_short]]