doc-src/TutorialI/Recdef/simplification.thy
author nipkow
Wed, 30 Aug 2000 18:09:20 +0200
changeset 9754 a123a64cadeb
parent 9541 d17c0b34d5c8
child 9792 bbefb6ce5cb2
permissions -rw-r--r--
*** empty log message ***
nipkow@8745
     1
(*<*)
nipkow@8745
     2
theory simplification = Main:;
nipkow@8745
     3
(*>*)
nipkow@8745
     4
nipkow@8745
     5
text{*
nipkow@8745
     6
Once we have succeeded in proving all termination conditions, the recursion
nipkow@8745
     7
equations become simplification rules, just as with
nipkow@8745
     8
\isacommand{primrec}. In most cases this works fine, but there is a subtle
nipkow@8745
     9
problem that must be mentioned: simplification may not
nipkow@9754
    10
terminate because of automatic splitting of @{name"if"}.
nipkow@8745
    11
Let us look at an example:
nipkow@8745
    12
*}
nipkow@8745
    13
nipkow@8745
    14
consts gcd :: "nat*nat \\<Rightarrow> nat";
nipkow@8745
    15
recdef gcd "measure (\\<lambda>(m,n).n)"
nipkow@8745
    16
  "gcd (m, n) = (if n=0 then m else gcd(n, m mod n))";
nipkow@8745
    17
nipkow@8745
    18
text{*\noindent
nipkow@8745
    19
According to the measure function, the second argument should decrease with
nipkow@8745
    20
each recursive call. The resulting termination condition
nipkow@9541
    21
\begin{quote}
nipkow@9541
    22
@{term[display]"n ~= 0 ==> m mod n < n"}
nipkow@9541
    23
\end{quote}
nipkow@8745
    24
is provded automatically because it is already present as a lemma in the
nipkow@8745
    25
arithmetic library. Thus the recursion equation becomes a simplification
nipkow@8745
    26
rule. Of course the equation is nonterminating if we are allowed to unfold
nipkow@9754
    27
the recursive call inside the @{name"if"} branch, which is why programming
nipkow@8745
    28
languages and our simplifier don't do that. Unfortunately the simplifier does
nipkow@9754
    29
something else which leads to the same problem: it splits @{name"if"}s if the
nipkow@9754
    30
condition simplifies to neither @{term"True"} nor @{term"False"}. For
nipkow@8745
    31
example, simplification reduces
nipkow@9541
    32
\begin{quote}
nipkow@9541
    33
@{term[display]"gcd(m,n) = k"}
nipkow@9541
    34
\end{quote}
nipkow@8745
    35
in one step to
nipkow@9541
    36
\begin{quote}
nipkow@9541
    37
@{term[display]"(if n=0 then m else gcd(n, m mod n)) = k"}
nipkow@9541
    38
\end{quote}
nipkow@8745
    39
where the condition cannot be reduced further, and splitting leads to
nipkow@9541
    40
\begin{quote}
nipkow@9541
    41
@{term[display]"(n=0 --> m=k) & (n ~= 0 --> gcd(n, m mod n)=k)"}
nipkow@9541
    42
\end{quote}
nipkow@9541
    43
Since the recursive call @{term"gcd(n, m mod n)"} is no longer protected by
nipkow@9754
    44
an @{name"if"}, it is unfolded again, which leads to an infinite chain of
nipkow@9541
    45
simplification steps. Fortunately, this problem can be avoided in many
nipkow@9541
    46
different ways.
nipkow@8745
    47
nipkow@9754
    48
The most radical solution is to disable the offending \@{name"split_if"} as
nipkow@9754
    49
shown in the section on case splits in \S\ref{sec:Simplification}.  However,
nipkow@9754
    50
we do not recommend this because it means you will often have to invoke the
nipkow@9754
    51
rule explicitly when @{name"if"} is involved.
nipkow@8745
    52
nipkow@8745
    53
If possible, the definition should be given by pattern matching on the left
nipkow@9754
    54
rather than @{name"if"} on the right. In the case of @{term"gcd"} the
nipkow@8745
    55
following alternative definition suggests itself:
nipkow@8745
    56
*}
nipkow@8745
    57
nipkow@8745
    58
consts gcd1 :: "nat*nat \\<Rightarrow> nat";
nipkow@8745
    59
recdef gcd1 "measure (\\<lambda>(m,n).n)"
nipkow@8745
    60
  "gcd1 (m, 0) = m"
nipkow@8745
    61
  "gcd1 (m, n) = gcd1(n, m mod n)";
nipkow@8745
    62
nipkow@8745
    63
nipkow@8745
    64
text{*\noindent
nipkow@8745
    65
Note that the order of equations is important and hides the side condition
nipkow@9754
    66
@{prop"n ~= 0"}. Unfortunately, in general the case distinction
nipkow@8745
    67
may not be expressible by pattern matching.
nipkow@8745
    68
nipkow@9754
    69
A very simple alternative is to replace @{name"if"} by @{name"case"}, which
nipkow@9754
    70
is also available for @{typ"bool"} but is not split automatically:
nipkow@8745
    71
*}
nipkow@8745
    72
nipkow@8745
    73
consts gcd2 :: "nat*nat \\<Rightarrow> nat";
nipkow@8745
    74
recdef gcd2 "measure (\\<lambda>(m,n).n)"
nipkow@8745
    75
  "gcd2(m,n) = (case n=0 of True \\<Rightarrow> m | False \\<Rightarrow> gcd2(n,m mod n))";
nipkow@8745
    76
nipkow@8745
    77
text{*\noindent
nipkow@8745
    78
In fact, this is probably the neatest solution next to pattern matching.
nipkow@8745
    79
nipkow@8745
    80
A final alternative is to replace the offending simplification rules by
nipkow@9754
    81
derived conditional ones. For @{term"gcd"} it means we have to prove
nipkow@8745
    82
*}
nipkow@8745
    83
nipkow@8745
    84
lemma [simp]: "gcd (m, 0) = m";
nipkow@9458
    85
by(simp);
nipkow@8745
    86
lemma [simp]: "n \\<noteq> 0 \\<Longrightarrow> gcd(m, n) = gcd(n, m mod n)";
nipkow@9458
    87
by(simp);
nipkow@8745
    88
nipkow@8745
    89
text{*\noindent
nipkow@8745
    90
after which we can disable the original simplification rule:
nipkow@8745
    91
*}
nipkow@8745
    92
nipkow@8745
    93
lemmas [simp del] = gcd.simps;
nipkow@8745
    94
nipkow@8745
    95
(*<*)
nipkow@8745
    96
end
nipkow@8745
    97
(*>*)