8749
|
1 |
\begin{isabelle}%
|
|
2 |
%
|
|
3 |
\begin{isamarkuptext}%
|
|
4 |
So far, all datatypes had the property that on the right-hand side of their
|
|
5 |
definition they occurred only at the top-level, i.e.\ directly below a
|
|
6 |
constructor. This is not the case any longer for the following model of terms
|
|
7 |
where function symbols can be applied to a list of arguments:%
|
|
8 |
\end{isamarkuptext}%
|
|
9 |
\isacommand{datatype}~('a,'b){"}term{"}~=~Var~'a~|~App~'b~{"}('a,'b)term~list{"}%
|
|
10 |
\begin{isamarkuptext}%
|
|
11 |
\noindent
|
|
12 |
Note that we need to quote \isa{term} on the left to avoid confusion with
|
|
13 |
the command \isacommand{term}.
|
|
14 |
Parameter \isa{'a} is the type of variables and \isa{'b} the type of
|
|
15 |
function symbols.
|
|
16 |
A mathematical term like $f(x,g(y))$ becomes \isa{App f [Var x, App g
|
|
17 |
[Var y]]}, where \isa{f}, \isa{g}, \isa{x}, \isa{y} are
|
|
18 |
suitable values, e.g.\ numbers or strings.
|
|
19 |
|
|
20 |
What complicates the definition of \isa{term} is the nested occurrence of
|
|
21 |
\isa{term} inside \isa{list} on the right-hand side. In principle,
|
|
22 |
nested recursion can be eliminated in favour of mutual recursion by unfolding
|
|
23 |
the offending datatypes, here \isa{list}. The result for \isa{term}
|
|
24 |
would be something like
|
8751
|
25 |
\medskip
|
|
26 |
|
|
27 |
\input{Datatype/document/unfoldnested.tex}
|
|
28 |
\medskip
|
|
29 |
|
|
30 |
\noindent
|
8749
|
31 |
Although we do not recommend this unfolding to the user, it shows how to
|
|
32 |
simulate nested recursion by mutual recursion.
|
|
33 |
Now we return to the initial definition of \isa{term} using
|
|
34 |
nested recursion.
|
|
35 |
|
|
36 |
Let us define a substitution function on terms. Because terms involve term
|
|
37 |
lists, we need to define two substitution functions simultaneously:%
|
|
38 |
\end{isamarkuptext}%
|
|
39 |
\isacommand{consts}\isanewline
|
|
40 |
subst~::~{"}('a{\isasymRightarrow}('a,'b)term)~{\isasymRightarrow}~('a,'b)term~~~~~~{\isasymRightarrow}~('a,'b)term{"}\isanewline
|
|
41 |
substs::~{"}('a{\isasymRightarrow}('a,'b)term)~{\isasymRightarrow}~('a,'b)term~list~{\isasymRightarrow}~('a,'b)term~list{"}\isanewline
|
|
42 |
\isanewline
|
|
43 |
\isacommand{primrec}\isanewline
|
|
44 |
~~{"}subst~s~(Var~x)~=~s~x{"}\isanewline
|
|
45 |
~~{"}subst~s~(App~f~ts)~=~App~f~(substs~s~ts){"}\isanewline
|
|
46 |
\isanewline
|
|
47 |
~~{"}substs~s~[]~=~[]{"}\isanewline
|
|
48 |
~~{"}substs~s~(t~\#~ts)~=~subst~s~t~\#~substs~s~ts{"}%
|
|
49 |
\begin{isamarkuptext}%
|
|
50 |
\noindent
|
|
51 |
Similarly, when proving a statement about terms inductively, we need
|
|
52 |
to prove a related statement about term lists simultaneously. For example,
|
|
53 |
the fact that the identity substitution does not change a term needs to be
|
|
54 |
strengthened and proved as follows:%
|
|
55 |
\end{isamarkuptext}%
|
|
56 |
\isacommand{lemma}~{"}subst~~Var~t~~=~(t~::('a,'b)term)~~{\isasymand}\isanewline
|
|
57 |
~~~~~~~~substs~Var~ts~=~(ts::('a,'b)term~list){"}\isanewline
|
|
58 |
\isacommand{apply}(induct\_tac~t~\isakeyword{and}~ts,~auto)\isacommand{.}%
|
|
59 |
\begin{isamarkuptext}%
|
|
60 |
\noindent
|
|
61 |
Note that \isa{Var} is the identity substitution because by definition it
|
|
62 |
leaves variables unchanged: \isa{subst Var (Var $x$) = Var $x$}. Note also
|
|
63 |
that the type annotations are necessary because otherwise there is nothing in
|
|
64 |
the goal to enforce that both halves of the goal talk about the same type
|
|
65 |
parameters \isa{('a,'b)}. As a result, induction would fail
|
|
66 |
because the two halves of the goal would be unrelated.
|
|
67 |
|
|
68 |
\begin{exercise}
|
|
69 |
The fact that substitution distributes over composition can be expressed
|
|
70 |
roughly as follows:
|
|
71 |
\begin{ttbox}
|
|
72 |
subst (f o g) t = subst f (subst g t)
|
|
73 |
\end{ttbox}
|
|
74 |
Correct this statement (you will find that it does not type-check),
|
|
75 |
strengthen it, and prove it. (Note: \ttindexbold{o} is function composition;
|
|
76 |
its definition is found in theorem \isa{o_def}).
|
|
77 |
\end{exercise}
|
|
78 |
|
|
79 |
Of course, you may also combine mutual and nested recursion. For example,
|
|
80 |
constructor \isa{Sum} in \S\ref{sec:datatype-mut-rec} could take a list of
|
|
81 |
expressions as its argument: \isa{Sum "'a aexp list"}.%
|
|
82 |
\end{isamarkuptext}%
|
|
83 |
\end{isabelle}%
|
9145
|
84 |
%%% Local Variables:
|
|
85 |
%%% mode: latex
|
|
86 |
%%% TeX-master: "root"
|
|
87 |
%%% End:
|