| 11235 |      1 | theory RECDEF = Main:
 | 
|  |      2 | 
 | 
|  |      3 | subsection{*Wellfounded Recursion*}
 | 
|  |      4 | 
 | 
|  |      5 | subsubsection{*Examples*}
 | 
|  |      6 | 
 | 
|  |      7 | consts fib :: "nat \<Rightarrow> nat";
 | 
|  |      8 | recdef fib "measure(\<lambda>n. n)"
 | 
|  |      9 |   "fib 0 = 0"
 | 
|  |     10 |   "fib 1 = 1"
 | 
|  |     11 |   "fib (Suc(Suc x)) = fib x + fib (Suc x)";
 | 
|  |     12 | 
 | 
|  |     13 | consts sep :: "'a \<times> 'a list \<Rightarrow> 'a list";
 | 
|  |     14 | recdef sep "measure (\<lambda>(a,xs). length xs)"
 | 
|  |     15 |   "sep(a, [])     = []"
 | 
|  |     16 |   "sep(a, [x])    = [x]"
 | 
|  |     17 |   "sep(a, x#y#zs) = x # a # sep(a,y#zs)";
 | 
|  |     18 | 
 | 
|  |     19 | consts last :: "'a list \<Rightarrow> 'a";
 | 
|  |     20 | recdef last "measure (\<lambda>xs. length xs)"
 | 
|  |     21 |   "last [x]      = x"
 | 
|  |     22 |   "last (x#y#zs) = last (y#zs)";
 | 
|  |     23 | 
 | 
|  |     24 | 
 | 
|  |     25 | consts sep1 :: "'a \<times> 'a list \<Rightarrow> 'a list";
 | 
|  |     26 | recdef sep1 "measure (\<lambda>(a,xs). length xs)"
 | 
|  |     27 |   "sep1(a, x#y#zs) = x # a # sep1(a,y#zs)"
 | 
|  |     28 |   "sep1(a, xs)     = xs";
 | 
|  |     29 | 
 | 
|  |     30 | thm sep1.simps
 | 
|  |     31 | 
 | 
|  |     32 | 
 | 
|  |     33 | consts swap12 :: "'a list \<Rightarrow> 'a list";
 | 
|  |     34 | recdef swap12 "{}"
 | 
|  |     35 |   "swap12 (x#y#zs) = y#x#zs"
 | 
|  |     36 |   "swap12 zs       = zs";
 | 
|  |     37 | 
 | 
|  |     38 | 
 | 
|  |     39 | subsubsection{*Beyond Measure*}
 | 
|  |     40 | 
 | 
|  |     41 | consts ack :: "nat\<times>nat \<Rightarrow> nat";
 | 
|  |     42 | recdef ack "measure(\<lambda>m. m) <*lex*> measure(\<lambda>n. n)"
 | 
|  |     43 |   "ack(0,n)         = Suc n"
 | 
|  |     44 |   "ack(Suc m,0)     = ack(m, 1)"
 | 
|  |     45 |   "ack(Suc m,Suc n) = ack(m,ack(Suc m,n))";
 | 
|  |     46 | 
 | 
|  |     47 | 
 | 
|  |     48 | subsubsection{*Induction*}
 | 
|  |     49 | 
 | 
|  |     50 | lemma "map f (sep(x,xs)) = sep(f x, map f xs)"
 | 
| 11293 |     51 | thm sep.induct
 | 
| 11235 |     52 | apply(induct_tac x xs rule: sep.induct)
 | 
|  |     53 | apply simp_all
 | 
|  |     54 | done
 | 
|  |     55 | 
 | 
|  |     56 | lemma ack_incr2: "n < ack(m,n)"
 | 
|  |     57 | apply(induct_tac m n rule: ack.induct)
 | 
|  |     58 | apply simp_all
 | 
|  |     59 | done
 | 
|  |     60 | 
 | 
|  |     61 | 
 | 
|  |     62 | subsubsection{*Recursion Over Nested Datatypes*}
 | 
|  |     63 | 
 | 
|  |     64 | datatype tree = C "tree list"
 | 
|  |     65 | 
 | 
|  |     66 | lemma termi_lem: "t \<in> set ts \<longrightarrow> size t < Suc(tree_list_size ts)"
 | 
|  |     67 | by(induct_tac ts, auto)
 | 
|  |     68 | 
 | 
|  |     69 | consts mirror :: "tree \<Rightarrow> tree"
 | 
|  |     70 | recdef mirror "measure size"
 | 
|  |     71 |  "mirror(C ts) = C(rev(map mirror ts))"
 | 
|  |     72 | (hints recdef_simp: termi_lem)
 | 
|  |     73 | 
 | 
|  |     74 | lemma "mirror(mirror t) = t"
 | 
|  |     75 | apply(induct_tac t rule: mirror.induct)
 | 
|  |     76 | apply(simp add:rev_map sym[OF map_compose] cong:map_cong)
 | 
|  |     77 | done
 | 
|  |     78 | 
 | 
| 11236 |     79 | text{*
 | 
|  |     80 | \begin{exercise}
 | 
|  |     81 | Define a function for merging two ordered lists (of natural numbers) and
 | 
|  |     82 | show that if the two input lists are ordered, so is the output.
 | 
|  |     83 | \end{exercise}
 | 
|  |     84 | *}
 | 
|  |     85 | 
 | 
| 11235 |     86 | end
 |