doc-src/TutorialI/fp.tex
author chaieb
Wed, 19 May 2004 11:23:59 +0200
changeset 14758 af3b71a46a1c
parent 13305 f88d0c363582
child 15358 26c501c5024d
permissions -rw-r--r--
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller. the tactic has also changed and allows the abstaction over fuction occurences whose type is nat or int.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
11419
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
     1
\chapter{Functional Programming in HOL}
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
     2
11450
1b02a6c4032f tweaks and indexing
paulson
parents: 11428
diff changeset
     3
This chapter describes how to write
1b02a6c4032f tweaks and indexing
paulson
parents: 11428
diff changeset
     4
functional programs in HOL and how to verify them.  However, 
1b02a6c4032f tweaks and indexing
paulson
parents: 11428
diff changeset
     5
most of the constructs and
1b02a6c4032f tweaks and indexing
paulson
parents: 11428
diff changeset
     6
proof procedures introduced are general and recur in any specification
1b02a6c4032f tweaks and indexing
paulson
parents: 11428
diff changeset
     7
or verification task.  We really should speak of functional
1b02a6c4032f tweaks and indexing
paulson
parents: 11428
diff changeset
     8
\emph{modelling} rather than functional \emph{programming}: 
1b02a6c4032f tweaks and indexing
paulson
parents: 11428
diff changeset
     9
our primary aim is not
11419
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
    10
to write programs but to design abstract models of systems.  HOL is
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
    11
a specification language that goes well beyond what can be expressed as a
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
    12
program. However, for the time being we concentrate on the computable.
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
    13
11450
1b02a6c4032f tweaks and indexing
paulson
parents: 11428
diff changeset
    14
If you are a purist functional programmer, please note that all functions
1b02a6c4032f tweaks and indexing
paulson
parents: 11428
diff changeset
    15
in HOL must be total:
1b02a6c4032f tweaks and indexing
paulson
parents: 11428
diff changeset
    16
they must terminate for all inputs.  Lazy data structures are not
11419
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
    17
directly available.
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
    18
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
    19
\section{An Introductory Theory}
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
    20
\label{sec:intro-theory}
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
    21
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
    22
Functional programming needs datatypes and functions. Both of them can be
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
    23
defined in a theory with a syntax reminiscent of languages like ML or
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
    24
Haskell. As an example consider the theory in figure~\ref{fig:ToyList}.
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
    25
We will now examine it line by line.
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
    26
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
    27
\begin{figure}[htbp]
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
    28
\begin{ttbox}\makeatother
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
    29
\input{ToyList2/ToyList1}\end{ttbox}
11450
1b02a6c4032f tweaks and indexing
paulson
parents: 11428
diff changeset
    30
\caption{A Theory of Lists}
11419
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
    31
\label{fig:ToyList}
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
    32
\end{figure}
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
    33
11457
279da0358aa9 additional revisions to chapters 1, 2
paulson
parents: 11456
diff changeset
    34
\index{*ToyList example|(}
11419
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
    35
{\makeatother\input{ToyList/document/ToyList.tex}}
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
    36
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
    37
The complete proof script is shown in Fig.\ts\ref{fig:ToyList-proofs}. The
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
    38
concatenation of Figs.\ts\ref{fig:ToyList} and~\ref{fig:ToyList-proofs}
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
    39
constitutes the complete theory \texttt{ToyList} and should reside in file
12327
5a4d78204492 *** empty log message ***
nipkow
parents: 11646
diff changeset
    40
\texttt{ToyList.thy}.
5a4d78204492 *** empty log message ***
nipkow
parents: 11646
diff changeset
    41
% It is good practice to present all declarations and
5a4d78204492 *** empty log message ***
nipkow
parents: 11646
diff changeset
    42
%definitions at the beginning of a theory to facilitate browsing.%
11457
279da0358aa9 additional revisions to chapters 1, 2
paulson
parents: 11456
diff changeset
    43
\index{*ToyList example|)}
11419
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
    44
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
    45
\begin{figure}[htbp]
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
    46
\begin{ttbox}\makeatother
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
    47
\input{ToyList2/ToyList2}\end{ttbox}
11450
1b02a6c4032f tweaks and indexing
paulson
parents: 11428
diff changeset
    48
\caption{Proofs about Lists}
11419
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
    49
\label{fig:ToyList-proofs}
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
    50
\end{figure}
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
    51
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
    52
\subsubsection*{Review}
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
    53
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
    54
This is the end of our toy proof. It should have familiarized you with
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
    55
\begin{itemize}
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
    56
\item the standard theorem proving procedure:
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
    57
state a goal (lemma or theorem); proceed with proof until a separate lemma is
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
    58
required; prove that lemma; come back to the original goal.
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
    59
\item a specific procedure that works well for functional programs:
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
    60
induction followed by all-out simplification via \isa{auto}.
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
    61
\item a basic repertoire of proof commands.
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
    62
\end{itemize}
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
    63
12332
aea72a834c85 *** empty log message ***
nipkow
parents: 12328
diff changeset
    64
\begin{warn}
aea72a834c85 *** empty log message ***
nipkow
parents: 12328
diff changeset
    65
It is tempting to think that all lemmas should have the \isa{simp} attribute
aea72a834c85 *** empty log message ***
nipkow
parents: 12328
diff changeset
    66
just because this was the case in the example above. However, in that example
aea72a834c85 *** empty log message ***
nipkow
parents: 12328
diff changeset
    67
all lemmas were equations, and the right-hand side was simpler than the
aea72a834c85 *** empty log message ***
nipkow
parents: 12328
diff changeset
    68
