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