src/Doc/Tutorial/Recdef/examples.thy
changeset 48985 5386df44a037
parent 16417 9bc16273c2d4
child 58860 fee7cfa69c50
equal deleted inserted replaced
48984:f51d4a302962 48985:5386df44a037
       
     1 (*<*)
       
     2 theory examples imports Main begin;
       
     3 (*>*)
       
     4 
       
     5 text{*
       
     6 Here is a simple example, the \rmindex{Fibonacci function}:
       
     7 *}
       
     8 
       
     9 consts fib :: "nat \<Rightarrow> nat";
       
    10 recdef fib "measure(\<lambda>n. n)"
       
    11   "fib 0 = 0"
       
    12   "fib (Suc 0) = 1"
       
    13   "fib (Suc(Suc x)) = fib x + fib (Suc x)";
       
    14 
       
    15 text{*\noindent
       
    16 \index{measure functions}%
       
    17 The definition of @{term"fib"} is accompanied by a \textbf{measure function}
       
    18 @{term"%n. n"} which maps the argument of @{term"fib"} to a
       
    19 natural number. The requirement is that in each equation the measure of the
       
    20 argument on the left-hand side is strictly greater than the measure of the
       
    21 argument of each recursive call. In the case of @{term"fib"} this is
       
    22 obviously true because the measure function is the identity and
       
    23 @{term"Suc(Suc x)"} is strictly greater than both @{term"x"} and
       
    24 @{term"Suc x"}.
       
    25 
       
    26 Slightly more interesting is the insertion of a fixed element
       
    27 between any two elements of a list:
       
    28 *}
       
    29 
       
    30 consts sep :: "'a \<times> 'a list \<Rightarrow> 'a list";
       
    31 recdef sep "measure (\<lambda>(a,xs). length xs)"
       
    32   "sep(a, [])     = []"
       
    33   "sep(a, [x])    = [x]"
       
    34   "sep(a, x#y#zs) = x # a # sep(a,y#zs)";
       
    35 
       
    36 text{*\noindent
       
    37 This time the measure is the length of the list, which decreases with the
       
    38 recursive call; the first component of the argument tuple is irrelevant.
       
    39 The details of tupled $\lambda$-abstractions @{text"\<lambda>(x\<^sub>1,\<dots>,x\<^sub>n)"} are
       
    40 explained in \S\ref{sec:products}, but for now your intuition is all you need.
       
    41 
       
    42 Pattern matching\index{pattern matching!and \isacommand{recdef}}
       
    43 need not be exhaustive:
       
    44 *}
       
    45 
       
    46 consts last :: "'a list \<Rightarrow> 'a";
       
    47 recdef last "measure (\<lambda>xs. length xs)"
       
    48   "last [x]      = x"
       
    49   "last (x#y#zs) = last (y#zs)";
       
    50 
       
    51 text{*
       
    52 Overlapping patterns are disambiguated by taking the order of equations into
       
    53 account, just as in functional programming:
       
    54 *}
       
    55 
       
    56 consts sep1 :: "'a \<times> 'a list \<Rightarrow> 'a list";
       
    57 recdef sep1 "measure (\<lambda>(a,xs). length xs)"
       
    58   "sep1(a, x#y#zs) = x # a # sep1(a,y#zs)"
       
    59   "sep1(a, xs)     = xs";
       
    60 
       
    61 text{*\noindent
       
    62 To guarantee that the second equation can only be applied if the first
       
    63 one does not match, Isabelle internally replaces the second equation
       
    64 by the two possibilities that are left: @{prop"sep1(a,[]) = []"} and
       
    65 @{prop"sep1(a, [x]) = [x]"}.  Thus the functions @{term sep} and
       
    66 @{const sep1} are identical.
       
    67 
       
    68 \begin{warn}
       
    69   \isacommand{recdef} only takes the first argument of a (curried)
       
    70   recursive function into account. This means both the termination measure
       
    71   and pattern matching can only use that first argument. In general, you will
       
    72   therefore have to combine several arguments into a tuple. In case only one
       
    73   argument is relevant for termination, you can also rearrange the order of
       
    74   arguments as in the following definition:
       
    75 \end{warn}
       
    76 *}
       
    77 consts sep2 :: "'a list \<Rightarrow> 'a \<Rightarrow> 'a list";
       
    78 recdef sep2 "measure length"
       
    79   "sep2 (x#y#zs) = (\<lambda>a. x # a # sep2 (y#zs) a)"
       
    80   "sep2 xs       = (\<lambda>a. xs)";
       
    81 
       
    82 text{*
       
    83 Because of its pattern-matching syntax, \isacommand{recdef} is also useful
       
    84 for the definition of non-recursive functions, where the termination measure
       
    85 degenerates to the empty set @{term"{}"}:
       
    86 *}
       
    87 
       
    88 consts swap12 :: "'a list \<Rightarrow> 'a list";
       
    89 recdef swap12 "{}"
       
    90   "swap12 (x#y#zs) = y#x#zs"
       
    91   "swap12 zs       = zs";
       
    92 
       
    93 (*<*)
       
    94 end
       
    95 (*>*)