doc-src/TutorialI/Recdef/termination.thy
author nipkow
Fri, 15 Sep 2000 19:59:05 +0200
changeset 9992 4281ccea43f0
parent 9933 9feb1e0c4cb3
child 10171 59d6633835fa
permissions -rw-r--r--
*** empty log message ***
     1 (*<*)
     2 theory termination = examples:
     3 (*>*)
     4 
     5 text{*
     6 When a function is defined via \isacommand{recdef}, Isabelle tries to prove
     7 its termination with the help of the user-supplied measure.  All of the above
     8 examples are simple enough that Isabelle can prove automatically that the
     9 measure of the argument goes down in each recursive call. As a result,
    10 $f$@{text".simps"} will contain the defining equations (or variants derived
    11 from them) as theorems. For example, look (via \isacommand{thm}) at
    12 @{thm[source]sep.simps} and @{thm[source]sep1.simps} to see that they define
    13 the same function. What is more, those equations are automatically declared as
    14 simplification rules.
    15 
    16 In general, Isabelle may not be able to prove all termination condition
    17 (there is one for each recursive call) automatically. For example,
    18 termination of the following artificial function
    19 *}
    20 
    21 consts f :: "nat\<times>nat \<Rightarrow> nat";
    22 recdef f "measure(\<lambda>(x,y). x-y)"
    23   "f(x,y) = (if x \<le> y then x else f(x,y+1))";
    24 
    25 text{*\noindent
    26 is not proved automatically (although maybe it should be). Isabelle prints a
    27 kind of error message showing you what it was unable to prove. You will then
    28 have to prove it as a separate lemma before you attempt the definition
    29 of your function once more. In our case the required lemma is the obvious one:
    30 *}
    31 
    32 lemma termi_lem: "\<not> x \<le> y \<Longrightarrow> x - Suc y < x - y";
    33 
    34 txt{*\noindent
    35 It was not proved automatically because of the special nature of @{text"-"}
    36 on @{typ"nat"}. This requires more arithmetic than is tried by default:
    37 *}
    38 
    39 by(arith);
    40 
    41 text{*\noindent
    42 Because \isacommand{recdef}'s termination prover involves simplification,
    43 we include with our second attempt the hint to use @{thm[source]termi_lem} as
    44 a simplification rule:
    45 *}
    46 
    47 consts g :: "nat\<times>nat \<Rightarrow> nat";
    48 recdef g "measure(\<lambda>(x,y). x-y)"
    49   "g(x,y) = (if x \<le> y then x else g(x,y+1))"
    50 (hints recdef_simp: termi_lem)
    51 
    52 text{*\noindent
    53 This time everything works fine. Now @{thm[source]g.simps} contains precisely
    54 the stated recursion equation for @{term"g"} and they are simplification
    55 rules. Thus we can automatically prove
    56 *}
    57 
    58 theorem "g(1,0) = g(1,1)";
    59 by(simp);
    60 
    61 text{*\noindent
    62 More exciting theorems require induction, which is discussed below.
    63 
    64 If the termination proof requires a new lemma that is of general use, you can
    65 turn it permanently into a simplification rule, in which case the above
    66 \isacommand{hint} is not necessary. But our @{thm[source]termi_lem} is not
    67 sufficiently general to warrant this distinction.
    68 
    69 The attentive reader may wonder why we chose to call our function @{term"g"}
    70 rather than @{term"f"} the second time around. The reason is that, despite
    71 the failed termination proof, the definition of @{term"f"} did not
    72 fail, and thus we could not define it a second time. However, all theorems
    73 about @{term"f"}, for example @{thm[source]f.simps}, carry as a precondition
    74 the unproved termination condition. Moreover, the theorems
    75 @{thm[source]f.simps} are not simplification rules. However, this mechanism
    76 allows a delayed proof of termination: instead of proving
    77 @{thm[source]termi_lem} up front, we could prove 
    78 it later on and then use it to remove the preconditions from the theorems
    79 about @{term"f"}. In most cases this is more cumbersome than proving things
    80 up front.
    81 %FIXME, with one exception: nested recursion.
    82 
    83 Although all the above examples employ measure functions, \isacommand{recdef}
    84 allows arbitrary wellfounded relations. For example, termination of
    85 Ackermann's function requires the lexicographic product @{text"<*lex*>"}:
    86 *}
    87 
    88 consts ack :: "nat\<times>nat \<Rightarrow> nat";
    89 recdef ack "measure(\<lambda>m. m) <*lex*> measure(\<lambda>n. n)"
    90   "ack(0,n)         = Suc n"
    91   "ack(Suc m,0)     = ack(m, 1)"
    92   "ack(Suc m,Suc n) = ack(m,ack(Suc m,n))";
    93 
    94 text{*\noindent
    95 For details see the manual~\cite{isabelle-HOL} and the examples in the
    96 library.
    97 *}
    98 
    99 (*<*)
   100 end
   101 (*>*)