doc-src/TutorialI/Recdef/simplification.thy
author nipkow
Wed, 12 Dec 2001 09:04:20 +0100
changeset 12473 f41e477576b9
parent 11458 09a6c44a48ea
child 16417 9bc16273c2d4
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{*
paulson@11458
     6
Once we have proved all the termination conditions, the \isacommand{recdef} 
paulson@11458
     7
recursion 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}.
paulson@11458
    11
\index{*if expressions!splitting of}
nipkow@8745
    12
Let us look at an example:
nipkow@8745
    13
*}
nipkow@8745
    14
nipkow@10171
    15
consts gcd :: "nat\<times>nat \<Rightarrow> nat";
nipkow@10171
    16
recdef gcd "measure (\<lambda>(m,n).n)"
nipkow@8745
    17
  "gcd (m, n) = (if n=0 then m else gcd(n, m mod n))";
nipkow@8745
    18
nipkow@8745
    19
text{*\noindent
nipkow@8745
    20
According to the measure function, the second argument should decrease with
nipkow@8745
    21
each recursive call. The resulting termination condition
nipkow@12473
    22
@{term[display]"n ~= (0::nat) ==> m mod n < n"}
paulson@10885
    23
is proved automatically because it is already present as a lemma in
paulson@10885
    24
HOL\@.  Thus the recursion equation becomes a simplification
nipkow@8745
    25
rule. Of course the equation is nonterminating if we are allowed to unfold
nipkow@10171
    26
the recursive call inside the @{text else} branch, which is why programming
nipkow@8745
    27
languages and our simplifier don't do that. Unfortunately the simplifier does
paulson@11458
    28
something else that leads to the same problem: it splits 
paulson@11458
    29
each @{text if}-expression unless its
paulson@11458
    30
condition simplifies to @{term True} or @{term False}.  For
nipkow@8745
    31
example, simplification reduces
nipkow@9541
    32
@{term[display]"gcd(m,n) = k"}
nipkow@8745
    33
in one step to
nipkow@9541
    34
@{term[display]"(if n=0 then m else gcd(n, m mod n)) = k"}
nipkow@8745
    35
where the condition cannot be reduced further, and splitting leads to
nipkow@9541
    36
@{term[display]"(n=0 --> m=k) & (n ~= 0 --> gcd(n, m mod n)=k)"}
nipkow@9541
    37
Since the recursive call @{term"gcd(n, m mod n)"} is no longer protected by
nipkow@10171
    38
an @{text if}, it is unfolded again, which leads to an infinite chain of
nipkow@9541
    39
simplification steps. Fortunately, this problem can be avoided in many
nipkow@9541
    40
different ways.
nipkow@8745
    41
paulson@11458
    42
The most radical solution is to disable the offending theorem
paulson@11458
    43
@{thm[source]split_if},
nipkow@11215
    44
as shown in \S\ref{sec:AutoCaseSplits}.  However, we do not recommend this
paulson@11458
    45
approach: you will often have to invoke the rule explicitly when
nipkow@11215
    46
@{text if} is involved.
nipkow@8745
    47
nipkow@8745
    48
If possible, the definition should be given by pattern matching on the left
nipkow@10171
    49
rather than @{text if} on the right. In the case of @{term gcd} the
nipkow@8745
    50
following alternative definition suggests itself:
nipkow@8745
    51
*}
nipkow@8745
    52
nipkow@10171
    53
consts gcd1 :: "nat\<times>nat \<Rightarrow> nat";
nipkow@10171
    54
recdef gcd1 "measure (\<lambda>(m,n).n)"
nipkow@8745
    55
  "gcd1 (m, 0) = m"
nipkow@8745
    56
  "gcd1 (m, n) = gcd1(n, m mod n)";
nipkow@8745
    57
nipkow@8745
    58
nipkow@8745
    59
text{*\noindent
paulson@11458
    60
The order of equations is important: it hides the side condition
nipkow@12473
    61
@{prop"n ~= (0::nat)"}.  Unfortunately, in general the case distinction
nipkow@8745
    62
may not be expressible by pattern matching.
nipkow@8745
    63
paulson@11458
    64
A simple alternative is to replace @{text if} by @{text case}, 
paulson@11458
    65
which is also available for @{typ bool} and is not split automatically:
nipkow@8745
    66
*}
nipkow@8745
    67
nipkow@10171
    68
consts gcd2 :: "nat\<times>nat \<Rightarrow> nat";
nipkow@10171
    69
recdef gcd2 "measure (\<lambda>(m,n).n)"
nipkow@10171
    70
  "gcd2(m,n) = (case n=0 of True \<Rightarrow> m | False \<Rightarrow> gcd2(n,m mod n))";
nipkow@8745
    71
nipkow@8745
    72
text{*\noindent
paulson@11458
    73
This is probably the neatest solution next to pattern matching, and it is
paulson@11458
    74
always available.
nipkow@8745
    75
nipkow@8745
    76
A final alternative is to replace the offending simplification rules by
nipkow@10171
    77
derived conditional ones. For @{term gcd} it means we have to prove
paulson@10795
    78
these lemmas:
nipkow@8745
    79
*}
nipkow@8745
    80
nipkow@8745
    81
lemma [simp]: "gcd (m, 0) = m";
nipkow@10171
    82
apply(simp);
nipkow@10171
    83
done
paulson@11458
    84
nipkow@10171
    85
lemma [simp]: "n \<noteq> 0 \<Longrightarrow> gcd(m, n) = gcd(n, m mod n)";
nipkow@10171
    86
apply(simp);
nipkow@10171
    87
done
nipkow@8745
    88
nipkow@8745
    89
text{*\noindent
paulson@11458
    90
Simplification terminates for these proofs because the condition of the @{text
paulson@11458
    91
if} simplifies to @{term True} or @{term False}.
paulson@10795
    92
Now we can disable the original simplification rule:
nipkow@8745
    93
*}
nipkow@8745
    94
nipkow@9933
    95
declare gcd.simps [simp del]
nipkow@8745
    96
nipkow@8745
    97
(*<*)
nipkow@8745
    98
end
nipkow@8745
    99
(*>*)