author | nipkow |
Fri, 01 Dec 2023 10:10:59 +0100 | |
changeset 79100 | e103e3cef3cb |
parent 76987 | 4c275405faae |
permissions | -rw-r--r-- |
50530 | 1 |
(* Title: Doc/Functions/Functions.thy |
21212 | 2 |
Author: Alexander Krauss, TU Muenchen |
3 |
||
4 |
Tutorial for function definitions with the new "function" package. |
|
5 |
*) |
|
6 |
||
7 |
theory Functions |
|
8 |
imports Main |
|
9 |
begin |
|
10 |
||
61419 | 11 |
section \<open>Function Definitions for Dummies\<close> |
21212 | 12 |
|
61419 | 13 |
text \<open> |
21212 | 14 |
In most cases, defining a recursive function is just as simple as other definitions: |
61419 | 15 |
\<close> |
21212 | 16 |
|
17 |
fun fib :: "nat \<Rightarrow> nat" |
|
18 |
where |
|
19 |
"fib 0 = 1" |
|
20 |
| "fib (Suc 0) = 1" |
|
21 |
| "fib (Suc (Suc n)) = fib n + fib (Suc n)" |
|
22 |
||
61419 | 23 |
text \<open> |
23003 | 24 |
The syntax is rather self-explanatory: We introduce a function by |
25091
a2ae7f71613d
Updated function tutorial: Types can be inferred and need not be given anymore
krauss
parents:
23805
diff
changeset
|
25 |
giving its name, its type, |
a2ae7f71613d
Updated function tutorial: Types can be inferred and need not be given anymore
krauss
parents:
23805
diff
changeset
|
26 |
and a set of defining recursive equations. |
a2ae7f71613d
Updated function tutorial: Types can be inferred and need not be given anymore
krauss
parents:
23805
diff
changeset
|
27 |
If we leave out the type, the most general type will be |
69597 | 28 |
inferred, which can sometimes lead to surprises: Since both \<^term>\<open>1::nat\<close> and \<open>+\<close> are overloaded, we would end up |
69505 | 29 |
with \<open>fib :: nat \<Rightarrow> 'a::{one,plus}\<close>. |
61419 | 30 |
\<close> |
23003 | 31 |
|
61419 | 32 |
text \<open> |
23003 | 33 |
The function always terminates, since its argument gets smaller in |
23188 | 34 |
every recursive call. |
35 |
Since HOL is a logic of total functions, termination is a |
|
36 |
fundamental requirement to prevent inconsistencies\footnote{From the |
|
69505 | 37 |
\qt{definition} \<open>f(n) = f(n) + 1\<close> we could prove |
38 |
\<open>0 = 1\<close> by subtracting \<open>f(n)\<close> on both sides.}. |
|
23188 | 39 |
Isabelle tries to prove termination automatically when a definition |
40 |
is made. In \S\ref{termination}, we will look at cases where this |
|
41 |
fails and see what to do then. |
|
61419 | 42 |
\<close> |
21212 | 43 |
|
61419 | 44 |
subsection \<open>Pattern matching\<close> |
21212 | 45 |
|
61419 | 46 |
text \<open>\label{patmatch} |
23003 | 47 |
Like in functional programming, we can use pattern matching to |
48 |
define functions. At the moment we will only consider \emph{constructor |
|
21212 | 49 |
patterns}, which only consist of datatype constructors and |
23805 | 50 |
variables. Furthermore, patterns must be linear, i.e.\ all variables |
51 |
on the left hand side of an equation must be distinct. In |
|
52 |
\S\ref{genpats} we discuss more general pattern matching. |
|
21212 | 53 |
|
54 |
If patterns overlap, the order of the equations is taken into |
|
55 |
account. The following function inserts a fixed element between any |
|
56 |
two elements of a list: |
|
61419 | 57 |
\<close> |
21212 | 58 |
|
59 |
fun sep :: "'a \<Rightarrow> 'a list \<Rightarrow> 'a list" |
|
60 |
where |
|
61 |
"sep a (x#y#xs) = x # a # sep a (y # xs)" |
|
62 |
| "sep a xs = xs" |
|
63 |
||
61419 | 64 |
text \<open> |
23188 | 65 |
Overlapping patterns are interpreted as \qt{increments} to what is |
21212 | 66 |
already there: The second equation is only meant for the cases where |
67 |
the first one does not match. Consequently, Isabelle replaces it |
|
22065 | 68 |
internally by the remaining cases, making the patterns disjoint: |
61419 | 69 |
\<close> |
21212 | 70 |
|
22065 | 71 |
thm sep.simps |
21212 | 72 |
|
61419 | 73 |
text \<open>@{thm [display] sep.simps[no_vars]}\<close> |
21212 | 74 |
|
61419 | 75 |
text \<open> |
23805 | 76 |
\noindent The equations from function definitions are automatically used in |
21212 | 77 |
simplification: |
61419 | 78 |
\<close> |
21212 | 79 |
|
23188 | 80 |
lemma "sep 0 [1, 2, 3] = [1, 0, 2, 0, 3]" |
21212 | 81 |
by simp |
82 |
||
61419 | 83 |
subsection \<open>Induction\<close> |
21212 | 84 |
|
61419 | 85 |
text \<open> |
22065 | 86 |
|
23805 | 87 |
Isabelle provides customized induction rules for recursive |
88 |
functions. These rules follow the recursive structure of the |
|
53107
57c7294eac0a
more document antiquotations (for proper theorem names);
Christian Sternagel
parents:
50530
diff
changeset
|
89 |
definition. Here is the rule @{thm [source] sep.induct} arising from the |
69597 | 90 |
above definition of \<^const>\<open>sep\<close>: |
23805 | 91 |
|
92 |
@{thm [display] sep.induct} |
|
93 |
||
94 |
We have a step case for list with at least two elements, and two |
|
95 |
base cases for the zero- and the one-element list. Here is a simple |
|
69597 | 96 |
proof about \<^const>\<open>sep\<close> and \<^const>\<open>map\<close> |
61419 | 97 |
\<close> |
23188 | 98 |
|
23805 | 99 |
lemma "map f (sep x ys) = sep (f x) (map f ys)" |
100 |
apply (induct x ys rule: sep.induct) |
|
101 |
||
61419 | 102 |
text \<open> |
23805 | 103 |
We get three cases, like in the definition. |
104 |
||
105 |
@{subgoals [display]} |
|
61419 | 106 |
\<close> |
23805 | 107 |
|
108 |
apply auto |
|
109 |
done |
|
61419 | 110 |
text \<open> |
23188 | 111 |
|
112 |
With the \cmd{fun} command, you can define about 80\% of the |
|
113 |
functions that occur in practice. The rest of this tutorial explains |
|
114 |
the remaining 20\%. |
|
61419 | 115 |
\<close> |
21212 | 116 |
|
117 |
||
61419 | 118 |
section \<open>fun vs.\ function\<close> |
21212 | 119 |
|
61419 | 120 |
text \<open> |
23188 | 121 |
The \cmd{fun} command provides a |
21212 | 122 |
convenient shorthand notation for simple function definitions. In |
123 |
this mode, Isabelle tries to solve all the necessary proof obligations |
|
27026 | 124 |
automatically. If any proof fails, the definition is |
22065 | 125 |
rejected. This can either mean that the definition is indeed faulty, |
126 |
or that the default proof procedures are just not smart enough (or |
|
127 |
rather: not designed) to handle the definition. |
|
128 |
||
23188 | 129 |
By expanding the abbreviation to the more verbose \cmd{function} command, these proof obligations become visible and can be analyzed or |
130 |
solved manually. The expansion from \cmd{fun} to \cmd{function} is as follows: |
|
22065 | 131 |
|
132 |
\end{isamarkuptext} |
|
133 |
||
134 |
||
23188 | 135 |
\[\left[\;\begin{minipage}{0.25\textwidth}\vspace{6pt} |
69505 | 136 |
\cmd{fun} \<open>f :: \<tau>\<close>\\% |
23188 | 137 |
\cmd{where}\\% |
138 |
\hspace*{2ex}{\it equations}\\% |
|
139 |
\hspace*{2ex}\vdots\vspace*{6pt} |
|
140 |
\end{minipage}\right] |
|
141 |
\quad\equiv\quad |
|
27026 | 142 |
\left[\;\begin{minipage}{0.48\textwidth}\vspace{6pt} |
69505 | 143 |
\cmd{function} \<open>(\<close>\cmd{sequential}\<open>) f :: \<tau>\<close>\\% |
23188 | 144 |
\cmd{where}\\% |
145 |
\hspace*{2ex}{\it equations}\\% |
|
146 |
\hspace*{2ex}\vdots\\% |
|
69505 | 147 |
\cmd{by} \<open>pat_completeness auto\<close>\\% |
148 |
\cmd{termination by} \<open>lexicographic_order\<close>\vspace{6pt} |
|
23188 | 149 |
\end{minipage} |
150 |
\right]\] |
|
21212 | 151 |
|
22065 | 152 |
\begin{isamarkuptext} |
153 |
\vspace*{1em} |
|
23188 | 154 |
\noindent Some details have now become explicit: |
21212 | 155 |
|
156 |
\begin{enumerate} |
|
22065 | 157 |
\item The \cmd{sequential} option enables the preprocessing of |
23805 | 158 |
pattern overlaps which we already saw. Without this option, the equations |
21212 | 159 |
must already be disjoint and complete. The automatic completion only |
23188 | 160 |
works with constructor patterns. |
21212 | 161 |
|
23188 | 162 |
\item A function definition produces a proof obligation which |
163 |
expresses completeness and compatibility of patterns (we talk about |
|
69505 | 164 |
this later). The combination of the methods \<open>pat_completeness\<close> and |
165 |
\<open>auto\<close> is used to solve this proof obligation. |
|
21212 | 166 |
|
167 |
\item A termination proof follows the definition, started by the |
|
23188 | 168 |
\cmd{termination} command. This will be explained in \S\ref{termination}. |
22065 | 169 |
\end{enumerate} |
170 |
Whenever a \cmd{fun} command fails, it is usually a good idea to |
|
171 |
expand the syntax to the more verbose \cmd{function} form, to see |
|
172 |
what is actually going on. |
|
61419 | 173 |
\<close> |
21212 | 174 |
|
175 |
||
61419 | 176 |
section \<open>Termination\<close> |
21212 | 177 |
|
61419 | 178 |
text \<open>\label{termination} |
69505 | 179 |
The method \<open>lexicographic_order\<close> is the default method for |
23805 | 180 |
termination proofs. It can prove termination of a |
23188 | 181 |
certain class of functions by searching for a suitable lexicographic |
182 |
combination of size measures. Of course, not all functions have such |
|
23805 | 183 |
a simple termination argument. For them, we can specify the termination |
184 |
relation manually. |
|
61419 | 185 |
\<close> |
23188 | 186 |
|
61419 | 187 |
subsection \<open>The {\tt relation} method\<close> |
188 |
text\<open> |
|
21212 | 189 |
Consider the following function, which sums up natural numbers up to |
69505 | 190 |
\<open>N\<close>, using a counter \<open>i\<close>: |
61419 | 191 |
\<close> |
21212 | 192 |
|
193 |
function sum :: "nat \<Rightarrow> nat \<Rightarrow> nat" |
|
194 |
where |
|
195 |
"sum i N = (if i > N then 0 else i + sum (Suc i) N)" |
|
22065 | 196 |
by pat_completeness auto |
21212 | 197 |
|
61419 | 198 |
text \<open> |
69505 | 199 |
\noindent The \<open>lexicographic_order\<close> method fails on this example, because none of the |
23805 | 200 |
arguments decreases in the recursive call, with respect to the standard size ordering. |
201 |
To prove termination manually, we must provide a custom wellfounded relation. |
|
21212 | 202 |
|
69505 | 203 |
The termination argument for \<open>sum\<close> is based on the fact that |
204 |
the \emph{difference} between \<open>i\<close> and \<open>N\<close> gets |
|
205 |
smaller in every step, and that the recursion stops when \<open>i\<close> |
|
206 |
is greater than \<open>N\<close>. Phrased differently, the expression |
|
207 |
\<open>N + 1 - i\<close> always decreases. |
|
21212 | 208 |
|
22065 | 209 |
We can use this expression as a measure function suitable to prove termination. |
61419 | 210 |
\<close> |
21212 | 211 |
|
27026 | 212 |
termination sum |
23188 | 213 |
apply (relation "measure (\<lambda>(i,N). N + 1 - i)") |
21212 | 214 |
|
61419 | 215 |
text \<open> |
23003 | 216 |
The \cmd{termination} command sets up the termination goal for the |
69505 | 217 |
specified function \<open>sum\<close>. If the function name is omitted, it |
23003 | 218 |
implicitly refers to the last function definition. |
219 |
||
69505 | 220 |
The \<open>relation\<close> method takes a relation of |
69597 | 221 |
type \<^typ>\<open>('a \<times> 'a) set\<close>, where \<^typ>\<open>'a\<close> is the argument type of |
22065 | 222 |
the function. If the function has multiple curried arguments, then |
223 |
these are packed together into a tuple, as it happened in the above |
|
224 |
example. |
|
21212 | 225 |
|
27026 | 226 |
The predefined function @{term[source] "measure :: ('a \<Rightarrow> nat) \<Rightarrow> ('a \<times> 'a) set"} constructs a |
23188 | 227 |
wellfounded relation from a mapping into the natural numbers (a |
228 |
\emph{measure function}). |
|
22065 | 229 |
|
69505 | 230 |
After the invocation of \<open>relation\<close>, we must prove that (a) |
22065 | 231 |
the relation we supplied is wellfounded, and (b) that the arguments |
232 |
of recursive calls indeed decrease with respect to the |
|
23188 | 233 |
relation: |
234 |
||
235 |
@{subgoals[display,indent=0]} |
|
22065 | 236 |
|
69505 | 237 |
These goals are all solved by \<open>auto\<close>: |
61419 | 238 |
\<close> |
23188 | 239 |
|
240 |
apply auto |
|
241 |
done |
|
242 |
||
61419 | 243 |
text \<open> |
22065 | 244 |
Let us complicate the function a little, by adding some more |
245 |
recursive calls: |
|
61419 | 246 |
\<close> |
21212 | 247 |
|
248 |
function foo :: "nat \<Rightarrow> nat \<Rightarrow> nat" |
|
249 |
where |
|
250 |
"foo i N = (if i > N |
|
251 |
then (if N = 0 then 0 else foo 0 (N - 1)) |
|
252 |
else i + foo (Suc i) N)" |
|
253 |
by pat_completeness auto |
|
254 |
||
61419 | 255 |
text \<open> |
69505 | 256 |
When \<open>i\<close> has reached \<open>N\<close>, it starts at zero again |
257 |
and \<open>N\<close> is decremented. |
|
21212 | 258 |
This corresponds to a nested |
259 |
loop where one index counts up and the other down. Termination can |
|
260 |
be proved using a lexicographic combination of two measures, namely |
|
69597 | 261 |
the value of \<open>N\<close> and the above difference. The \<^const>\<open>measures\<close> combinator generalizes \<open>measure\<close> by taking a |
22065 | 262 |
list of measure functions. |
61419 | 263 |
\<close> |
21212 | 264 |
|
265 |
termination |
|
22065 | 266 |
by (relation "measures [\<lambda>(i, N). N, \<lambda>(i,N). N + 1 - i]") auto |
21212 | 267 |
|
69505 | 268 |
subsection \<open>How \<open>lexicographic_order\<close> works\<close> |
23188 | 269 |
|
270 |
(*fun fails :: "nat \<Rightarrow> nat list \<Rightarrow> nat" |
|
271 |
where |
|
272 |
"fails a [] = a" |
|
273 |
| "fails a (x#xs) = fails (x + a) (x # xs)" |
|
274 |
*) |
|
23003 | 275 |
|
61419 | 276 |
text \<open> |
23188 | 277 |
To see how the automatic termination proofs work, let's look at an |
278 |
example where it fails\footnote{For a detailed discussion of the |
|
76987 | 279 |
termination prover, see \<^cite>\<open>bulwahnKN07\<close>}: |
23188 | 280 |
|
281 |
\end{isamarkuptext} |
|
69505 | 282 |
\cmd{fun} \<open>fails :: "nat \<Rightarrow> nat list \<Rightarrow> nat"\<close>\\% |
23188 | 283 |
\cmd{where}\\% |
69505 | 284 |
\hspace*{2ex}\<open>"fails a [] = a"\<close>\\% |
285 |
|\hspace*{1.5ex}\<open>"fails a (x#xs) = fails (x + a) (x#xs)"\<close>\\ |
|
23188 | 286 |
\begin{isamarkuptext} |
287 |
||
288 |
\noindent Isabelle responds with the following error: |
|
289 |
||
290 |
\begin{isabelle} |
|
23805 | 291 |
*** Unfinished subgoals:\newline |
292 |
*** (a, 1, <):\newline |
|
69505 | 293 |
*** \ 1.~\<open>\<And>x. x = 0\<close>\newline |
23805 | 294 |
*** (a, 1, <=):\newline |
295 |
*** \ 1.~False\newline |
|
296 |
*** (a, 2, <):\newline |
|
297 |
*** \ 1.~False\newline |
|
23188 | 298 |
*** Calls:\newline |
69505 | 299 |
*** a) \<open>(a, x # xs) -->> (x + a, x # xs)\<close>\newline |
23188 | 300 |
*** Measures:\newline |
69505 | 301 |
*** 1) \<open>\<lambda>x. size (fst x)\<close>\newline |
302 |
*** 2) \<open>\<lambda>x. size (snd x)\<close>\newline |
|
23805 | 303 |
*** Result matrix:\newline |
304 |
*** \ \ \ \ 1\ \ 2 \newline |
|
305 |
*** a: ? <= \newline |
|
306 |
*** Could not find lexicographic termination order.\newline |
|
23188 | 307 |
*** At command "fun".\newline |
308 |
\end{isabelle} |
|
61419 | 309 |
\<close> |
310 |
text \<open> |
|
28040 | 311 |
The key to this error message is the matrix at the bottom. The rows |
23188 | 312 |
of that matrix correspond to the different recursive calls (In our |
313 |
case, there is just one). The columns are the function's arguments |
|
314 |
(expressed through different measure functions, which map the |
|
315 |
argument tuple to a natural number). |
|
316 |
||
317 |
The contents of the matrix summarize what is known about argument |
|
69505 | 318 |
descents: The second argument has a weak descent (\<open><=\<close>) at the |
23188 | 319 |
recursive call, and for the first argument nothing could be proved, |
69505 | 320 |
which is expressed by \<open>?\<close>. In general, there are the values |
321 |
\<open><\<close>, \<open><=\<close> and \<open>?\<close>. |
|
23188 | 322 |
|
323 |
For the failed proof attempts, the unfinished subgoals are also |
|
23805 | 324 |
printed. Looking at these will often point to a missing lemma. |
61419 | 325 |
\<close> |
23188 | 326 |
|
69505 | 327 |
subsection \<open>The \<open>size_change\<close> method\<close> |
23003 | 328 |
|
61419 | 329 |
text \<open> |
33856 | 330 |
Some termination goals that are beyond the powers of |
69505 | 331 |
\<open>lexicographic_order\<close> can be solved automatically by the |
332 |
more powerful \<open>size_change\<close> method, which uses a variant of |
|
33856 | 333 |
the size-change principle, together with some other |
334 |
techniques. While the details are discussed |
|
76987 | 335 |
elsewhere \<^cite>\<open>krauss_phd\<close>, |
33856 | 336 |
here are a few typical situations where |
69505 | 337 |
\<open>lexicographic_order\<close> has difficulties and \<open>size_change\<close> |
33856 | 338 |
may be worth a try: |
339 |
\begin{itemize} |
|
340 |
\item Arguments are permuted in a recursive call. |
|
341 |
\item Several mutually recursive functions with multiple arguments. |
|
342 |
\item Unusual control flow (e.g., when some recursive calls cannot |
|
343 |
occur in sequence). |
|
344 |
\end{itemize} |
|
23003 | 345 |
|
69505 | 346 |
Loading the theory \<open>Multiset\<close> makes the \<open>size_change\<close> |
33856 | 347 |
method a bit stronger: it can then use multiset orders internally. |
61419 | 348 |
\<close> |
21212 | 349 |
|
70276 | 350 |
subsection \<open>Configuring simplification rules for termination proofs\<close> |
351 |
||
352 |
text \<open> |
|
353 |
Since both \<open>lexicographic_order\<close> and \<open>size_change\<close> rely on the simplifier internally, |
|
354 |
there can sometimes be the need for adding additional simp rules to them. |
|
355 |
This can be done either as arguments to the methods themselves, or globally via the |
|
356 |
theorem attribute \<open>termination_simp\<close>, which is useful in rare cases. |
|
357 |
\<close> |
|
358 |
||
61419 | 359 |
section \<open>Mutual Recursion\<close> |
21212 | 360 |
|
61419 | 361 |
text \<open> |
21212 | 362 |
If two or more functions call one another mutually, they have to be defined |
69505 | 363 |
in one step. Here are \<open>even\<close> and \<open>odd\<close>: |
61419 | 364 |
\<close> |
21212 | 365 |
|
22065 | 366 |
function even :: "nat \<Rightarrow> bool" |
367 |
and odd :: "nat \<Rightarrow> bool" |
|
21212 | 368 |
where |
369 |
"even 0 = True" |
|
370 |
| "odd 0 = False" |
|
371 |
| "even (Suc n) = odd n" |
|
372 |
| "odd (Suc n) = even n" |
|
22065 | 373 |
by pat_completeness auto |
21212 | 374 |
|
61419 | 375 |
text \<open> |
23188 | 376 |
To eliminate the mutual dependencies, Isabelle internally |
21212 | 377 |
creates a single function operating on the sum |
69597 | 378 |
type \<^typ>\<open>nat + nat\<close>. Then, \<^const>\<open>even\<close> and \<^const>\<open>odd\<close> are |
23188 | 379 |
defined as projections. Consequently, termination has to be proved |
21212 | 380 |
simultaneously for both functions, by specifying a measure on the |
381 |
sum type: |
|
61419 | 382 |
\<close> |
21212 | 383 |
|
384 |
termination |
|
23188 | 385 |
by (relation "measure (\<lambda>x. case x of Inl n \<Rightarrow> n | Inr n \<Rightarrow> n)") auto |
386 |
||
61419 | 387 |
text \<open> |
69505 | 388 |
We could also have used \<open>lexicographic_order\<close>, which |
23188 | 389 |
supports mutual recursive termination proofs to a certain extent. |
61419 | 390 |
\<close> |
22065 | 391 |
|
61419 | 392 |
subsection \<open>Induction for mutual recursion\<close> |
22065 | 393 |
|
61419 | 394 |
text \<open> |
22065 | 395 |
|
396 |
When functions are mutually recursive, proving properties about them |
|
53107
57c7294eac0a
more document antiquotations (for proper theorem names);
Christian Sternagel
parents:
50530
diff
changeset
|
397 |
generally requires simultaneous induction. The induction rule @{thm [source] "even_odd.induct"} |
23188 | 398 |
generated from the above definition reflects this. |
22065 | 399 |
|
69597 | 400 |
Let us prove something about \<^const>\<open>even\<close> and \<^const>\<open>odd\<close>: |
61419 | 401 |
\<close> |
22065 | 402 |
|
23188 | 403 |
lemma even_odd_mod2: |
22065 | 404 |
"even n = (n mod 2 = 0)" |
405 |
"odd n = (n mod 2 = 1)" |
|
406 |
||
61419 | 407 |
text \<open> |
22065 | 408 |
We apply simultaneous induction, specifying the induction variable |
61419 | 409 |
for both goals, separated by \cmd{and}:\<close> |
22065 | 410 |
|
411 |
apply (induct n and n rule: even_odd.induct) |
|
412 |
||
61419 | 413 |
text \<open> |
22065 | 414 |
We get four subgoals, which correspond to the clauses in the |
69597 | 415 |
definition of \<^const>\<open>even\<close> and \<^const>\<open>odd\<close>: |
22065 | 416 |
@{subgoals[display,indent=0]} |
417 |
Simplification solves the first two goals, leaving us with two |
|
69505 | 418 |
statements about the \<open>mod\<close> operation to prove: |
61419 | 419 |
\<close> |
22065 | 420 |
|
421 |
apply simp_all |
|
422 |
||
61419 | 423 |
text \<open> |
22065 | 424 |
@{subgoals[display,indent=0]} |
425 |
||
23805 | 426 |
\noindent These can be handled by Isabelle's arithmetic decision procedures. |
22065 | 427 |
|
61419 | 428 |
\<close> |
22065 | 429 |
|
23805 | 430 |
apply arith |
431 |
apply arith |
|
22065 | 432 |
done |
21212 | 433 |
|
61419 | 434 |
text \<open> |
23188 | 435 |
In proofs like this, the simultaneous induction is really essential: |
436 |
Even if we are just interested in one of the results, the other |
|
437 |
one is necessary to strengthen the induction hypothesis. If we leave |
|
69597 | 438 |
out the statement about \<^const>\<open>odd\<close> and just write \<^term>\<open>True\<close> instead, |
27026 | 439 |
the same proof fails: |
61419 | 440 |
\<close> |
22065 | 441 |
|
23188 | 442 |
lemma failed_attempt: |
22065 | 443 |
"even n = (n mod 2 = 0)" |
444 |
"True" |
|
445 |
apply (induct n rule: even_odd.induct) |
|
446 |
||
61419 | 447 |
text \<open> |
22065 | 448 |
\noindent Now the third subgoal is a dead end, since we have no |
23188 | 449 |
useful induction hypothesis available: |
22065 | 450 |
|
451 |
@{subgoals[display,indent=0]} |
|
61419 | 452 |
\<close> |
22065 | 453 |
|
454 |
oops |
|
455 |
||
61419 | 456 |
section \<open>Elimination\<close> |
54017
2a3c07f49615
basic documentation for function elimination rules and fun_cases
krauss
parents:
53107
diff
changeset
|
457 |
|
61419 | 458 |
text \<open> |
69505 | 459 |
A definition of function \<open>f\<close> gives rise to two kinds of elimination rules. Rule \<open>f.cases\<close> |
54017
2a3c07f49615
basic documentation for function elimination rules and fun_cases
krauss
parents:
53107
diff
changeset
|
460 |
simply describes case analysis according to the patterns used in the definition: |
61419 | 461 |
\<close> |
54017
2a3c07f49615
basic documentation for function elimination rules and fun_cases
krauss
parents:
53107
diff
changeset
|
462 |
|
2a3c07f49615
basic documentation for function elimination rules and fun_cases
krauss
parents:
53107
diff
changeset
|
463 |
fun list_to_option :: "'a list \<Rightarrow> 'a option" |
2a3c07f49615
basic documentation for function elimination rules and fun_cases
krauss
parents:
53107
diff
changeset
|
464 |
where |
2a3c07f49615
basic documentation for function elimination rules and fun_cases
krauss
parents:
53107
diff
changeset
|
465 |
"list_to_option [x] = Some x" |
2a3c07f49615
basic documentation for function elimination rules and fun_cases
krauss
parents:
53107
diff
changeset
|
466 |
| "list_to_option _ = None" |
2a3c07f49615
basic documentation for function elimination rules and fun_cases
krauss
parents:
53107
diff
changeset
|
467 |
|
2a3c07f49615
basic documentation for function elimination rules and fun_cases
krauss
parents:
53107
diff
changeset
|
468 |
thm list_to_option.cases |
61419 | 469 |
text \<open> |
54017
2a3c07f49615
basic documentation for function elimination rules and fun_cases
krauss
parents:
53107
diff
changeset
|
470 |
@{thm[display] list_to_option.cases} |
2a3c07f49615
basic documentation for function elimination rules and fun_cases
krauss
parents:
53107
diff
changeset
|
471 |
|
2a3c07f49615
basic documentation for function elimination rules and fun_cases
krauss
parents:
53107
diff
changeset
|
472 |
Note that this rule does not mention the function at all, but only describes the cases used for |
2a3c07f49615
basic documentation for function elimination rules and fun_cases
krauss
parents:
53107
diff
changeset
|
473 |
defining it. In contrast, the rule @{thm[source] list_to_option.elims} also tell us what the function |
2a3c07f49615
basic documentation for function elimination rules and fun_cases
krauss
parents:
53107
diff
changeset
|
474 |
value will be in each case: |
61419 | 475 |
\<close> |
54017
2a3c07f49615
basic documentation for function elimination rules and fun_cases
krauss
parents:
53107
diff
changeset
|
476 |
thm list_to_option.elims |
61419 | 477 |
text \<open> |
54017
2a3c07f49615
basic documentation for function elimination rules and fun_cases
krauss
parents:
53107
diff
changeset
|
478 |
@{thm[display] list_to_option.elims} |
2a3c07f49615
basic documentation for function elimination rules and fun_cases
krauss
parents:
53107
diff
changeset
|
479 |
|
2a3c07f49615
basic documentation for function elimination rules and fun_cases
krauss
parents:
53107
diff
changeset
|
480 |
\noindent |
69597 | 481 |
This lets us eliminate an assumption of the form \<^prop>\<open>list_to_option xs = y\<close> and replace it |
54017
2a3c07f49615
basic documentation for function elimination rules and fun_cases
krauss
parents:
53107
diff
changeset
|
482 |
with the two cases, e.g.: |
61419 | 483 |
\<close> |
54017
2a3c07f49615
basic documentation for function elimination rules and fun_cases
krauss
parents:
53107
diff
changeset
|
484 |
|
2a3c07f49615
basic documentation for function elimination rules and fun_cases
krauss
parents:
53107
diff
changeset
|
485 |
lemma "list_to_option xs = y \<Longrightarrow> P" |
2a3c07f49615
basic documentation for function elimination rules and fun_cases
krauss
parents:
53107
diff
changeset
|
486 |
proof (erule list_to_option.elims) |
2a3c07f49615
basic documentation for function elimination rules and fun_cases
krauss
parents:
53107
diff
changeset
|
487 |
fix x assume "xs = [x]" "y = Some x" thus P sorry |
2a3c07f49615
basic documentation for function elimination rules and fun_cases
krauss
parents:
53107
diff
changeset
|
488 |
next |
2a3c07f49615
basic documentation for function elimination rules and fun_cases
krauss
parents:
53107
diff
changeset
|
489 |
assume "xs = []" "y = None" thus P sorry |
2a3c07f49615
basic documentation for function elimination rules and fun_cases
krauss
parents:
53107
diff
changeset
|
490 |
next |
2a3c07f49615
basic documentation for function elimination rules and fun_cases
krauss
parents:
53107
diff
changeset
|
491 |
fix a b xs' assume "xs = a # b # xs'" "y = None" thus P sorry |
2a3c07f49615
basic documentation for function elimination rules and fun_cases
krauss
parents:
53107
diff
changeset
|
492 |
qed |
2a3c07f49615
basic documentation for function elimination rules and fun_cases
krauss
parents:
53107
diff
changeset
|
493 |
|
2a3c07f49615
basic documentation for function elimination rules and fun_cases
krauss
parents:
53107
diff
changeset
|
494 |
|
61419 | 495 |
text \<open> |
69505 | 496 |
Sometimes it is convenient to derive specialized versions of the \<open>elim\<close> rules above and |
54017
2a3c07f49615
basic documentation for function elimination rules and fun_cases
krauss
parents:
53107
diff
changeset
|
497 |
keep them around as facts explicitly. For example, it is natural to show that if |
69597 | 498 |
\<^prop>\<open>list_to_option xs = Some y\<close>, then \<^term>\<open>xs\<close> must be a singleton. The command |
54017
2a3c07f49615
basic documentation for function elimination rules and fun_cases
krauss
parents:
53107
diff
changeset
|
499 |
\cmd{fun\_cases} derives such facts automatically, by instantiating and simplifying the general |
2a3c07f49615
basic documentation for function elimination rules and fun_cases
krauss
parents:
53107
diff
changeset
|
500 |
elimination rules given some pattern: |
61419 | 501 |
\<close> |
54017
2a3c07f49615
basic documentation for function elimination rules and fun_cases
krauss
parents:
53107
diff
changeset
|
502 |
|
2a3c07f49615
basic documentation for function elimination rules and fun_cases
krauss
parents:
53107
diff
changeset
|
503 |
fun_cases list_to_option_SomeE[elim]: "list_to_option xs = Some y" |
2a3c07f49615
basic documentation for function elimination rules and fun_cases
krauss
parents:
53107
diff
changeset
|
504 |
|
2a3c07f49615
basic documentation for function elimination rules and fun_cases
krauss
parents:
53107
diff
changeset
|
505 |
thm list_to_option_SomeE |
61419 | 506 |
text \<open> |
54017
2a3c07f49615
basic documentation for function elimination rules and fun_cases
krauss
parents:
53107
diff
changeset
|
507 |
@{thm[display] list_to_option_SomeE} |
61419 | 508 |
\<close> |
54017
2a3c07f49615
basic documentation for function elimination rules and fun_cases
krauss
parents:
53107
diff
changeset
|
509 |
|
2a3c07f49615
basic documentation for function elimination rules and fun_cases
krauss
parents:
53107
diff
changeset
|
510 |
|
61419 | 511 |
section \<open>General pattern matching\<close> |
512 |
text\<open>\label{genpats}\<close> |
|
22065 | 513 |
|
61419 | 514 |
subsection \<open>Avoiding automatic pattern splitting\<close> |
22065 | 515 |
|
61419 | 516 |
text \<open> |
22065 | 517 |
|
518 |
Up to now, we used pattern matching only on datatypes, and the |
|
519 |
patterns were always disjoint and complete, and if they weren't, |
|
520 |
they were made disjoint automatically like in the definition of |
|
69597 | 521 |
\<^const>\<open>sep\<close> in \S\ref{patmatch}. |
22065 | 522 |
|
23188 | 523 |
This automatic splitting can significantly increase the number of |
524 |
equations involved, and this is not always desirable. The following |
|
525 |
example shows the problem: |
|
22065 | 526 |
|
23805 | 527 |
Suppose we are modeling incomplete knowledge about the world by a |
69597 | 528 |
three-valued datatype, which has values \<^term>\<open>T\<close>, \<^term>\<open>F\<close> |
529 |
and \<^term>\<open>X\<close> for true, false and uncertain propositions, respectively. |
|
61419 | 530 |
\<close> |
22065 | 531 |
|
58310 | 532 |
datatype P3 = T | F | X |
22065 | 533 |
|
61419 | 534 |
text \<open>\noindent Then the conjunction of such values can be defined as follows:\<close> |
22065 | 535 |
|
536 |
fun And :: "P3 \<Rightarrow> P3 \<Rightarrow> P3" |
|
537 |
where |
|
538 |
"And T p = p" |
|
23003 | 539 |
| "And p T = p" |
540 |
| "And p F = F" |
|
541 |
| "And F p = F" |
|
542 |
| "And X X = X" |
|
21212 | 543 |
|
544 |
||
61419 | 545 |
text \<open> |
22065 | 546 |
This definition is useful, because the equations can directly be used |
28040 | 547 |
as simplification rules. But the patterns overlap: For example, |
69597 | 548 |
the expression \<^term>\<open>And T T\<close> is matched by both the first and |
23188 | 549 |
the second equation. By default, Isabelle makes the patterns disjoint by |
22065 | 550 |
splitting them up, producing instances: |
61419 | 551 |
\<close> |
22065 | 552 |
|
553 |
thm And.simps |
|
554 |
||
61419 | 555 |
text \<open> |
22065 | 556 |
@{thm[indent=4] And.simps} |
557 |
||
558 |
\vspace*{1em} |
|
23003 | 559 |
\noindent There are several problems with this: |
22065 | 560 |
|
561 |
\begin{enumerate} |
|
23188 | 562 |
\item If the datatype has many constructors, there can be an |
69597 | 563 |
explosion of equations. For \<^const>\<open>And\<close>, we get seven instead of |
23003 | 564 |
five equations, which can be tolerated, but this is just a small |
22065 | 565 |
example. |
566 |
||
23188 | 567 |
\item Since splitting makes the equations \qt{less general}, they |
69597 | 568 |
do not always match in rewriting. While the term \<^term>\<open>And x F\<close> |
569 |
can be simplified to \<^term>\<open>F\<close> with the original equations, a |
|
570 |
(manual) case split on \<^term>\<open>x\<close> is now necessary. |
|
22065 | 571 |
|
53107
57c7294eac0a
more document antiquotations (for proper theorem names);
Christian Sternagel
parents:
50530
diff
changeset
|
572 |
\item The splitting also concerns the induction rule @{thm [source] |
22065 | 573 |
"And.induct"}. Instead of five premises it now has seven, which |
574 |
means that our induction proofs will have more cases. |
|
575 |
||
576 |
\item In general, it increases clarity if we get the same definition |
|
577 |
back which we put in. |
|
578 |
\end{enumerate} |
|
579 |
||
23188 | 580 |
If we do not want the automatic splitting, we can switch it off by |
581 |
leaving out the \cmd{sequential} option. However, we will have to |
|
582 |
prove that our pattern matching is consistent\footnote{This prevents |
|
69597 | 583 |
us from defining something like \<^term>\<open>f x = True\<close> and \<^term>\<open>f x |
584 |
= False\<close> simultaneously.}: |
|
61419 | 585 |
\<close> |
22065 | 586 |
|
587 |
function And2 :: "P3 \<Rightarrow> P3 \<Rightarrow> P3" |
|
588 |
where |
|
589 |
"And2 T p = p" |
|
23003 | 590 |
| "And2 p T = p" |
591 |
| "And2 p F = F" |
|
592 |
| "And2 F p = F" |
|
593 |
| "And2 X X = X" |
|
22065 | 594 |
|
61419 | 595 |
text \<open> |
23188 | 596 |
\noindent Now let's look at the proof obligations generated by a |
22065 | 597 |
function definition. In this case, they are: |
598 |
||
23188 | 599 |
@{subgoals[display,indent=0]}\vspace{-1.2em}\hspace{3cm}\vdots\vspace{1.2em} |
22065 | 600 |
|
601 |
The first subgoal expresses the completeness of the patterns. It has |
|
69597 | 602 |
the form of an elimination rule and states that every \<^term>\<open>x\<close> of |
23188 | 603 |
the function's input type must match at least one of the patterns\footnote{Completeness could |
22065 | 604 |
be equivalently stated as a disjunction of existential statements: |
69597 | 605 |
\<^term>\<open>(\<exists>p. x = (T, p)) \<or> (\<exists>p. x = (p, T)) \<or> (\<exists>p. x = (p, F)) \<or> |
606 |
(\<exists>p. x = (F, p)) \<or> (x = (X, X))\<close>, and you can use the method \<open>atomize_elim\<close> to get that form instead.}. If the patterns just involve |
|
69505 | 607 |
datatypes, we can solve it with the \<open>pat_completeness\<close> |
23188 | 608 |
method: |
61419 | 609 |
\<close> |
22065 | 610 |
|
611 |
apply pat_completeness |
|
612 |
||
61419 | 613 |
text \<open> |
22065 | 614 |
The remaining subgoals express \emph{pattern compatibility}. We do |
23188 | 615 |
allow that an input value matches multiple patterns, but in this |
22065 | 616 |
case, the result (i.e.~the right hand sides of the equations) must |
617 |
also be equal. For each pair of two patterns, there is one such |
|
618 |
subgoal. Usually this needs injectivity of the constructors, which |
|
69505 | 619 |
is used automatically by \<open>auto\<close>. |
61419 | 620 |
\<close> |
22065 | 621 |
|
622 |
by auto |
|
43042
0f9534b7ea75
function tutorial: do not omit termination proof, even when discussing other things
krauss
parents:
41847
diff
changeset
|
623 |
termination by (relation "{}") simp |
21212 | 624 |
|
625 |
||
61419 | 626 |
subsection \<open>Non-constructor patterns\<close> |
21212 | 627 |
|
61419 | 628 |
text \<open> |
23805 | 629 |
Most of Isabelle's basic types take the form of inductive datatypes, |
630 |
and usually pattern matching works on the constructors of such types. |
|
631 |
However, this need not be always the case, and the \cmd{function} |
|
632 |
command handles other kind of patterns, too. |
|
23188 | 633 |
|
23805 | 634 |
One well-known instance of non-constructor patterns are |
23188 | 635 |
so-called \emph{$n+k$-patterns}, which are a little controversial in |
636 |
the functional programming world. Here is the initial fibonacci |
|
637 |
example with $n+k$-patterns: |
|
61419 | 638 |
\<close> |
23188 | 639 |
|
640 |
function fib2 :: "nat \<Rightarrow> nat" |
|
641 |
where |
|
642 |
"fib2 0 = 1" |
|
643 |
| "fib2 1 = 1" |
|
644 |
| "fib2 (n + 2) = fib2 n + fib2 (Suc n)" |
|
645 |
||
61419 | 646 |
text \<open> |
23805 | 647 |
This kind of matching is again justified by the proof of pattern |
648 |
completeness and compatibility. |
|
23188 | 649 |
The proof obligation for pattern completeness states that every natural number is |
69597 | 650 |
either \<^term>\<open>0::nat\<close>, \<^term>\<open>1::nat\<close> or \<^term>\<open>n + |
651 |
(2::nat)\<close>: |
|
23188 | 652 |
|
39752
06fc1a79b4bf
removed unnecessary reference poking (cf. f45d332a90e3)
krauss
parents:
33856
diff
changeset
|
653 |
@{subgoals[display,indent=0,goals_limit=1]} |
23188 | 654 |
|
655 |
This is an arithmetic triviality, but unfortunately the |
|
69505 | 656 |
\<open>arith\<close> method cannot handle this specific form of an |
657 |
elimination rule. However, we can use the method \<open>atomize_elim\<close> to do an ad-hoc conversion to a disjunction of |
|
28040 | 658 |
existentials, which can then be solved by the arithmetic decision procedure. |
23805 | 659 |
Pattern compatibility and termination are automatic as usual. |
61419 | 660 |
\<close> |
26580
c3e597a476fd
Generic conversion and tactic "atomize_elim" to convert elimination rules
krauss
parents:
25688
diff
changeset
|
661 |
apply atomize_elim |
23805 | 662 |
apply arith |
23188 | 663 |
apply auto |
664 |
done |
|
665 |
termination by lexicographic_order |
|
61419 | 666 |
text \<open> |
23188 | 667 |
We can stretch the notion of pattern matching even more. The |
668 |
following function is not a sensible functional program, but a |
|
669 |
perfectly valid mathematical definition: |
|
61419 | 670 |
\<close> |
23188 | 671 |
|
672 |
function ev :: "nat \<Rightarrow> bool" |
|
673 |
where |
|
674 |
"ev (2 * n) = True" |
|
675 |
| "ev (2 * n + 1) = False" |
|
26580
c3e597a476fd
Generic conversion and tactic "atomize_elim" to convert elimination rules
krauss
parents:
25688
diff
changeset
|
676 |
apply atomize_elim |
23805 | 677 |
by arith+ |
23188 | 678 |
termination by (relation "{}") simp |
679 |
||
61419 | 680 |
text \<open> |
27026 | 681 |
This general notion of pattern matching gives you a certain freedom |
682 |
in writing down specifications. However, as always, such freedom should |
|
23188 | 683 |
be used with care: |
684 |
||
685 |
If we leave the area of constructor |
|
686 |
patterns, we have effectively departed from the world of functional |
|
687 |
programming. This means that it is no longer possible to use the |
|
688 |
code generator, and expect it to generate ML code for our |
|
689 |
definitions. Also, such a specification might not work very well together with |
|
690 |
simplification. Your mileage may vary. |
|
61419 | 691 |
\<close> |
23188 | 692 |
|
693 |
||
61419 | 694 |
subsection \<open>Conditional equations\<close> |
23188 | 695 |
|
61419 | 696 |
text \<open> |
23188 | 697 |
The function package also supports conditional equations, which are |
698 |
similar to guards in a language like Haskell. Here is Euclid's |
|
699 |
algorithm written with conditional patterns\footnote{Note that the |
|
700 |
patterns are also overlapping in the base case}: |
|
61419 | 701 |
\<close> |
23188 | 702 |
|
703 |
function gcd :: "nat \<Rightarrow> nat \<Rightarrow> nat" |
|
704 |
where |
|
705 |
"gcd x 0 = x" |
|
706 |
| "gcd 0 y = y" |
|
707 |
| "x < y \<Longrightarrow> gcd (Suc x) (Suc y) = gcd (Suc x) (y - x)" |
|
708 |
| "\<not> x < y \<Longrightarrow> gcd (Suc x) (Suc y) = gcd (x - y) (Suc y)" |
|
26580
c3e597a476fd
Generic conversion and tactic "atomize_elim" to convert elimination rules
krauss
parents:
25688
diff
changeset
|
709 |
by (atomize_elim, auto, arith) |
23188 | 710 |
termination by lexicographic_order |
711 |
||
61419 | 712 |
text \<open> |
23188 | 713 |
By now, you can probably guess what the proof obligations for the |
714 |
pattern completeness and compatibility look like. |
|
715 |
||
716 |
Again, functions with conditional patterns are not supported by the |
|
717 |
code generator. |
|
61419 | 718 |
\<close> |
23188 | 719 |
|
720 |
||
61419 | 721 |
subsection \<open>Pattern matching on strings\<close> |
23188 | 722 |
|
61419 | 723 |
text \<open> |
23805 | 724 |
As strings (as lists of characters) are normal datatypes, pattern |
23188 | 725 |
matching on them is possible, but somewhat problematic. Consider the |
726 |
following definition: |
|
727 |
||
728 |
\end{isamarkuptext} |
|
69505 | 729 |
\noindent\cmd{fun} \<open>check :: "string \<Rightarrow> bool"\<close>\\% |
23188 | 730 |
\cmd{where}\\% |
69505 | 731 |
\hspace*{2ex}\<open>"check (''good'') = True"\<close>\\% |
732 |
\<open>| "check s = False"\<close> |
|
23188 | 733 |
\begin{isamarkuptext} |
734 |
||
23805 | 735 |
\noindent An invocation of the above \cmd{fun} command does not |
23188 | 736 |
terminate. What is the problem? Strings are lists of characters, and |
23805 | 737 |
characters are a datatype with a lot of constructors. Splitting the |
23188 | 738 |
catch-all pattern thus leads to an explosion of cases, which cannot |
739 |
be handled by Isabelle. |
|
740 |
||
741 |
There are two things we can do here. Either we write an explicit |
|
69505 | 742 |
\<open>if\<close> on the right hand side, or we can use conditional patterns: |
61419 | 743 |
\<close> |
23188 | 744 |
|
745 |
function check :: "string \<Rightarrow> bool" |
|
746 |
where |
|
747 |
"check (''good'') = True" |
|
748 |
| "s \<noteq> ''good'' \<Longrightarrow> check s = False" |
|
749 |
by auto |
|
43042
0f9534b7ea75
function tutorial: do not omit termination proof, even when discussing other things
krauss
parents:
41847
diff
changeset
|
750 |
termination by (relation "{}") simp |
22065 | 751 |
|
752 |
||
76649
9a6cb5ecc183
Added section about code generation for partial functions
nipkow
parents:
70276
diff
changeset
|
753 |
section \<open>Partiality \label{sec:partiality}\<close> |
22065 | 754 |
|
61419 | 755 |
text \<open> |
69597 | 756 |
In HOL, all functions are total. A function \<^term>\<open>f\<close> applied to |
757 |
\<^term>\<open>x\<close> always has the value \<^term>\<open>f x\<close>, and there is no notion |
|
22065 | 758 |
of undefinedness. |
23188 | 759 |
This is why we have to do termination |
760 |
proofs when defining functions: The proof justifies that the |
|
761 |
function can be defined by wellfounded recursion. |
|
22065 | 762 |
|
23188 | 763 |
However, the \cmd{function} package does support partiality to a |
764 |
certain extent. Let's look at the following function which looks |
|
765 |
for a zero of a given function f. |
|
61419 | 766 |
\<close> |
23003 | 767 |
|
41846
b368a7aee46a
removed support for tail-recursion from function package (now implemented by partial_function)
krauss
parents:
39754
diff
changeset
|
768 |
function (*<*)(domintros)(*>*)findzero :: "(nat \<Rightarrow> nat) \<Rightarrow> nat \<Rightarrow> nat" |
23003 | 769 |
where |
770 |
"findzero f n = (if f n = 0 then n else findzero f (Suc n))" |
|
771 |
by pat_completeness auto |
|
772 |
||
61419 | 773 |
text \<open> |
23805 | 774 |
\noindent Clearly, any attempt of a termination proof must fail. And without |
69505 | 775 |
that, we do not get the usual rules \<open>findzero.simps\<close> and |
776 |
\<open>findzero.induct\<close>. So what was the definition good for at all? |
|
61419 | 777 |
\<close> |
23003 | 778 |
|
61419 | 779 |
subsection \<open>Domain predicates\<close> |
23003 | 780 |
|
61419 | 781 |
text \<open> |
69597 | 782 |
The trick is that Isabelle has not only defined the function \<^const>\<open>findzero\<close>, but also |
783 |
a predicate \<^term>\<open>findzero_dom\<close> that characterizes the values where the function |
|
23188 | 784 |
terminates: the \emph{domain} of the function. If we treat a |
785 |
partial function just as a total function with an additional domain |
|
786 |
predicate, we can derive simplification and |
|
787 |
induction rules as we do for total functions. They are guarded |
|
69505 | 788 |
by domain conditions and are called \<open>psimps\<close> and \<open>pinduct\<close>: |
61419 | 789 |
\<close> |
23003 | 790 |
|
61419 | 791 |
text \<open> |
23805 | 792 |
\noindent\begin{minipage}{0.79\textwidth}@{thm[display,margin=85] findzero.psimps}\end{minipage} |
53107
57c7294eac0a
more document antiquotations (for proper theorem names);
Christian Sternagel
parents:
50530
diff
changeset
|
793 |
\hfill(@{thm [source] "findzero.psimps"}) |
23805 | 794 |
\vspace{1em} |
23003 | 795 |
|
23805 | 796 |
\noindent\begin{minipage}{0.79\textwidth}@{thm[display,margin=85] findzero.pinduct}\end{minipage} |
53107
57c7294eac0a
more document antiquotations (for proper theorem names);
Christian Sternagel
parents:
50530
diff
changeset
|
797 |
\hfill(@{thm [source] "findzero.pinduct"}) |
61419 | 798 |
\<close> |
23003 | 799 |
|
61419 | 800 |
text \<open> |
23188 | 801 |
Remember that all we |
802 |
are doing here is use some tricks to make a total function appear |
|
69597 | 803 |
as if it was partial. We can still write the term \<^term>\<open>findzero |
804 |
(\<lambda>x. 1) 0\<close> and like any other term of type \<^typ>\<open>nat\<close> it is equal |
|
23003 | 805 |
to some natural number, although we might not be able to find out |
23188 | 806 |
which one. The function is \emph{underdefined}. |
23003 | 807 |
|
23805 | 808 |
But it is defined enough to prove something interesting about it. We |
69597 | 809 |
can prove that if \<^term>\<open>findzero f n\<close> |
810 |
terminates, it indeed returns a zero of \<^term>\<open>f\<close>: |
|
61419 | 811 |
\<close> |
23003 | 812 |
|
813 |
lemma findzero_zero: "findzero_dom (f, n) \<Longrightarrow> f (findzero f n) = 0" |
|
814 |
||
61419 | 815 |
text \<open>\noindent We apply induction as usual, but using the partial induction |
816 |
rule:\<close> |
|
23003 | 817 |
|
818 |
apply (induct f n rule: findzero.pinduct) |
|
819 |
||
61419 | 820 |
text \<open>\noindent This gives the following subgoals: |
23003 | 821 |
|
822 |
@{subgoals[display,indent=0]} |
|
823 |
||
23805 | 824 |
\noindent The hypothesis in our lemma was used to satisfy the first premise in |
69597 | 825 |
the induction rule. However, we also get \<^term>\<open>findzero_dom (f, n)\<close> as a local assumption in the induction step. This |
826 |
allows unfolding \<^term>\<open>findzero f n\<close> using the \<open>psimps\<close> |
|
39754 | 827 |
rule, and the rest is trivial. |
61419 | 828 |
\<close> |
39754 | 829 |
apply (simp add: findzero.psimps) |
23003 | 830 |
done |
831 |
||
61419 | 832 |
text \<open> |
23003 | 833 |
Proofs about partial functions are often not harder than for total |
834 |
functions. Fig.~\ref{findzero_isar} shows a slightly more |
|
835 |
complicated proof written in Isar. It is verbose enough to show how |
|
836 |
partiality comes into play: From the partial induction, we get an |
|
837 |
additional domain condition hypothesis. Observe how this condition |
|
69597 | 838 |
is applied when calls to \<^term>\<open>findzero\<close> are unfolded. |
61419 | 839 |
\<close> |
23003 | 840 |
|
61419 | 841 |
text_raw \<open> |
23003 | 842 |
\begin{figure} |
23188 | 843 |
\hrule\vspace{6pt} |
23003 | 844 |
\begin{minipage}{0.8\textwidth} |
845 |
\isabellestyle{it} |
|
846 |
\isastyle\isamarkuptrue |
|
61419 | 847 |
\<close> |
23003 | 848 |
lemma "\<lbrakk>findzero_dom (f, n); x \<in> {n ..< findzero f n}\<rbrakk> \<Longrightarrow> f x \<noteq> 0" |
849 |
proof (induct rule: findzero.pinduct) |
|
850 |
fix f n assume dom: "findzero_dom (f, n)" |
|
23188 | 851 |
and IH: "\<lbrakk>f n \<noteq> 0; x \<in> {Suc n ..< findzero f (Suc n)}\<rbrakk> \<Longrightarrow> f x \<noteq> 0" |
852 |
and x_range: "x \<in> {n ..< findzero f n}" |
|
23003 | 853 |
have "f n \<noteq> 0" |
854 |
proof |
|
855 |
assume "f n = 0" |
|
39754 | 856 |
with dom have "findzero f n = n" by (simp add: findzero.psimps) |
23003 | 857 |
with x_range show False by auto |
858 |
qed |
|
859 |
||
860 |
from x_range have "x = n \<or> x \<in> {Suc n ..< findzero f n}" by auto |
|
861 |
thus "f x \<noteq> 0" |
|
862 |
proof |
|
863 |
assume "x = n" |
|
61419 | 864 |
with \<open>f n \<noteq> 0\<close> show ?thesis by simp |
23003 | 865 |
next |
23188 | 866 |
assume "x \<in> {Suc n ..< findzero f n}" |
61419 | 867 |
with dom and \<open>f n \<noteq> 0\<close> have "x \<in> {Suc n ..< findzero f (Suc n)}" by (simp add: findzero.psimps) |
868 |
with IH and \<open>f n \<noteq> 0\<close> |
|
23003 | 869 |
show ?thesis by simp |
870 |
qed |
|
871 |
qed |
|
61419 | 872 |
text_raw \<open> |
23003 | 873 |
\isamarkupfalse\isabellestyle{tt} |
23188 | 874 |
\end{minipage}\vspace{6pt}\hrule |
23003 | 875 |
\caption{A proof about a partial function}\label{findzero_isar} |
876 |
\end{figure} |
|
61419 | 877 |
\<close> |
23003 | 878 |
|
61419 | 879 |
subsection \<open>Partial termination proofs\<close> |
23003 | 880 |
|
61419 | 881 |
text \<open> |
23003 | 882 |
Now that we have proved some interesting properties about our |
883 |
function, we should turn to the domain predicate and see if it is |
|
884 |
actually true for some values. Otherwise we would have just proved |
|
69597 | 885 |
lemmas with \<^term>\<open>False\<close> as a premise. |
23003 | 886 |
|
69505 | 887 |
Essentially, we need some introduction rules for \<open>findzero_dom\<close>. The function package can prove such domain |
23003 | 888 |
introduction rules automatically. But since they are not used very |
23188 | 889 |
often (they are almost never needed if the function is total), this |
890 |
functionality is disabled by default for efficiency reasons. So we have to go |
|
69505 | 891 |
back and ask for them explicitly by passing the \<open>(domintros)\<close> option to the function package: |
23003 | 892 |
|
23188 | 893 |
\vspace{1ex} |
69505 | 894 |
\noindent\cmd{function} \<open>(domintros) findzero :: "(nat \<Rightarrow> nat) \<Rightarrow> nat \<Rightarrow> nat"\<close>\\% |
23003 | 895 |
\cmd{where}\isanewline% |
896 |
\ \ \ldots\\ |
|
897 |
||
69505 | 898 |
\noindent Now the package has proved an introduction rule for \<open>findzero_dom\<close>: |
61419 | 899 |
\<close> |
23003 | 900 |
|
901 |
thm findzero.domintros |
|
902 |
||
61419 | 903 |
text \<open> |
23003 | 904 |
@{thm[display] findzero.domintros} |
905 |
||
906 |
Domain introduction rules allow to show that a given value lies in the |
|
907 |
domain of a function, if the arguments of all recursive calls |
|
908 |
are in the domain as well. They allow to do a \qt{single step} in a |
|
909 |
termination proof. Usually, you want to combine them with a suitable |
|
910 |
induction principle. |
|
911 |
||
912 |
Since our function increases its argument at recursive calls, we |
|
913 |
need an induction principle which works \qt{backwards}. We will use |
|
53107
57c7294eac0a
more document antiquotations (for proper theorem names);
Christian Sternagel
parents:
50530
diff
changeset
|
914 |
@{thm [source] inc_induct}, which allows to do induction from a fixed number |
23003 | 915 |
\qt{downwards}: |
916 |
||
53107
57c7294eac0a
more document antiquotations (for proper theorem names);
Christian Sternagel
parents:
50530
diff
changeset
|
917 |
\begin{center}@{thm inc_induct}\hfill(@{thm [source] "inc_induct"})\end{center} |
23003 | 918 |
|
23188 | 919 |
Figure \ref{findzero_term} gives a detailed Isar proof of the fact |
69505 | 920 |
that \<open>findzero\<close> terminates if there is a zero which is greater |
69597 | 921 |
or equal to \<^term>\<open>n\<close>. First we derive two useful rules which will |
23003 | 922 |
solve the base case and the step case of the induction. The |
23805 | 923 |
induction is then straightforward, except for the unusual induction |
23003 | 924 |
principle. |
925 |
||
61419 | 926 |
\<close> |
23003 | 927 |
|
61419 | 928 |
text_raw \<open> |
23003 | 929 |
\begin{figure} |
23188 | 930 |
\hrule\vspace{6pt} |
23003 | 931 |
\begin{minipage}{0.8\textwidth} |
932 |
\isabellestyle{it} |
|
933 |
\isastyle\isamarkuptrue |
|
61419 | 934 |
\<close> |
23003 | 935 |
lemma findzero_termination: |
23188 | 936 |
assumes "x \<ge> n" and "f x = 0" |
23003 | 937 |
shows "findzero_dom (f, n)" |
938 |
proof - |
|
939 |
have base: "findzero_dom (f, x)" |
|
61419 | 940 |
by (rule findzero.domintros) (simp add:\<open>f x = 0\<close>) |
23003 | 941 |
|
942 |
have step: "\<And>i. findzero_dom (f, Suc i) |
|
943 |
\<Longrightarrow> findzero_dom (f, i)" |
|
944 |
by (rule findzero.domintros) simp |
|
945 |
||
61419 | 946 |
from \<open>x \<ge> n\<close> show ?thesis |
23003 | 947 |
proof (induct rule:inc_induct) |
23188 | 948 |
show "findzero_dom (f, x)" by (rule base) |
23003 | 949 |
next |
950 |
fix i assume "findzero_dom (f, Suc i)" |
|
23188 | 951 |
thus "findzero_dom (f, i)" by (rule step) |
23003 | 952 |
qed |
953 |
qed |
|
61419 | 954 |
text_raw \<open> |
23003 | 955 |
\isamarkupfalse\isabellestyle{tt} |
23188 | 956 |
\end{minipage}\vspace{6pt}\hrule |
69505 | 957 |
\caption{Termination proof for \<open>findzero\<close>}\label{findzero_term} |
23003 | 958 |
\end{figure} |
61419 | 959 |
\<close> |
23003 | 960 |
|
61419 | 961 |
text \<open> |
23003 | 962 |
Again, the proof given in Fig.~\ref{findzero_term} has a lot of |
963 |
detail in order to explain the principles. Using more automation, we |
|
964 |
can also have a short proof: |
|
61419 | 965 |
\<close> |
23003 | 966 |
|
967 |
lemma findzero_termination_short: |
|
968 |
assumes zero: "x >= n" |
|
969 |
assumes [simp]: "f x = 0" |
|
970 |
shows "findzero_dom (f, n)" |
|
23805 | 971 |
using zero |
972 |
by (induct rule:inc_induct) (auto intro: findzero.domintros) |
|
23003 | 973 |
|
61419 | 974 |
text \<open> |
23188 | 975 |
\noindent It is simple to combine the partial correctness result with the |
23003 | 976 |
termination lemma: |
61419 | 977 |
\<close> |
23003 | 978 |
|
979 |
lemma findzero_total_correctness: |
|
980 |
"f x = 0 \<Longrightarrow> f (findzero f 0) = 0" |
|
981 |
by (blast intro: findzero_zero findzero_termination) |
|
982 |
||
61419 | 983 |
subsection \<open>Definition of the domain predicate\<close> |
23003 | 984 |
|
61419 | 985 |
text \<open> |
23003 | 986 |
Sometimes it is useful to know what the definition of the domain |
69505 | 987 |
predicate looks like. Actually, \<open>findzero_dom\<close> is just an |
23003 | 988 |
abbreviation: |
989 |
||
990 |
@{abbrev[display] findzero_dom} |
|
991 |
||
69597 | 992 |
The domain predicate is the \emph{accessible part} of a relation \<^const>\<open>findzero_rel\<close>, which was also created internally by the function |
993 |
package. \<^const>\<open>findzero_rel\<close> is just a normal |
|
23188 | 994 |
inductive predicate, so we can inspect its definition by |
53107
57c7294eac0a
more document antiquotations (for proper theorem names);
Christian Sternagel
parents:
50530
diff
changeset
|
995 |
looking at the introduction rules @{thm [source] findzero_rel.intros}. |
23003 | 996 |
In our case there is just a single rule: |
997 |
||
998 |
@{thm[display] findzero_rel.intros} |
|
999 |
||
69597 | 1000 |
The predicate \<^const>\<open>findzero_rel\<close> |
23003 | 1001 |
describes the \emph{recursion relation} of the function |
1002 |
definition. The recursion relation is a binary relation on |
|
1003 |
the arguments of the function that relates each argument to its |
|
1004 |
recursive calls. In general, there is one introduction rule for each |
|
1005 |
recursive call. |
|
1006 |
||
69597 | 1007 |
The predicate \<^term>\<open>Wellfounded.accp findzero_rel\<close> is the accessible part of |
23003 | 1008 |
that relation. An argument belongs to the accessible part, if it can |
69505 | 1009 |
be reached in a finite number of steps (cf.~its definition in \<open>Wellfounded.thy\<close>). |
23003 | 1010 |
|
1011 |
Since the domain predicate is just an abbreviation, you can use |
|
69597 | 1012 |
lemmas for \<^const>\<open>Wellfounded.accp\<close> and \<^const>\<open>findzero_rel\<close> directly. Some |
53107
57c7294eac0a
more document antiquotations (for proper theorem names);
Christian Sternagel
parents:
50530
diff
changeset
|
1013 |
lemmas which are occasionally useful are @{thm [source] accpI}, @{thm [source] |
23805 | 1014 |
accp_downward}, and of course the introduction and elimination rules |
53107
57c7294eac0a
more document antiquotations (for proper theorem names);
Christian Sternagel
parents:
50530
diff
changeset
|
1015 |
for the recursion relation @{thm [source] "findzero_rel.intros"} and @{thm |
57c7294eac0a
more document antiquotations (for proper theorem names);
Christian Sternagel
parents:
50530
diff
changeset
|
1016 |
[source] "findzero_rel.cases"}. |
61419 | 1017 |
\<close> |
23003 | 1018 |
|
61419 | 1019 |
section \<open>Nested recursion\<close> |
23003 | 1020 |
|
61419 | 1021 |
text \<open> |
23003 | 1022 |
Recursive calls which are nested in one another frequently cause |
1023 |
complications, since their termination proof can depend on a partial |
|
1024 |
correctness property of the function itself. |
|
1025 |
||
1026 |
As a small example, we define the \qt{nested zero} function: |
|
61419 | 1027 |
\<close> |
23003 | 1028 |
|
1029 |
function nz :: "nat \<Rightarrow> nat" |
|
1030 |
where |
|
1031 |
"nz 0 = 0" |
|
1032 |
| "nz (Suc n) = nz (nz n)" |
|
1033 |
by pat_completeness auto |
|
1034 |
||
61419 | 1035 |
text \<open> |
23003 | 1036 |
If we attempt to prove termination using the identity measure on |
1037 |
naturals, this fails: |
|
61419 | 1038 |
\<close> |
23003 | 1039 |
|
1040 |
termination |
|
1041 |
apply (relation "measure (\<lambda>n. n)") |
|
1042 |
apply auto |
|
1043 |
||
61419 | 1044 |
text \<open> |
23003 | 1045 |
We get stuck with the subgoal |
1046 |
||
1047 |
@{subgoals[display]} |
|
1048 |
||
69597 | 1049 |
Of course this statement is true, since we know that \<^const>\<open>nz\<close> is |
23003 | 1050 |
the zero function. And in fact we have no problem proving this |
1051 |
property by induction. |
|
61419 | 1052 |
\<close> |
23805 | 1053 |
(*<*)oops(*>*) |
23003 | 1054 |
lemma nz_is_zero: "nz_dom n \<Longrightarrow> nz n = 0" |
39754 | 1055 |
by (induct rule:nz.pinduct) (auto simp: nz.psimps) |
23003 | 1056 |
|
61419 | 1057 |
text \<open> |
23003 | 1058 |
We formulate this as a partial correctness lemma with the condition |
69597 | 1059 |
\<^term>\<open>nz_dom n\<close>. This allows us to prove it with the \<open>pinduct\<close> rule before we have proved termination. With this lemma, |
23003 | 1060 |
the termination proof works as expected: |
61419 | 1061 |
\<close> |
23003 | 1062 |
|
1063 |
termination |
|
1064 |
by (relation "measure (\<lambda>n. n)") (auto simp: nz_is_zero) |
|
1065 |
||
61419 | 1066 |
text \<open> |
23003 | 1067 |
As a general strategy, one should prove the statements needed for |
1068 |
termination as a partial property first. Then they can be used to do |
|
1069 |
the termination proof. This also works for less trivial |
|
23188 | 1070 |
examples. Figure \ref{f91} defines the 91-function, a well-known |
1071 |
challenge problem due to John McCarthy, and proves its termination. |
|
61419 | 1072 |
\<close> |
23003 | 1073 |
|
61419 | 1074 |
text_raw \<open> |
23003 | 1075 |
\begin{figure} |
23188 | 1076 |
\hrule\vspace{6pt} |
23003 | 1077 |
\begin{minipage}{0.8\textwidth} |
1078 |
\isabellestyle{it} |
|
1079 |
\isastyle\isamarkuptrue |
|
61419 | 1080 |
\<close> |
23003 | 1081 |
|
23188 | 1082 |
function f91 :: "nat \<Rightarrow> nat" |
23003 | 1083 |
where |
1084 |
"f91 n = (if 100 < n then n - 10 else f91 (f91 (n + 11)))" |
|
1085 |
by pat_completeness auto |
|
1086 |
||
1087 |
lemma f91_estimate: |
|
1088 |
assumes trm: "f91_dom n" |
|
1089 |
shows "n < f91 n + 11" |
|
39754 | 1090 |
using trm by induct (auto simp: f91.psimps) |
23003 | 1091 |
|
1092 |
termination |
|
1093 |
proof |
|
1094 |
let ?R = "measure (\<lambda>x. 101 - x)" |
|
1095 |
show "wf ?R" .. |
|
1096 |
||
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
67406
diff
changeset
|
1097 |
fix n :: nat assume "\<not> 100 < n" \<comment> \<open>Assumptions for both calls\<close> |
23003 | 1098 |
|
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
67406
diff
changeset
|
1099 |
thus "(n + 11, n) \<in> ?R" by simp \<comment> \<open>Inner call\<close> |
23003 | 1100 |
|
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
67406
diff
changeset
|
1101 |
assume inner_trm: "f91_dom (n + 11)" \<comment> \<open>Outer call\<close> |
23003 | 1102 |
with f91_estimate have "n + 11 < f91 (n + 11) + 11" . |
61419 | 1103 |
with \<open>\<not> 100 < n\<close> show "(f91 (n + 11), n) \<in> ?R" by simp |
23003 | 1104 |
qed |
1105 |
||
61419 | 1106 |
text_raw \<open> |
23003 | 1107 |
\isamarkupfalse\isabellestyle{tt} |
23188 | 1108 |
\end{minipage} |
1109 |
\vspace{6pt}\hrule |
|
23003 | 1110 |
\caption{McCarthy's 91-function}\label{f91} |
1111 |
\end{figure} |
|
61419 | 1112 |
\<close> |
22065 | 1113 |
|
1114 |
||
61419 | 1115 |
section \<open>Higher-Order Recursion\<close> |
22065 | 1116 |
|
61419 | 1117 |
text \<open> |
23003 | 1118 |
Higher-order recursion occurs when recursive calls |
69597 | 1119 |
are passed as arguments to higher-order combinators such as \<^const>\<open>map\<close>, \<^term>\<open>filter\<close> etc. |
23805 | 1120 |
As an example, imagine a datatype of n-ary trees: |
61419 | 1121 |
\<close> |
23003 | 1122 |
|
58310 | 1123 |
datatype 'a tree = |
23003 | 1124 |
Leaf 'a |
1125 |
| Branch "'a tree list" |
|
1126 |
||
1127 |
||
61419 | 1128 |
text \<open>\noindent We can define a function which swaps the left and right subtrees recursively, using the |
69597 | 1129 |
list functions \<^const>\<open>rev\<close> and \<^const>\<open>map\<close>:\<close> |
25688
6c24a82a98f1
temporarily fixed documentation due to changed size functions
krauss
parents:
25278
diff
changeset
|
1130 |
|
27026 | 1131 |
fun mirror :: "'a tree \<Rightarrow> 'a tree" |
23003 | 1132 |
where |
25278 | 1133 |
"mirror (Leaf n) = Leaf n" |
1134 |
| "mirror (Branch l) = Branch (rev (map mirror l))" |
|
22065 | 1135 |
|
61419 | 1136 |
text \<open> |
27026 | 1137 |
Although the definition is accepted without problems, let us look at the termination proof: |
61419 | 1138 |
\<close> |
23003 | 1139 |
|
1140 |
termination proof |
|
61419 | 1141 |
text \<open> |
23003 | 1142 |
|
1143 |
As usual, we have to give a wellfounded relation, such that the |
|
1144 |
arguments of the recursive calls get smaller. But what exactly are |
|
27026 | 1145 |
the arguments of the recursive calls when mirror is given as an |
69597 | 1146 |
argument to \<^const>\<open>map\<close>? Isabelle gives us the |
23003 | 1147 |
subgoals |
1148 |
||
1149 |
@{subgoals[display,indent=0]} |
|
1150 |
||
69597 | 1151 |
So the system seems to know that \<^const>\<open>map\<close> only |
1152 |
applies the recursive call \<^term>\<open>mirror\<close> to elements |
|
1153 |
of \<^term>\<open>l\<close>, which is essential for the termination proof. |
|
23003 | 1154 |
|
69597 | 1155 |
This knowledge about \<^const>\<open>map\<close> is encoded in so-called congruence rules, |
23003 | 1156 |
which are special theorems known to the \cmd{function} command. The |
69597 | 1157 |
rule for \<^const>\<open>map\<close> is |
23003 | 1158 |
|
1159 |
@{thm[display] map_cong} |
|
1160 |
||
69597 | 1161 |
You can read this in the following way: Two applications of \<^const>\<open>map\<close> are equal, if the list arguments are equal and the functions |
23003 | 1162 |
coincide on the elements of the list. This means that for the value |
69597 | 1163 |
\<^term>\<open>map f l\<close> we only have to know how \<^term>\<open>f\<close> behaves on |
1164 |
the elements of \<^term>\<open>l\<close>. |
|
23003 | 1165 |
|
1166 |
Usually, one such congruence rule is |
|
1167 |
needed for each higher-order construct that is used when defining |
|
69597 | 1168 |
new functions. In fact, even basic functions like \<^const>\<open>If\<close> and \<^const>\<open>Let\<close> are handled by this mechanism. The congruence |
1169 |
rule for \<^const>\<open>If\<close> states that the \<open>then\<close> branch is only |
|
69505 | 1170 |
relevant if the condition is true, and the \<open>else\<close> branch only if it |
23003 | 1171 |
is false: |
1172 |
||
1173 |
@{thm[display] if_cong} |
|
1174 |
||
1175 |
Congruence rules can be added to the |
|
69597 | 1176 |
function package by giving them the \<^term>\<open>fundef_cong\<close> attribute. |
23003 | 1177 |
|
23805 | 1178 |
The constructs that are predefined in Isabelle, usually |
1179 |
come with the respective congruence rules. |
|
27026 | 1180 |
But if you define your own higher-order functions, you may have to |
1181 |
state and prove the required congruence rules yourself, if you want to use your |
|
23805 | 1182 |
functions in recursive definitions. |
61419 | 1183 |
\<close> |
27026 | 1184 |
(*<*)oops(*>*) |
23805 | 1185 |
|
61419 | 1186 |
subsection \<open>Congruence Rules and Evaluation Order\<close> |
23805 | 1187 |
|
61419 | 1188 |
text \<open> |
23805 | 1189 |
Higher order logic differs from functional programming languages in |
1190 |
that it has no built-in notion of evaluation order. A program is |
|
1191 |
just a set of equations, and it is not specified how they must be |
|
1192 |
evaluated. |
|
1193 |
||
1194 |
However for the purpose of function definition, we must talk about |
|
1195 |
evaluation order implicitly, when we reason about termination. |
|
1196 |
Congruence rules express that a certain evaluation order is |
|
1197 |
consistent with the logical definition. |
|
1198 |
||
1199 |
Consider the following function. |
|
61419 | 1200 |
\<close> |
23805 | 1201 |
|
1202 |
function f :: "nat \<Rightarrow> bool" |
|
1203 |
where |
|
1204 |
"f n = (n = 0 \<or> f (n - 1))" |
|
1205 |
(*<*)by pat_completeness auto(*>*) |
|
1206 |
||
61419 | 1207 |
text \<open> |
27026 | 1208 |
For this definition, the termination proof fails. The default configuration |
23805 | 1209 |
specifies no congruence rule for disjunction. We have to add a |
1210 |
congruence rule that specifies left-to-right evaluation order: |
|
1211 |
||
1212 |
\vspace{1ex} |
|
53107
57c7294eac0a
more document antiquotations (for proper theorem names);
Christian Sternagel
parents:
50530
diff
changeset
|
1213 |
\noindent @{thm disj_cong}\hfill(@{thm [source] "disj_cong"}) |
23805 | 1214 |
\vspace{1ex} |
1215 |
||
1216 |
Now the definition works without problems. Note how the termination |
|
1217 |
proof depends on the extra condition that we get from the congruence |
|
1218 |
rule. |
|
1219 |
||
1220 |
However, as evaluation is not a hard-wired concept, we |
|
1221 |
could just turn everything around by declaring a different |
|
1222 |
congruence rule. Then we can make the reverse definition: |
|
61419 | 1223 |
\<close> |
23805 | 1224 |
|
1225 |
lemma disj_cong2[fundef_cong]: |
|
1226 |
"(\<not> Q' \<Longrightarrow> P = P') \<Longrightarrow> (Q = Q') \<Longrightarrow> (P \<or> Q) = (P' \<or> Q')" |
|
1227 |
by blast |
|
1228 |
||
1229 |
fun f' :: "nat \<Rightarrow> bool" |
|
1230 |
where |
|
1231 |
"f' n = (f' (n - 1) \<or> n = 0)" |
|
1232 |
||
61419 | 1233 |
text \<open> |
23805 | 1234 |
\noindent These examples show that, in general, there is no \qt{best} set of |
1235 |
congruence rules. |
|
1236 |
||
1237 |
However, such tweaking should rarely be necessary in |
|
1238 |
practice, as most of the time, the default set of congruence rules |
|
1239 |
works well. |
|
61419 | 1240 |
\<close> |
23805 | 1241 |
|
21212 | 1242 |
end |