| 9645 |      1 | (*<*)
 | 
| 11626 |      2 | theory Nested1 = Nested0:
 | 
| 9645 |      3 | (*>*)
 | 
|  |      4 | 
 | 
|  |      5 | text{*\noindent
 | 
| 11277 |      6 | Although the definition of @{term trev} below is quite natural, we will have
 | 
| 10885 |      7 | to overcome a minor difficulty in convincing Isabelle of its termination.
 | 
| 9754 |      8 | It is precisely this difficulty that is the \textit{raison d'\^etre} of
 | 
| 9645 |      9 | this subsection.
 | 
|  |     10 | 
 | 
| 10171 |     11 | Defining @{term trev} by \isacommand{recdef} rather than \isacommand{primrec}
 | 
| 9645 |     12 | simplifies matters because we are now free to use the recursion equation
 | 
|  |     13 | suggested at the end of \S\ref{sec:nested-datatype}:
 | 
| 11626 |     14 | *}
 | 
| 9645 |     15 | 
 | 
| 11636 |     16 | recdef (*<*)(permissive)(*>*)trev "measure size"
 | 
| 9754 |     17 |  "trev (Var x)    = Var x"
 | 
| 11626 |     18 |  "trev (App f ts) = App f (rev(map trev ts))"
 | 
| 9689 |     19 | 
 | 
| 9754 |     20 | text{*\noindent
 | 
| 10171 |     21 | Remember that function @{term size} is defined for each \isacommand{datatype}.
 | 
| 9754 |     22 | However, the definition does not succeed. Isabelle complains about an
 | 
|  |     23 | unproved termination condition
 | 
| 10242 |     24 | @{prop[display]"t : set ts --> size t < Suc (term_list_size ts)"}
 | 
| 10171 |     25 | where @{term set} returns the set of elements of a list
 | 
| 9933 |     26 | and @{text"term_list_size :: term list \<Rightarrow> nat"} is an auxiliary
 | 
|  |     27 | function automatically defined by Isabelle
 | 
| 11494 |     28 | (while processing the declaration of @{text term}).  Why does the
 | 
|  |     29 | recursive call of @{term trev} lead to this
 | 
|  |     30 | condition?  Because \isacommand{recdef} knows that @{term map}
 | 
| 10885 |     31 | will apply @{term trev} only to elements of @{term ts}. Thus the 
 | 
| 10171 |     32 | condition expresses that the size of the argument @{prop"t : set ts"} of any
 | 
| 10885 |     33 | recursive call of @{term trev} is strictly less than @{term"size(App f ts)"},
 | 
|  |     34 | which equals @{term"Suc(term_list_size ts)"}.  We will now prove the termination condition and
 | 
| 9754 |     35 | continue with our definition.  Below we return to the question of how
 | 
| 11494 |     36 | \isacommand{recdef} knows about @{term map}.
 | 
| 12491 |     37 | 
 | 
|  |     38 | The termination condition is easily proved by induction:
 | 
| 11626 |     39 | *}
 | 
| 9645 |     40 | 
 | 
|  |     41 | (*<*)
 | 
| 11626 |     42 | end
 | 
| 9754 |     43 | (*>*)
 |