25260
|
1 |
%
|
|
2 |
\begin{isabellebody}%
|
|
3 |
\def\isabellecontext{fun{\isadigit{0}}}%
|
|
4 |
%
|
|
5 |
\isadelimtheory
|
|
6 |
%
|
|
7 |
\endisadelimtheory
|
|
8 |
%
|
|
9 |
\isatagtheory
|
|
10 |
%
|
|
11 |
\endisatagtheory
|
|
12 |
{\isafoldtheory}%
|
|
13 |
%
|
|
14 |
\isadelimtheory
|
|
15 |
%
|
|
16 |
\endisadelimtheory
|
|
17 |
%
|
|
18 |
\begin{isamarkuptext}%
|
|
19 |
\subsection{Definition}
|
|
20 |
\label{sec:fun-examples}
|
|
21 |
|
|
22 |
Here is a simple example, the \rmindex{Fibonacci function}:%
|
|
23 |
\end{isamarkuptext}%
|
|
24 |
\isamarkuptrue%
|
|
25 |
\isacommand{fun}\isamarkupfalse%
|
|
26 |
\ fib\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\ \isakeyword{where}\isanewline
|
|
27 |
\ \ {\isachardoublequoteopen}fib\ {\isadigit{0}}\ {\isacharequal}\ {\isadigit{0}}{\isachardoublequoteclose}\ {\isacharbar}\isanewline
|
|
28 |
\ \ {\isachardoublequoteopen}fib\ {\isacharparenleft}Suc\ {\isadigit{0}}{\isacharparenright}\ {\isacharequal}\ {\isadigit{1}}{\isachardoublequoteclose}\ {\isacharbar}\isanewline
|
|
29 |
\ \ {\isachardoublequoteopen}fib\ {\isacharparenleft}Suc{\isacharparenleft}Suc\ x{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ fib\ x\ {\isacharplus}\ fib\ {\isacharparenleft}Suc\ x{\isacharparenright}{\isachardoublequoteclose}%
|
|
30 |
\begin{isamarkuptext}%
|
|
31 |
\noindent
|
|
32 |
This resembles ordinary functional programming languages. Note the obligatory
|
|
33 |
\isacommand{where} and \isa{|}. Command \isacommand{fun} declares and
|
|
34 |
defines the function in one go. Isabelle establishes termination automatically
|
|
35 |
because \isa{fib}'s argument decreases in every recursive call.
|
|
36 |
|
|
37 |
Slightly more interesting is the insertion of a fixed element
|
|
38 |
between any two elements of a list:%
|
|
39 |
\end{isamarkuptext}%
|
|
40 |
\isamarkuptrue%
|
|
41 |
\isacommand{fun}\isamarkupfalse%
|
|
42 |
\ sep\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list{\isachardoublequoteclose}\ \isakeyword{where}\isanewline
|
|
43 |
\ \ {\isachardoublequoteopen}sep\ a\ {\isacharbrackleft}{\isacharbrackright}\ \ \ \ \ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\ {\isacharbar}\isanewline
|
|
44 |
\ \ {\isachardoublequoteopen}sep\ a\ {\isacharbrackleft}x{\isacharbrackright}\ \ \ \ {\isacharequal}\ {\isacharbrackleft}x{\isacharbrackright}{\isachardoublequoteclose}\ {\isacharbar}\isanewline
|
|
45 |
\ \ {\isachardoublequoteopen}sep\ a\ {\isacharparenleft}x{\isacharhash}y{\isacharhash}zs{\isacharparenright}\ {\isacharequal}\ x\ {\isacharhash}\ a\ {\isacharhash}\ sep\ a\ {\isacharparenleft}y{\isacharhash}zs{\isacharparenright}{\isachardoublequoteclose}%
|
|
46 |
\begin{isamarkuptext}%
|
|
47 |
\noindent
|
|
48 |
This time the length of the list decreases with the
|
|
49 |
recursive call; the first argument is irrelevant for termination.
|
|
50 |
|
|
51 |
Pattern matching\index{pattern matching!and \isacommand{fun}}
|
|
52 |
need not be exhaustive and may employ wildcards:%
|
|
53 |
\end{isamarkuptext}%
|
|
54 |
\isamarkuptrue%
|
|
55 |
\isacommand{fun}\isamarkupfalse%
|
|
56 |
\ last\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequoteclose}\ \isakeyword{where}\isanewline
|
|
57 |
\ \ {\isachardoublequoteopen}last\ {\isacharbrackleft}x{\isacharbrackright}\ \ \ \ \ \ {\isacharequal}\ x{\isachardoublequoteclose}\ {\isacharbar}\isanewline
|
|
58 |
\ \ {\isachardoublequoteopen}last\ {\isacharparenleft}{\isacharunderscore}{\isacharhash}y{\isacharhash}zs{\isacharparenright}\ {\isacharequal}\ last\ {\isacharparenleft}y{\isacharhash}zs{\isacharparenright}{\isachardoublequoteclose}%
|
|
59 |
\begin{isamarkuptext}%
|
|
60 |
Overlapping patterns are disambiguated by taking the order of equations into
|
|
61 |
account, just as in functional programming:%
|
|
62 |
\end{isamarkuptext}%
|
|
63 |
\isamarkuptrue%
|
|
64 |
\isacommand{fun}\isamarkupfalse%
|
|
65 |
\ sep{\isadigit{1}}\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list{\isachardoublequoteclose}\ \isakeyword{where}\isanewline
|
|
66 |
\ \ {\isachardoublequoteopen}sep{\isadigit{1}}\ a\ {\isacharparenleft}x{\isacharhash}y{\isacharhash}zs{\isacharparenright}\ {\isacharequal}\ x\ {\isacharhash}\ a\ {\isacharhash}\ sep{\isadigit{1}}\ a\ {\isacharparenleft}y{\isacharhash}zs{\isacharparenright}{\isachardoublequoteclose}\ {\isacharbar}\isanewline
|
|
67 |
\ \ {\isachardoublequoteopen}sep{\isadigit{1}}\ {\isacharunderscore}\ xs\ \ \ \ \ \ \ {\isacharequal}\ xs{\isachardoublequoteclose}%
|
|
68 |
\begin{isamarkuptext}%
|
|
69 |
\noindent
|
|
70 |
To guarantee that the second equation can only be applied if the first
|
|
71 |
one does not match, Isabelle internally replaces the second equation
|
|
72 |
by the two possibilities that are left: \isa{sep{\isadigit{1}}\ a\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}} and
|
|
73 |
\isa{sep{\isadigit{1}}\ a\ {\isacharbrackleft}x{\isacharbrackright}\ {\isacharequal}\ {\isacharbrackleft}x{\isacharbrackright}}. Thus the functions \isa{sep} and
|
|
74 |
\isa{sep{\isadigit{1}}} are identical.
|
|
75 |
|
25263
|
76 |
Because of its pattern matching syntax, \isacommand{fun} is also useful
|
25260
|
77 |
for the definition of non-recursive functions:%
|
|
78 |
\end{isamarkuptext}%
|
|
79 |
\isamarkuptrue%
|
|
80 |
\isacommand{fun}\isamarkupfalse%
|
|
81 |
\ swap{\isadigit{1}}{\isadigit{2}}\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list{\isachardoublequoteclose}\ \isakeyword{where}\isanewline
|
|
82 |
\ \ {\isachardoublequoteopen}swap{\isadigit{1}}{\isadigit{2}}\ {\isacharparenleft}x{\isacharhash}y{\isacharhash}zs{\isacharparenright}\ {\isacharequal}\ y{\isacharhash}x{\isacharhash}zs{\isachardoublequoteclose}\ {\isacharbar}\isanewline
|
|
83 |
\ \ {\isachardoublequoteopen}swap{\isadigit{1}}{\isadigit{2}}\ zs\ \ \ \ \ \ \ {\isacharequal}\ zs{\isachardoublequoteclose}%
|
|
84 |
\begin{isamarkuptext}%
|
|
85 |
After a function~$f$ has been defined via \isacommand{fun},
|
|
86 |
its defining equations (or variants derived from them) are available
|
|
87 |
under the name $f$\isa{{\isachardot}simps} as theorems.
|
|
88 |
For example, look (via \isacommand{thm}) at
|
|
89 |
\isa{sep{\isachardot}simps} and \isa{sep{\isadigit{1}}{\isachardot}simps} to see that they define
|
|
90 |
the same function. What is more, those equations are automatically declared as
|
|
91 |
simplification rules.
|
|
92 |
|
|
93 |
\subsection{Termination}
|
|
94 |
|
|
95 |
Isabelle's automatic termination prover for \isacommand{fun} has a
|
|
96 |
fixed notion of the \emph{size} (of type \isa{nat}) of an
|
|
97 |
argument. The size of a natural number is the number itself. The size
|
|
98 |
of a list is its length. In general, every datatype \isa{t} comes
|
|
99 |
with its own size function (named \isa{t{\isachardot}size}) which counts the
|
|
100 |
number of constructors in a term of type \isa{t} --- more precisely
|
|
101 |
those constructors where one of the arguments is of the type itself:
|
|
102 |
\isa{Suc} in the case of \isa{nat} and \isa{op\ {\isacharhash}} in the case
|
|
103 |
of lists. A recursive function is accepted if \isacommand{fun} can
|
|
104 |
show that the size of one fixed argument becomes smaller with each
|
|
105 |
recursive call.
|
|
106 |
|
25261
|
107 |
More generally, \isacommand{fun} allows any \emph{lexicographic
|
25260
|
108 |
combination} of size measures in case there are multiple
|
25261
|
109 |
arguments. For example, the following version of \rmindex{Ackermann's
|
25260
|
110 |
function} is accepted:%
|
|
111 |
\end{isamarkuptext}%
|
|
112 |
\isamarkuptrue%
|
|
113 |
\isacommand{fun}\isamarkupfalse%
|
|
114 |
\ ack{\isadigit{2}}\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\ \isakeyword{where}\isanewline
|
|
115 |
\ \ {\isachardoublequoteopen}ack{\isadigit{2}}\ n\ {\isadigit{0}}\ {\isacharequal}\ Suc\ n{\isachardoublequoteclose}\ {\isacharbar}\isanewline
|
|
116 |
\ \ {\isachardoublequoteopen}ack{\isadigit{2}}\ {\isadigit{0}}\ {\isacharparenleft}Suc\ m{\isacharparenright}\ {\isacharequal}\ ack{\isadigit{2}}\ {\isacharparenleft}Suc\ {\isadigit{0}}{\isacharparenright}\ m{\isachardoublequoteclose}\ {\isacharbar}\isanewline
|
|
117 |
\ \ {\isachardoublequoteopen}ack{\isadigit{2}}\ {\isacharparenleft}Suc\ n{\isacharparenright}\ {\isacharparenleft}Suc\ m{\isacharparenright}\ {\isacharequal}\ ack{\isadigit{2}}\ {\isacharparenleft}ack{\isadigit{2}}\ n\ {\isacharparenleft}Suc\ m{\isacharparenright}{\isacharparenright}\ m{\isachardoublequoteclose}%
|
|
118 |
\begin{isamarkuptext}%
|
25263
|
119 |
The order of arguments has no influence on whether
|
25260
|
120 |
\isacommand{fun} can prove termination of a function. For more details
|
|
121 |
see elsewhere~\cite{bulwahnKN07}.
|
|
122 |
|
|
123 |
\subsection{Simplification}
|
|
124 |
\label{sec:fun-simplification}
|
|
125 |
|
25265
|
126 |
Upon a successful termination proof, the recursion equations become
|
25260
|
127 |
simplification rules, just as with \isacommand{primrec}.
|
|
128 |
In most cases this works fine, but there is a subtle
|
|
129 |
problem that must be mentioned: simplification may not
|
|
130 |
terminate because of automatic splitting of \isa{if}.
|
|
131 |
\index{*if expressions!splitting of}
|
|
132 |
Let us look at an example:%
|
|
133 |
\end{isamarkuptext}%
|
|
134 |
\isamarkuptrue%
|
|
135 |
\isacommand{fun}\isamarkupfalse%
|
25261
|
136 |
\ gcd\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\ \isakeyword{where}\isanewline
|
|
137 |
\ \ {\isachardoublequoteopen}gcd\ m\ n\ {\isacharequal}\ {\isacharparenleft}if\ n{\isacharequal}{\isadigit{0}}\ then\ m\ else\ gcd\ n\ {\isacharparenleft}m\ mod\ n{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}%
|
25260
|
138 |
\begin{isamarkuptext}%
|
|
139 |
\noindent
|
|
140 |
The second argument decreases with each recursive call.
|
|
141 |
The termination condition
|
|
142 |
\begin{isabelle}%
|
|
143 |
\ \ \ \ \ n\ {\isasymnoteq}\ {\isadigit{0}}\ {\isasymLongrightarrow}\ m\ mod\ n\ {\isacharless}\ n%
|
|
144 |
\end{isabelle}
|
|
145 |
is proved automatically because it is already present as a lemma in
|
|
146 |
HOL\@. Thus the recursion equation becomes a simplification
|
|
147 |
rule. Of course the equation is nonterminating if we are allowed to unfold
|
|
148 |
the recursive call inside the \isa{else} branch, which is why programming
|
|
149 |
languages and our simplifier don't do that. Unfortunately the simplifier does
|
|
150 |
something else that leads to the same problem: it splits
|
|
151 |
each \isa{if}-expression unless its
|
|
152 |
condition simplifies to \isa{True} or \isa{False}. For
|
|
153 |
example, simplification reduces
|
|
154 |
\begin{isabelle}%
|
25261
|
155 |
\ \ \ \ \ gcd\ m\ n\ {\isacharequal}\ k%
|
25260
|
156 |
\end{isabelle}
|
|
157 |
in one step to
|
|
158 |
\begin{isabelle}%
|
25261
|
159 |
\ \ \ \ \ {\isacharparenleft}if\ n\ {\isacharequal}\ {\isadigit{0}}\ then\ m\ else\ gcd\ n\ {\isacharparenleft}m\ mod\ n{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ k%
|
25260
|
160 |
\end{isabelle}
|
|
161 |
where the condition cannot be reduced further, and splitting leads to
|
|
162 |
\begin{isabelle}%
|
25261
|
163 |
\ \ \ \ \ {\isacharparenleft}n\ {\isacharequal}\ {\isadigit{0}}\ {\isasymlongrightarrow}\ m\ {\isacharequal}\ k{\isacharparenright}\ {\isasymand}\ {\isacharparenleft}n\ {\isasymnoteq}\ {\isadigit{0}}\ {\isasymlongrightarrow}\ gcd\ n\ {\isacharparenleft}m\ mod\ n{\isacharparenright}\ {\isacharequal}\ k{\isacharparenright}%
|
25260
|
164 |
\end{isabelle}
|
25261
|
165 |
Since the recursive call \isa{gcd\ n\ {\isacharparenleft}m\ mod\ n{\isacharparenright}} is no longer protected by
|
25260
|
166 |
an \isa{if}, it is unfolded again, which leads to an infinite chain of
|
|
167 |
simplification steps. Fortunately, this problem can be avoided in many
|
|
168 |
different ways.
|
|
169 |
|
|
170 |
The most radical solution is to disable the offending theorem
|
|
171 |
\isa{split{\isacharunderscore}if},
|
|
172 |
as shown in \S\ref{sec:AutoCaseSplits}. However, we do not recommend this
|
|
173 |
approach: you will often have to invoke the rule explicitly when
|
|
174 |
\isa{if} is involved.
|
|
175 |
|
|
176 |
If possible, the definition should be given by pattern matching on the left
|
|
177 |
rather than \isa{if} on the right. In the case of \isa{gcd} the
|
|
178 |
following alternative definition suggests itself:%
|
|
179 |
\end{isamarkuptext}%
|
|
180 |
\isamarkuptrue%
|
|
181 |
\isacommand{fun}\isamarkupfalse%
|
|
182 |
\ gcd{\isadigit{1}}\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\ \isakeyword{where}\isanewline
|
|
183 |
\ \ {\isachardoublequoteopen}gcd{\isadigit{1}}\ m\ {\isadigit{0}}\ {\isacharequal}\ m{\isachardoublequoteclose}\ {\isacharbar}\isanewline
|
|
184 |
\ \ {\isachardoublequoteopen}gcd{\isadigit{1}}\ m\ n\ {\isacharequal}\ gcd{\isadigit{1}}\ n\ {\isacharparenleft}m\ mod\ n{\isacharparenright}{\isachardoublequoteclose}%
|
|
185 |
\begin{isamarkuptext}%
|
|
186 |
\noindent
|
|
187 |
The order of equations is important: it hides the side condition
|
25263
|
188 |
\isa{n\ {\isasymnoteq}\ {\isadigit{0}}}. Unfortunately, not all conditionals can be
|
|
189 |
expressed by pattern matching.
|
25260
|
190 |
|
|
191 |
A simple alternative is to replace \isa{if} by \isa{case},
|
|
192 |
which is also available for \isa{bool} and is not split automatically:%
|
|
193 |
\end{isamarkuptext}%
|
|
194 |
\isamarkuptrue%
|
|
195 |
\isacommand{fun}\isamarkupfalse%
|
|
196 |
\ gcd{\isadigit{2}}\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\ \isakeyword{where}\isanewline
|
|
197 |
\ \ {\isachardoublequoteopen}gcd{\isadigit{2}}\ m\ n\ {\isacharequal}\ {\isacharparenleft}case\ n{\isacharequal}{\isadigit{0}}\ of\ True\ {\isasymRightarrow}\ m\ {\isacharbar}\ False\ {\isasymRightarrow}\ gcd{\isadigit{2}}\ n\ {\isacharparenleft}m\ mod\ n{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}%
|
|
198 |
\begin{isamarkuptext}%
|
|
199 |
\noindent
|
|
200 |
This is probably the neatest solution next to pattern matching, and it is
|
|
201 |
always available.
|
|
202 |
|
|
203 |
A final alternative is to replace the offending simplification rules by
|
|
204 |
derived conditional ones. For \isa{gcd} it means we have to prove
|
|
205 |
these lemmas:%
|
|
206 |
\end{isamarkuptext}%
|
|
207 |
\isamarkuptrue%
|
|
208 |
\isacommand{lemma}\isamarkupfalse%
|
25261
|
209 |
\ {\isacharbrackleft}simp{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}gcd\ m\ {\isadigit{0}}\ {\isacharequal}\ m{\isachardoublequoteclose}\isanewline
|
25260
|
210 |
%
|
|
211 |
\isadelimproof
|
|
212 |
%
|
|
213 |
\endisadelimproof
|
|
214 |
%
|
|
215 |
\isatagproof
|
|
216 |
\isacommand{apply}\isamarkupfalse%
|
|
217 |
{\isacharparenleft}simp{\isacharparenright}\isanewline
|
|
218 |
\isacommand{done}\isamarkupfalse%
|
|
219 |
%
|
|
220 |
\endisatagproof
|
|
221 |
{\isafoldproof}%
|
|
222 |
%
|
|
223 |
\isadelimproof
|
|
224 |
\isanewline
|
|
225 |
%
|
|
226 |
\endisadelimproof
|
|
227 |
\isanewline
|
|
228 |
\isacommand{lemma}\isamarkupfalse%
|
25261
|
229 |
\ {\isacharbrackleft}simp{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}n\ {\isasymnoteq}\ {\isadigit{0}}\ {\isasymLongrightarrow}\ gcd\ m\ n\ {\isacharequal}\ gcd\ n\ {\isacharparenleft}m\ mod\ n{\isacharparenright}{\isachardoublequoteclose}\isanewline
|
25260
|
230 |
%
|
|
231 |
\isadelimproof
|
|
232 |
%
|
|
233 |
\endisadelimproof
|
|
234 |
%
|
|
235 |
\isatagproof
|
|
236 |
\isacommand{apply}\isamarkupfalse%
|
|
237 |
{\isacharparenleft}simp{\isacharparenright}\isanewline
|
|
238 |
\isacommand{done}\isamarkupfalse%
|
|
239 |
%
|
|
240 |
\endisatagproof
|
|
241 |
{\isafoldproof}%
|
|
242 |
%
|
|
243 |
\isadelimproof
|
|
244 |
%
|
|
245 |
\endisadelimproof
|
|
246 |
%
|
|
247 |
\begin{isamarkuptext}%
|
|
248 |
\noindent
|
|
249 |
Simplification terminates for these proofs because the condition of the \isa{if} simplifies to \isa{True} or \isa{False}.
|
|
250 |
Now we can disable the original simplification rule:%
|
|
251 |
\end{isamarkuptext}%
|
|
252 |
\isamarkuptrue%
|
|
253 |
\isacommand{declare}\isamarkupfalse%
|
|
254 |
\ gcd{\isachardot}simps\ {\isacharbrackleft}simp\ del{\isacharbrackright}%
|
|
255 |
\begin{isamarkuptext}%
|
|
256 |
\index{induction!recursion|(}
|
|
257 |
\index{recursion induction|(}
|
|
258 |
|
|
259 |
\subsection{Induction}
|
|
260 |
\label{sec:fun-induction}
|
|
261 |
|
|
262 |
Having defined a function we might like to prove something about it.
|
|
263 |
Since the function is recursive, the natural proof principle is
|
|
264 |
again induction. But this time the structural form of induction that comes
|
|
265 |
with datatypes is unlikely to work well --- otherwise we could have defined the
|
|
266 |
function by \isacommand{primrec}. Therefore \isacommand{fun} automatically
|
|
267 |
proves a suitable induction rule $f$\isa{{\isachardot}induct} that follows the
|
|
268 |
recursion pattern of the particular function $f$. We call this
|
|
269 |
\textbf{recursion induction}. Roughly speaking, it
|
|
270 |
requires you to prove for each \isacommand{fun} equation that the property
|
|
271 |
you are trying to establish holds for the left-hand side provided it holds
|
|
272 |
for all recursive calls on the right-hand side. Here is a simple example
|
|
273 |
involving the predefined \isa{map} functional on lists:%
|
|
274 |
\end{isamarkuptext}%
|
|
275 |
\isamarkuptrue%
|
|
276 |
\isacommand{lemma}\isamarkupfalse%
|
|
277 |
\ {\isachardoublequoteopen}map\ f\ {\isacharparenleft}sep\ x\ xs{\isacharparenright}\ {\isacharequal}\ sep\ {\isacharparenleft}f\ x{\isacharparenright}\ {\isacharparenleft}map\ f\ xs{\isacharparenright}{\isachardoublequoteclose}%
|
|
278 |
\isadelimproof
|
|
279 |
%
|
|
280 |
\endisadelimproof
|
|
281 |
%
|
|
282 |
\isatagproof
|
|
283 |
%
|
|
284 |
\begin{isamarkuptxt}%
|
|
285 |
\noindent
|
|
286 |
Note that \isa{map\ f\ xs}
|
|
287 |
is the result of applying \isa{f} to all elements of \isa{xs}. We prove
|
|
288 |
this lemma by recursion induction over \isa{sep}:%
|
|
289 |
\end{isamarkuptxt}%
|
|
290 |
\isamarkuptrue%
|
|
291 |
\isacommand{apply}\isamarkupfalse%
|
|
292 |
{\isacharparenleft}induct{\isacharunderscore}tac\ x\ xs\ rule{\isacharcolon}\ sep{\isachardot}induct{\isacharparenright}%
|
|
293 |
\begin{isamarkuptxt}%
|
|
294 |
\noindent
|
|
295 |
The resulting proof state has three subgoals corresponding to the three
|
|
296 |
clauses for \isa{sep}:
|
|
297 |
\begin{isabelle}%
|
|
298 |
\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}a{\isachardot}\ map\ f\ {\isacharparenleft}sep\ a\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\ {\isacharequal}\ sep\ {\isacharparenleft}f\ a{\isacharparenright}\ {\isacharparenleft}map\ f\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\isanewline
|
|
299 |
\ {\isadigit{2}}{\isachardot}\ {\isasymAnd}a\ x{\isachardot}\ map\ f\ {\isacharparenleft}sep\ a\ {\isacharbrackleft}x{\isacharbrackright}{\isacharparenright}\ {\isacharequal}\ sep\ {\isacharparenleft}f\ a{\isacharparenright}\ {\isacharparenleft}map\ f\ {\isacharbrackleft}x{\isacharbrackright}{\isacharparenright}\isanewline
|
|
300 |
\ {\isadigit{3}}{\isachardot}\ {\isasymAnd}a\ x\ y\ zs{\isachardot}\isanewline
|
|
301 |
\isaindent{\ {\isadigit{3}}{\isachardot}\ \ \ \ }map\ f\ {\isacharparenleft}sep\ a\ {\isacharparenleft}y\ {\isacharhash}\ zs{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ sep\ {\isacharparenleft}f\ a{\isacharparenright}\ {\isacharparenleft}map\ f\ {\isacharparenleft}y\ {\isacharhash}\ zs{\isacharparenright}{\isacharparenright}\ {\isasymLongrightarrow}\isanewline
|
|
302 |
\isaindent{\ {\isadigit{3}}{\isachardot}\ \ \ \ }map\ f\ {\isacharparenleft}sep\ a\ {\isacharparenleft}x\ {\isacharhash}\ y\ {\isacharhash}\ zs{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ sep\ {\isacharparenleft}f\ a{\isacharparenright}\ {\isacharparenleft}map\ f\ {\isacharparenleft}x\ {\isacharhash}\ y\ {\isacharhash}\ zs{\isacharparenright}{\isacharparenright}%
|
|
303 |
\end{isabelle}
|
|
304 |
The rest is pure simplification:%
|
|
305 |
\end{isamarkuptxt}%
|
|
306 |
\isamarkuptrue%
|
|
307 |
\isacommand{apply}\isamarkupfalse%
|
|
308 |
\ simp{\isacharunderscore}all\isanewline
|
|
309 |
\isacommand{done}\isamarkupfalse%
|
|
310 |
%
|
|
311 |
\endisatagproof
|
|
312 |
{\isafoldproof}%
|
|
313 |
%
|
|
314 |
\isadelimproof
|
|
315 |
%
|
|
316 |
\endisadelimproof
|
|
317 |
%
|
|
318 |
\begin{isamarkuptext}%
|
25263
|
319 |
\noindent The proof goes smoothly because the induction rule
|
|
320 |
follows the recursion of \isa{sep}. Try proving the above lemma by
|
|
321 |
structural induction, and you find that you need an additional case
|
|
322 |
distinction.
|
25260
|
323 |
|
|
324 |
In general, the format of invoking recursion induction is
|
|
325 |
\begin{quote}
|
|
326 |
\isacommand{apply}\isa{{\isacharparenleft}induct{\isacharunderscore}tac} $x@1 \dots x@n$ \isa{rule{\isacharcolon}} $f$\isa{{\isachardot}induct{\isacharparenright}}
|
|
327 |
\end{quote}\index{*induct_tac (method)}%
|
|
328 |
where $x@1~\dots~x@n$ is a list of free variables in the subgoal and $f$ the
|
25263
|
329 |
name of a function that takes an $n$ arguments. Usually the subgoal will
|
|
330 |
contain the term $f x@1 \dots x@n$ but this need not be the case. The
|
25260
|
331 |
induction rules do not mention $f$ at all. Here is \isa{sep{\isachardot}induct}:
|
|
332 |
\begin{isabelle}
|
|
333 |
{\isasymlbrakk}~{\isasymAnd}a.~P~a~[];\isanewline
|
|
334 |
~~{\isasymAnd}a~x.~P~a~[x];\isanewline
|
|
335 |
~~{\isasymAnd}a~x~y~zs.~P~a~(y~\#~zs)~{\isasymLongrightarrow}~P~a~(x~\#~y~\#~zs){\isasymrbrakk}\isanewline
|
|
336 |
{\isasymLongrightarrow}~P~u~v%
|
|
337 |
\end{isabelle}
|
|
338 |
It merely says that in order to prove a property \isa{P} of \isa{u} and
|
|
339 |
\isa{v} you need to prove it for the three cases where \isa{v} is the
|
|
340 |
empty list, the singleton list, and the list with at least two elements.
|
|
341 |
The final case has an induction hypothesis: you may assume that \isa{P}
|
|
342 |
holds for the tail of that list.
|
|
343 |
\index{induction!recursion|)}
|
|
344 |
\index{recursion induction|)}%
|
|
345 |
\end{isamarkuptext}%
|
|
346 |
\isamarkuptrue%
|
|
347 |
%
|
|
348 |
\isadelimtheory
|
|
349 |
%
|
|
350 |
\endisadelimtheory
|
|
351 |
%
|
|
352 |
\isatagtheory
|
|
353 |
%
|
|
354 |
\endisatagtheory
|
|
355 |
{\isafoldtheory}%
|
|
356 |
%
|
|
357 |
\isadelimtheory
|
|
358 |
%
|
|
359 |
\endisadelimtheory
|
|
360 |
\end{isabellebody}%
|
|
361 |
%%% Local Variables:
|
|
362 |
%%% mode: latex
|
|
363 |
%%% TeX-master: "root"
|
|
364 |
%%% End:
|