doc-src/TutorialI/Datatype/document/Fundata.tex
author nipkow
Fri, 18 Aug 2000 10:34:08 +0200
changeset 9644 6b0b6b471855
parent 9541 d17c0b34d5c8
child 9673 1b2d4f995b13
permissions -rw-r--r--
*** empty log message ***

\begin{isabelle}%
\isacommand{datatype}\ ('a,'i)bigtree\ =\ Tip\ |\ Branch\ 'a\ {"}'i\ {\isasymRightarrow}\ ('a,'i)bigtree{"}%
\begin{isamarkuptext}%
\noindent Parameter \isa{'a} is the type of values stored in
the \isa{Branch}es of the tree, whereas \isa{'i} is the index
type over which the tree branches. If \isa{'i} is instantiated to
\isa{bool}, the result is a binary tree; if it is instantiated to
\isa{nat}, we have an infinitely branching tree because each node
has as many subtrees as there are natural numbers. How can we possibly
write down such a tree? Using functional notation! For example, the term
\begin{quote}

\begin{isabelle}%
Branch\ 0\ ({\isasymlambda}\mbox{i}.\ Branch\ \mbox{i}\ ({\isasymlambda}\mbox{n}.\ Tip))
\end{isabelle}%

\end{quote}
of type \isa{(nat,\ nat)\ bigtree} is the tree whose
root is labeled with 0 and whose $i$th subtree is labeled with $i$ and
has merely \isa{Tip}s as further subtrees.

Function \isa{map_bt} applies a function to all labels in a \isa{bigtree}:%
\end{isamarkuptext}%
\isacommand{consts}\ map\_bt\ ::\ {"}('a\ {\isasymRightarrow}\ 'b)\ {\isasymRightarrow}\ ('a,'i)bigtree\ {\isasymRightarrow}\ ('b,'i)bigtree{"}\isanewline
\isacommand{primrec}\isanewline
{"}map\_bt\ f\ Tip\ \ \ \ \ \ \ \ \ \ =\ Tip{"}\isanewline
{"}map\_bt\ f\ (Branch\ a\ F)\ =\ Branch\ (f\ a)\ ({\isasymlambda}i.\ map\_bt\ f\ (F\ i)){"}%
\begin{isamarkuptext}%
\noindent This is a valid \isacommand{primrec} definition because the
recursive calls of \isa{map_bt} involve only subtrees obtained from
\isa{F}, i.e.\ the left-hand side. Thus termination is assured.  The
seasoned functional programmer might have written \isa{map\_bt\ \mbox{f}\ {\isasymcirc}\ \mbox{F}}
instead of \isa{{\isasymlambda}\mbox{i}.\ map\_bt\ \mbox{f}\ (\mbox{F}\ \mbox{i})}, but the former is not accepted by
Isabelle because the termination proof is not as obvious since
\isa{map_bt} is only partially applied.

The following lemma has a canonical proof%
\end{isamarkuptext}%
\isacommand{lemma}\ {"}map\_bt\ (g\ o\ f)\ T\ =\ map\_bt\ g\ (map\_bt\ f\ T){"}\isanewline
\isacommand{by}(induct\_tac\ {"}T{"},\ auto)%
\begin{isamarkuptext}%
\noindent
but it is worth taking a look at the proof state after the induction step
to understand what the presence of the function type entails:
\begin{isabellepar}%
~1.~map\_bt~g~(map\_bt~f~Tip)~=~map\_bt~(g~{\isasymcirc}~f)~Tip\isanewline
~2.~{\isasymAnd}a~F.\isanewline
~~~~~~{\isasymforall}x.~map\_bt~g~(map\_bt~f~(F~x))~=~map\_bt~(g~{\isasymcirc}~f)~(F~x)~{\isasymLongrightarrow}\isanewline
~~~~~~map\_bt~g~(map\_bt~f~(Branch~a~F))~=~map\_bt~(g~{\isasymcirc}~f)~(Branch~a~F)%
\end{isabellepar}%%
\end{isamarkuptext}%
\end{isabelle}%
%%% Local Variables:
%%% mode: latex
%%% TeX-master: "root"
%%% End: