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