| author | wenzelm | 
| Tue, 19 Feb 2013 21:44:37 +0100 | |
| changeset 51224 | c3e99efacb67 | 
| parent 48985 | 5386df44a037 | 
| child 58620 | 7435b6a3f72e | 
| permissions | -rw-r--r-- | 
| 17914 | 1 | (*<*)theory WFrec imports Main begin(*>*) | 
| 10187 | 2 | |
| 3 | text{*\noindent
 | |
| 11161 | 4 | So far, all recursive definitions were shown to terminate via measure | 
| 11494 | 5 | functions. Sometimes this can be inconvenient or | 
| 10187 | 6 | impossible. Fortunately, \isacommand{recdef} supports much more
 | 
| 7 | general definitions. For example, termination of Ackermann's function | |
| 10654 | 8 | can be shown by means of the \rmindex{lexicographic product} @{text"<*lex*>"}:
 | 
| 10187 | 9 | *} | 
| 10 | ||
| 11626 | 11 | consts ack :: "nat\<times>nat \<Rightarrow> nat" | 
| 10187 | 12 | recdef ack "measure(\<lambda>m. m) <*lex*> measure(\<lambda>n. n)" | 
| 13 | "ack(0,n) = Suc n" | |
| 14 | "ack(Suc m,0) = ack(m, 1)" | |
| 11626 | 15 | "ack(Suc m,Suc n) = ack(m,ack(Suc m,n))" | 
| 10187 | 16 | |
| 17 | text{*\noindent
 | |
| 18 | The lexicographic product decreases if either its first component | |
| 19 | decreases (as in the second equation and in the outer call in the | |
| 20 | third equation) or its first component stays the same and the second | |
| 21 | component decreases (as in the inner call in the third equation). | |
| 22 | ||
| 23 | In general, \isacommand{recdef} supports termination proofs based on
 | |
| 10396 | 24 | arbitrary well-founded relations as introduced in \S\ref{sec:Well-founded}.
 | 
| 25 | This is called \textbf{well-founded
 | |
| 11494 | 26 | recursion}\indexbold{recursion!well-founded}.  A function definition
 | 
| 27 | is total if and only if the set of | |
| 28 | all pairs $(r,l)$, where $l$ is the argument on the | |
| 10396 | 29 | left-hand side of an equation and $r$ the argument of some recursive call on | 
| 30 | the corresponding right-hand side, induces a well-founded relation. For a | |
| 31 | systematic account of termination proofs via well-founded relations see, for | |
| 10885 | 32 | example, Baader and Nipkow~\cite{Baader-Nipkow}.
 | 
| 10187 | 33 | |
| 11494 | 34 | Each \isacommand{recdef} definition should be accompanied (after the function's
 | 
| 35 | name) by a well-founded relation on the function's argument type. | |
| 36 | Isabelle/HOL formalizes some of the most important | |
| 10396 | 37 | constructions of well-founded relations (see \S\ref{sec:Well-founded}). For
 | 
| 11494 | 38 | example, @{term"measure f"} is always well-founded.   The lexicographic
 | 
| 10396 | 39 | product of two well-founded relations is again well-founded, which we relied | 
| 40 | on when defining Ackermann's function above. | |
| 11308 | 41 | Of course the lexicographic product can also be iterated: | 
| 10189 | 42 | *} | 
| 10187 | 43 | |
| 10189 | 44 | consts contrived :: "nat \<times> nat \<times> nat \<Rightarrow> nat" | 
| 45 | recdef contrived | |
| 46 | "measure(\<lambda>i. i) <*lex*> measure(\<lambda>j. j) <*lex*> measure(\<lambda>k. k)" | |
| 47 | "contrived(i,j,Suc k) = contrived(i,j,k)" | |
| 48 | "contrived(i,Suc j,0) = contrived(i,j,j)" | |
| 49 | "contrived(Suc i,0,0) = contrived(i,i,i)" | |
| 50 | "contrived(0,0,0) = 0" | |
| 51 | ||
| 52 | text{*
 | |
| 10396 | 53 | Lexicographic products of measure functions already go a long | 
| 10885 | 54 | way. Furthermore, you may embed a type in an | 
| 10396 | 55 | existing well-founded relation via the inverse image construction @{term
 | 
| 56 | inv_image}. All these constructions are known to \isacommand{recdef}. Thus you
 | |
| 10241 | 57 | will never have to prove well-foundedness of any relation composed | 
| 10189 | 58 | solely of these building blocks. But of course the proof of | 
| 11494 | 59 | termination of your function definition --- that the arguments | 
| 60 | decrease with every recursive call --- may still require you to provide | |
| 10189 | 61 | additional lemmas. | 
| 62 | ||
| 10841 | 63 | It is also possible to use your own well-founded relations with | 
| 64 | \isacommand{recdef}.  For example, the greater-than relation can be made
 | |
| 65 | well-founded by cutting it off at a certain point. Here is an example | |
| 66 | of a recursive function that calls itself with increasing values up to ten: | |
| 10187 | 67 | *} | 
| 10189 | 68 | |
| 69 | consts f :: "nat \<Rightarrow> nat" | |
| 11705 | 70 | recdef (*<*)(permissive)(*>*)f "{(i,j). j<i \<and> i \<le> (10::nat)}"
 | 
