doc-src/TutorialI/fp.tex
author berghofe
Sun, 10 Jan 2010 18:43:45 +0100
changeset 34915 7894c7dab132
parent 27712 007a339b9e7d
child 44048 64f574163ca2
permissions -rw-r--r--
Adapted to changes in induct method.
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|(}
17182
ae84287f44e3 tune spacing where a generated theory text is included directly;
wenzelm
parents: 16523
diff changeset
    35
{\makeatother\medskip\input{ToyList/document/ToyList.tex}}
11419
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}
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
    93
\item[Modifying the order of subgoals:]
11419
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
    94
\commdx{defer} moves the first subgoal to the end and
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
    95
\commdx{prefer}~$n$ moves subgoal $n$ to the front.
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
    96
\item[Printing theorems:]
11419
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
    97
  \commdx{thm}~\textit{name}$@1$~\dots~\textit{name}$@n$
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
    98
  prints the named theorems.
11419
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
    99
\item[Reading terms and types:] \commdx{term}
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   100
  \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
   101
  the current context; the inferred type is output as well.
11419
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   102
  \commdx{typ} \textit{string} reads and prints the given
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   103
  string as a type in the current context.
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   104
\end{description}
12327
5a4d78204492 *** empty log message ***
nipkow
parents: 11646
diff changeset
   105
Further commands are found in the Isabelle/Isar Reference
5a4d78204492 *** empty log message ***
nipkow
parents: 11646
diff changeset
   106
Manual~\cite{isabelle-isar-ref}.
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   107
16412
50eab0183aea *** empty log message ***
nipkow
parents: 15358
diff changeset
   108
\begin{pgnote}
16523
f8a734dc0fbc *** empty log message ***
nipkow
parents: 16412
diff changeset
   109
Clicking on the \pgmenu{State} button redisplays the current proof state.
16412
50eab0183aea *** empty log message ***
nipkow
parents: 15358
diff changeset
   110
This is helpful in case commands like \isacommand{thm} have overwritten it.
50eab0183aea *** empty log message ***
nipkow
parents: 15358
diff changeset
   111
\end{pgnote}
50eab0183aea *** empty log message ***
nipkow
parents: 15358
diff changeset
   112
8771
026f37a86ea7 *** empty log message ***
nipkow
parents: 8743
diff changeset
   113
We now examine Isabelle's functional programming constructs systematically,
026f37a86ea7 *** empty log message ***
nipkow
parents: 8743
diff changeset
   114
starting with inductive datatypes.
026f37a86ea7 *** empty log message ***
nipkow
parents: 8743
diff changeset
   115
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   116
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   117
\section{Datatypes}
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   118
\label{sec:datatype}
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   119
11456
7eb63f63e6c6 revisions and indexing
paulson
parents: 11450
diff changeset
   120
\index{datatypes|(}%
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   121
Inductive datatypes are part of almost every non-trivial application of HOL.
11458
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11457
diff changeset
   122
First we take another look at an important example, the datatype of
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   123
lists, before we turn to datatypes in general. The section closes with a
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   124
case study.
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   125
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   126
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   127
\subsection{Lists}
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   128
11428
332347b9b942 tidying the index
paulson
parents: 11419
diff changeset
   129
\index{*list (type)}%
11457
279da0358aa9 additional revisions to chapters 1, 2
paulson
parents: 11456
diff changeset
   130
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
   131
are already familiar with their basic operations.
8771
026f37a86ea7 *** empty log message ***
nipkow
parents: 8743
diff changeset
   132
Theory \isa{ToyList} is only a small fragment of HOL's predefined theory
11428
332347b9b942 tidying the index
paulson
parents: 11419
diff changeset
   133
\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
   134
The latter contains many further operations. For example, the functions
11419
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   135
\cdx{hd} (``head'') and \cdx{tl} (``tail'') return the first
25263
b54744935036 *** empty log message ***
nipkow
parents: 25261
diff changeset
   136
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
   137
preferable to \isa{hd} and \isa{tl}.)  
10800
2d4c058749a7 *** empty log message ***
nipkow
parents: 10795
diff changeset
   138
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
   139
