71930
|
1 |
section \<open>A Tail-Recursive, Stack-Based Ackermann's Function\<close>
|
|
2 |
|
|
3 |
theory Ackermann imports Main
|
|
4 |
|
|
5 |
begin
|
|
6 |
|
|
7 |
text\<open>This theory investigates a stack-based implementation of Ackermann's function.
|
|
8 |
Let's recall the traditional definition,
|
|
9 |
as modified by R{\'o}zsa P\'eter and Raphael Robinson.\<close>
|
|
10 |
|
|
11 |
fun ack :: "[nat,nat] \<Rightarrow> nat" where
|
|
12 |
"ack 0 n = Suc n"
|
|
13 |
| "ack (Suc m) 0 = ack m 1"
|
|
14 |
| "ack (Suc m) (Suc n) = ack m (ack (Suc m) n)"
|
|
15 |
|
|
16 |
text\<open>Here is the stack-based version, which uses lists.\<close>
|
|
17 |
|
|
18 |
function (domintros) ackloop :: "nat list \<Rightarrow> nat" where
|
|
19 |
"ackloop (n # 0 # l) = ackloop (Suc n # l)"
|
|
20 |
| "ackloop (0 # Suc m # l) = ackloop (1 # m # l)"
|
|
21 |
| "ackloop (Suc n # Suc m # l) = ackloop (n # Suc m # m # l)"
|
|
22 |
| "ackloop [m] = m"
|
|
23 |
| "ackloop [] = 0"
|
|
24 |
by pat_completeness auto
|
|
25 |
|
|
26 |
text\<open>
|
|
27 |
The key task is to prove termination. In the first recursive call, the head of the list gets bigger
|
|
28 |
while the list gets shorter, suggesting that the length of the list should be the primary
|
|
29 |
termination criterion. But in the third recursive call, the list gets longer. The idea of trying
|
|
30 |
a multiset-based termination argument is frustrated by the second recursive call when m = 0:
|
|
31 |
the list elements are simply permuted.
|
|
32 |
|
|
33 |
Fortunately, the function definition package allows us to define a function and only later identify its domain of termination.
|
|
34 |
Instead, it makes all the recursion equations conditional on satisfying
|
|
35 |
the function's domain predicate. Here we shall eventually be able
|
|
36 |
to show that the predicate is always satisfied.\<close>
|
|
37 |
|
|
38 |
text\<open>@{thm [display] ackloop.domintros[no_vars]}\<close>
|
|
39 |
declare ackloop.domintros [simp]
|
|
40 |
|
|
41 |
text \<open>Termination is trivial if the length of the list is less then two.
|
|
42 |
The following lemma is the key to proving termination for longer lists.\<close>
|
|
43 |
lemma "\<And>n l. ackloop_dom (ack m n # l) \<Longrightarrow> ackloop_dom (n # m # l)"
|
|
44 |
proof (induction m)
|
|
45 |
case 0
|
|
46 |
then show ?case
|
|
47 |
by auto
|
|
48 |
next
|
|
49 |
case (Suc m)
|
|
50 |
note IH = Suc
|
|
51 |
have "\<And>l. ackloop_dom (ack (Suc m) n # l) \<Longrightarrow> ackloop_dom (n # Suc m # l)"
|
|
52 |
proof (induction n)
|
|
53 |
case 0
|
|
54 |
then show ?case
|
|
55 |
by (simp add: IH)
|
|
56 |
next
|
|
57 |
case (Suc n)
|
|
58 |
then show ?case
|
|
59 |
by (auto simp: IH)
|
|
60 |
qed
|
|
61 |
then show ?case
|
|
62 |
using Suc.prems by blast
|
|
63 |
qed
|
|
64 |
|
|
65 |
text \<open>The proof above (which actually is unused) can be expressed concisely as follows.\<close>
|
|
66 |
lemma ackloop_dom_longer:
|
|
67 |
"ackloop_dom (ack m n # l) \<Longrightarrow> ackloop_dom (n # m # l)"
|
|
68 |
by (induction m n arbitrary: l rule: ack.induct) auto
|
|
69 |
|
|
70 |
lemma "ackloop_dom (ack m n # l) \<Longrightarrow> ackloop_dom (n # m # l)"
|
|
71 |
by (induction m n arbitrary: l rule: ack.induct) auto
|
|
72 |
|
|
73 |
text\<open>This function codifies what @{term ackloop} is designed to do.
|
|
74 |
Proving the two functions equivalent also shows that @{term ackloop} can be used
|
|
75 |
to compute Ackermann's function.\<close>
|
|
76 |
fun acklist :: "nat list \<Rightarrow> nat" where
|
|
77 |
"acklist (n#m#l) = acklist (ack m n # l)"
|
|
78 |
| "acklist [m] = m"
|
|
79 |
| "acklist [] = 0"
|
|
80 |
|
|
81 |
text\<open>The induction rule for @{term acklist} is @{thm [display] acklist.induct[no_vars]}.\<close>
|
|
82 |
|
|
83 |
lemma ackloop_dom: "ackloop_dom l"
|
|
84 |
by (induction l rule: acklist.induct) (auto simp: ackloop_dom_longer)
|
|
85 |
|
|
86 |
termination ackloop
|
|
87 |
by (simp add: ackloop_dom)
|
|
88 |
|
|
89 |
text\<open>This result is trivial even by inspection of the function definitions
|
|
90 |
(which faithfully follow the definition of Ackermann's function).
|
|
91 |
All that we needed was termination.\<close>
|
|
92 |
lemma ackloop_acklist: "ackloop l = acklist l"
|
|
93 |
by (induction l rule: ackloop.induct) auto
|
|
94 |
|
|
95 |
theorem ack: "ack m n = ackloop [n,m]"
|
|
96 |
by (simp add: ackloop_acklist)
|
|
97 |
|
|
98 |
end
|