| 71 | "f i = (if 10 \<le> i then 0 else i * f(Suc i))" | |
| 10189 | 72 | |
| 10396 | 73 | text{*\noindent
 | 
| 10841 | 74 | Since \isacommand{recdef} is not prepared for the relation supplied above,
 | 
| 75 | Isabelle rejects the definition. We should first have proved that | |
| 76 | our relation was well-founded: | |
| 10189 | 77 | *} | 
| 78 | ||
| 10841 | 79 | lemma wf_greater: "wf {(i,j). j<i \<and> i \<le> (N::nat)}"
 | 
| 80 | ||
| 11196 | 81 | txt{*\noindent
 | 
| 10841 | 82 | The proof is by showing that our relation is a subset of another well-founded | 
| 11494 | 83 | relation: one given by a measure function.\index{*wf_subset (theorem)}
 | 
| 11626 | 84 | *} | 
| 10841 | 85 | |
| 11626 | 86 | apply (rule wf_subset [of "measure (\<lambda>k::nat. N-k)"], blast) | 
| 10841 | 87 | |
| 88 | txt{*
 | |
| 89 | @{subgoals[display,indent=0,margin=65]}
 | |
| 90 | ||
| 91 | \noindent | |
| 92 | The inclusion remains to be proved. After unfolding some definitions, | |
| 20217 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 webertj parents: 
17914diff
changeset | 93 | we are left with simple arithmetic that is dispatched automatically. | 
| 11626 | 94 | *} | 
| 10841 | 95 | |
| 20217 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 webertj parents: 
17914diff
changeset | 96 | by (clarify, simp add: measure_def inv_image_def) | 
| 10189 | 97 | |
| 98 | text{*\noindent
 | |
| 10841 | 99 | |
| 11429 | 100 | Armed with this lemma, we use the \attrdx{recdef_wf} attribute to attach a
 | 
| 13111 | 101 | crucial hint\cmmdx{hints} to our definition:
 | 
| 10189 | 102 | *} | 
| 103 | (*<*) | |
| 104 | consts g :: "nat \<Rightarrow> nat" | |
| 11705 | 105 | recdef g "{(i,j). j<i \<and> i \<le> (10::nat)}"
 | 
| 106 | "g i = (if 10 \<le> i then 0 else i * g(Suc i))" | |
| 10189 | 107 | (*>*) | 
| 11626 | 108 | (hints recdef_wf: wf_greater) | 
| 10841 | 109 | |
| 110 | text{*\noindent
 | |
| 11705 | 111 | Alternatively, we could have given @{text "measure (\<lambda>k::nat. 10-k)"} for the
 | 
| 10841 | 112 | well-founded relation in our \isacommand{recdef}.  However, the arithmetic
 | 
| 113 | goal in the lemma above would have arisen instead in the \isacommand{recdef}
 | |
| 114 | termination proof, where we have less control. A tailor-made termination | |
| 115 | relation makes even more sense when it can be used in several function | |
| 116 | declarations. | |
| 117 | *} | |
| 118 | ||
| 10396 | 119 | (*<*)end(*>*) |