10187
|
1 |
(*<*)theory WFrec = Main:(*>*)
|
|
2 |
|
|
3 |
text{*\noindent
|
|
4 |
So far, all recursive definitions where shown to terminate via measure
|
|
5 |
functions. Sometimes this can be quite inconvenient or even
|
|
6 |
impossible. Fortunately, \isacommand{recdef} supports much more
|
|
7 |
general definitions. For example, termination of Ackermann's function
|
|
8 |
can be shown by means of the lexicographic product @{text"<*lex*>"}:
|
|
9 |
*}
|
|
10 |
|
|
11 |
consts ack :: "nat\<times>nat \<Rightarrow> nat";
|
|
12 |
recdef ack "measure(\<lambda>m. m) <*lex*> measure(\<lambda>n. n)"
|
|
13 |
"ack(0,n) = Suc n"
|
|
14 |
"ack(Suc m,0) = ack(m, 1)"
|
|
15 |
"ack(Suc m,Suc n) = ack(m,ack(Suc m,n))";
|
|
16 |
|
|
17 |
text{*\noindent
|
|
18 |
The lexicographic product decreases if either its first component
|
|
19 |
decreases (as in the second equation and in the outer call in the
|
|
20 |
third equation) or its first component stays the same and the second
|
|
21 |
component decreases (as in the inner call in the third equation).
|
|
22 |
|
|
23 |
In general, \isacommand{recdef} supports termination proofs based on
|
10241
|
24 |
arbitrary \emph{well-founded relations}, i.e.\ \emph{well-founded
|
|
25 |
recursion}\indexbold{recursion!well-founded}\index{well-founded
|
|
26 |
recursion|see{recursion, well-founded}}. A relation $<$ is
|
|
27 |
\bfindex{well-founded} if it has no infinite descending chain $\cdots <
|
10187
|
28 |
a@2 < a@1 < a@0$. Clearly, a function definition is total iff the set
|
10189
|
29 |
of all pairs $(r,l)$, where $l$ is the argument on the left-hand side
|
|
30 |
of an equation and $r$ the argument of some recursive call on the
|
10241
|
31 |
corresponding right-hand side, induces a well-founded relation. For a
|
|
32 |
systematic account of termination proofs via well-founded relations
|
10189
|
33 |
see, for example, \cite{Baader-Nipkow}. The HOL library formalizes
|
10241
|
34 |
some of the theory of well-founded relations. For example
|
10189
|
35 |
@{prop"wf r"}\index{*wf|bold} means that relation @{term[show_types]"r::('a*'a)set"} is
|
10241
|
36 |
well-founded.
|
10187
|
37 |
|
|
38 |
Each \isacommand{recdef} definition should be accompanied (after the
|
10241
|
39 |
name of the function) by a well-founded relation on the argument type
|
10190
|
40 |
of the function. For example, \isaindexbold{measure} is defined by
|
10187
|
41 |
@{prop[display]"measure(f::'a \<Rightarrow> nat) \<equiv> {(y,x). f y < f x}"}
|
10241
|
42 |
and it has been proved that @{term"measure f"} is always well-founded.
|
10187
|
43 |
|
|
44 |
In addition to @{term measure}, the library provides
|
10241
|
45 |
a number of further constructions for obtaining well-founded relations.
|
10189
|
46 |
Above we have already met @{text"<*lex*>"} of type
|
|
47 |
@{typ[display,source]"('a \<times> 'a)set \<Rightarrow> ('b \<times> 'b)set \<Rightarrow> (('a \<times> 'b) \<times> ('a \<times> 'b))set"}
|
|
48 |
Of course the lexicographic product can also be interated, as in the following
|
|
49 |
function definition:
|
|
50 |
*}
|
10187
|
51 |
|
10189
|
52 |
consts contrived :: "nat \<times> nat \<times> nat \<Rightarrow> nat"
|
|
53 |
recdef contrived
|
|
54 |
"measure(\<lambda>i. i) <*lex*> measure(\<lambda>j. j) <*lex*> measure(\<lambda>k. k)"
|
|
55 |
"contrived(i,j,Suc k) = contrived(i,j,k)"
|
|
56 |
"contrived(i,Suc j,0) = contrived(i,j,j)"
|
|
57 |
"contrived(Suc i,0,0) = contrived(i,i,i)"
|
|
58 |
"contrived(0,0,0) = 0"
|
|
59 |
|
|
60 |
text{*
|
|
61 |
Lexicographic products of measure functions already go a long way. A
|
|
62 |
further useful construction is the embedding of some type in an
|
10241
|
63 |
existing well-founded relation via the inverse image of a function:
|
10189
|
64 |
@{thm[display,show_types]inv_image_def[no_vars]}
|
|
65 |
\begin{sloppypar}
|
|
66 |
\noindent
|
|
67 |
For example, @{term measure} is actually defined as @{term"inv_mage less_than"}, where
|
|
68 |
@{term less_than} of type @{typ"(nat \<times> nat)set"} is the less-than relation on type @{typ nat}
|
|
69 |
(as opposed to @{term"op <"}, which is of type @{typ"nat \<Rightarrow> nat \<Rightarrow> bool"}).
|
|
70 |
\end{sloppypar}
|
|
71 |
|
|
72 |
%Finally there is also {finite_psubset} the proper subset relation on finite sets
|
|
73 |
|
|
74 |
All the above constructions are known to \isacommand{recdef}. Thus you
|
10241
|
75 |
will never have to prove well-foundedness of any relation composed
|
10189
|
76 |
solely of these building blocks. But of course the proof of
|
|
77 |
termination of your function definition, i.e.\ that the arguments
|
|
78 |
decrease with every recursive call, may still require you to provide
|
|
79 |
additional lemmas.
|
|
80 |
|
10241
|
81 |
It is also possible to use your own well-founded relations with \isacommand{recdef}.
|
10189
|
82 |
Here is a simplistic example:
|
10187
|
83 |
*}
|
10189
|
84 |
|
|
85 |
consts f :: "nat \<Rightarrow> nat"
|
|
86 |
recdef f "id(less_than)"
|
|
87 |
"f 0 = 0"
|
|
88 |
"f (Suc n) = f n"
|
|
89 |
|
|
90 |
text{*
|
|
91 |
Since \isacommand{recdef} is not prepared for @{term id}, the identity
|
|
92 |
function, this leads to the complaint that it could not prove
|
10241
|
93 |
@{prop"wf (id less_than)"}, the well-foundedness of @{term"id
|
|
94 |
less_than"}. We should first have proved that @{term id} preserves well-foundedness
|
10189
|
95 |
*}
|
|
96 |
|
|
97 |
lemma wf_id: "wf r \<Longrightarrow> wf(id r)"
|
|
98 |
by simp;
|
|
99 |
|
|
100 |
text{*\noindent
|
|
101 |
and should have added the following hint to our above definition:
|
|
102 |
*}
|
|
103 |
(*<*)
|
|
104 |
consts g :: "nat \<Rightarrow> nat"
|
|
105 |
recdef g "id(less_than)"
|
|
106 |
"g 0 = 0"
|
|
107 |
"g (Suc n) = g n"
|
|
108 |
(*>*)
|
|
109 |
(hints recdef_wf add: wf_id)
|
10187
|
110 |
(*<*)end(*>*) |