author | paulson <lp15@cam.ac.uk> |
Fri, 30 Dec 2022 17:48:41 +0000 | |
changeset 76832 | ab08604729a2 |
parent 76304 | e5162a8baa24 |
child 77569 | a8fa53c086a4 |
permissions | -rw-r--r-- |
72029 | 1 |
(* Title: HOL/Examples/Ackermann.thy |
2 |
Author: Larry Paulson |
|
3 |
*) |
|
4 |
||
71930 | 5 |
section \<open>A Tail-Recursive, Stack-Based Ackermann's Function\<close> |
6 |
||
76302
8d2bf9ce5302
Added the multiset termination proof
paulson <lp15@cam.ac.uk>
parents:
75013
diff
changeset
|
7 |
theory Ackermann imports "HOL-Library.Multiset_Order" "HOL-Library.Product_Lexorder" |
71930 | 8 |
|
9 |
begin |
|
10 |
||
11 |
text\<open>This theory investigates a stack-based implementation of Ackermann's function. |
|
12 |
Let's recall the traditional definition, |
|
13 |
as modified by R{\'o}zsa P\'eter and Raphael Robinson.\<close> |
|
14 |
||
15 |
fun ack :: "[nat,nat] \<Rightarrow> nat" where |
|
16 |
"ack 0 n = Suc n" |
|
17 |
| "ack (Suc m) 0 = ack m 1" |
|
18 |
| "ack (Suc m) (Suc n) = ack m (ack (Suc m) n)" |
|
19 |
||
76302
8d2bf9ce5302
Added the multiset termination proof
paulson <lp15@cam.ac.uk>
parents:
75013
diff
changeset
|
20 |
subsection \<open>Example of proving termination by reasoning about the domain\<close> |
8d2bf9ce5302
Added the multiset termination proof
paulson <lp15@cam.ac.uk>
parents:
75013
diff
changeset
|
21 |
|
8d2bf9ce5302
Added the multiset termination proof
paulson <lp15@cam.ac.uk>
parents:
75013
diff
changeset
|
22 |
text\<open>The stack-based version uses lists.\<close> |
71930 | 23 |
|
24 |
function (domintros) ackloop :: "nat list \<Rightarrow> nat" where |
|
25 |
"ackloop (n # 0 # l) = ackloop (Suc n # l)" |
|
26 |
| "ackloop (0 # Suc m # l) = ackloop (1 # m # l)" |
|
27 |
| "ackloop (Suc n # Suc m # l) = ackloop (n # Suc m # m # l)" |
|
28 |
| "ackloop [m] = m" |
|
29 |
| "ackloop [] = 0" |
|
30 |
by pat_completeness auto |
|
31 |
||
32 |
text\<open> |
|
33 |
The key task is to prove termination. In the first recursive call, the head of the list gets bigger |
|
34 |
while the list gets shorter, suggesting that the length of the list should be the primary |
|
35 |
termination criterion. But in the third recursive call, the list gets longer. The idea of trying |
|
36 |
a multiset-based termination argument is frustrated by the second recursive call when m = 0: |
|
37 |
the list elements are simply permuted. |
|
38 |
||
39 |
Fortunately, the function definition package allows us to define a function and only later identify its domain of termination. |
|
40 |
Instead, it makes all the recursion equations conditional on satisfying |
|
41 |
the function's domain predicate. Here we shall eventually be able |
|
42 |
to show that the predicate is always satisfied.\<close> |
|
43 |
||
44 |
text\<open>@{thm [display] ackloop.domintros[no_vars]}\<close> |
|
45 |
declare ackloop.domintros [simp] |
|
46 |
||
47 |
text \<open>Termination is trivial if the length of the list is less then two. |
|
48 |
The following lemma is the key to proving termination for longer lists.\<close> |
|
73531
c89922715bf5
Cosmetic: no !! in the lemma statement
paulson <lp15@cam.ac.uk>
parents:
72029
diff
changeset
|
49 |
lemma "ackloop_dom (ack m n # l) \<Longrightarrow> ackloop_dom (n # m # l)" |
c89922715bf5
Cosmetic: no !! in the lemma statement
paulson <lp15@cam.ac.uk>
parents:
72029
diff
changeset
|
50 |
proof (induction m arbitrary: n l) |
71930 | 51 |
case 0 |
52 |
then show ?case |
|
53 |
by auto |
|
54 |
next |
|
55 |
case (Suc m) |
|
73531
c89922715bf5
Cosmetic: no !! in the lemma statement
paulson <lp15@cam.ac.uk>
parents:
72029
diff
changeset
|
56 |
show ?case |
c89922715bf5
Cosmetic: no !! in the lemma statement
paulson <lp15@cam.ac.uk>
parents:
72029
diff
changeset
|
57 |
using Suc.prems |
c89922715bf5
Cosmetic: no !! in the lemma statement
paulson <lp15@cam.ac.uk>
parents:
72029
diff
changeset
|
58 |
by (induction n arbitrary: l) (simp_all add: Suc) |
71930 | 59 |
qed |
60 |
||
61 |
text \<open>The proof above (which actually is unused) can be expressed concisely as follows.\<close> |
|
62 |
lemma ackloop_dom_longer: |
|
63 |
"ackloop_dom (ack m n # l) \<Longrightarrow> ackloop_dom (n # m # l)" |
|
64 |
by (induction m n arbitrary: l rule: ack.induct) auto |
|
65 |
||
66 |
text\<open>This function codifies what @{term ackloop} is designed to do. |
|
67 |
Proving the two functions equivalent also shows that @{term ackloop} can be used |
|
68 |
to compute Ackermann's function.\<close> |
|
69 |
fun acklist :: "nat list \<Rightarrow> nat" where |
|
70 |
"acklist (n#m#l) = acklist (ack m n # l)" |
|
71 |
| "acklist [m] = m" |
|
72 |
| "acklist [] = 0" |
|
73 |
||
74 |
text\<open>The induction rule for @{term acklist} is @{thm [display] acklist.induct[no_vars]}.\<close> |
|
75 |
||
76 |
lemma ackloop_dom: "ackloop_dom l" |
|
77 |
by (induction l rule: acklist.induct) (auto simp: ackloop_dom_longer) |
|
78 |
||
79 |
termination ackloop |
|
80 |
by (simp add: ackloop_dom) |
|
81 |
||
82 |
text\<open>This result is trivial even by inspection of the function definitions |
|
83 |
(which faithfully follow the definition of Ackermann's function). |
|
84 |
All that we needed was termination.\<close> |
|
85 |
lemma ackloop_acklist: "ackloop l = acklist l" |
|
86 |
by (induction l rule: ackloop.induct) auto |
|
87 |
||
88 |
theorem ack: "ack m n = ackloop [n,m]" |
|
89 |
by (simp add: ackloop_acklist) |
|
90 |
||
76302
8d2bf9ce5302
Added the multiset termination proof
paulson <lp15@cam.ac.uk>
parents:
75013
diff
changeset
|
91 |
subsection \<open>Example of proving termination using a multiset ordering\<close> |
8d2bf9ce5302
Added the multiset termination proof
paulson <lp15@cam.ac.uk>
parents:
75013
diff
changeset
|
92 |
|
8d2bf9ce5302
Added the multiset termination proof
paulson <lp15@cam.ac.uk>
parents:
75013
diff
changeset
|
93 |
text \<open>This termination proof uses the argument from |
8d2bf9ce5302
Added the multiset termination proof
paulson <lp15@cam.ac.uk>
parents:
75013
diff
changeset
|
94 |
Nachum Dershowitz and Zohar Manna. Proving termination with multiset orderings. |
8d2bf9ce5302
Added the multiset termination proof
paulson <lp15@cam.ac.uk>
parents:
75013
diff
changeset
|
95 |
Communications of the ACM 22 (8) 1979, 465–476.\<close> |
8d2bf9ce5302
Added the multiset termination proof
paulson <lp15@cam.ac.uk>
parents:
75013
diff
changeset
|
96 |
|
8d2bf9ce5302
Added the multiset termination proof
paulson <lp15@cam.ac.uk>
parents:
75013
diff
changeset
|
97 |
text\<open>Setting up the termination proof. Note that Dershowitz had @{term z} as a global variable. |
8d2bf9ce5302
Added the multiset termination proof
paulson <lp15@cam.ac.uk>
parents:
75013
diff
changeset
|
98 |
The top two stack elements are treated differently from the rest.\<close> |
8d2bf9ce5302
Added the multiset termination proof
paulson <lp15@cam.ac.uk>
parents:
75013
diff
changeset
|
99 |
|
8d2bf9ce5302
Added the multiset termination proof
paulson <lp15@cam.ac.uk>
parents:
75013
diff
changeset
|
100 |
fun ack_mset :: "nat list \<Rightarrow> (nat\<times>nat) multiset" where |
8d2bf9ce5302
Added the multiset termination proof
paulson <lp15@cam.ac.uk>
parents:
75013
diff
changeset
|
101 |
"ack_mset [] = {#}" |
8d2bf9ce5302
Added the multiset termination proof
paulson <lp15@cam.ac.uk>
parents:
75013
diff
changeset
|
102 |
| "ack_mset [x] = {#}" |
8d2bf9ce5302
Added the multiset termination proof
paulson <lp15@cam.ac.uk>
parents:
75013
diff
changeset
|
103 |
| "ack_mset (z#y#l) = mset ((y,z) # map (\<lambda>x. (Suc x, 0)) l)" |
8d2bf9ce5302
Added the multiset termination proof
paulson <lp15@cam.ac.uk>
parents:
75013
diff
changeset
|
104 |
|
76304 | 105 |
lemma case1: "ack_mset (Suc n # l) < add_mset (0,n) {# (Suc x, 0). x \<in># mset l #}" |
76302
8d2bf9ce5302
Added the multiset termination proof
paulson <lp15@cam.ac.uk>
parents:
75013
diff
changeset
|
106 |
proof (cases l) |
8d2bf9ce5302
Added the multiset termination proof
paulson <lp15@cam.ac.uk>
parents:
75013
diff
changeset
|
107 |
case (Cons m list) |
8d2bf9ce5302
Added the multiset termination proof
paulson <lp15@cam.ac.uk>
parents:
75013
diff
changeset
|
108 |
have "{#(m, Suc n)#} < {#(Suc m, 0)#}" |
8d2bf9ce5302
Added the multiset termination proof
paulson <lp15@cam.ac.uk>
parents:
75013
diff
changeset
|
109 |
by auto |
8d2bf9ce5302
Added the multiset termination proof
paulson <lp15@cam.ac.uk>
parents:
75013
diff
changeset
|
110 |
also have "\<dots> \<le> {#(Suc m, 0), (0,n)#}" |
8d2bf9ce5302
Added the multiset termination proof
paulson <lp15@cam.ac.uk>
parents:
75013
diff
changeset
|
111 |
by auto |
8d2bf9ce5302
Added the multiset termination proof
paulson <lp15@cam.ac.uk>
parents:
75013
diff
changeset
|
112 |
finally show ?thesis |
8d2bf9ce5302
Added the multiset termination proof
paulson <lp15@cam.ac.uk>
parents:
75013
diff
changeset
|
113 |
by (simp add: Cons) |
8d2bf9ce5302
Added the multiset termination proof
paulson <lp15@cam.ac.uk>
parents:
75013
diff
changeset
|
114 |
qed auto |
8d2bf9ce5302
Added the multiset termination proof
paulson <lp15@cam.ac.uk>
parents:
75013
diff
changeset
|
115 |
|
8d2bf9ce5302
Added the multiset termination proof
paulson <lp15@cam.ac.uk>
parents:
75013
diff
changeset
|
116 |
text\<open>The stack-based version again. We need a fresh copy because |
8d2bf9ce5302
Added the multiset termination proof
paulson <lp15@cam.ac.uk>
parents:
75013
diff
changeset
|
117 |
we've already proved the termination of @{term ackloop}.\<close> |
8d2bf9ce5302
Added the multiset termination proof
paulson <lp15@cam.ac.uk>
parents:
75013
diff
changeset
|
118 |
|
8d2bf9ce5302
Added the multiset termination proof
paulson <lp15@cam.ac.uk>
parents:
75013
diff
changeset
|
119 |
function Ackloop :: "nat list \<Rightarrow> nat" where |
8d2bf9ce5302
Added the multiset termination proof
paulson <lp15@cam.ac.uk>
parents:
75013
diff
changeset
|
120 |
"Ackloop (n # 0 # l) = Ackloop (Suc n # l)" |
8d2bf9ce5302
Added the multiset termination proof
paulson <lp15@cam.ac.uk>
parents:
75013
diff
changeset
|
121 |
| "Ackloop (0 # Suc m # l) = Ackloop (1 # m # l)" |
8d2bf9ce5302
Added the multiset termination proof
paulson <lp15@cam.ac.uk>
parents:
75013
diff
changeset
|
122 |
| "Ackloop (Suc n # Suc m # l) = Ackloop (n # Suc m # m # l)" |
8d2bf9ce5302
Added the multiset termination proof
paulson <lp15@cam.ac.uk>
parents:
75013
diff
changeset
|
123 |
| "Ackloop [m] = m" |
8d2bf9ce5302
Added the multiset termination proof
paulson <lp15@cam.ac.uk>
parents:
75013
diff
changeset
|
124 |
| "Ackloop [] = 0" |
8d2bf9ce5302
Added the multiset termination proof
paulson <lp15@cam.ac.uk>
parents:
75013
diff
changeset
|
125 |
by pat_completeness auto |
8d2bf9ce5302
Added the multiset termination proof
paulson <lp15@cam.ac.uk>
parents:
75013
diff
changeset
|
126 |
|
8d2bf9ce5302
Added the multiset termination proof
paulson <lp15@cam.ac.uk>
parents:
75013
diff
changeset
|
127 |
|
8d2bf9ce5302
Added the multiset termination proof
paulson <lp15@cam.ac.uk>
parents:
75013
diff
changeset
|
128 |
text \<open>In each recursive call, the function @{term ack_mset} decreases according to the multiset |
8d2bf9ce5302
Added the multiset termination proof
paulson <lp15@cam.ac.uk>
parents:
75013
diff
changeset
|
129 |
ordering.\<close> |
8d2bf9ce5302
Added the multiset termination proof
paulson <lp15@cam.ac.uk>
parents:
75013
diff
changeset
|
130 |
termination |
76304 | 131 |
by (relation "inv_image {(x,y). x<y} ack_mset") (auto simp: wf case1) |
76302
8d2bf9ce5302
Added the multiset termination proof
paulson <lp15@cam.ac.uk>
parents:
75013
diff
changeset
|
132 |
|
8d2bf9ce5302
Added the multiset termination proof
paulson <lp15@cam.ac.uk>
parents:
75013
diff
changeset
|
133 |
text \<open>Another shortcut compared with before: equivalence follows directly from this lemma.\<close> |
8d2bf9ce5302
Added the multiset termination proof
paulson <lp15@cam.ac.uk>
parents:
75013
diff
changeset
|
134 |
lemma Ackloop_ack: "Ackloop (n # m # l) = Ackloop (ack m n # l)" |
8d2bf9ce5302
Added the multiset termination proof
paulson <lp15@cam.ac.uk>
parents:
75013
diff
changeset
|
135 |
by (induction m n arbitrary: l rule: ack.induct) auto |
8d2bf9ce5302
Added the multiset termination proof
paulson <lp15@cam.ac.uk>
parents:
75013
diff
changeset
|
136 |
|
8d2bf9ce5302
Added the multiset termination proof
paulson <lp15@cam.ac.uk>
parents:
75013
diff
changeset
|
137 |
theorem "ack m n = Ackloop [n,m]" |
8d2bf9ce5302
Added the multiset termination proof
paulson <lp15@cam.ac.uk>
parents:
75013
diff
changeset
|
138 |
by (simp add: Ackloop_ack) |
8d2bf9ce5302
Added the multiset termination proof
paulson <lp15@cam.ac.uk>
parents:
75013
diff
changeset
|
139 |
|
71930 | 140 |
end |