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