left-hand side --- an ideal situation for simplification purposes. Unless
aea72a834c85 *** empty log message ***
nipkow
parents: 12328
diff changeset
    69
this is clearly the case, novices should refrain from awarding a lemma the
aea72a834c85 *** empty log message ***
nipkow
parents: 12328
diff changeset
    70
\isa{simp} attribute, which has a global effect. Instead, lemmas can be
aea72a834c85 *** empty log message ***
nipkow
parents: 12328
diff changeset
    71
applied locally where they are needed, which is discussed in the following
aea72a834c85 *** empty log message ***
nipkow
parents: 12328
diff changeset
    72
chapter.
aea72a834c85 *** empty log message ***
nipkow
parents: 12328
diff changeset
    73
\end{warn}
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
    74
10885
90695f46440b lcp's pass over the book, chapters 1-8
paulson
parents: 10824
diff changeset
    75
\section{Some Helpful Commands}
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
    76
\label{sec:commands-and-hints}
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
    77
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
    78
This section discusses a few basic commands for manipulating the proof state
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
    79
and can be skipped by casual readers.
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
    80
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
    81
There are two kinds of commands used during a proof: the actual proof
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
    82
commands and auxiliary commands for examining the proof state and controlling
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
    83
the display. Simple proof commands are of the form
12327
5a4d78204492 *** empty log message ***
nipkow
parents: 11646
diff changeset
    84
\commdx{apply}(\textit{method}), where \textit{method} is typically 
11419
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
    85
\isa{induct_tac} or \isa{auto}.  All such theorem proving operations
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
    86
are referred to as \bfindex{methods}, and further ones are
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
    87
introduced throughout the tutorial.  Unless stated otherwise, you may
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
    88
assume that a method attacks merely the first subgoal. An exception is
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
    89
\isa{auto}, which tries to solve all subgoals.
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
    90
11419
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
    91
The most useful auxiliary commands are as follows:
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
    92
\begin{description}
11419
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
    93
\item[Undoing:] \commdx{undo} undoes the effect of
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
    94
the
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
    95
  last command; \isacommand{undo} can be undone by
11419
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
    96
  \commdx{redo}.  Both are only needed at the shell
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
    97
  level and should not occur in the final theory.
11419
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
    98
\item[Printing the current state:] \commdx{pr}
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
    99
redisplays
11302
9e0708bb15f7 minor revisons
paulson
parents: 11294
diff changeset
   100
  the current proof state, for example when it has scrolled past the top of
9e0708bb15f7 minor revisons
paulson
parents: 11294
diff changeset
   101
  the screen.
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   102
\item[Limiting the number of subgoals:] \isacommand{pr}~$n$ tells Isabelle to
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   103
  print only the first $n$ subgoals from now on and redisplays the current
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   104
  proof state. This is helpful when there are many subgoals.
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   105
\item[Modifying the order of subgoals:]
11419
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   106
\commdx{defer} moves the first subgoal to the end and
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   107
\commdx{prefer}~$n$ moves subgoal $n$ to the front.
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   108
\item[Printing theorems:]
11419
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   109
  \commdx{thm}~\textit{name}$@1$~\dots~\textit{name}$@n$
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   110
  prints the named theorems.
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   111
\item[Displaying types:] We have already mentioned the flag
11456
7eb63f63e6c6 revisions and indexing
paulson
parents: 11450
diff changeset
   112
  \texttt{show_types} above.\index{*show_types (flag)}
7eb63f63e6c6 revisions and indexing
paulson
parents: 11450
diff changeset
   113
  It can also be useful for detecting misspellings in
7eb63f63e6c6 revisions and indexing
paulson
parents: 11450
diff changeset
   114
  formulae. For example, if \texttt{show_types} is set and the goal
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   115
  \isa{rev(rev xs) = xs} is started, Isabelle prints the additional output
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   116
\par\noindent
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   117
\begin{isabelle}%
11646
6a7d80a139c6 updated output;
wenzelm
parents: 11458
diff changeset
   118
variables:\isanewline
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   119
~~xs~::~'a~list
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   120
\end{isabelle}%
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   121
\par\noindent
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   122
which tells us that Isabelle has correctly inferred that
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   123
\isa{xs} is a variable of list type. On the other hand, had we
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   124
made a typo as in \isa{rev(re xs) = xs}, the response
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   125
\par\noindent
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   126
\begin{isabelle}%
11646
6a7d80a139c6 updated output;
wenzelm
parents: 11458
diff changeset
   127
variables:\isanewline
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   128
~~re~::~'a~list~{\isasymRightarrow}~'a~list\isanewline
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   129
~~xs~::~'a~list%
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   130
\end{isabelle}%
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   131
\par\noindent
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   132
would have alerted us because of the unexpected variable \isa{re}.
11419
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   133
\item[Reading terms and types:] \commdx{term}
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   134
  \textit{string} reads, type-checks and prints the given string as a term in
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   135
  the current context; the inferred type is output as well.
11419
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   136
  \commdx{typ} \textit{string} reads and prints the given
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   137
  string as a type in the current context.
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   138
\item[(Re)loading theories:] When you start your interaction you must open a
11456
7eb63f63e6c6 revisions and indexing
paulson
parents: 11450
diff changeset
   139
  named theory with the line \commdx{theory}~\isa{T~=~\dots~:}. Isabelle
8771
026f37a86ea7 *** empty log message ***
nipkow
parents: 8743
diff changeset
   140
  automatically loads all the required parent theories from their respective
