| 8745 |      1 | (*<*)
 | 
| 9792 |      2 | theory termination = examples:
 | 
| 8745 |      3 | (*>*)
 | 
|  |      4 | 
 | 
|  |      5 | text{*
 | 
| 11309 |      6 | When a function~$f$ is defined via \isacommand{recdef}, Isabelle tries to prove
 | 
|  |      7 | its termination with the help of the user-supplied measure.  Each of the examples
 | 
|  |      8 | above is simple enough that Isabelle can automatically prove that the
 | 
|  |      9 | argument's measure decreases in each recursive call. As a result,
 | 
| 9792 |     10 | $f$@{text".simps"} will contain the defining equations (or variants derived
 | 
|  |     11 | from them) as theorems. For example, look (via \isacommand{thm}) at
 | 
|  |     12 | @{thm[source]sep.simps} and @{thm[source]sep1.simps} to see that they define
 | 
|  |     13 | the same function. What is more, those equations are automatically declared as
 | 
| 8745 |     14 | simplification rules.
 | 
|  |     15 | 
 | 
| 11458 |     16 | Isabelle may fail to prove the termination condition for some
 | 
|  |     17 | recursive call.  Let us try the following artificial function:*}
 | 
| 8745 |     18 | 
 | 
| 9933 |     19 | consts f :: "nat\<times>nat \<Rightarrow> nat";
 | 
|  |     20 | recdef f "measure(\<lambda>(x,y). x-y)"
 | 
|  |     21 |   "f(x,y) = (if x \<le> y then x else f(x,y+1))";
 | 
| 8745 |     22 | 
 | 
|  |     23 | text{*\noindent
 | 
| 11458 |     24 | Isabelle prints a
 | 
|  |     25 | \REMARK{error or warning?  change this part?  rename g to f?}
 | 
| 10795 |     26 | message showing you what it was unable to prove. You will then
 | 
| 8745 |     27 | have to prove it as a separate lemma before you attempt the definition
 | 
|  |     28 | of your function once more. In our case the required lemma is the obvious one:
 | 
|  |     29 | *}
 | 
|  |     30 | 
 | 
| 9933 |     31 | lemma termi_lem: "\<not> x \<le> y \<Longrightarrow> x - Suc y < x - y";
 | 
| 8745 |     32 | 
 | 
|  |     33 | txt{*\noindent
 | 
| 11458 |     34 | It was not proved automatically because of the awkward behaviour of subtraction
 | 
|  |     35 | on type @{typ"nat"}. This requires more arithmetic than is tried by default:
 | 
| 8745 |     36 | *}
 | 
|  |     37 | 
 | 
| 10171 |     38 | apply(arith);
 | 
|  |     39 | done
 | 
| 8745 |     40 | 
 | 
|  |     41 | text{*\noindent
 | 
| 8771 |     42 | Because \isacommand{recdef}'s termination prover involves simplification,
 | 
| 11429 |     43 | we include in our second attempt a hint: the \attrdx{recdef_simp} attribute 
 | 
|  |     44 | says to use @{thm[source]termi_lem} as
 | 
|  |     45 | a simplification rule.  
 | 
| 8745 |     46 | *}
 | 
|  |     47 | 
 | 
| 9933 |     48 | consts g :: "nat\<times>nat \<Rightarrow> nat";
 | 
|  |     49 | recdef g "measure(\<lambda>(x,y). x-y)"
 | 
|  |     50 |   "g(x,y) = (if x \<le> y then x else g(x,y+1))"
 | 
| 9992 |     51 | (hints recdef_simp: termi_lem)
 | 
| 8745 |     52 | 
 | 
|  |     53 | text{*\noindent
 | 
| 9792 |     54 | This time everything works fine. Now @{thm[source]g.simps} contains precisely
 | 
| 11458 |     55 | the stated recursion equation for @{term g}, which has been stored as a
 | 
|  |     56 | simplification rule.  Thus we can automatically prove results such as this one:
 | 
| 8745 |     57 | *}
 | 
|  |     58 | 
 | 
| 9933 |     59 | theorem "g(1,0) = g(1,1)";
 | 
| 10171 |     60 | apply(simp);
 | 
|  |     61 | done
 | 
| 8745 |     62 | 
 | 
|  |     63 | text{*\noindent
 | 
|  |     64 | More exciting theorems require induction, which is discussed below.
 | 
|  |     65 | 
 | 
| 9933 |     66 | If the termination proof requires a new lemma that is of general use, you can
 | 
|  |     67 | turn it permanently into a simplification rule, in which case the above
 | 
|  |     68 | \isacommand{hint} is not necessary. But our @{thm[source]termi_lem} is not
 | 
|  |     69 | sufficiently general to warrant this distinction.
 | 
| 8745 |     70 | 
 | 
| 10171 |     71 | The attentive reader may wonder why we chose to call our function @{term g}
 | 
|  |     72 | rather than @{term f} the second time around. The reason is that, despite
 | 
|  |     73 | the failed termination proof, the definition of @{term f} did not
 | 
| 9792 |     74 | fail, and thus we could not define it a second time. However, all theorems
 | 
| 10171 |     75 | about @{term f}, for example @{thm[source]f.simps}, carry as a precondition
 | 
| 9792 |     76 | the unproved termination condition. Moreover, the theorems
 | 
| 11458 |     77 | @{thm[source]f.simps} are not stored as simplification rules. 
 | 
|  |     78 | However, this mechanism
 | 
| 9792 |     79 | allows a delayed proof of termination: instead of proving
 | 
|  |     80 | @{thm[source]termi_lem} up front, we could prove 
 | 
| 8745 |     81 | it later on and then use it to remove the preconditions from the theorems
 | 
| 10171 |     82 | about @{term f}. In most cases this is more cumbersome than proving things
 | 
| 9792 |     83 | up front.
 | 
| 11458 |     84 | \REMARK{FIXME, with one exception: nested recursion.}
 | 
| 8745 |     85 | *}
 | 
|  |     86 | 
 | 
|  |     87 | (*<*)
 | 
|  |     88 | end
 | 
|  |     89 | (*>*)
 |