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