21212
|
1 |
(* Title: Doc/IsarAdvanced/Functions/Thy/Fundefs.thy
|
|
2 |
ID: $Id$
|
|
3 |
Author: Alexander Krauss, TU Muenchen
|
|
4 |
|
|
5 |
Tutorial for function definitions with the new "function" package.
|
|
6 |
*)
|
|
7 |
|
|
8 |
theory Functions
|
|
9 |
imports Main
|
|
10 |
begin
|
|
11 |
|
|
12 |
chapter {* Defining Recursive Functions in Isabelle/HOL *}
|
|
13 |
text {* \cite{isa-tutorial} *}
|
|
14 |
|
|
15 |
section {* Function Definition for Dummies *}
|
|
16 |
|
|
17 |
text {*
|
|
18 |
In most cases, defining a recursive function is just as simple as other definitions:
|
|
19 |
*}
|
|
20 |
|
|
21 |
fun fib :: "nat \<Rightarrow> nat"
|
|
22 |
where
|
|
23 |
"fib 0 = 1"
|
|
24 |
| "fib (Suc 0) = 1"
|
|
25 |
| "fib (Suc (Suc n)) = fib n + fib (Suc n)"
|
|
26 |
|
|
27 |
(*<*)termination by lexicographic_order(*>*)
|
|
28 |
|
|
29 |
text {*
|
|
30 |
The function always terminates, since the argument of gets smaller in every
|
|
31 |
recursive call. Termination is an
|
|
32 |
important requirement, since it prevents inconsistencies: From
|
|
33 |
the "definition" @{text "f(n) = f(n) + 1"} we could prove
|
|
34 |
@{text "0 = 1"} by subtracting @{text "f(n)"} on both sides.
|
|
35 |
|
|
36 |
Isabelle tries to prove termination automatically when a function is
|
|
37 |
defined. We will later look at cases where this fails and see what to
|
|
38 |
do then.
|
|
39 |
*}
|
|
40 |
|
|
41 |
subsection {* Pattern matching *}
|
|
42 |
|
|
43 |
text {*
|
|
44 |
Like in functional programming, functions can be defined by pattern
|
|
45 |
matching. At the moment we will only consider \emph{datatype
|
|
46 |
patterns}, which only consist of datatype constructors and
|
|
47 |
variables.
|
|
48 |
|
|
49 |
If patterns overlap, the order of the equations is taken into
|
|
50 |
account. The following function inserts a fixed element between any
|
|
51 |
two elements of a list:
|
|
52 |
*}
|
|
53 |
|
|
54 |
fun sep :: "'a \<Rightarrow> 'a list \<Rightarrow> 'a list"
|
|
55 |
where
|
|
56 |
"sep a (x#y#xs) = x # a # sep a (y # xs)"
|
|
57 |
| "sep a xs = xs"
|
|
58 |
(*<*)termination by lexicographic_order(*>*)
|
|
59 |
|
|
60 |
text {*
|
|
61 |
Overlapping patterns are interpreted as "increments" to what is
|
|
62 |
already there: The second equation is only meant for the cases where
|
|
63 |
the first one does not match. Consequently, Isabelle replaces it
|
|
64 |
internally by the remaining cases, making the patterns disjoint.
|
|
65 |
This results in the equations @{thm [display] sep.simps[no_vars]}
|
|
66 |
*}
|
|
67 |
|
|
68 |
|
|
69 |
|
|
70 |
|
|
71 |
|
|
72 |
|
|
73 |
|
|
74 |
text {*
|
|
75 |
The equations from function definitions are automatically used in
|
|
76 |
simplification:
|
|
77 |
*}
|
|
78 |
|
|
79 |
lemma "fib (Suc (Suc 0)) = (Suc (Suc 0))"
|
|
80 |
by simp
|
|
81 |
|
|
82 |
|
|
83 |
|
|
84 |
|
|
85 |
subsection {* Induction *}
|
|
86 |
|
|
87 |
text {* FIXME *}
|
|
88 |
|
|
89 |
|
|
90 |
section {* If it does not work *}
|
|
91 |
|
|
92 |
text {*
|
|
93 |
Up to now, we were using the \cmd{fun} command, which provides a
|
|
94 |
convenient shorthand notation for simple function definitions. In
|
|
95 |
this mode, Isabelle tries to solve all the necessary proof obligations
|
|
96 |
automatically. If a proof does not go through, the definition is
|
|
97 |
rejected. This can mean that the definition is indeed faulty, like,
|
|
98 |
or that the default proof procedures are not smart enough (or
|
|
99 |
rather: not designed) to handle the specific definition.
|
|
100 |
.
|
|
101 |
By expanding the abbreviation to the full \cmd{function} command, the
|
|
102 |
proof obligations become visible and can be analyzed and solved manually.
|
|
103 |
*}
|
|
104 |
|
|
105 |
(*<*)types "\<tau>" = "nat \<Rightarrow> nat"
|
|
106 |
(*>*)
|
|
107 |
fun f :: "\<tau>"
|
|
108 |
where
|
|
109 |
(*<*)"f x = x" (*>*)text {* \vspace{-0.8cm}\emph{equations} *}
|
|
110 |
|
|
111 |
text {* \noindent abbreviates *}
|
|
112 |
|
|
113 |
function (sequential) fo :: "\<tau>"
|
|
114 |
where
|
|
115 |
(*<*)"fo x = x" (*>*)txt {* \vspace{-0.8cm}\emph{equations} *}
|
|
116 |
by pat_completeness auto
|
|
117 |
termination by lexicographic_order
|
|
118 |
|
|
119 |
text {*
|
|
120 |
Some declarations and proofs have now become explicit:
|
|
121 |
|
|
122 |
\begin{enumerate}
|
|
123 |
\item The "sequential" option enables the preprocessing of
|
|
124 |
pattern overlaps we already saw. Without this option, the equations
|
|
125 |
must already be disjoint and complete. The automatic completion only
|
|
126 |
works with datatype patterns.
|
|
127 |
|
|
128 |
\item A function definition now produces a proof obligation which
|
|
129 |
expresses completeness and compatibility of patterns (We talk about
|
|
130 |
this later). The combination of the methods {\tt pat\_completeness} and
|
|
131 |
{\tt auto} is used to solve this proof obligation.
|
|
132 |
|
|
133 |
\item A termination proof follows the definition, started by the
|
|
134 |
\cmd{termination} command. The {\tt lexicographic\_order} method can prove termination of a
|
|
135 |
certain class of functions by searching for a suitable lexicographic combination of size
|
|
136 |
measures.
|
|
137 |
\end{enumerate}
|
|
138 |
*}
|
|
139 |
|
|
140 |
|
|
141 |
section {* Proving termination *}
|
|
142 |
|
|
143 |
text {*
|
|
144 |
Consider the following function, which sums up natural numbers up to
|
|
145 |
@{text "N"}, using a counter @{text "i"}
|
|
146 |
*}
|
|
147 |
|
|
148 |
function sum :: "nat \<Rightarrow> nat \<Rightarrow> nat"
|
|
149 |
where
|
|
150 |
"sum i N = (if i > N then 0 else i + sum (Suc i) N)"
|
|
151 |
by pat_completeness auto
|
|
152 |
|
|
153 |
text {*
|
|
154 |
The {\tt lexicographic\_order} method fails on this example, because none of the
|
|
155 |
arguments decreases in the recursive call.
|
|
156 |
|
|
157 |
A more general method for termination proofs is to supply a wellfounded
|
|
158 |
relation on the argument type, and to show that the argument
|
|
159 |
decreases in every recursive call.
|
|
160 |
|
|
161 |
The termination argument for @{text "sum"} is based on the fact that
|
|
162 |
the \emph{difference} between @{text "i"} and @{text "N"} gets
|
|
163 |
smaller in every step, and that the recursion stops when @{text "i"}
|
|
164 |
is greater then @{text "n"}. Phrased differently, the expression
|
|
165 |
@{text "N + 1 - i"} decreases in every recursive call.
|
|
166 |
|
|
167 |
We can use this expression as a measure function to construct a
|
|
168 |
wellfounded relation, which is a proof of termination:
|
|
169 |
*}
|
|
170 |
|
|
171 |
termination
|
|
172 |
by (auto_term "measure (\<lambda>(i,N). N + 1 - i)")
|
|
173 |
|
|
174 |
text {*
|
|
175 |
Note that the two (curried) function arguments appear as a pair in
|
|
176 |
the measure function. The \cmd{function} command packs together curried
|
|
177 |
arguments into a tuple to simplify its internal operation. Hence,
|
|
178 |
measure functions and termination relations always work on the
|
|
179 |
tupled type.
|
|
180 |
|
|
181 |
Let us complicate the function a little, by adding some more recursive calls:
|
|
182 |
*}
|
|
183 |
|
|
184 |
function foo :: "nat \<Rightarrow> nat \<Rightarrow> nat"
|
|
185 |
where
|
|
186 |
"foo i N = (if i > N
|
|
187 |
then (if N = 0 then 0 else foo 0 (N - 1))
|
|
188 |
else i + foo (Suc i) N)"
|
|
189 |
by pat_completeness auto
|
|
190 |
|
|
191 |
text {*
|
|
192 |
When @{text "i"} has reached @{text "N"}, it starts at zero again
|
|
193 |
and @{text "N"} is decremented.
|
|
194 |
This corresponds to a nested
|
|
195 |
loop where one index counts up and the other down. Termination can
|
|
196 |
be proved using a lexicographic combination of two measures, namely
|
|
197 |
the value of @{text "N"} and the above difference. The @{text "measures"}
|
|
198 |
combinator generalizes @{text "measure"} by taking a list of measure functions.
|
|
199 |
*}
|
|
200 |
|
|
201 |
termination
|
|
202 |
by (auto_term "measures [\<lambda>(i, N). N, \<lambda>(i,N). N + 1 - i]")
|
|
203 |
|
|
204 |
|
|
205 |
section {* Mutual Recursion *}
|
|
206 |
|
|
207 |
text {*
|
|
208 |
If two or more functions call one another mutually, they have to be defined
|
|
209 |
in one step. The simplest example are probably @{text "even"} and @{text "odd"}:
|
|
210 |
*}
|
|
211 |
(*<*)hide const even odd(*>*)
|
|
212 |
|
|
213 |
function even odd :: "nat \<Rightarrow> bool"
|
|
214 |
where
|
|
215 |
"even 0 = True"
|
|
216 |
| "odd 0 = False"
|
|
217 |
| "even (Suc n) = odd n"
|
|
218 |
| "odd (Suc n) = even n"
|
|
219 |
by pat_completeness auto
|
|
220 |
|
|
221 |
text {*
|
|
222 |
To solve the problem of mutual dependencies, Isabelle internally
|
|
223 |
creates a single function operating on the sum
|
|
224 |
type. Then the original functions are defined as
|
|
225 |
projections. Consequently, termination has to be proved
|
|
226 |
simultaneously for both functions, by specifying a measure on the
|
|
227 |
sum type:
|
|
228 |
*}
|
|
229 |
|
|
230 |
termination
|
|
231 |
by (auto_term "measure (sum_case (%n. n) (%n. n))")
|
|
232 |
|
|
233 |
|
|
234 |
|
|
235 |
(* FIXME: Mutual induction *)
|
|
236 |
|
|
237 |
|
|
238 |
|
|
239 |
section {* Nested recursion *}
|
|
240 |
|
|
241 |
text {* FIXME *}
|
|
242 |
|
|
243 |
section {* More general patterns *}
|
|
244 |
|
|
245 |
text {* FIXME *}
|
|
246 |
|
|
247 |
|
|
248 |
|
|
249 |
|
|
250 |
end
|