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