| author | wenzelm | 
| Sat, 13 Nov 2021 17:26:35 +0100 | |
| changeset 74779 | 5fca489a6ac1 | 
| parent 73531 | c89922715bf5 | 
| child 75013 | ccf203c9b2db | 
| 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 | ||
| 7 | theory Ackermann imports Main | |
| 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 | ||
| 20 | text\<open>Here is the stack-based version, which uses lists.\<close> | |
| 21 | ||
| 22 | function (domintros) ackloop :: "nat list \<Rightarrow> nat" where | |
| 23 | "ackloop (n # 0 # l) = ackloop (Suc n # l)" | |
| 24 | | "ackloop (0 # Suc m # l) = ackloop (1 # m # l)" | |
| 25 | | "ackloop (Suc n # Suc m # l) = ackloop (n # Suc m # m # l)" | |
| 26 | | "ackloop [m] = m" | |
| 27 | | "ackloop [] = 0" | |
| 28 | by pat_completeness auto | |
| 29 | ||
| 30 | text\<open> | |
| 31 | The key task is to prove termination. In the first recursive call, the head of the list gets bigger | |
| 32 | while the list gets shorter, suggesting that the length of the list should be the primary | |
| 33 | termination criterion. But in the third recursive call, the list gets longer. The idea of trying | |
| 34 | a multiset-based termination argument is frustrated by the second recursive call when m = 0: | |
| 35 | the list elements are simply permuted. | |
| 36 | ||
| 37 | Fortunately, the function definition package allows us to define a function and only later identify its domain of termination. | |
| 38 | Instead, it makes all the recursion equations conditional on satisfying | |
| 39 | the function's domain predicate. Here we shall eventually be able | |
| 40 | to show that the predicate is always satisfied.\<close> | |
| 41 | ||
| 42 | text\<open>@{thm [display] ackloop.domintros[no_vars]}\<close>
 | |
| 43 | declare ackloop.domintros [simp] | |
| 44 | ||
| 45 | text \<open>Termination is trivial if the length of the list is less then two. | |
| 46 | 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: 
72029diff
changeset | 47 | 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: 
72029diff
changeset | 48 | proof (induction m arbitrary: n l) | 
| 71930 | 49 | case 0 | 
| 50 | then show ?case | |
| 51 | by auto | |
| 52 | next | |
| 53 | case (Suc m) | |
| 73531 
c89922715bf5
Cosmetic: no !! in the lemma statement
 paulson <lp15@cam.ac.uk> parents: 
72029diff
changeset | 54 | show ?case | 
| 
c89922715bf5
Cosmetic: no !! in the lemma statement
 paulson <lp15@cam.ac.uk> parents: 
72029diff
changeset | 55 | using Suc.prems | 
| 
c89922715bf5
Cosmetic: no !! in the lemma statement
 paulson <lp15@cam.ac.uk> parents: 
72029diff
changeset | 56 | by (induction n arbitrary: l) (simp_all add: Suc) | 
| 71930 | 57 | qed | 
| 58 | ||
| 59 | text \<open>The proof above (which actually is unused) can be expressed concisely as follows.\<close> | |
| 60 | lemma ackloop_dom_longer: | |
| 61 | "ackloop_dom (ack m n # l) \<Longrightarrow> ackloop_dom (n # m # l)" | |
| 62 | by (induction m n arbitrary: l rule: ack.induct) auto | |
| 63 | ||
| 64 | lemma "ackloop_dom (ack m n # l) \<Longrightarrow> ackloop_dom (n # m # l)" | |
| 65 | by (induction m n arbitrary: l rule: ack.induct) auto | |
| 66 | ||
| 67 | text\<open>This function codifies what @{term ackloop} is designed to do.
 | |
| 68 | Proving the two functions equivalent also shows that @{term ackloop} can be used
 | |
| 69 | to compute Ackermann's function.\<close> | |
| 70 | fun acklist :: "nat list \<Rightarrow> nat" where | |
| 71 | "acklist (n#m#l) = acklist (ack m n # l)" | |
| 72 | | "acklist [m] = m" | |
| 73 | | "acklist [] = 0" | |
| 74 | ||
| 75 | text\<open>The induction rule for @{term acklist} is @{thm [display] acklist.induct[no_vars]}.\<close>
 | |
| 76 | ||
| 77 | lemma ackloop_dom: "ackloop_dom l" | |
| 78 | by (induction l rule: acklist.induct) (auto simp: ackloop_dom_longer) | |
| 79 | ||
| 80 | termination ackloop | |
| 81 | by (simp add: ackloop_dom) | |
| 82 | ||
| 83 | text\<open>This result is trivial even by inspection of the function definitions | |
| 84 | (which faithfully follow the definition of Ackermann's function). | |
| 85 | All that we needed was termination.\<close> | |
| 86 | lemma ackloop_acklist: "ackloop l = acklist l" | |
| 87 | by (induction l rule: ackloop.induct) auto | |
| 88 | ||
| 89 | theorem ack: "ack m n = ackloop [n,m]" | |
| 90 | by (simp add: ackloop_acklist) | |
| 91 | ||
| 92 | end |