doc-src/TutorialI/Recdef/termination.thy
author wenzelm
Fri, 28 Sep 2001 20:08:28 +0200
changeset 11636 0bec857c9871
parent 11626 0dbfb578bf75
child 12332 aea72a834c85
permissions -rw-r--r--
oops;
     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
    24 Isabelle prints a
    25 \REMARK{error or warning?  change this part?  rename g to f?}
    26 message showing you what it was unable to prove. You will then
    27 have to prove it as a separate lemma before you attempt the definition
    28 of your function once more. In our case the required lemma is the obvious one:
    29 *}
    30 
    31 lemma termi_lem: "\<not> x \<le> y \<Longrightarrow> x - Suc y < x - y"
    32 
    33 txt{*\noindent
    34 It was not proved automatically because of the awkward behaviour of subtraction
    35 on type @{typ"nat"}. This requires more arithmetic than is tried by default:
    36 *}
    37 
    38 apply(arith)
    39 done
    40 
    41 text{*\noindent
    42 Because \isacommand{recdef}'s termination prover involves simplification,
    43 we include in our second attempt a hint: the \attrdx{recdef_simp} attribute 
    44 says to use @{thm[source]termi_lem} as
    45 a simplification rule.  
    46 *}
    47 
    48 consts g :: "nat\<times>nat \<Rightarrow> nat"
    49 recdef g "measure(\<lambda>(x,y). x-y)"
    50   "g(x,y) = (if x \<le> y then x else g(x,y+1))"
    51 (hints recdef_simp: termi_lem)
    52 
    53 text{*\noindent
    54 This time everything works fine. Now @{thm[source]g.simps} contains precisely
    55 the stated recursion equation for @{term g}, which has been stored as a
    56 simplification rule.  Thus we can automatically prove results such as this one:
    57 *}
    58 
    59 theorem "g(1,0) = g(1,1)"
    60 apply(simp)
    61 done
    62 
    63 text{*\noindent
    64 More exciting theorems require induction, which is discussed below.
    65 
    66 If the termination proof requires a new lemma that is of general use, you can
    67 turn it permanently into a simplification rule, in which case the above
    68 \isacommand{hint} is not necessary. But our @{thm[source]termi_lem} is not
    69 sufficiently general to warrant this distinction.
    70 
    71 The attentive reader may wonder why we chose to call our function @{term g}
    72 rather than @{term f} the second time around. The reason is that, despite
    73 the failed termination proof, the definition of @{term f} did not
    74 fail, and thus we could not define it a second time. However, all theorems
    75 about @{term f}, for example @{thm[source]f.simps}, carry as a precondition
    76 the unproved termination condition. Moreover, the theorems
    77 @{thm[source]f.simps} are not stored as simplification rules. 
    78 However, this mechanism
    79 allows a delayed proof of termination: instead of proving
    80 @{thm[source]termi_lem} up front, we could prove 
    81 it later on and then use it to remove the preconditions from the theorems
    82 about @{term f}. In most cases this is more cumbersome than proving things
    83 up front.
    84 \REMARK{FIXME, with one exception: nested recursion.}
    85 *}
    86 
    87 (*<*)
    88 end
    89 (*>*)