doc-src/TutorialI/Recdef/simplification.thy
author paulson
Fri, 12 Jan 2001 16:32:01 +0100
changeset 10885 90695f46440b
parent 10795 9e888d60d3e5
child 11215 b44ad7e4c4d2
permissions -rw-r--r--
lcp's pass over the book, chapters 1-8
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@10171
    10
terminate because of automatic splitting of @{text if}.
nipkow@8745
    11
Let us look at an example:
nipkow@8745
    12
*}
nipkow@8745
    13
nipkow@10171
    14
consts gcd :: "nat\<times>nat \<Rightarrow> nat";
nipkow@10171
    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
@{term[display]"n ~= 0 ==> m mod n < n"}
paulson@10885
    22
is proved automatically because it is already present as a lemma in
paulson@10885
    23
HOL\@.  Thus the recursion equation becomes a simplification
nipkow@8745
    24
rule. Of course the equation is nonterminating if we are allowed to unfold
nipkow@10171
    25
the recursive call inside the @{text else} branch, which is why programming
nipkow@8745
    26
languages and our simplifier don't do that. Unfortunately the simplifier does
nipkow@10171
    27
something else which leads to the same problem: it splits @{text if}s if the
nipkow@10171
    28
condition simplifies to neither @{term True} nor @{term False}. For
nipkow@8745
    29
example, simplification reduces
nipkow@9541
    30
@{term[display]"gcd(m,n) = k"}
nipkow@8745
    31
in one step to
nipkow@9541
    32
@{term[display]"(if n=0 then m else gcd(n, m mod n)) = k"}
nipkow@8745
    33
where the condition cannot be reduced further, and splitting leads to
nipkow@9541
    34
@{term[display]"(n=0 --> m=k) & (n ~= 0 --> gcd(n, m mod n)=k)"}
nipkow@9541
    35
Since the recursive call @{term"gcd(n, m mod n)"} is no longer protected by
nipkow@10171
    36
an @{text if}, it is unfolded again, which leads to an infinite chain of
nipkow@9541
    37
simplification steps. Fortunately, this problem can be avoided in many
nipkow@9541
    38
different ways.
nipkow@8745
    39
nipkow@9792
    40
The most radical solution is to disable the offending
nipkow@9792
    41
@{thm[source]split_if} as shown in the section on case splits in
nipkow@9792
    42
\S\ref{sec:Simplification}.  However, we do not recommend this because it
nipkow@10171
    43
means you will often have to invoke the rule explicitly when @{text if} is
nipkow@9792
    44
involved.
nipkow@8745
    45
nipkow@8745
    46
If possible, the definition should be given by pattern matching on the left
nipkow@10171
    47
rather than @{text if} on the right. In the case of @{term gcd} the
nipkow@8745
    48
following alternative definition suggests itself:
nipkow@8745
    49
*}
nipkow@8745
    50
nipkow@10171
    51
consts gcd1 :: "nat\<times>nat \<Rightarrow> nat";
nipkow@10171
    52
recdef gcd1 "measure (\<lambda>(m,n).n)"
nipkow@8745
    53
  "gcd1 (m, 0) = m"
nipkow@8745
    54
  "gcd1 (m, n) = gcd1(n, m mod n)";
nipkow@8745
    55
nipkow@8745
    56
nipkow@8745
    57
text{*\noindent
nipkow@8745
    58
Note that the order of equations is important and hides the side condition
nipkow@9754
    59
@{prop"n ~= 0"}. Unfortunately, in general the case distinction
nipkow@8745
    60
may not be expressible by pattern matching.
nipkow@8745
    61
nipkow@10171
    62
A very simple alternative is to replace @{text if} by @{text case}, which
nipkow@10178
    63
is also available for @{typ bool} but is not split automatically:
nipkow@8745
    64
*}
nipkow@8745
    65
nipkow@10171
    66
consts gcd2 :: "nat\<times>nat \<Rightarrow> nat";
nipkow@10171
    67
recdef gcd2 "measure (\<lambda>(m,n).n)"
nipkow@10171
    68
  "gcd2(m,n) = (case n=0 of True \<Rightarrow> m | False \<Rightarrow> gcd2(n,m mod n))";
nipkow@8745
    69
nipkow@8745
    70
text{*\noindent
nipkow@8745
    71
In fact, this is probably the neatest solution next to pattern matching.
nipkow@8745
    72
nipkow@8745
    73
A final alternative is to replace the offending simplification rules by
nipkow@10171
    74
derived conditional ones. For @{term gcd} it means we have to prove
paulson@10795
    75
these lemmas:
nipkow@8745
    76
*}
nipkow@8745
    77
nipkow@8745
    78
lemma [simp]: "gcd (m, 0) = m";
nipkow@10171
    79
apply(simp);
nipkow@10171
    80
done
nipkow@10171
    81
lemma [simp]: "n \<noteq> 0 \<Longrightarrow> gcd(m, n) = gcd(n, m mod n)";
nipkow@10171
    82
apply(simp);
nipkow@10171
    83
done
nipkow@8745
    84
nipkow@8745
    85
text{*\noindent
paulson@10795
    86
Now we can disable the original simplification rule:
nipkow@8745
    87
*}
nipkow@8745
    88
nipkow@9933
    89
declare gcd.simps [simp del]
nipkow@8745
    90
nipkow@8745
    91
(*<*)
nipkow@8745
    92
end
nipkow@8745
    93
(*>*)