doc-src/TutorialI/Recdef/examples.thy
author wenzelm
Mon, 08 Oct 2001 12:13:34 +0200
changeset 11705 ac8ca15c556c
parent 11480 0fba0357c04c
child 15905 0a4cc9b113c7
permissions -rw-r--r--
fixed numerals;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     1
(*<*)
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     2
theory examples = Main:;
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     3
(*>*)
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     4
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     5
text{*
11458
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11231
diff changeset
     6
Here is a simple example, the \rmindex{Fibonacci function}:
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     7
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     8
9933
9feb1e0c4cb3 *** empty log message ***
nipkow
parents: 9792
diff changeset
     9
consts fib :: "nat \<Rightarrow> nat";
9feb1e0c4cb3 *** empty log message ***
nipkow
parents: 9792
diff changeset
    10
recdef fib "measure(\<lambda>n. n)"
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    11
  "fib 0 = 0"
11705
ac8ca15c556c fixed numerals;
wenzelm
parents: 11480
diff changeset
    12
  "fib (Suc 0) = 1"
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    13
  "fib (Suc(Suc x)) = fib x + fib (Suc x)";
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    14
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    15
text{*\noindent
11458
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11231
diff changeset
    16
\index{measure functions}%
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11231
diff changeset
    17
The definition of @{term"fib"} is accompanied by a \textbf{measure function}
9792
bbefb6ce5cb2 *** empty log message ***
nipkow
parents: 9541
diff changeset
    18
@{term"%n. n"} which maps the argument of @{term"fib"} to a
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    19
natural number. The requirement is that in each equation the measure of the
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    20
argument on the left-hand side is strictly greater than the measure of the
9792
bbefb6ce5cb2 *** empty log message ***
nipkow
parents: 9541
diff changeset
    21
argument of each recursive call. In the case of @{term"fib"} this is
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    22
obviously true because the measure function is the identity and
9792
bbefb6ce5cb2 *** empty log message ***
nipkow
parents: 9541
diff changeset
    23
@{term"Suc(Suc x)"} is strictly greater than both @{term"x"} and
9541
d17c0b34d5c8 *** empty log message ***
nipkow
parents: 8771
diff changeset
    24
@{term"Suc x"}.
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    25
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    26
Slightly more interesting is the insertion of a fixed element
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    27
between any two elements of a list:
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    28
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    29
9933
9feb1e0c4cb3 *** empty log message ***
nipkow
parents: 9792
diff changeset
    30
consts sep :: "'a \<times> 'a list \<Rightarrow> 'a list";
9feb1e0c4cb3 *** empty log message ***
nipkow
parents: 9792
diff changeset
    31
recdef sep "measure (\<lambda>(a,xs). length xs)"
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    32
  "sep(a, [])     = []"
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    33
  "sep(a, [x])    = [x]"
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    34
  "sep(a, x#y#zs) = x # a # sep(a,y#zs)";
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    35
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    36
text{*\noindent
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    37
This time the measure is the length of the list, which decreases with the
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    38
recursive call; the first component of the argument tuple is irrelevant.
10654
458068404143 *** empty log message ***
nipkow
parents: 10362
diff changeset
    39
The details of tupled $\lambda$-abstractions @{text"\<lambda>(x\<^sub>1,\<dots>,x\<^sub>n)"} are
458068404143 *** empty log message ***
nipkow
parents: 10362
diff changeset
    40
explained in \S\ref{sec:products}, but for now your intuition is all you need.
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    41
11458
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11231
diff changeset
    42
Pattern matching\index{pattern matching!and \isacommand{recdef}}
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11231
diff changeset
    43
need not be exhaustive:
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    44
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    45
9933
9feb1e0c4cb3 *** empty log message ***
nipkow
parents: 9792
diff changeset
    46
consts last :: "'a list \<Rightarrow> 'a";
9feb1e0c4cb3 *** empty log message ***
nipkow
parents: 9792
diff changeset
    47
recdef last "measure (\<lambda>xs. length xs)"
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    48
  "last [x]      = x"
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    49
  "last (x#y#zs) = last (y#zs)";
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    50
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    51
text{*
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    52
Overlapping patterns are disambiguated by taking the order of equations into
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    53
account, just as in functional programming:
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    54
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    55
9933
9feb1e0c4cb3 *** empty log message ***
nipkow
parents: 9792
diff changeset
    56
consts sep1 :: "'a \<times> 'a list \<Rightarrow> 'a list";
9feb1e0c4cb3 *** empty log message ***
nipkow
parents: 9792
diff changeset
    57
recdef sep1 "measure (\<lambda>(a,xs). length xs)"
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    58
  "sep1(a, x#y#zs) = x # a # sep1(a,y#zs)"
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    59
  "sep1(a, xs)     = xs";
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    60
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    61
text{*\noindent
11160
e0ab13bec5c8 *** empty log message ***
nipkow
parents: 10654
diff changeset
    62
To guarantee that the second equation can only be applied if the first
e0ab13bec5c8 *** empty log message ***
nipkow
parents: 10654
diff changeset
    63
one does not match, Isabelle internally replaces the second equation
e0ab13bec5c8 *** empty log message ***
nipkow
parents: 10654
diff changeset
    64
by the two possibilities that are left: @{prop"sep1(a,[]) = []"} and
e0ab13bec5c8 *** empty log message ***
nipkow
parents: 10654
diff changeset
    65
@{prop"sep1(a, [x]) = [x]"}.  Thus the functions @{term sep} and
e0ab13bec5c8 *** empty log message ***
nipkow
parents: 10654
diff changeset
    66
@{term sep1} are identical.
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    67
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    68
\begin{warn}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    69
  \isacommand{recdef} only takes the first argument of a (curried)
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    70
  recursive function into account. This means both the termination measure
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    71
  and pattern matching can only use that first argument. In general, you will
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    72
  therefore have to combine several arguments into a tuple. In case only one
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    73
  argument is relevant for termination, you can also rearrange the order of
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    74
  arguments as in the following definition:
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    75
\end{warn}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    76
*}
9933
9feb1e0c4cb3 *** empty log message ***
nipkow
parents: 9792
diff changeset
    77
consts sep2 :: "'a list \<Rightarrow> 'a \<Rightarrow> 'a list";
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    78
recdef sep2 "measure length"
10362
c6b197ccf1f1 *** empty log message ***
nipkow
parents: 9933
diff changeset
    79
  "sep2 (x#y#zs) = (\<lambda>a. x # a # sep2 (y#zs) a)"
9933
9feb1e0c4cb3 *** empty log message ***
nipkow
parents: 9792
diff changeset
    80
  "sep2 xs       = (\<lambda>a. xs)";
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    81
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    82
text{*
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    83
Because of its pattern-matching syntax, \isacommand{recdef} is also useful
11231
30d96882f915 *** empty log message ***
nipkow
parents: 11160
diff changeset
    84
for the definition of non-recursive functions, where the termination measure
30d96882f915 *** empty log message ***
nipkow
parents: 11160
diff changeset
    85
degenerates to the empty set @{term"{}"}:
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    86
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    87
9933
9feb1e0c4cb3 *** empty log message ***
nipkow
parents: 9792
diff changeset
    88
consts swap12 :: "'a list \<Rightarrow> 'a list";
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    89
recdef swap12 "{}"
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    90
  "swap12 (x#y#zs) = y#x#zs"
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    91
  "swap12 zs       = zs";
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    92
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    93
(*<*)
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    94
end
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    95
(*>*)