9690
|
1 |
\begin{isabelle}%
|
|
2 |
%
|
|
3 |
\begin{isamarkuptext}%
|
|
4 |
\noindent
|
|
5 |
The termintion condition is easily proved by induction:%
|
|
6 |
\end{isamarkuptext}%
|
|
7 |
\isacommand{lemma}\ [simp]:\ {"}t\ {\isasymin}\ set\ ts\ {\isasymlongrightarrow}\ size\ t\ <\ Suc(term\_size\ ts){"}\isanewline
|
|
8 |
\isacommand{by}(induct\_tac\ ts,\ auto)%
|
|
9 |
\begin{isamarkuptext}%
|
|
10 |
\noindent
|
|
11 |
By making this theorem a simplification rule, \isacommand{recdef}
|
|
12 |
applies it automatically and the above definition of \isa{trev}
|
|
13 |
succeeds now. As a reward for our effort, we can now prove the desired
|
|
14 |
lemma directly. The key is the fact that we no longer need the verbose
|
|
15 |
induction schema for type \isa{term} but the simpler one arising from
|
|
16 |
\isa{trev}:%
|
|
17 |
\end{isamarkuptext}%
|
|
18 |
\isacommand{lemmas}\ [cong]\ =\ map\_cong\isanewline
|
|
19 |
\isacommand{lemma}\ {"}trev(trev\ t)\ =\ t{"}\isanewline
|
|
20 |
\isacommand{apply}(induct\_tac\ t\ rule:trev.induct)%
|
|
21 |
\begin{isamarkuptxt}%
|
|
22 |
\noindent
|
|
23 |
This leaves us with a trivial base case \isa{trev\ (trev\ (Var\ \mbox{x}))\ =\ Var\ \mbox{x}} and the step case
|
|
24 |
\begin{quote}
|
|
25 |
|
|
26 |
\begin{isabelle}%
|
|
27 |
{\isasymforall}\mbox{t}.\ \mbox{t}\ {\isasymin}\ set\ \mbox{ts}\ {\isasymlongrightarrow}\ trev\ (trev\ \mbox{t})\ =\ \mbox{t}\ {\isasymLongrightarrow}\isanewline
|
|
28 |
trev\ (trev\ (App\ \mbox{f}\ \mbox{ts}))\ =\ App\ \mbox{f}\ \mbox{ts}
|
|
29 |
\end{isabelle}%
|
|
30 |
|
|
31 |
\end{quote}
|
|
32 |
both of which are solved by simplification:%
|
|
33 |
\end{isamarkuptxt}%
|
|
34 |
\isacommand{by}(simp\_all\ del:map\_compose\ add:sym[OF\ map\_compose]\ rev\_map)%
|
|
35 |
\begin{isamarkuptext}%
|
|
36 |
\noindent
|
|
37 |
If this surprises you, see Datatype/Nested2......
|
|
38 |
|
|
39 |
The above definition of \isa{trev} is superior to the one in \S\ref{sec:nested-datatype}
|
|
40 |
because it brings \isa{rev} into play, about which already know a lot, in particular
|
|
41 |
\isa{rev\ (rev\ \mbox{xs})\ =\ \mbox{xs}}.
|
|
42 |
Thus this proof is a good example of an important principle:
|
|
43 |
\begin{quote}
|
|
44 |
\emph{Chose your definitions carefully\\
|
|
45 |
because they determine the complexity of your proofs.}
|
|
46 |
\end{quote}
|
|
47 |
|
|
48 |
Let us now return to the question of how \isacommand{recdef} can come up with sensible termination
|
|
49 |
conditions in the presence of higher-order functions like \isa{map}. For a start, if nothing
|
|
50 |
were known about \isa{map}, \isa{map\ trev\ \mbox{ts}} might apply \isa{trev} to arbitrary terms,
|
|
51 |
and thus \isacommand{recdef} would try to prove the unprovable
|
|
52 |
\isa{size\ \mbox{t}\ <\ Suc\ (term\_size\ \mbox{ts})}, without any assumption about \isa{t}.
|
|
53 |
Therefore \isacommand{recdef} has been supplied with the congruence theorem \isa{map\_cong}:
|
|
54 |
\begin{quote}
|
|
55 |
|
|
56 |
\begin{isabelle}%
|
|
57 |
{\isasymlbrakk}\mbox{xs}\ =\ \mbox{ys};\ {\isasymAnd}\mbox{x}.\ \mbox{x}\ {\isasymin}\ set\ \mbox{ys}\ {\isasymLongrightarrow}\ \mbox{f}\ \mbox{x}\ =\ \mbox{g}\ \mbox{x}{\isasymrbrakk}\isanewline
|
|
58 |
{\isasymLongrightarrow}\ map\ \mbox{f}\ \mbox{xs}\ =\ map\ \mbox{g}\ \mbox{ys}
|
|
59 |
\end{isabelle}%
|
|
60 |
|
|
61 |
\end{quote}
|
|
62 |
Its second premise expresses (indirectly) that the second argument of \isa{map} is only applied
|
|
63 |
to elements of its third argument. Congruence rules for other higher-order functions on lists would
|
|
64 |
look very similar but have not been proved yet because they were never needed.
|
|
65 |
If you get into a situation where you need to supply \isacommand{recdef} with new congruence
|
|
66 |
rules, you can either append the line
|
|
67 |
\begin{ttbox}
|
|
68 |
congs <congruence rules>
|
|
69 |
\end{ttbox}
|
|
70 |
to the specific occurrence of \isacommand{recdef} or declare them globally:
|
|
71 |
\begin{ttbox}
|
|
72 |
lemmas [????????] = <congruence rules>
|
|
73 |
\end{ttbox}
|
|
74 |
|
|
75 |
Note that \isacommand{recdef} feeds on exactly the same \emph{kind} of
|
|
76 |
congruence rules as the simplifier (\S\ref{sec:simp-cong}) but that
|
|
77 |
declaring a congruence rule for the simplifier does not make it
|
|
78 |
available to \isacommand{recdef}, and vice versa. This is intentional.%
|
|
79 |
\end{isamarkuptext}%
|
|
80 |
\end{isabelle}%
|
|
81 |
%%% Local Variables:
|
|
82 |
%%% mode: latex
|
|
83 |
%%% TeX-master: "root"
|
|
84 |
%%% End:
|