3 \def\isabellecontext{Nested2}%
7 The termintion condition is easily proved by induction:%
9 \isacommand{lemma}\ {\isacharbrackleft}simp{\isacharbrackright}{\isacharcolon}\ {\isachardoublequote}t\ {\isasymin}\ set\ ts\ {\isasymlongrightarrow}\ size\ t\ {\isacharless}\ Suc{\isacharparenleft}term{\isacharunderscore}list{\isacharunderscore}size\ ts{\isacharparenright}{\isachardoublequote}\isanewline
10 \isacommand{by}{\isacharparenleft}induct{\isacharunderscore}tac\ ts{\isacharcomma}\ auto{\isacharparenright}%
11 \begin{isamarkuptext}%
13 By making this theorem a simplification rule, \isacommand{recdef}
14 applies it automatically and the above definition of \isa{trev}
15 succeeds now. As a reward for our effort, we can now prove the desired
16 lemma directly. The key is the fact that we no longer need the verbose
17 induction schema for type \isa{term} but the simpler one arising from
20 \isacommand{lemma}\ {\isachardoublequote}trev{\isacharparenleft}trev\ t{\isacharparenright}\ {\isacharequal}\ t{\isachardoublequote}\isanewline
21 \isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ t\ rule{\isacharcolon}trev{\isachardot}induct{\isacharparenright}%
24 This leaves us with a trivial base case \isa{trev\ {\isacharparenleft}trev\ {\isacharparenleft}Var\ x{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ Var\ x} and the step case
26 \ \ \ \ \ {\isasymforall}t{\isachardot}\ t\ {\isasymin}\ set\ ts\ {\isasymlongrightarrow}\ trev\ {\isacharparenleft}trev\ t{\isacharparenright}\ {\isacharequal}\ t\ {\isasymLongrightarrow}\isanewline
27 \ \ \ \ \ trev\ {\isacharparenleft}trev\ {\isacharparenleft}App\ f\ ts{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ App\ f\ ts%
29 both of which are solved by simplification:%
31 \isacommand{by}{\isacharparenleft}simp{\isacharunderscore}all\ add{\isacharcolon}rev{\isacharunderscore}map\ sym{\isacharbrackleft}OF\ map{\isacharunderscore}compose{\isacharbrackright}\ cong{\isacharcolon}map{\isacharunderscore}cong{\isacharparenright}%
32 \begin{isamarkuptext}%
34 If the proof of the induction step mystifies you, we recommend to go through
35 the chain of simplification steps in detail; you will probably need the help of
36 \isa{trace{\isacharunderscore}simp}. Theorem \isa{map{\isacharunderscore}cong} is discussed below.
38 %{term[display]"trev(trev(App f ts))"}\\
39 %{term[display]"App f (rev(map trev (rev(map trev ts))))"}\\
40 %{term[display]"App f (map trev (rev(rev(map trev ts))))"}\\
41 %{term[display]"App f (map trev (map trev ts))"}\\
42 %{term[display]"App f (map (trev o trev) ts)"}\\
43 %{term[display]"App f (map (%x. x) ts)"}\\
44 %{term[display]"App f ts"}
47 The above definition of \isa{trev} is superior to the one in
48 \S\ref{sec:nested-datatype} because it brings \isa{rev} into play, about
49 which already know a lot, in particular \isa{rev\ {\isacharparenleft}rev\ xs{\isacharparenright}\ {\isacharequal}\ xs}.
50 Thus this proof is a good example of an important principle:
52 \emph{Chose your definitions carefully\\
53 because they determine the complexity of your proofs.}
56 Let us now return to the question of how \isacommand{recdef} can come up with
57 sensible termination conditions in the presence of higher-order functions
58 like \isa{map}. For a start, if nothing were known about \isa{map},
59 \isa{map\ trev\ ts} might apply \isa{trev} to arbitrary terms, and thus
60 \isacommand{recdef} would try to prove the unprovable \isa{size\ t\ {\isacharless}\ Suc\ {\isacharparenleft}term{\isacharunderscore}list{\isacharunderscore}size\ ts{\isacharparenright}}, without any assumption about \isa{t}. Therefore
61 \isacommand{recdef} has been supplied with the congruence theorem
62 \isa{map{\isacharunderscore}cong}:
64 \ \ \ \ \ {\isasymlbrakk}xs\ {\isacharequal}\ ys{\isacharsemicolon}\ {\isasymAnd}x{\isachardot}\ x\ {\isasymin}\ set\ ys\ {\isasymLongrightarrow}\ f\ x\ {\isacharequal}\ g\ x{\isasymrbrakk}\isanewline
65 \ \ \ \ \ {\isasymLongrightarrow}\ map\ f\ xs\ {\isacharequal}\ map\ g\ ys%
67 Its second premise expresses (indirectly) that the second argument of
68 \isa{map} is only applied to elements of its third argument. Congruence
69 rules for other higher-order functions on lists would look very similar but
70 have not been proved yet because they were never needed. If you get into a
71 situation where you need to supply \isacommand{recdef} with new congruence
72 rules, you can either append the line
74 congs <congruence rules>
76 to the specific occurrence of \isacommand{recdef} or declare them globally:
78 lemmas [????????] = <congruence rules>
81 Note that \isacommand{recdef} feeds on exactly the same \emph{kind} of
82 congruence rules as the simplifier (\S\ref{sec:simp-cong}) but that
83 declaring a congruence rule for the simplifier does not make it
84 available to \isacommand{recdef}, and vice versa. This is intentional.
85 %The simplifier's congruence rules cannot be used by recdef.
86 %For example the weak congruence rules for if and case would prevent
87 %recdef from generating sensible termination conditions.%
92 %%% TeX-master: "root"