doc-src/TutorialI/Recdef/termination.thy
author nipkow
Tue, 07 May 2002 15:03:50 +0200
changeset 13111 2d6782e71702
parent 12489 c92e38c3cbaa
child 15270 8b3f707a78a7
permissions -rw-r--r--
*** empty log message ***
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@11458
    16
Isabelle may fail to prove the termination condition for some
nipkow@12489
    17
recursive call.  Let us try to define Quicksort:*}
nipkow@8745
    18
nipkow@12489
    19
consts qs :: "nat list \<Rightarrow> nat list"
nipkow@12489
    20
recdef(*<*)(permissive)(*>*) qs "measure length"
nipkow@12489
    21
 "qs [] = []"
nipkow@12489
    22
 "qs(x#xs) = qs(filter (\<lambda>y. y\<le>x) xs) @ [x] @ qs(filter (\<lambda>y. x<y) xs)"
nipkow@8745
    23
nipkow@12489
    24
text{*\noindent where @{term filter} is predefined and @{term"filter P xs"}
nipkow@12489
    25
is the list of elements of @{term xs} satisfying @{term P}.
nipkow@12489
    26
This definition of @{term qs} fails, and Isabelle prints an error message
nipkow@12489
    27
showing you what it was unable to prove:
nipkow@12489
    28
@{text[display]"length (filter ... xs) < Suc (length xs)"}
nipkow@12489
    29
We can either prove this as a separate lemma, or try to figure out which
nipkow@12489
    30
existing lemmas may help. We opt for the second alternative. The theory of
nipkow@12489
    31
lists contains the simplification rule @{thm length_filter[no_vars]},
nipkow@12489
    32
which is already
nipkow@12489
    33
close to what we need, except that we still need to turn \mbox{@{text"< Suc"}}
nipkow@12489
    34
into
nipkow@12489
    35
@{text"\<le>"} for the simplification rule to apply. Lemma
nipkow@12489
    36
@{thm[source]less_Suc_eq_le} does just that: @{thm less_Suc_eq_le[no_vars]}.
nipkow@8745
    37
nipkow@12489
    38
Now we retry the above definition but supply the lemma(s) just found (or
nipkow@12489
    39
proved). Because \isacommand{recdef}'s termination prover involves
nipkow@12489
    40
simplification, we include in our second attempt a hint: the
nipkow@12489
    41
\attrdx{recdef_simp} attribute says to use @{thm[source]less_Suc_eq_le} as a
nipkow@13111
    42
simplification rule.\cmmdx{hints}  *}
nipkow@8745
    43
nipkow@12489
    44
(*<*)global consts qs :: "nat list \<Rightarrow> nat list" (*>*)
nipkow@12489
    45
recdef qs "measure length"
nipkow@12489
    46
 "qs [] = []"
nipkow@12489
    47
 "qs(x#xs) = qs(filter (\<lambda>y. y\<le>x) xs) @ [x] @ qs(filter (\<lambda>y. x<y) xs)"
nipkow@12489
    48
(hints recdef_simp: less_Suc_eq_le)
nipkow@12489
    49
(*<*)local(*>*)
nipkow@12489
    50
text{*\noindent
nipkow@12489
    51
This time everything works fine. Now @{thm[source]qs.simps} contains precisely
nipkow@12489
    52
the stated recursion equations for @{text qs} and they have become
nipkow@12489
    53
simplification rules.
nipkow@12489
    54
Thus we can automatically prove results such as this one:
nipkow@8745
    55
*}
nipkow@8745
    56
nipkow@12489
    57
theorem "qs[2,3,0] = qs[3,0,2]"
wenzelm@11626
    58
apply(simp)
nipkow@10171
    59
done
nipkow@8745
    60
nipkow@8745
    61
text{*\noindent
nipkow@8745
    62
More exciting theorems require induction, which is discussed below.
nipkow@8745
    63
nipkow@12489
    64
If the termination proof requires a lemma that is of general use, you can
nipkow@9933
    65
turn it permanently into a simplification rule, in which case the above
nipkow@12489
    66
\isacommand{hint} is not necessary. But in the case of
nipkow@12489
    67
@{thm[source]less_Suc_eq_le} this would be of dubious value.
nipkow@8745
    68
*}
nipkow@8745
    69
(*<*)
nipkow@8745
    70
end
nipkow@8745
    71
(*>*)