Theory \isa{List} also contains
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   140
more syntactic sugar: \isa{[}$x@1$\isa{,}\dots\isa{,}$x@n$\isa{]} abbreviates
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   141
$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
   142
always use HOL's predefined lists by building on theory \isa{Main}.
27712
007a339b9e7d added hint about writing "x : set xs".
nipkow
parents: 27015
diff changeset
   143
\begin{warn}
007a339b9e7d added hint about writing "x : set xs".
nipkow
parents: 27015
diff changeset
   144
Looking ahead to sets and quanifiers in Part II:
007a339b9e7d added hint about writing "x : set xs".
nipkow
parents: 27015
diff changeset
   145
The best way to express that some element \isa{x} is in a list \isa{xs} is
007a339b9e7d added hint about writing "x : set xs".
nipkow
parents: 27015
diff changeset
   146
\isa{x $\in$ set xs}, where \isa{set} is a function that turns a list into the
007a339b9e7d added hint about writing "x : set xs".
nipkow
parents: 27015
diff changeset
   147
set of its elements.
007a339b9e7d added hint about writing "x : set xs".
nipkow
parents: 27015
diff changeset
   148
By the same device you can also write bounded quantifiers like
007a339b9e7d added hint about writing "x : set xs".
nipkow
parents: 27015
diff changeset
   149
\isa{$\forall$x $\in$ set xs} or embed lists in other set expressions.
007a339b9e7d added hint about writing "x : set xs".
nipkow
parents: 27015
diff changeset
   150
\end{warn}
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   151
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   152
10885
90695f46440b lcp's pass over the book, chapters 1-8
paulson
parents: 10824
diff changeset
   153
\subsection{The General Format}
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   154
\label{sec:general-datatype}
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   155
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   156
The general HOL \isacommand{datatype} definition is of the form
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   157
\[
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   158
\isacommand{datatype}~(\alpha@1, \dots, \alpha@n) \, t ~=~
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   159
C@1~\tau@{11}~\dots~\tau@{1k@1} ~\mid~ \dots ~\mid~
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   160
C@m~\tau@{m1}~\dots~\tau@{mk@m}
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   161
\]
8771
026f37a86ea7 *** empty log message ***
nipkow
parents: 8743
diff changeset
   162
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
   163
