| 
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>
  | 
| 
 | 
    47  | 
lemma "\<And>n l. ackloop_dom (ack m n # l) \<Longrightarrow> ackloop_dom (n # m # l)"
  | 
| 
 | 
    48  | 
proof (induction m)
  | 
| 
 | 
    49  | 
  case 0
  | 
| 
 | 
    50  | 
  then show ?case
  | 
| 
 | 
    51  | 
    by auto
  | 
| 
 | 
    52  | 
next
  | 
| 
 | 
    53  | 
  case (Suc m)
  | 
| 
 | 
    54  | 
  note IH = Suc
  | 
| 
 | 
    55  | 
  have "\<And>l. ackloop_dom (ack (Suc m) n # l) \<Longrightarrow> ackloop_dom (n # Suc m # l)"
  | 
| 
 | 
    56  | 
  proof (induction n)
  | 
| 
 | 
    57  | 
    case 0
  | 
| 
 | 
    58  | 
    then show ?case
  | 
| 
 | 
    59  | 
      by (simp add: IH)
  | 
| 
 | 
    60  | 
  next
  | 
| 
 | 
    61  | 
    case (Suc n)
  | 
| 
 | 
    62  | 
    then show ?case
  | 
| 
 | 
    63  | 
      by (auto simp: IH)
  | 
| 
 | 
    64  | 
  qed
  | 
| 
 | 
    65  | 
  then show ?case
  | 
| 
 | 
    66  | 
    using Suc.prems by blast
  | 
| 
 | 
    67  | 
qed
  | 
| 
 | 
    68  | 
  | 
| 
 | 
    69  | 
text \<open>The proof above (which actually is unused) can be expressed concisely as follows.\<close>
  | 
| 
 | 
    70  | 
lemma ackloop_dom_longer:
  | 
| 
 | 
    71  | 
  "ackloop_dom (ack m n # l) \<Longrightarrow> ackloop_dom (n # m # l)"
  | 
| 
 | 
    72  | 
  by (induction m n arbitrary: l rule: ack.induct) auto
  | 
| 
 | 
    73  | 
  | 
| 
 | 
    74  | 
lemma "ackloop_dom (ack m n # l) \<Longrightarrow> ackloop_dom (n # m # l)"
  | 
| 
 | 
    75  | 
  by (induction m n arbitrary: l rule: ack.induct) auto
  | 
| 
 | 
    76  | 
  | 
| 
 | 
    77  | 
text\<open>This function codifies what @{term ackloop} is designed to do.
 | 
| 
 | 
    78  | 
Proving the two functions equivalent also shows that @{term ackloop} can be used
 | 
| 
 | 
    79  | 
to compute Ackermann's function.\<close>
  | 
| 
 | 
    80  | 
fun acklist :: "nat list \<Rightarrow> nat" where
  | 
| 
 | 
    81  | 
  "acklist (n#m#l) = acklist (ack m n # l)"
  | 
| 
 | 
    82  | 
| "acklist [m] = m"
  | 
| 
 | 
    83  | 
| "acklist [] =  0"
  | 
| 
 | 
    84  | 
  | 
| 
 | 
    85  | 
text\<open>The induction rule for @{term acklist} is @{thm [display] acklist.induct[no_vars]}.\<close>
 | 
| 
 | 
    86  | 
  | 
| 
 | 
    87  | 
lemma ackloop_dom: "ackloop_dom l"
  | 
| 
 | 
    88  | 
  by (induction l rule: acklist.induct) (auto simp: ackloop_dom_longer)
  | 
| 
 | 
    89  | 
  | 
| 
 | 
    90  | 
termination ackloop
  | 
| 
 | 
    91  | 
  by (simp add: ackloop_dom)
  | 
| 
 | 
    92  | 
  | 
| 
 | 
    93  | 
text\<open>This result is trivial even by inspection of the function definitions
  | 
| 
 | 
    94  | 
(which faithfully follow the definition of Ackermann's function).
  | 
| 
 | 
    95  | 
All that we needed was termination.\<close>
  | 
| 
 | 
    96  | 
lemma ackloop_acklist: "ackloop l = acklist l"
  | 
| 
 | 
    97  | 
  by (induction l rule: ackloop.induct) auto
  | 
| 
 | 
    98  | 
  | 
| 
 | 
    99  | 
theorem ack: "ack m n = ackloop [n,m]"
  | 
| 
 | 
   100  | 
  by (simp add: ackloop_acklist)
  | 
| 
 | 
   101  | 
  | 
| 
 | 
   102  | 
end
  |