doc-src/TutorialI/Datatype/Fundata.thy
author wenzelm
Thu, 07 Sep 2000 20:57:22 +0200
changeset 9901 09a142decdda
parent 9834 109b11c4e77e
child 9933 9feb1e0c4cb3
permissions -rw-r--r--
avoid handle_error (better msgs);
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     1
(*<*)
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     2
theory Fundata = Main:
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     3
(*>*)
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     4
datatype ('a,'i)bigtree = Tip | Branch 'a "'i \\<Rightarrow> ('a,'i)bigtree"
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     5
9792
bbefb6ce5cb2 *** empty log message ***
nipkow
parents: 9723
diff changeset
     6
text{*\noindent
bbefb6ce5cb2 *** empty log message ***
nipkow
parents: 9723
diff changeset
     7
Parameter @{typ"'a"} is the type of values stored in
bbefb6ce5cb2 *** empty log message ***
nipkow
parents: 9723
diff changeset
     8
the @{term"Branch"}es of the tree, whereas @{typ"'i"} is the index
bbefb6ce5cb2 *** empty log message ***
nipkow
parents: 9723
diff changeset
     9
type over which the tree branches. If @{typ"'i"} is instantiated to
bbefb6ce5cb2 *** empty log message ***
nipkow
parents: 9723
diff changeset
    10
@{typ"bool"}, the result is a binary tree; if it is instantiated to
bbefb6ce5cb2 *** empty log message ***
nipkow
parents: 9723
diff changeset
    11
@{typ"nat"}, we have an infinitely branching tree because each node
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    12
has as many subtrees as there are natural numbers. How can we possibly
9541
d17c0b34d5c8 *** empty log message ***
nipkow
parents: 8771
diff changeset
    13
write down such a tree? Using functional notation! For example, the term
d17c0b34d5c8 *** empty log message ***
nipkow
parents: 8771
diff changeset
    14
@{term[display]"Branch 0 (%i. Branch i (%n. Tip))"}
d17c0b34d5c8 *** empty log message ***
nipkow
parents: 8771
diff changeset
    15
of type @{typ"(nat,nat)bigtree"} is the tree whose
8771
026f37a86ea7 *** empty log message ***
nipkow
parents: 8745
diff changeset
    16
root is labeled with 0 and whose $i$th subtree is labeled with $i$ and
9792
bbefb6ce5cb2 *** empty log message ***
nipkow
parents: 9723
diff changeset
    17
has merely @{term"Tip"}s as further subtrees.
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    18
9792
bbefb6ce5cb2 *** empty log message ***
nipkow
parents: 9723
diff changeset
    19
Function @{term"map_bt"} applies a function to all labels in a @{text"bigtree"}:
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    20
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    21
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    22
consts map_bt :: "('a \\<Rightarrow> 'b) \\<Rightarrow> ('a,'i)bigtree \\<Rightarrow> ('b,'i)bigtree"
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    23
primrec
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    24
"map_bt f Tip          = Tip"
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    25
"map_bt f (Branch a F) = Branch (f a) (\\<lambda>i. map_bt f (F i))"
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    26
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    27
text{*\noindent This is a valid \isacommand{primrec} definition because the
9792
bbefb6ce5cb2 *** empty log message ***
nipkow
parents: 9723
diff changeset
    28
recursive calls of @{term"map_bt"} involve only subtrees obtained from
bbefb6ce5cb2 *** empty log message ***
nipkow
parents: 9723
diff changeset
    29
@{term"F"}, i.e.\ the left-hand side. Thus termination is assured.  The
9541
d17c0b34d5c8 *** empty log message ***
nipkow
parents: 8771
diff changeset
    30
seasoned functional programmer might have written @{term"map_bt f o F"}
d17c0b34d5c8 *** empty log message ***
nipkow
parents: 8771
diff changeset
    31
instead of @{term"%i. map_bt f (F i)"}, but the former is not accepted by
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    32
Isabelle because the termination proof is not as obvious since
9792
bbefb6ce5cb2 *** empty log message ***
nipkow
parents: 9723
diff changeset
    33
@{term"map_bt"} is only partially applied.
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    34
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    35
The following lemma has a canonical proof  *}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    36
8771
026f37a86ea7 *** empty log message ***
nipkow
parents: 8745
diff changeset
    37
lemma "map_bt (g o f) T = map_bt g (map_bt f T)";
9689
751fde5307e4 *** empty log message ***
nipkow
parents: 9644
diff changeset
    38
by(induct_tac "T", simp_all)
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    39
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    40
text{*\noindent
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    41
but it is worth taking a look at the proof state after the induction step
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    42
to understand what the presence of the function type entails:
9723
a977245dfc8a *** empty log message ***
nipkow
parents: 9689
diff changeset
    43
\begin{isabelle}
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    44
~1.~map\_bt~g~(map\_bt~f~Tip)~=~map\_bt~(g~{\isasymcirc}~f)~Tip\isanewline
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    45
~2.~{\isasymAnd}a~F.\isanewline
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    46
~~~~~~{\isasymforall}x.~map\_bt~g~(map\_bt~f~(F~x))~=~map\_bt~(g~{\isasymcirc}~f)~(F~x)~{\isasymLongrightarrow}\isanewline
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    47
~~~~~~map\_bt~g~(map\_bt~f~(Branch~a~F))~=~map\_bt~(g~{\isasymcirc}~f)~(Branch~a~F)%
9723
a977245dfc8a *** empty log message ***
nipkow
parents: 9689
diff changeset
    48
\end{isabelle}
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    49
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    50
(*<*)
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    51
end
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    52
(*>*)