constructor names and $\tau@{ij}$ are types; it is customary to capitalize
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   164
the first letter in constructor names. There are a number of
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   165
restrictions (such as that the type should not be empty) detailed
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   166
elsewhere~\cite{isabelle-HOL}. Isabelle notifies you if you violate them.
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   167
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   168
Laws about datatypes, such as \isa{[] \isasymnoteq~x\#xs} and
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   169
\isa{(x\#xs = y\#ys) = (x=y \isasymand~xs=ys)}, are used automatically
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   170
during proofs by simplification.  The same is true for the equations in
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   171
primitive recursive function definitions.
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   172
12327
5a4d78204492 *** empty log message ***
nipkow
parents: 11646
diff changeset
   173
Every\footnote{Except for advanced datatypes where the recursion involves
5a4d78204492 *** empty log message ***
nipkow
parents: 11646
diff changeset
   174
``\isasymRightarrow'' as in {\S}\ref{sec:nested-fun-datatype}.} datatype $t$
5a4d78204492 *** empty log message ***
nipkow
parents: 11646
diff changeset
   175
comes equipped with a \isa{size} function from $t$ into the natural numbers
5a4d78204492 *** empty log message ***
nipkow
parents: 11646
diff changeset
   176
(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
   177
\isa{size [] = 0} and \isa{size(x \# xs) = size xs + 1}.  In general,
5a4d78204492 *** empty log message ***
nipkow
parents: 11646
diff changeset
   178
\cdx{size} returns
11456
7eb63f63e6c6 revisions and indexing
paulson
parents: 11450
diff changeset
   179
\begin{itemize}
27015
f8537d69f514 *** empty log message ***
nipkow
parents: 25281
diff changeset
   180
\item zero for all constructors that do not have an argument of type $t$,
11457
279da0358aa9 additional revisions to chapters 1, 2
paulson
parents: 11456
diff changeset
   181
\item one plus the sum of the sizes of all arguments of type~$t$,
27015
f8537d69f514 *** empty log message ***
nipkow
parents: 25281
diff changeset
   182
for all other constructors.
11456
7eb63f63e6c6 revisions and indexing
paulson
parents: 11450
diff changeset
   183
\end{itemize}
7eb63f63e6c6 revisions and indexing
paulson
parents: 11450
diff changeset
   184
Note that because
9644
6b0b6b471855 *** empty log message ***
nipkow
parents: 9541
diff changeset
   185
\isa{size} is defined on every datatype, it is overloaded; on lists
11419
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   186
\isa{size} is also called \sdx{length}, which is not overloaded.
10795
9e888d60d3e5 minor edits to Chapters 1-3
paulson
parents: 10654
diff changeset
   187
Isabelle will always show \isa{size} on lists as \isa{length}.
9644
6b0b6b471855 *** empty log message ***
nipkow
parents: 9541
diff changeset
   188
6b0b6b471855 *** empty log message ***
nipkow
parents: 9541
diff changeset
   189
10885
90695f46440b lcp's pass over the book, chapters 1-8
paulson
parents: 10824
diff changeset
   190
\subsection{Primitive Recursion}
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   191
11456
7eb63f63e6c6 revisions and indexing
paulson
parents: 11450
diff changeset
   192
\index{recursion!primitive}%
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   193
Functions on datatypes are usually defined by recursion. In fact, most of the
27015
f8537d69f514 *** empty log message ***
nipkow
parents: 25281
diff changeset
   194
time they are defined by what is called \textbf{primitive recursion} over some
f8537d69f514 *** empty log message ***
nipkow
parents: 25281
diff changeset
   195
datatype $t$. This means that the recursion equations must be of the form
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   196
\[ f \, x@1 \, \dots \, (C \, y@1 \, \dots \, y@k)\, \dots \, x@n = r \]
27015
f8537d69f514 *** empty log message ***
nipkow
parents: 25281
diff changeset
   197
such that $C$ is a constructor of $t$ and all recursive calls of
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   198
$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
   199
Isabelle immediately sees that $f$ terminates because one (fixed!) argument
10654
458068404143 *** empty log message ***
nipkow
parents: 10608
diff changeset
   200
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
   201
for each constructor.  Their order is immaterial.
8771
026f37a86ea7 *** empty log message ***
nipkow
parents: 8743
diff changeset
   202
A more general method for defining total recursive functions is introduced in
25258
22d16596c306 recdef -> fun
nipkow
parents: 17182
diff changeset
   203
{\S}\ref{sec:fun}.
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   204
9493
494f8cd34df7 *** empty log message ***
nipkow
parents: 9458
diff changeset
   205
\begin{exercise}\label{ex:Tree}
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   206
\input{Misc/document/Tree.tex}%
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   207
\end{exercise}
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   208
9721
7e51c9f3d5a0 *** empty log message ***
nipkow
parents: 9689
diff changeset
   209
\input{Misc/document/case_exprs.tex}
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   210
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   211
\input{Ifexpr/document/Ifexpr.tex}
11456
7eb63f63e6c6 revisions and indexing
paulson
parents: 11450
diff changeset
   212
\index{datatypes|)}
7eb63f63e6c6 revisions and indexing
paulson
parents: 11450
diff changeset
   213
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   214
10885
90695f46440b lcp's pass over the book, chapters 1-8
paulson
parents: 10824
diff changeset
   215
\section{Some Basic Types}
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   216
11457
279da0358aa9 additional revisions to chapters 1, 2
paulson
parents: 11456
diff changeset
   217
This section introduces the types of natural numbers and ordered pairs.  Also
279da0358aa9 additional revisions to chapters 1, 2
paulson
parents: 11456
diff changeset
   218
described is type \isa{option}, which is useful for modelling exceptional
279da0358aa9 additional revisions to chapters 1, 2
paulson
parents: 11456
diff changeset
   219
cases. 
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   220
10885
90695f46440b lcp's pass over the book, chapters 1-8
paulson
parents: 10824
diff changeset
   221
\subsection{Natural Numbers}
11456
7eb63f63e6c6 revisions and indexing
paulson
parents: 11450
diff changeset
   222
\label{sec:nat}\index{natural numbers}%
11428
332347b9b942 tidying the index
paulson
parents: 11419
diff changeset
   223
\index{linear arithmetic|(}
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   224
17182
ae84287f44e3 tune spacing where a generated theory text is included directly;
wenzelm
parents: 16523
diff changeset
   225
\input{Misc/document/fakenat.tex}\medskip
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   226
\input{Misc/document/natsum.tex}
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   227
11428
332347b9b942 tidying the index
paulson
parents: 11419
diff changeset
   228
\index{linear arithmetic|)}
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   229
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   230
10396
5ab08609e6c8 *** empty log message ***
nipkow
parents: 10237
diff changeset
   231
\subsection{Pairs}
9541
d17c0b34d5c8 *** empty log message ***
nipkow
parents: 9494
diff changeset
   232
\input{Misc/document/pairs.tex}
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   233
10608
620647438780 *** empty log message ***
nipkow
parents: 10543
diff changeset
   234
\subsection{Datatype {\tt\slshape option}}
10543
8e4307d1207a *** empty log message ***
nipkow
parents: 10538
diff changeset
   235
\label{sec:option}
8e4307d1207a *** empty log message ***
nipkow
parents: 10538
diff changeset
   236
\input{Misc/document/Option2.tex}
8e4307d1207a *** empty log message ***
nipkow
parents: 10538
diff changeset
   237
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   238
\section{Definitions}
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   239
\label{sec:Definitions}
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   240
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   241
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
   242
construction. In particular, definitions cannot be recursive. Isabelle offers
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   243
definitions on the level of types and terms. Those on the type level are
11456
7eb63f63e6c6 revisions and indexing
paulson
parents: 11450
diff changeset
   244
called \textbf{type synonyms}; those on the term level are simply called 
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   245
definitions.
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   246
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   247
10885
90695f46440b lcp's pass over the book, chapters 1-8
paulson
parents: 10824
diff changeset
   248
\subsection{Type Synonyms}
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   249
12327
5a4d78204492 *** empty log message ***
nipkow
parents: 11646
diff changeset
   250
\index{type synonyms}%
11456
7eb63f63e6c6 revisions and indexing
paulson
parents: 11450
diff changeset
   251
Type synonyms are similar to those found in ML\@. They are created by a 
11428
332347b9b942 tidying the index
paulson
parents: 11419
diff changeset
   252
\commdx{types} command:
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   253
17182
ae84287f44e3 tune spacing where a generated theory text is included directly;
wenzelm
parents: 16523
diff changeset
   254
\medskip
12327
5a4d78204492 *** empty log message ***
nipkow
parents: 11646
diff changeset
   255
\input{Misc/document/types.tex}
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   256
9844
8016321c7de1 *** empty log message ***
nipkow
parents: 9792
diff changeset
   257
\input{Misc/document/prime_def.tex}
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   258
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   259
11201
eddc69b55fac *** empty log message ***
nipkow
parents: 11159
diff changeset
   260
\section{The Definitional Approach}
eddc69b55fac *** empty log message ***
nipkow
parents: 11159
diff changeset
   261
\label{sec:definitional}
eddc69b55fac *** empty log message ***
nipkow
parents: 11159
diff changeset
   262
11457
279da0358aa9 additional revisions to chapters 1, 2
paulson
parents: 11456
diff changeset
   263
\index{Definitional Approach}%
11201
eddc69b55fac *** empty log message ***
nipkow
parents: 11159
diff changeset
   264
As we pointed out at the beginning of the chapter, asserting arbitrary
11456
7eb63f63e6c6 revisions and indexing
paulson
parents: 11450
diff changeset
   265
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
   266
to avoid this danger, we advocate the definitional rather than
11456
7eb63f63e6c6 revisions and indexing
paulson
parents: 11450
diff changeset
   267
the axiomatic approach: introduce new concepts by definitions. However,  Isabelle/HOL seems to
7eb63f63e6c6 revisions and indexing
paulson
parents: 11450
diff changeset
   268
support many richer definitional constructs, such as
7eb63f63e6c6 revisions and indexing
paulson
parents: 11450
diff changeset
   269
\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
   270
\isacommand{primrec} function definition is turned into a proper
7eb63f63e6c6 revisions and indexing
paulson
parents: 11450
diff changeset
   271
(nonrecursive!) definition from which the user-supplied recursion equations are
11457
279da0358aa9 additional revisions to chapters 1, 2
paulson
parents: 11456
diff changeset
   272
automatically proved.  This process is
11456
7eb63f63e6c6 revisions and indexing
paulson
parents: 11450
diff changeset
   273
hidden from the user, who does not have to understand the details.  Other commands described
25258
22d16596c306 recdef -> fun
nipkow
parents: 17182
diff changeset
   274
later, like \isacommand{fun} and \isacommand{inductive}, work similarly.  
11457
279da0358aa9 additional revisions to chapters 1, 2
paulson
parents: 11456
diff changeset
   275
This strict adherence to the definitional approach reduces the risk of 
279da0358aa9 additional revisions to chapters 1, 2
paulson
parents: 11456
diff changeset
   276
soundness errors.
11201
eddc69b55fac *** empty log message ***
nipkow
parents: 11159
diff changeset
   277
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   278
\chapter{More Functional Programming}
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   279
11458
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11457
diff changeset
   280
The purpose of this chapter is to deepen your understanding of the
8771
026f37a86ea7 *** empty log message ***
nipkow
parents: 8743
diff changeset
   281
concepts encountered so far and to introduce advanced forms of datatypes and
026f37a86ea7 *** empty log message ***
nipkow
parents: 8743
diff changeset
   282
recursive functions. The first two sections give a structured presentation of
10538
d1bf9ca9008d *** empty log message ***
nipkow
parents: 10522
diff changeset
   283
theorem proving by simplification ({\S}\ref{sec:Simplification}) and discuss
11458
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11457
diff changeset
   284
important heuristics for induction ({\S}\ref{sec:InductionHeuristics}).  You can
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11457
diff changeset
   285
skip them if you are not planning to perform proofs yourself.
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11457
diff changeset
   286
We then present a case
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11457
diff changeset
   287
study: a compiler for expressions ({\S}\ref{sec:ExprCompiler}). Advanced
8771
026f37a86ea7 *** empty log message ***
nipkow
parents: 8743
diff changeset
   288
datatypes, including those involving function spaces, are covered in
11458
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11457
diff changeset
   289
{\S}\ref{sec:advanced-datatypes}; it closes with another case study, search
25258
22d16596c306 recdef -> fun
nipkow
parents: 17182
diff changeset
   290
trees (``tries'').  Finally we introduce \isacommand{fun}, a general
11458
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11457
diff changeset
   291
form of recursive function definition that goes well beyond 
25258
22d16596c306 recdef -> fun
nipkow
parents: 17182
diff changeset
   292
\isacommand{primrec} ({\S}\ref{sec:fun}).
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   293
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   294
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   295
\section{Simplification}
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   296
\label{sec:Simplification}
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   297
\index{simplification|(}
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   298
10795
9e888d60d3e5 minor edits to Chapters 1-3
paulson
parents: 10654
diff changeset
   299
So far we have proved our theorems by \isa{auto}, which simplifies
11458
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11457
diff changeset
   300
all subgoals. In fact, \isa{auto} can do much more than that. 
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11457
diff changeset
   301
To go beyond toy examples, you
9541
d17c0b34d5c8 *** empty log message ***
nipkow
parents: 9494
diff changeset
   302
need to understand the ingredients of \isa{auto}.  This section covers the
10971
6852682eaf16 *** empty log message ***
nipkow
parents: 10885
diff changeset
   303
method that \isa{auto} always applies first, simplification.
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   304
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   305
Simplification is one of the central theorem proving tools in Isabelle and
11458
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11457
diff changeset
   306
many other systems. The tool itself is called the \textbf{simplifier}. 
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11457
diff changeset
   307
This section introduces the many features of the simplifier
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11457
diff changeset
   308
and is required reading if you intend to perform proofs.  Later on,
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11457
diff changeset
   309
{\S}\ref{sec:simplification-II} explains some more advanced features and a
9754
a123a64cadeb *** empty log message ***
nipkow
parents: 9742
diff changeset
   310
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
   311
section as well, in particular to understand why the simplifier did
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11457
diff changeset
   312
something unexpected.
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   313
11458
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11457
diff changeset
   314
\subsection{What is Simplification?}
9458
c613cd06d5cf apply. -> by
nipkow
parents: 8771
diff changeset
   315
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   316
In its most basic form, simplification means repeated application of
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   317
equations from left to right. For example, taking the rules for \isa{\at}
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   318
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
   319
simplification steps:
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   320
\begin{ttbox}\makeatother
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   321
(0#1#[]) @ []  \(\leadsto\)  0#((1#[]) @ [])  \(\leadsto\)  0#(1#([] @ []))  \(\leadsto\)  0#1#[]
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   322
\end{ttbox}
9933
9feb1e0c4cb3 *** empty log message ***
nipkow
parents: 9844
diff changeset
   323
This is also known as \bfindex{term rewriting}\indexbold{rewriting} and the
11458
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11457
diff changeset
   324
equations are referred to as \bfindex{rewrite rules}.
9933
9feb1e0c4cb3 *** empty log message ***
nipkow
parents: 9844
diff changeset
   325
``Rewriting'' is more honest than ``simplification'' because the terms do not
9feb1e0c4cb3 *** empty log message ***
nipkow
parents: 9844
diff changeset
   326
necessarily become simpler in the process.
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   327
11458
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11457
diff changeset
   328
The simplifier proves arithmetic goals as described in
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11457
diff changeset
   329
{\S}\ref{sec:nat} above.  Arithmetic expressions are simplified using built-in
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11457
diff changeset
   330
procedures that go beyond mere rewrite rules.  New simplification procedures
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11457
diff changeset
   331
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
   332
tutorial. 
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11457
diff changeset
   333
9844
8016321c7de1 *** empty log message ***
nipkow
parents: 9792
diff changeset
   334
\input{Misc/document/simp.tex}
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   335
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   336
\index{simplification|)}
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   337
9844
8016321c7de1 *** empty log message ***
nipkow
parents: 9792
diff changeset
   338
\input{Misc/document/Itrev.tex}
13305
f88d0c363582 *** empty log message ***
nipkow
parents: 12582
diff changeset
   339
\begin{exercise}
f88d0c363582 *** empty log message ***
nipkow
parents: 12582
diff changeset
   340
\input{Misc/document/Plus.tex}%
f88d0c363582 *** empty log message ***
nipkow
parents: 12582
diff changeset
   341
\end{exercise}
9493
494f8cd34df7 *** empty log message ***
nipkow
parents: 9458
diff changeset
   342
\begin{exercise}
494f8cd34df7 *** empty log message ***
nipkow
parents: 9458
diff changeset
   343
\input{Misc/document/Tree2.tex}%
494f8cd34df7 *** empty log message ***
nipkow
parents: 9458
diff changeset
   344
\end{exercise}
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   345
9844
8016321c7de1 *** empty log message ***
nipkow
parents: 9792
diff changeset
   346
\input{CodeGen/document/CodeGen.tex}
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   347
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   348
10885
90695f46440b lcp's pass over the book, chapters 1-8
paulson
parents: 10824
diff changeset
   349
\section{Advanced Datatypes}
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   350
\label{sec:advanced-datatypes}
11428
332347b9b942 tidying the index
paulson
parents: 11419
diff changeset
   351
\index{datatype@\isacommand {datatype} (command)|(}
332347b9b942 tidying the index
paulson
parents: 11419
diff changeset
   352
\index{primrec@\isacommand {primrec} (command)|(}
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   353
%|)
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   354
11428
332347b9b942 tidying the index
paulson
parents: 11419
diff changeset
   355
This section presents advanced forms of datatypes: mutual and nested
332347b9b942 tidying the index
paulson
parents: 11419
diff changeset
   356
recursion.  A series of examples will culminate in a treatment of the trie
332347b9b942 tidying the index
paulson
parents: 11419
diff changeset
   357
data structure.
332347b9b942 tidying the index
paulson
parents: 11419
diff changeset
   358
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   359
10885
90695f46440b lcp's pass over the book, chapters 1-8
paulson
parents: 10824
diff changeset
   360
\subsection{Mutual Recursion}
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   361
\label{sec:datatype-mut-rec}
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   362
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   363
\input{Datatype/document/ABexpr.tex}
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   364
10885
90695f46440b lcp's pass over the book, chapters 1-8
paulson
parents: 10824
diff changeset
   365
\subsection{Nested Recursion}
9644
6b0b6b471855 *** empty log message ***
nipkow
parents: 9541
diff changeset
   366
\label{sec:nested-datatype}
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   367
9644
6b0b6b471855 *** empty log message ***
nipkow
parents: 9541
diff changeset
   368
{\makeatother\input{Datatype/document/Nested.tex}}
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   369
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   370
11419
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   371
\subsection{The Limits of Nested Recursion}
12327
5a4d78204492 *** empty log message ***
nipkow
parents: 11646
diff changeset
   372
\label{sec:nested-fun-datatype}
11419
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   373
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   374
How far can we push nested recursion? By the unfolding argument above, we can
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   375
reduce nested to mutual recursion provided the nested recursion only involves
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   376
previously defined datatypes. This does not include functions:
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   377
\begin{isabelle}
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   378
\isacommand{datatype} t = C "t \isasymRightarrow\ bool"
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   379
\end{isabelle}
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   380
This declaration is a real can of worms.
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   381
In HOL it must be ruled out because it requires a type
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   382
\isa{t} such that \isa{t} and its power set \isa{t \isasymFun\ bool} have the
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   383
same cardinality --- an impossibility. For the same reason it is not possible
12328
7c4ec77a8715 *** empty log message ***
nipkow
parents: 12327
diff changeset
   384
to allow recursion involving the type \isa{t set}, which is isomorphic to
11419
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   385
\isa{t \isasymFun\ bool}.
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   386
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   387
Fortunately, a limited form of recursion
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   388
involving function spaces is permitted: the recursive type may occur on the
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   389
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
   390
is ruled out but the following example of a potentially 
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11457
diff changeset
   391
\index{infinitely branching trees}%
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11457
diff changeset
   392
infinitely branching tree is accepted:
11419
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   393
\smallskip
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   394
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   395
\input{Datatype/document/Fundata.tex}
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   396
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   397
If you need nested recursion on the left of a function arrow, there are
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   398
alternatives to pure HOL\@.  In the Logic for Computable Functions 
11458
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11457
diff changeset
   399
(\rmindex{LCF}), types like
11419
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   400
\begin{isabelle}
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   401
\isacommand{datatype} lam = C "lam \isasymrightarrow\ lam"
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   402
\end{isabelle}
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   403
do indeed make sense~\cite{paulson87}.  Note the different arrow,
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   404
\isa{\isasymrightarrow} instead of \isa{\isasymRightarrow},
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   405
expressing the type of \emph{continuous} functions. 
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   406
There is even a version of LCF on top of HOL,
11458
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11457
diff changeset
   407
called \rmindex{HOLCF}~\cite{MuellerNvOS99}.
11428
332347b9b942 tidying the index
paulson
parents: 11419
diff changeset
   408
\index{datatype@\isacommand {datatype} (command)|)}
332347b9b942 tidying the index
paulson
parents: 11419
diff changeset
   409
\index{primrec@\isacommand {primrec} (command)|)}
11419
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   410
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   411
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   412
\subsection{Case Study: Tries}
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   413
\label{sec:Trie}
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   414
11458
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11457
diff changeset
   415
\index{tries|(}%
11419
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   416
Tries are a classic search tree data structure~\cite{Knuth3-75} for fast
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   417
indexing with strings. Figure~\ref{fig:trie} gives a graphical example of a
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   418
trie containing the words ``all'', ``an'', ``ape'', ``can'', ``car'' and
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   419
``cat''.  When searching a string in a trie, the letters of the string are
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   420
examined sequentially. Each letter determines which subtrie to search next.
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   421
In this case study we model tries as a datatype, define a lookup and an
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   422
update function, and prove that they behave as expected.
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   423
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   424
\begin{figure}[htbp]
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   425
\begin{center}
8743
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   426
\unitlength1mm
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   427
\begin{picture}(60,30)
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   428
\put( 5, 0){\makebox(0,0)[b]{l}}
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   429
\put(25, 0){\makebox(0,0)[b]{e}}
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   430
\put(35, 0){\makebox(0,0)[b]{n}}
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   431
\put(45, 0){\makebox(0,0)[b]{r}}
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   432
\put(55, 0){\makebox(0,0)[b]{t}}
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   433
%
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   434
\put( 5, 9){\line(0,-1){5}}
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   435
\put(25, 9){\line(0,-1){5}}
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   436
\put(44, 9){\line(-3,-2){9}}
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   437
\put(45, 9){\line(0,-1){5}}
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   438
\put(46, 9){\line(3,-2){9}}
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   439
%
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   440
\put( 5,10){\makebox(0,0)[b]{l}}
3253c6046d57 I wonder if that's all?
nipkow
parents:
diff changeset
   441
\put(15,10){\makebox(0,0)[b]{n}}
11419
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   442
\put(25,10){\makebox(0,0)[b]{p}}
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   443
\put(45,10){\makebox(0,0)[b]{a}}
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   444
%
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   445
\put(14,19){\line(-3,-2){9}}
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   446
\put(15,19){\line(0,-1){5}}
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   447
\put(16,19){\line(3,-2){9}}
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   448
\put(45,19){\line(0,-1){5}}
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   449
%
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   450
\put(15,20){\makebox(0,0)[b]{a}}
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   451
\put(45,20){\makebox(0,0)[b]{c}}
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   452
%
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   453
\put(30,30){\line(-3,-2){13}}
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   454
\put(30,30){\line(3,-2){13}}
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   455
\end{picture}
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   456
\end{center}
11450
1b02a6c4032f tweaks and indexing
paulson
parents: 11428
diff changeset
   457
\caption{A Sample Trie}
11419
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   458
\label{fig:trie}
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   459
\end{figure}
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   460
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   461
Proper tries associate some value with each string. Since the
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   462
information is stored only in the final node associated with the string, many
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   463
nodes do not carry any value. This distinction is modeled with the help
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   464
of the predefined datatype \isa{option} (see {\S}\ref{sec:option}).
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   465
\input{Trie/document/Trie.tex}
11458
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11457
diff changeset
   466
\index{tries|)}
11419
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   467
25261
3dc292be0b54 recdef -> fun
nipkow
parents: 25258
diff changeset
   468
\section{Total Recursive Functions: \isacommand{fun}}
25258
22d16596c306 recdef -> fun
nipkow
parents: 17182
diff changeset
   469
\label{sec:fun}
22d16596c306 recdef -> fun
nipkow
parents: 17182
diff changeset
   470
\index{fun@\isacommand {fun} (command)|(}\index{functions!total|(}
11419
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   471
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   472
Although many total functions have a natural primitive recursive definition,
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   473
this is not always the case. Arbitrary total recursive functions can be
25263
b54744935036 *** empty log message ***
nipkow
parents: 25261
diff changeset
   474
defined by means of \isacommand{fun}: you can use full pattern matching,
11419
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   475
recursion need not involve datatypes, and termination is proved by showing
25258
22d16596c306 recdef -> fun
nipkow
parents: 17182
diff changeset
   476
that the arguments of all recursive calls are smaller in a suitable sense.
22d16596c306 recdef -> fun
nipkow
parents: 17182
diff changeset
   477
In this section we restrict ourselves to functions where Isabelle can prove
25281
8d309beb66d6 removed advanced recdef section and replaced it by citation of Alex's tutorial.
nipkow
parents: 25263
diff changeset
   478
termination automatically. More advanced function definitions, including user
8d309beb66d6 removed advanced recdef section and replaced it by citation of Alex's tutorial.
nipkow
parents: 25263
diff changeset
   479
supplied termination proofs, nested recursion and partiality, are discussed
8d309beb66d6 removed advanced recdef section and replaced it by citation of Alex's tutorial.
nipkow
parents: 25263
diff changeset
   480
in a separate tutorial~\cite{isabelle-function}.
11419
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   481
25258
22d16596c306 recdef -> fun
nipkow
parents: 17182
diff changeset
   482
\input{Fun/document/fun0.tex}
11419
9577530e8a5a more indexing
paulson
parents: 11309
diff changeset
   483
25258
22d16596c306 recdef -> fun
nipkow
parents: 17182
diff changeset
   484
\index{fun@\isacommand {fun} (command)|)}\index{functions!total|)}