| 9645 |      1 | (*<*)
 | 
| 16417 |      2 | theory Nested0 imports Main begin
 | 
| 9645 |      3 | (*>*)
 | 
|  |      4 | 
 | 
|  |      5 | text{*
 | 
| 11494 |      6 | \index{datatypes!nested}%
 | 
| 9645 |      7 | In \S\ref{sec:nested-datatype} we defined the datatype of terms
 | 
|  |      8 | *}
 | 
|  |      9 | 
 | 
|  |     10 | datatype ('a,'b)"term" = Var 'a | App 'b "('a,'b)term list"
 | 
|  |     11 | 
 | 
|  |     12 | text{*\noindent
 | 
|  |     13 | and closed with the observation that the associated schema for the definition
 | 
|  |     14 | of primitive recursive functions leads to overly verbose definitions. Moreover,
 | 
|  |     15 | if you have worked exercise~\ref{ex:trev-trev} you will have noticed that
 | 
| 10885 |     16 | you needed to declare essentially the same function as @{term"rev"}
 | 
| 11196 |     17 | and prove many standard properties of list reversal all over again. 
 | 
| 10885 |     18 | We will now show you how \isacommand{recdef} can simplify
 | 
| 9645 |     19 | definitions and proofs about nested recursive datatypes. As an example we
 | 
| 9754 |     20 | choose exercise~\ref{ex:trev-trev}:
 | 
| 9645 |     21 | *}
 | 
| 10186 |     22 | 
 | 
|  |     23 | consts trev  :: "('a,'b)term \<Rightarrow> ('a,'b)term"
 | 
|  |     24 | (*<*)end(*>*)
 |