doc-src/TutorialI/Recdef/termination.thy
author nipkow
Wed, 12 Dec 2001 09:04:20 +0100
changeset 12473 f41e477576b9
parent 12332 aea72a834c85
child 12489 c92e38c3cbaa
permissions -rw-r--r--
*** empty log message ***
     1 (*<*)
     2 theory termination = examples:
     3 (*>*)
     4 
     5 text{*
     6 When a function~$f$ is defined via \isacommand{recdef}, Isabelle tries to prove
     7 its termination with the help of the user-supplied measure.  Each of the examples
     8 above is simple enough that Isabelle can automatically prove that the
     9 argument's measure decreases 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 Isabelle may fail to prove the termination condition for some
    17 recursive call.  Let us try the following artificial function:*}
    18 
    19 consts f :: "nat\<times>nat \<Rightarrow> nat"
    20 recdef (*<*)(permissive)(*>*)f "measure(\<lambda>(x,y). x-y)"
    21   "f(x,y) = (if x \<le> y then x else f(x,y+1))"
    22 
    23 text{*\noindent This definition fails, and Isabelle prints an error message
    24 showing you what it was unable to prove. You will then have to prove it as a
    25 separate lemma before you attempt the definition of your function once
    26 more. In our case the required lemma is the obvious one: *}
    27 
    28 lemma termi_lem: "\<not> x \<le> y \<Longrightarrow> x - Suc y < x - y"
    29 
    30 txt{*\noindent
    31 It was not proved automatically because of the awkward behaviour of subtraction
    32 on type @{typ"nat"}. This requires more arithmetic than is tried by default:
    33 *}
    34 
    35 apply(arith)
    36 done
    37 
    38 text{*\noindent
    39 Because \isacommand{recdef}'s termination prover involves simplification,
    40 we include in our second attempt a hint: the \attrdx{recdef_simp} attribute 
    41 says to use @{thm[source]termi_lem} as a simplification rule.  
    42 *}
    43 
    44 (*<*)global consts f :: "nat\<times>nat \<Rightarrow> nat" (*>*)
    45 recdef f "measure(\<lambda>(x,y). x-y)"
    46   "f(x,y) = (if x \<le> y then x else f(x,y+1))"
    47 (hints recdef_simp: termi_lem)
    48 (*<*)local(*>*)
    49 text{*\noindent
    50 This time everything works fine. Now @{thm[source]f.simps} contains precisely
    51 the stated recursion equation for @{text f}, which has been turned into a
    52 simplification rule.  Thus we can automatically prove results such as this one:
    53 *}
    54 
    55 theorem "f(1,0) = f(1,1)"
    56 apply(simp)
    57 done
    58 
    59 text{*\noindent
    60 More exciting theorems require induction, which is discussed below.
    61 
    62 If the termination proof requires a new lemma that is of general use, you can
    63 turn it permanently into a simplification rule, in which case the above
    64 \isacommand{hint} is not necessary. But our @{thm[source]termi_lem} is not
    65 sufficiently general to warrant this distinction.
    66 *}
    67 (*<*)
    68 end
    69 (*>*)