026f37a86ea7 *** empty log message ***
nipkow
parents: 8743
diff changeset
   141
  files (which may take a moment, unless the theories are already loaded and
9541
d17c0b34d5c8 *** empty log message ***
nipkow
parents: 9494
diff changeset
   142
  the files have not been modified).
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   143
  
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   144
  If you suddenly discover that you need to modify a parent theory of your
11428
332347b9b942 tidying the index
paulson
parents: 11419
diff changeset
   145
  current theory, you must first abandon your current theory%
332347b9b942 tidying the index
paulson
parents: 11419
diff changeset
   146
  \indexbold{abandoning a theory}\indexbold{theories!abandoning} 
332347b9b942 tidying the index
paulson
parents: 11419
diff changeset
   147
  (at the shell
11419
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   148
  level type \commdx{kill}). After the parent theory has
10971
6852682eaf16 *** empty log message ***
nipkow
parents: 10885
diff changeset
   149
  been modified, you go back to your original theory. When its first line
6852682eaf16 *** empty log message ***
nipkow
parents: 10885
diff changeset
   150
  \isa{\isacommand{theory}~T~=~\dots~:} is processed, the
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   151
  modified parent is reloaded automatically.
9541
d17c0b34d5c8 *** empty log message ***
nipkow
parents: 9494
diff changeset
   152
  
10978
5eebea8f359f *** empty log message ***
nipkow
parents: 10971
diff changeset
   153
%  The only time when you need to load a theory by hand is when you simply
5eebea8f359f *** empty log message ***
nipkow
parents: 10971
diff changeset
   154
%  want to check if it loads successfully without wanting to make use of the
5eebea8f359f *** empty log message ***
nipkow
parents: 10971
diff changeset
   155
%  theory itself. This you can do by typing
11419
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   156
%  \isa{\commdx{use\_thy}~"T"}.
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   157
\end{description}
12327
5a4d78204492 *** empty log message ***
nipkow
parents: 11646
diff changeset
   158
Further commands are found in the Isabelle/Isar Reference
5a4d78204492 *** empty log message ***
nipkow
parents: 11646
diff changeset
   159
Manual~\cite{isabelle-isar-ref}.
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   160
8771
026f37a86ea7 *** empty log message ***
nipkow
parents: 8743
diff changeset
   161
We now examine Isabelle's functional programming constructs systematically,
026f37a86ea7 *** empty log message ***
nipkow
parents: 8743
diff changeset
   162
starting with inductive datatypes.
026f37a86ea7 *** empty log message ***
nipkow
parents: 8743
diff changeset
   163
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   164
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   165
\section{Datatypes}
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   166
\label{sec:datatype}
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   167
11456
7eb63f63e6c6 revisions and indexing
paulson
parents: 11450
diff changeset
   168
\index{datatypes|(}%
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   169
Inductive datatypes are part of almost every non-trivial application of HOL.
11458
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11457
diff changeset
   170
First we take another look at an important example, the datatype of
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   171
lists, before we turn to datatypes in general. The section closes with a
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   172
case study.
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   173
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   174
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   175
\subsection{Lists}
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   176
11428
332347b9b942 tidying the index
paulson
parents: 11419
diff changeset
   177
\index{*list (type)}%
11457
279da0358aa9 additional revisions to chapters 1, 2
paulson
parents: 11456
diff changeset
   178
Lists are one of the essential datatypes in computing.  We expect that you
279da0358aa9 additional revisions to chapters 1, 2
paulson
parents: 11456
diff changeset
   179
are already familiar with their basic operations.
8771
026f37a86ea7 *** empty log message ***
nipkow
parents: 8743
diff changeset
   180
Theory \isa{ToyList} is only a small fragment of HOL's predefined theory
11428
332347b9b942 tidying the index
paulson
parents: 11419
diff changeset
   181
\thydx{List}\footnote{\url{http://isabelle.in.tum.de/library/HOL/List.html}}.
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   182
The latter contains many further operations. For example, the functions
11419
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   183
\cdx{hd} (``head'') and \cdx{tl} (``tail'') return the first
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   184
element and the remainder of a list. (However, pattern-matching is usually
10795
9e888d60d3e5 minor edits to Chapters 1-3
paulson
parents: 10654
diff changeset
   185
preferable to \isa{hd} and \isa{tl}.)  
10800
2d4c058749a7 *** empty log message ***
nipkow
parents: 10795
diff changeset
   186
Also available are higher-order functions like \isa{map} and \isa{filter}.
10795
9e888d60d3e5 minor edits to Chapters 1-3
paulson
parents: 10654
diff changeset
   187
Theory \isa{List} also contains
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   188
more syntactic sugar: \isa{[}$x@1$\isa{,}\dots\isa{,}$x@n$\isa{]} abbreviates
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   189
$x@1$\isa{\#}\dots\isa{\#}$x@n$\isa{\#[]}.  In the rest of the tutorial we
12327
5a4d78204492 *** empty log message ***
nipkow
parents: 11646
diff changeset
   190
always use HOL's predefined lists by building on theory \isa{Main}.
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   191
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   192
10885
90695f46440b lcp's pass over the book, chapters 1-8
paulson
parents: 10824
diff changeset
   193
\subsection{The General Format}
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   194
\label{sec:general-datatype}
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   195
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   196
The general HOL \isacommand{datatype} definition is of the form
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   197
\[
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   198
\isacommand{datatype}~(\alpha@1, \dots, \alpha@n) \, t ~=~
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   199
C@1~\tau@{11}~\dots~\tau@{1k@1} ~\mid~ \dots ~\mid~
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   200
C@m~\tau@{m1}~\dots~\tau@{mk@m}
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   201
\]
8771
026f37a86ea7 *** empty log message ***
nipkow
parents: 8743
diff changeset
   202
where $\alpha@i$ are distinct type variables (the parameters), $C@i$ are distinct
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   203
constructor names and $\tau@{ij}$ are types; it is customary to capitalize
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   204
the first letter in constructor names. There are a number of
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   205
restrictions (such as that the type should not be empty) detailed
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   206
elsewhere~\cite{isabelle-HOL}. Isabelle notifies you if you violate them.
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   207
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   208
Laws about datatypes, such as \isa{[] \isasymnoteq~x\#xs} and
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   209
\isa{(x\#xs = y\#ys) = (x=y \isasymand~xs=ys)}, are used automatically
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   210
during proofs by simplification.  The same is true for the equations in
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   211
primitive recursive function definitions.
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   212
12327
5a4d78204492 *** empty log message ***
nipkow
parents: 11646
diff changeset
   213
Every\footnote{Except for advanced datatypes where the recursion involves
5a4d78204492 *** empty log message ***
nipkow
parents: 11646
diff changeset
   214
``\isasymRightarrow'' as in {\S}\ref{sec:nested-fun-datatype}.} datatype $t$
5a4d78204492 *** empty log message ***
nipkow
parents: 11646
diff changeset
   215
comes equipped with a \isa{size} function from $t$ into the natural numbers
5a4d78204492 *** empty log message ***
nipkow
parents: 11646
diff changeset
   216
(see~{\S}\ref{sec:nat} below). For lists, \isa{size} is just the length, i.e.\
5a4d78204492 *** empty log message ***
nipkow
parents: 11646
diff changeset
   217
\isa{size [] = 0} and \isa{size(x \# xs) = size xs + 1}.  In general,
5a4d78204492 *** empty log message ***
nipkow
parents: 11646
diff changeset
   218
\cdx{size} returns
11456
7eb63f63e6c6 revisions and indexing
paulson
parents: 11450
diff changeset
   219
\begin{itemize}
11457
279da0358aa9 additional revisions to chapters 1, 2
paulson
parents: 11456
diff changeset
   220
\item zero for all constructors
11456
7eb63f63e6c6 revisions and indexing
paulson
parents: 11450
diff changeset
   221
that do not have an argument of type $t$\\
11457
279da0358aa9 additional revisions to chapters 1, 2
paulson
parents: 11456
diff changeset
   222
\item one plus the sum of the sizes of all arguments of type~$t$,
279da0358aa9 additional revisions to chapters 1, 2
paulson
parents: 11456
diff changeset
   223
for all other constructors
11456
7eb63f63e6c6 revisions and indexing
paulson
parents: 11450
diff changeset
   224
\end{itemize}
7eb63f63e6c6 revisions and indexing
paulson
parents: 11450
diff changeset
   225
Note that because
9644
6b0b6b471855 *** empty log message ***
nipkow
parents: 9541
diff changeset
   226
\isa{size} is defined on every datatype, it is overloaded; on lists
11419
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   227
\isa{size} is also called \sdx{length}, which is not overloaded.
10795
9e888d60d3e5 minor edits to Chapters 1-3
paulson
parents: 10654
diff changeset
   228
Isabelle will always show \isa{size} on lists as \isa{length}.
9644
6b0b6b471855 *** empty log message ***
nipkow
parents: 9541
diff changeset
   229
6b0b6b471855 *** empty log message ***
nipkow
parents: 9541
diff changeset
   230
10885
90695f46440b lcp's pass over the book, chapters 1-8
paulson
parents: 10824
diff changeset
   231
\subsection{Primitive Recursion}
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   232
11456
7eb63f63e6c6 revisions and indexing
paulson
parents: 11450
diff changeset
   233
\index{recursion!primitive}%
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   234
Functions on datatypes are usually defined by recursion. In fact, most of the
11428
332347b9b942 tidying the index
paulson
parents: 11419
diff changeset
   235
time they are defined by what is called \textbf{primitive recursion}.
332347b9b942 tidying the index
paulson
parents: 11419
diff changeset
   236
The keyword \commdx{primrec} is followed by a list of
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   237
equations
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   238
\[ f \, x@1 \, \dots \, (C \, y@1 \, \dots \, y@k)\, \dots \, x@n = r \]
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   239
such that $C$ is a constructor of the datatype $t$ and all recursive calls of
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   240
$f$ in $r$ are of the form $f \, \dots \, y@i \, \dots$ for some $i$. Thus
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   241
Isabelle immediately sees that $f$ terminates because one (fixed!) argument
10654
458068404143 *** empty log message ***
nipkow
parents: 10608
diff changeset
   242
becomes smaller with every recursive call. There must be at most one equation
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   243
for each constructor.  Their order is immaterial.
8771
026f37a86ea7 *** empty log message ***
nipkow
parents: 8743
diff changeset
   244
A more general method for defining total recursive functions is introduced in
10538
d1bf9ca9008d *** empty log message ***
nipkow
parents: 10522
diff changeset
   245
{\S}\ref{sec:recdef}.
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   246
9493
494f8cd34df7 *** empty log message ***
nipkow
parents: 9458
diff changeset
   247
\begin{exercise}\label{ex:Tree}
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   248
\input{Misc/document/Tree.tex}%
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   249
\end{exercise}
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   250
9721
7e51c9f3d5a0 *** empty log message ***
nipkow
parents: 9689
diff changeset
   251
\input{Misc/document/case_exprs.tex}
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   252
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   253
\input{Ifexpr/document/Ifexpr.tex}
11456
7eb63f63e6c6 revisions and indexing
paulson
parents: 11450
diff changeset
   254
\index{datatypes|)}
7eb63f63e6c6 revisions and indexing
paulson
parents: 11450
diff changeset
   255
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   256
10885
90695f46440b lcp's pass over the book, chapters 1-8
paulson
parents: 10824
diff changeset
   257
\section{Some Basic Types}
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   258
11457
279da0358aa9 additional revisions to chapters 1, 2
paulson
parents: 11456
diff changeset
   259
This section introduces the types of natural numbers and ordered pairs.  Also
279da0358aa9 additional revisions to chapters 1, 2
paulson
parents: 11456
diff changeset
   260
described is type \isa{option}, which is useful for modelling exceptional
279da0358aa9 additional revisions to chapters 1, 2
paulson
parents: 11456
diff changeset
   261
cases. 
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   262
10885
90695f46440b lcp's pass over the book, chapters 1-8
paulson
parents: 10824
diff changeset
   263
\subsection{Natural Numbers}
11456
7eb63f63e6c6 revisions and indexing
paulson
parents: 11450
diff changeset
   264
\label{sec:nat}\index{natural numbers}%
11428
332347b9b942 tidying the index
paulson
parents: 11419
diff changeset
   265
\index{linear arithmetic|(}
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   266
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   267
\input{Misc/document/fakenat.tex}
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   268
\input{Misc/document/natsum.tex}
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   269
11428
332347b9b942 tidying the index
paulson
parents: 11419
diff changeset
   270
\index{linear arithmetic|)}
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   271
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   272
10396
5ab08609e6c8 *** empty log message ***
nipkow
parents: 10237
diff changeset
   273
\subsection{Pairs}
9541
d17c0b34d5c8 *** empty log message ***
nipkow
parents: 9494
diff changeset
   274
\input{Misc/document/pairs.tex}
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   275
10608
620647438780 *** empty log message ***
nipkow
parents: 10543
diff changeset
   276
\subsection{Datatype {\tt\slshape option}}
10543
8e4307d1207a *** empty log message ***
nipkow
parents: 10538
diff changeset
   277
\label{sec:option}
8e4307d1207a *** empty log message ***
nipkow
parents: 10538
diff changeset
   278
\input{Misc/document/Option2.tex}
8e4307d1207a *** empty log message ***
nipkow
parents: 10538
diff changeset
   279
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   280
\section{Definitions}
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   281
\label{sec:Definitions}
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   282
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   283
A definition is simply an abbreviation, i.e.\ a new name for an existing
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   284
construction. In particular, definitions cannot be recursive. Isabelle offers
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   285
definitions on the level of types and terms. Those on the type level are
11456
7eb63f63e6c6 revisions and indexing
paulson
parents: 11450
diff changeset
   286
called \textbf{type synonyms}; those on the term level are simply called 
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   287
definitions.
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   288
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   289
10885
90695f46440b lcp's pass over the book, chapters 1-8
paulson
parents: 10824
diff changeset
   290
\subsection{Type Synonyms}
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   291
12327
5a4d78204492 *** empty log message ***
nipkow
parents: 11646
diff changeset
   292
\index{type synonyms}%
11456
7eb63f63e6c6 revisions and indexing
paulson
parents: 11450
diff changeset
   293
Type synonyms are similar to those found in ML\@. They are created by a 
11428
332347b9b942 tidying the index
paulson
parents: 11419
diff changeset
   294
\commdx{types} command:
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   295
12327
5a4d78204492 *** empty log message ***
nipkow
parents: 11646
diff changeset
   296
\input{Misc/document/types.tex}
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   297
9844
8016321c7de1 *** empty log message ***
nipkow
parents: 9792
diff changeset
   298
\input{Misc/document/prime_def.tex}
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   299
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   300
11201
eddc69b55fac *** empty log message ***
nipkow
parents: 11159
diff changeset
   301
\section{The Definitional Approach}
eddc69b55fac *** empty log message ***
nipkow
parents: 11159
diff changeset
   302
\label{sec:definitional}
eddc69b55fac *** empty log message ***
nipkow
parents: 11159
diff changeset
   303
11457
279da0358aa9 additional revisions to chapters 1, 2
paulson
parents: 11456
diff changeset
   304
\index{Definitional Approach}%
11201
eddc69b55fac *** empty log message ***
nipkow
parents: 11159
diff changeset
   305
As we pointed out at the beginning of the chapter, asserting arbitrary
11456
7eb63f63e6c6 revisions and indexing
paulson
parents: 11450
diff changeset
   306
axioms such as $f(n) = f(n) + 1$ can easily lead to contradictions. In order
11457
279da0358aa9 additional revisions to chapters 1, 2
paulson
parents: 11456
diff changeset
   307
to avoid this danger, we advocate the definitional rather than
11456
7eb63f63e6c6 revisions and indexing
paulson
parents: 11450
diff changeset
   308
the axiomatic approach: introduce new concepts by definitions. However,  Isabelle/HOL seems to
7eb63f63e6c6 revisions and indexing
paulson
parents: 11450
diff changeset
   309
support many richer definitional constructs, such as
7eb63f63e6c6 revisions and indexing
paulson
parents: 11450
diff changeset
   310
\isacommand{primrec}. The point is that Isabelle reduces such constructs to first principles. For example, each
7eb63f63e6c6 revisions and indexing
paulson
parents: 11450
diff changeset
   311
\isacommand{primrec} function definition is turned into a proper
7eb63f63e6c6 revisions and indexing
paulson
parents: 11450
diff changeset
   312
(nonrecursive!) definition from which the user-supplied recursion equations are
11457
279da0358aa9 additional revisions to chapters 1, 2
paulson
parents: 11456
diff changeset
   313
automatically proved.  This process is
11456
7eb63f63e6c6 revisions and indexing
paulson
parents: 11450
diff changeset
   314
hidden from the user, who does not have to understand the details.  Other commands described
11457
279da0358aa9 additional revisions to chapters 1, 2
paulson
parents: 11456
diff changeset
   315
later, like \isacommand{recdef} and \isacommand{inductive}, work similarly.  
279da0358aa9 additional revisions to chapters 1, 2
paulson
parents: 11456
diff changeset
   316
This strict adherence to the definitional approach reduces the risk of 
279da0358aa9 additional revisions to chapters 1, 2
paulson
parents: 11456
diff changeset
   317
soundness errors.
11201
eddc69b55fac *** empty log message ***
nipkow
parents: 11159
diff changeset
   318
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   319
\chapter{More Functional Programming}
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   320
11458
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11457
diff changeset
   321
The purpose of this chapter is to deepen your understanding of the
8771
026f37a86ea7 *** empty log message ***
nipkow
parents: 8743
diff changeset
   322
concepts encountered so far and to introduce advanced forms of datatypes and
026f37a86ea7 *** empty log message ***
nipkow
parents: 8743
diff changeset
   323
recursive functions. The first two sections give a structured presentation of
10538
d1bf9ca9008d *** empty log message ***
nipkow
parents: 10522
diff changeset
   324
theorem proving by simplification ({\S}\ref{sec:Simplification}) and discuss
11458
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11457
diff changeset
   325
important heuristics for induction ({\S}\ref{sec:InductionHeuristics}).  You can
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11457
diff changeset
   326
skip them if you are not planning to perform proofs yourself.
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11457
diff changeset
   327
We then present a case
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11457
diff changeset
   328
study: a compiler for expressions ({\S}\ref{sec:ExprCompiler}). Advanced
8771
026f37a86ea7 *** empty log message ***
nipkow
parents: 8743
diff changeset
   329
datatypes, including those involving function spaces, are covered in
11458
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11457
diff changeset
   330
{\S}\ref{sec:advanced-datatypes}; it closes with another case study, search
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11457
diff changeset
   331
trees (``tries'').  Finally we introduce \isacommand{recdef}, a general
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11457
diff changeset
   332
form of recursive function definition that goes well beyond 
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11457
diff changeset
   333
\isacommand{primrec} ({\S}\ref{sec:recdef}).
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   334
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   335
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   336
\section{Simplification}
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   337
\label{sec:Simplification}
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   338
\index{simplification|(}
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   339
10795
9e888d60d3e5 minor edits to Chapters 1-3
paulson
parents: 10654
diff changeset
   340
So far we have proved our theorems by \isa{auto}, which simplifies
11458
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11457
diff changeset
   341
all subgoals. In fact, \isa{auto} can do much more than that. 
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11457
diff changeset
   342
To go beyond toy examples, you
9541
d17c0b34d5c8 *** empty log message ***
nipkow
parents: 9494
diff changeset
   343
need to understand the ingredients of \isa{auto}.  This section covers the
10971
6852682eaf16 *** empty log message ***
nipkow
parents: 10885
diff changeset
   344
method that \isa{auto} always applies first, simplification.
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   345
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   346
Simplification is one of the central theorem proving tools in Isabelle and
11458
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11457
diff changeset
   347
many other systems. The tool itself is called the \textbf{simplifier}. 
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11457
diff changeset
   348
This section introduces the many features of the simplifier
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11457
diff changeset
   349
and is required reading if you intend to perform proofs.  Later on,
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11457
diff changeset
   350
{\S}\ref{sec:simplification-II} explains some more advanced features and a
9754
a123a64cadeb *** empty log message ***
nipkow
parents: 9742
diff changeset
   351
little bit of how the simplifier works. The serious student should read that
11458
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11457
diff changeset
   352
section as well, in particular to understand why the simplifier did
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11457
diff changeset
   353
something unexpected.
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   354
11458
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11457
diff changeset
   355
\subsection{What is Simplification?}
9458
c613cd06d5cf apply. -> by
nipkow
parents: 8771
diff changeset
   356
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   357
In its most basic form, simplification means repeated application of
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   358
equations from left to right. For example, taking the rules for \isa{\at}
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   359
and applying them to the term \isa{[0,1] \at\ []} results in a sequence of
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   360
simplification steps:
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   361
\begin{ttbox}\makeatother
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   362
(0#1#[]) @ []  \(\leadsto\)  0#((1#[]) @ [])  \(\leadsto\)  0#(1#([] @ []))  \(\leadsto\)  0#1#[]
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   363
\end{ttbox}
9933
9feb1e0c4cb3 *** empty log message ***
nipkow
parents: 9844
diff changeset
   364
This is also known as \bfindex{term rewriting}\indexbold{rewriting} and the
11458
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11457
diff changeset
   365
equations are referred to as \bfindex{rewrite rules}.
9933
9feb1e0c4cb3 *** empty log message ***
nipkow
parents: 9844
diff changeset
   366
``Rewriting'' is more honest than ``simplification'' because the terms do not
9feb1e0c4cb3 *** empty log message ***
nipkow
parents: 9844
diff changeset
   367
necessarily become simpler in the process.
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   368
11458
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11457
diff changeset
   369
The simplifier proves arithmetic goals as described in
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11457
diff changeset
   370
{\S}\ref{sec:nat} above.  Arithmetic expressions are simplified using built-in
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11457
diff changeset
   371
procedures that go beyond mere rewrite rules.  New simplification procedures
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11457
diff changeset
   372
can be coded and installed, but they are definitely not a matter for this
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11457
diff changeset
   373
tutorial. 
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11457
diff changeset
   374
9844
8016321c7de1 *** empty log message ***
nipkow
parents: 9792
diff changeset
   375
\input{Misc/document/simp.tex}
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   376
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   377
\index{simplification|)}
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   378
9844
8016321c7de1 *** empty log message ***
nipkow
parents: 9792
diff changeset
   379
\input{Misc/document/Itrev.tex}
13305
f88d0c363582 *** empty log message ***
nipkow
parents: 12582
diff changeset
   380
\begin{exercise}
f88d0c363582 *** empty log message ***
nipkow
parents: 12582
diff changeset
   381
\input{Misc/document/Plus.tex}%
f88d0c363582 *** empty log message ***
nipkow
parents: 12582
diff changeset
   382
\end{exercise}
9493
494f8cd34df7 *** empty log message ***
nipkow
parents: 9458
diff changeset
   383
\begin{exercise}
494f8cd34df7 *** empty log message ***
nipkow
parents: 9458
diff changeset
   384
\input{Misc/document/Tree2.tex}%
494f8cd34df7 *** empty log message ***
nipkow
parents: 9458
diff changeset
   385
\end{exercise}
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   386
9844
8016321c7de1 *** empty log message ***
nipkow
parents: 9792
diff changeset
   387
\input{CodeGen/document/CodeGen.tex}
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   388
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   389
10885
90695f46440b lcp's pass over the book, chapters 1-8
paulson
parents: 10824
diff changeset
   390
\section{Advanced Datatypes}
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   391
\label{sec:advanced-datatypes}
11428
332347b9b942 tidying the index
paulson
parents: 11419
diff changeset
   392
\index{datatype@\isacommand {datatype} (command)|(}
332347b9b942 tidying the index
paulson
parents: 11419
diff changeset
   393
\index{primrec@\isacommand {primrec} (command)|(}
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   394
%|)
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   395
11428
332347b9b942 tidying the index
paulson
parents: 11419
diff changeset
   396
This section presents advanced forms of datatypes: mutual and nested
332347b9b942 tidying the index
paulson
parents: 11419
diff changeset
   397
recursion.  A series of examples will culminate in a treatment of the trie
332347b9b942 tidying the index
paulson
parents: 11419
diff changeset
   398
data structure.
332347b9b942 tidying the index
paulson
parents: 11419
diff changeset
   399
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   400
10885
90695f46440b lcp's pass over the book, chapters 1-8
paulson
parents: 10824
diff changeset
   401
\subsection{Mutual Recursion}
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   402
\label{sec:datatype-mut-rec}
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   403
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   404
\input{Datatype/document/ABexpr.tex}
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   405
10885
90695f46440b lcp's pass over the book, chapters 1-8
paulson
parents: 10824
diff changeset
   406
\subsection{Nested Recursion}
9644
6b0b6b471855 *** empty log message ***
nipkow
parents: 9541
diff changeset
   407
\label{sec:nested-datatype}
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   408
9644
6b0b6b471855 *** empty log message ***
nipkow
parents: 9541
diff changeset
   409
{\makeatother\input{Datatype/document/Nested.tex}}
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   410
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   411
11419
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   412
\subsection{The Limits of Nested Recursion}
12327
5a4d78204492 *** empty log message ***
nipkow
parents: 11646
diff changeset
   413
\label{sec:nested-fun-datatype}
11419
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   414
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   415
How far can we push nested recursion? By the unfolding argument above, we can
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   416
reduce nested to mutual recursion provided the nested recursion only involves
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   417
previously defined datatypes. This does not include functions:
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   418
\begin{isabelle}
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   419
\isacommand{datatype} t = C "t \isasymRightarrow\ bool"
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   420
\end{isabelle}
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   421
This declaration is a real can of worms.
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   422
In HOL it must be ruled out because it requires a type
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   423
\isa{t} such that \isa{t} and its power set \isa{t \isasymFun\ bool} have the
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   424
same cardinality --- an impossibility. For the same reason it is not possible
12328
7c4ec77a8715 *** empty log message ***
nipkow
parents: 12327
diff changeset
   425
to allow recursion involving the type \isa{t set}, which is isomorphic to
11419
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   426
\isa{t \isasymFun\ bool}.
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   427
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   428
Fortunately, a limited form of recursion
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   429
involving function spaces is permitted: the recursive type may occur on the
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   430
right of a function arrow, but never on the left. Hence the above can of worms
11458
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11457
diff changeset
   431
is ruled out but the following example of a potentially 
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11457
diff changeset
   432
\index{infinitely branching trees}%
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11457
diff changeset
   433
infinitely branching tree is accepted:
11419
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   434
\smallskip
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   435
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   436
\input{Datatype/document/Fundata.tex}
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   437
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   438
If you need nested recursion on the left of a function arrow, there are
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   439
alternatives to pure HOL\@.  In the Logic for Computable Functions 
11458
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11457
diff changeset
   440
(\rmindex{LCF}), types like
11419
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   441
\begin{isabelle}
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   442
\isacommand{datatype} lam = C "lam \isasymrightarrow\ lam"
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   443
\end{isabelle}
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   444
do indeed make sense~\cite{paulson87}.  Note the different arrow,
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   445
\isa{\isasymrightarrow} instead of \isa{\isasymRightarrow},
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   446
expressing the type of \emph{continuous} functions. 
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   447
There is even a version of LCF on top of HOL,
11458
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11457
diff changeset
   448
called \rmindex{HOLCF}~\cite{MuellerNvOS99}.
11428
332347b9b942 tidying the index
paulson
parents: 11419
diff changeset
   449
\index{datatype@\isacommand {datatype} (command)|)}
332347b9b942 tidying the index
paulson
parents: 11419
diff changeset
   450
\index{primrec@\isacommand {primrec} (command)|)}
11419
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   451
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   452
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   453
\subsection{Case Study: Tries}
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   454
\label{sec:Trie}
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   455
11458
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11457
diff changeset
   456
\index{tries|(}%
11419
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   457
Tries are a classic search tree data structure~\cite{Knuth3-75} for fast
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   458
indexing with strings. Figure~\ref{fig:trie} gives a graphical example of a
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   459
trie containing the words ``all'', ``an'', ``ape'', ``can'', ``car'' and
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   460
``cat''.  When searching a string in a trie, the letters of the string are
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   461
examined sequentially. Each letter determines which subtrie to search next.
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   462
In this case study we model tries as a datatype, define a lookup and an
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   463
update function, and prove that they behave as expected.
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   464
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   465
\begin{figure}[htbp]
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   466
\begin{center}
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   467
\unitlength1mm
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   468
\begin{picture}(60,30)
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   469
\put( 5, 0){\makebox(0,0)[b]{l}}
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   470
\put(25, 0){\makebox(0,0)[b]{e}}
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   471
\put(35, 0){\makebox(0,0)[b]{n}}
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   472
\put(45, 0){\makebox(0,0)[b]{r}}
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   473
\put(55, 0){\makebox(0,0)[b]{t}}
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   474
%
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   475
\put( 5, 9){\line(0,-1){5}}
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   476
\put(25, 9){\line(0,-1){5}}
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   477
\put(44, 9){\line(-3,-2){9}}
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   478
\put(45, 9){\line(0,-1){5}}
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   479
\put(46, 9){\line(3,-2){9}}
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   480
%
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   481
\put( 5,10){\makebox(0,0)[b]{l}}
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   482
\put(15,10){\makebox(0,0)[b]{n}}
11419
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   483
\put(25,10){\makebox(0,0)[b]{p}}
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   484
\put(45,10){\makebox(0,0)[b]{a}}
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   485
%
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   486
\put(14,19){\line(-3,-2){9}}
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   487
\put(15,19){\line(0,-1){5}}
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   488
\put(16,19){\line(3,-2){9}}
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   489
\put(45,19){\line(0,-1){5}}
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   490
%
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   491
\put(15,20){\makebox(0,0)[b]{a}}
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   492
\put(45,20){\makebox(0,0)[b]{c}}
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   493
%
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   494
\put(30,30){\line(-3,-2){13}}
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   495
\put(30,30){\line(3,-2){13}}
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   496
\end{picture}
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   497
\end{center}
11450
1b02a6c4032f tweaks and indexing
paulson
parents: 11428
diff changeset
   498
\caption{A Sample Trie}
11419
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   499
\label{fig:trie}
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   500
\end{figure}
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   501
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   502
Proper tries associate some value with each string. Since the
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   503
information is stored only in the final node associated with the string, many
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   504
nodes do not carry any value. This distinction is modeled with the help
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   505
of the predefined datatype \isa{option} (see {\S}\ref{sec:option}).
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   506
\input{Trie/document/Trie.tex}
11458
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11457
diff changeset
   507
\index{tries|)}
11419
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   508
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   509
\section{Total Recursive Functions}
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   510
\label{sec:recdef}
11428
332347b9b942 tidying the index
paulson
parents: 11419
diff changeset
   511
\index{recdef@\isacommand {recdef} (command)|(}\index{functions!total|(}
11419
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   512
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   513
Although many total functions have a natural primitive recursive definition,
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   514
this is not always the case. Arbitrary total recursive functions can be
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   515
defined by means of \isacommand{recdef}: you can use full pattern-matching,
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   516
recursion need not involve datatypes, and termination is proved by showing
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   517
that the arguments of all recursive calls are smaller in a suitable (user
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   518
supplied) sense. In this section we restrict ourselves to measure functions;
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   519
more advanced termination proofs are discussed in {\S}\ref{sec:beyond-measure}.
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   520
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   521
\subsection{Defining Recursive Functions}
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   522
\label{sec:recdef-examples}
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   523
\input{Recdef/document/examples.tex}
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   524
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   525
\subsection{Proving Termination}
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   526
\input{Recdef/document/termination.tex}
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   527
11458
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11457
diff changeset
   528
\subsection{Simplification and Recursive Functions}
11419
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   529
\label{sec:recdef-simplification}
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   530
\input{Recdef/document/simplification.tex}
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   531
11458
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11457
diff changeset
   532
\subsection{Induction and Recursive Functions}
11419
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   533
\index{induction!recursion|(}
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   534
\index{recursion induction|(}
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   535
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   536
\input{Recdef/document/Induction.tex}
9644
6b0b6b471855 *** empty log message ***
nipkow
parents: 9541
diff changeset
   537
\label{sec:recdef-induction}
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   538
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   539
\index{induction!recursion|)}
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   540
\index{recursion induction|)}
11428
332347b9b942 tidying the index
paulson
parents: 11419
diff changeset
   541
\index{recdef@\isacommand {recdef} (command)|)}\index{functions!total|)}