src/Doc/Functions/Functions.thy
author wenzelm
Tue, 18 May 2021 15:57:49 +0200
changeset 73725 cd9afbd0ccb9
parent 70276 910dc065b869
child 76649 9a6cb5ecc183
permissions -rw-r--r--
redundant: copy produced from session document_files;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
50530
6266e44b3396 updated some headers;
wenzelm
parents: 49322
diff changeset
     1
(*  Title:      Doc/Functions/Functions.thy
21212
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
     2
    Author:     Alexander Krauss, TU Muenchen
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
     3
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
     4
Tutorial for function definitions with the new "function" package.
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
     5
*)
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
     6
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
     7
theory Functions
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
     8
imports Main
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
     9
begin
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
    10
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
    11
section \<open>Function Definitions for Dummies\<close>
21212
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
    12
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
    13
text \<open>
21212
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
    14
  In most cases, defining a recursive function is just as simple as other definitions:
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
    15
\<close>
21212
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
    16
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
    17
fun fib :: "nat \<Rightarrow> nat"
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
    18
where
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
    19
  "fib 0 = 1"
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
    20
| "fib (Suc 0) = 1"
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
    21
| "fib (Suc (Suc n)) = fib n + fib (Suc n)"
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
    22
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
    23
text \<open>
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
    24
  The syntax is rather self-explanatory: We introduce a function by
25091
a2ae7f71613d Updated function tutorial: Types can be inferred and need not be given anymore
krauss
parents: 23805
diff changeset
    25
  giving its name, its type, 
a2ae7f71613d Updated function tutorial: Types can be inferred and need not be given anymore
krauss
parents: 23805
diff changeset
    26
  and a set of defining recursive equations.
a2ae7f71613d Updated function tutorial: Types can be inferred and need not be given anymore
krauss
parents: 23805
diff changeset
    27
  If we leave out the type, the most general type will be
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
    28
  inferred, which can sometimes lead to surprises: Since both \<^term>\<open>1::nat\<close> and \<open>+\<close> are overloaded, we would end up
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67443
diff changeset
    29
  with \<open>fib :: nat \<Rightarrow> 'a::{one,plus}\<close>.
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
    30
\<close>
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
    31
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
    32
text \<open>
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
    33
  The function always terminates, since its argument gets smaller in
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
    34
  every recursive call. 
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
    35
  Since HOL is a logic of total functions, termination is a
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
    36
  fundamental requirement to prevent inconsistencies\footnote{From the
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67443
diff changeset
    37
  \qt{definition} \<open>f(n) = f(n) + 1\<close> we could prove 
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67443
diff changeset
    38
  \<open>0 = 1\<close> by subtracting \<open>f(n)\<close> on both sides.}.
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
    39
  Isabelle tries to prove termination automatically when a definition
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
    40
  is made. In \S\ref{termination}, we will look at cases where this
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
    41
  fails and see what to do then.
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
    42
\<close>
21212
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
    43
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
    44
subsection \<open>Pattern matching\<close>
21212
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
    45
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
    46
text \<open>\label{patmatch}
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
    47
  Like in functional programming, we can use pattern matching to
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
    48
  define functions. At the moment we will only consider \emph{constructor
21212
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
    49
  patterns}, which only consist of datatype constructors and
23805
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
    50
  variables. Furthermore, patterns must be linear, i.e.\ all variables
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
    51
  on the left hand side of an equation must be distinct. In
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
    52
  \S\ref{genpats} we discuss more general pattern matching.
21212
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
    53
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
    54
  If patterns overlap, the order of the equations is taken into
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
    55
  account. The following function inserts a fixed element between any
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
    56
  two elements of a list:
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
    57
\<close>
21212
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
    58
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
    59
fun sep :: "'a \<Rightarrow> 'a list \<Rightarrow> 'a list"
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
    60
where
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
    61
  "sep a (x#y#xs) = x # a # sep a (y # xs)"
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
    62
| "sep a xs       = xs"
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
    63
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
    64
text \<open>
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
    65
  Overlapping patterns are interpreted as \qt{increments} to what is
21212
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
    66
  already there: The second equation is only meant for the cases where
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
    67
  the first one does not match. Consequently, Isabelle replaces it
22065
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
    68
  internally by the remaining cases, making the patterns disjoint:
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
    69
\<close>
21212
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
    70
22065
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
    71
thm sep.simps
21212
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
    72
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
    73
text \<open>@{thm [display] sep.simps[no_vars]}\<close>
21212
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
    74
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
    75
text \<open>
23805
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
    76
  \noindent The equations from function definitions are automatically used in
21212
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
    77
  simplification:
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
    78
\<close>
21212
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
    79
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
    80
lemma "sep 0 [1, 2, 3] = [1, 0, 2, 0, 3]"
21212
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
    81
by simp
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
    82
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
    83
subsection \<open>Induction\<close>
21212
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
    84
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
    85
text \<open>
22065
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
    86
23805
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
    87
  Isabelle provides customized induction rules for recursive
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
    88
  functions. These rules follow the recursive structure of the
53107
57c7294eac0a more document antiquotations (for proper theorem names);
Christian Sternagel
parents: 50530
diff changeset
    89
  definition. Here is the rule @{thm [source] sep.induct} arising from the
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
    90
  above definition of \<^const>\<open>sep\<close>:
23805
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
    91
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
    92
  @{thm [display] sep.induct}
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
    93
  
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
    94
  We have a step case for list with at least two elements, and two
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
    95
  base cases for the zero- and the one-element list. Here is a simple
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
    96
  proof about \<^const>\<open>sep\<close> and \<^const>\<open>map\<close>
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
    97
\<close>
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
    98
23805
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
    99
lemma "map f (sep x ys) = sep (f x) (map f ys)"
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
   100
apply (induct x ys rule: sep.induct)
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
   101
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   102
text \<open>
23805
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
   103
  We get three cases, like in the definition.
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
   104
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
   105
  @{subgoals [display]}
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   106
\<close>
23805
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
   107
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
   108
apply auto 
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
   109
done
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   110
text \<open>
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   111
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   112
  With the \cmd{fun} command, you can define about 80\% of the
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   113
  functions that occur in practice. The rest of this tutorial explains
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   114
  the remaining 20\%.
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   115
\<close>
21212
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
   116
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
   117
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   118
section \<open>fun vs.\ function\<close>
21212
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
   119
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   120
text \<open>
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   121
  The \cmd{fun} command provides a
21212
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
   122
  convenient shorthand notation for simple function definitions. In
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
   123
  this mode, Isabelle tries to solve all the necessary proof obligations
27026
3602b81665b5 Updated function tutorial.
krauss
parents: 26749
diff changeset
   124
  automatically. If any proof fails, the definition is
22065
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   125
  rejected. This can either mean that the definition is indeed faulty,
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   126
  or that the default proof procedures are just not smart enough (or
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   127
  rather: not designed) to handle the definition.
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   128
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   129
  By expanding the abbreviation to the more verbose \cmd{function} command, these proof obligations become visible and can be analyzed or
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   130
  solved manually. The expansion from \cmd{fun} to \cmd{function} is as follows:
22065
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   131
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   132
\end{isamarkuptext}
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   133
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   134
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   135
\[\left[\;\begin{minipage}{0.25\textwidth}\vspace{6pt}
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67443
diff changeset
   136
\cmd{fun} \<open>f :: \<tau>\<close>\\%
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   137
\cmd{where}\\%
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   138
\hspace*{2ex}{\it equations}\\%
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   139
\hspace*{2ex}\vdots\vspace*{6pt}
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   140
\end{minipage}\right]
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   141
\quad\equiv\quad
27026
3602b81665b5 Updated function tutorial.
krauss
parents: 26749
diff changeset
   142
\left[\;\begin{minipage}{0.48\textwidth}\vspace{6pt}
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67443
diff changeset
   143
\cmd{function} \<open>(\<close>\cmd{sequential}\<open>) f :: \<tau>\<close>\\%
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   144
\cmd{where}\\%
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   145
\hspace*{2ex}{\it equations}\\%
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   146
\hspace*{2ex}\vdots\\%
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67443
diff changeset
   147
\cmd{by} \<open>pat_completeness auto\<close>\\%
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67443
diff changeset
   148
\cmd{termination by} \<open>lexicographic_order\<close>\vspace{6pt}
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   149
\end{minipage}
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   150
\right]\]
21212
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
   151
22065
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   152
\begin{isamarkuptext}
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   153
  \vspace*{1em}
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   154
  \noindent Some details have now become explicit:
21212
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
   155
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
   156
  \begin{enumerate}
22065
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   157
  \item The \cmd{sequential} option enables the preprocessing of
23805
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
   158
  pattern overlaps which we already saw. Without this option, the equations
21212
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
   159
  must already be disjoint and complete. The automatic completion only
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   160
  works with constructor patterns.
21212
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
   161
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   162
  \item A function definition produces a proof obligation which
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   163
  expresses completeness and compatibility of patterns (we talk about
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67443
diff changeset
   164
  this later). The combination of the methods \<open>pat_completeness\<close> and
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67443
diff changeset
   165
  \<open>auto\<close> is used to solve this proof obligation.
21212
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
   166
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
   167
  \item A termination proof follows the definition, started by the
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   168
  \cmd{termination} command. This will be explained in \S\ref{termination}.
22065
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   169
 \end{enumerate}
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   170
  Whenever a \cmd{fun} command fails, it is usually a good idea to
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   171
  expand the syntax to the more verbose \cmd{function} form, to see
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   172
  what is actually going on.
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   173
\<close>
21212
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
   174
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
   175
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   176
section \<open>Termination\<close>
21212
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
   177
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   178
text \<open>\label{termination}
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67443
diff changeset
   179
  The method \<open>lexicographic_order\<close> is the default method for
23805
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
   180
  termination proofs. It can prove termination of a
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   181
  certain class of functions by searching for a suitable lexicographic
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   182
  combination of size measures. Of course, not all functions have such
23805
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
   183
  a simple termination argument. For them, we can specify the termination
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
   184
  relation manually.
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   185
\<close>
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   186
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   187
subsection \<open>The {\tt relation} method\<close>
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   188
text\<open>
21212
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
   189
  Consider the following function, which sums up natural numbers up to
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67443
diff changeset
   190
  \<open>N\<close>, using a counter \<open>i\<close>:
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   191
\<close>
21212
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
   192
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
   193
function sum :: "nat \<Rightarrow> nat \<Rightarrow> nat"
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
   194
where
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
   195
  "sum i N = (if i > N then 0 else i + sum (Suc i) N)"
22065
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   196
by pat_completeness auto
21212
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
   197
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   198
text \<open>
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67443
diff changeset
   199
  \noindent The \<open>lexicographic_order\<close> method fails on this example, because none of the
23805
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
   200
  arguments decreases in the recursive call, with respect to the standard size ordering.
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
   201
  To prove termination manually, we must provide a custom wellfounded relation.
21212
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
   202
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67443
diff changeset
   203
  The termination argument for \<open>sum\<close> is based on the fact that
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67443
diff changeset
   204
  the \emph{difference} between \<open>i\<close> and \<open>N\<close> gets
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67443
diff changeset
   205
  smaller in every step, and that the recursion stops when \<open>i\<close>
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67443
diff changeset
   206
  is greater than \<open>N\<close>. Phrased differently, the expression 
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67443
diff changeset
   207
  \<open>N + 1 - i\<close> always decreases.
21212
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
   208
22065
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   209
  We can use this expression as a measure function suitable to prove termination.
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   210
\<close>
21212
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
   211
27026
3602b81665b5 Updated function tutorial.
krauss
parents: 26749
diff changeset
   212
termination sum
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   213
apply (relation "measure (\<lambda>(i,N). N + 1 - i)")
21212
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
   214
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   215
text \<open>
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   216
  The \cmd{termination} command sets up the termination goal for the
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67443
diff changeset
   217
  specified function \<open>sum\<close>. If the function name is omitted, it
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   218
  implicitly refers to the last function definition.
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   219
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67443
diff changeset
   220
  The \<open>relation\<close> method takes a relation of
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   221
  type \<^typ>\<open>('a \<times> 'a) set\<close>, where \<^typ>\<open>'a\<close> is the argument type of
22065
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   222
  the function. If the function has multiple curried arguments, then
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   223
  these are packed together into a tuple, as it happened in the above
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   224
  example.
21212
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
   225
27026
3602b81665b5 Updated function tutorial.
krauss
parents: 26749
diff changeset
   226
  The predefined function @{term[source] "measure :: ('a \<Rightarrow> nat) \<Rightarrow> ('a \<times> 'a) set"} constructs a
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   227
  wellfounded relation from a mapping into the natural numbers (a
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   228
  \emph{measure function}). 
22065
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   229
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67443
diff changeset
   230
  After the invocation of \<open>relation\<close>, we must prove that (a)
22065
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   231
  the relation we supplied is wellfounded, and (b) that the arguments
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   232
  of recursive calls indeed decrease with respect to the
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   233
  relation:
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   234
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   235
  @{subgoals[display,indent=0]}
22065
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   236
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67443
diff changeset
   237
  These goals are all solved by \<open>auto\<close>:
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   238
\<close>
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   239
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   240
apply auto
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   241
done
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   242
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   243
text \<open>
22065
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   244
  Let us complicate the function a little, by adding some more
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   245
  recursive calls: 
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   246
\<close>
21212
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
   247
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
   248
function foo :: "nat \<Rightarrow> nat \<Rightarrow> nat"
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
   249
where
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
   250
  "foo i N = (if i > N 
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
   251
              then (if N = 0 then 0 else foo 0 (N - 1))
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
   252
              else i + foo (Suc i) N)"
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
   253
by pat_completeness auto
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
   254
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   255
text \<open>
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67443
diff changeset
   256
  When \<open>i\<close> has reached \<open>N\<close>, it starts at zero again
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67443
diff changeset
   257
  and \<open>N\<close> is decremented.
21212
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
   258
  This corresponds to a nested
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
   259
  loop where one index counts up and the other down. Termination can
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
   260
  be proved using a lexicographic combination of two measures, namely
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   261
  the value of \<open>N\<close> and the above difference. The \<^const>\<open>measures\<close> combinator generalizes \<open>measure\<close> by taking a
22065
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   262
  list of measure functions.  
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   263
\<close>
21212
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
   264
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
   265
termination 
22065
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   266
by (relation "measures [\<lambda>(i, N). N, \<lambda>(i,N). N + 1 - i]") auto
21212
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
   267
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67443
diff changeset
   268
subsection \<open>How \<open>lexicographic_order\<close> works\<close>
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   269
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   270
(*fun fails :: "nat \<Rightarrow> nat list \<Rightarrow> nat"
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   271
where
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   272
  "fails a [] = a"
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   273
| "fails a (x#xs) = fails (x + a) (x # xs)"
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   274
*)
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   275
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   276
text \<open>
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   277
  To see how the automatic termination proofs work, let's look at an
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   278
  example where it fails\footnote{For a detailed discussion of the
58620
7435b6a3f72e more antiquotations;
wenzelm
parents: 58310
diff changeset
   279
  termination prover, see @{cite bulwahnKN07}}:
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   280
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   281
\end{isamarkuptext}  
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67443
diff changeset
   282
\cmd{fun} \<open>fails :: "nat \<Rightarrow> nat list \<Rightarrow> nat"\<close>\\%
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   283
\cmd{where}\\%
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67443
diff changeset
   284
\hspace*{2ex}\<open>"fails a [] = a"\<close>\\%
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67443
diff changeset
   285
|\hspace*{1.5ex}\<open>"fails a (x#xs) = fails (x + a) (x#xs)"\<close>\\
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   286
\begin{isamarkuptext}
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   287
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   288
\noindent Isabelle responds with the following error:
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   289
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   290
\begin{isabelle}
23805
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
   291
*** Unfinished subgoals:\newline
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
   292
*** (a, 1, <):\newline
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67443
diff changeset
   293
*** \ 1.~\<open>\<And>x. x = 0\<close>\newline
23805
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
   294
*** (a, 1, <=):\newline
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
   295
*** \ 1.~False\newline
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
   296
*** (a, 2, <):\newline
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
   297
*** \ 1.~False\newline
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   298
*** Calls:\newline
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67443
diff changeset
   299
*** a) \<open>(a, x # xs) -->> (x + a, x # xs)\<close>\newline
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   300
*** Measures:\newline
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67443
diff changeset
   301
*** 1) \<open>\<lambda>x. size (fst x)\<close>\newline
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67443
diff changeset
   302
*** 2) \<open>\<lambda>x. size (snd x)\<close>\newline
23805
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
   303
*** Result matrix:\newline
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
   304
*** \ \ \ \ 1\ \ 2  \newline
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
   305
*** a:  ?   <= \newline
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
   306
*** Could not find lexicographic termination order.\newline
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   307
*** At command "fun".\newline
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   308
\end{isabelle}
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   309
\<close>
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   310
text \<open>
28040
f47b4af3716a corrected some typos
krauss
parents: 27026
diff changeset
   311
  The key to this error message is the matrix at the bottom. The rows
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   312
  of that matrix correspond to the different recursive calls (In our
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   313
  case, there is just one). The columns are the function's arguments 
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   314
  (expressed through different measure functions, which map the
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   315
  argument tuple to a natural number). 
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   316
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   317
  The contents of the matrix summarize what is known about argument
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67443
diff changeset
   318
  descents: The second argument has a weak descent (\<open><=\<close>) at the
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   319
  recursive call, and for the first argument nothing could be proved,
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67443
diff changeset
   320
  which is expressed by \<open>?\<close>. In general, there are the values
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67443
diff changeset
   321
  \<open><\<close>, \<open><=\<close> and \<open>?\<close>.
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   322
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   323
  For the failed proof attempts, the unfinished subgoals are also
23805
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
   324
  printed. Looking at these will often point to a missing lemma.
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   325
\<close>
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   326
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67443
diff changeset
   327
subsection \<open>The \<open>size_change\<close> method\<close>
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   328
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   329
text \<open>
33856
14a658faadb6 mentioned method size_change in function tutorial
krauss
parents: 30226
diff changeset
   330
  Some termination goals that are beyond the powers of
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67443
diff changeset
   331
  \<open>lexicographic_order\<close> can be solved automatically by the
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67443
diff changeset
   332
  more powerful \<open>size_change\<close> method, which uses a variant of
33856
14a658faadb6 mentioned method size_change in function tutorial
krauss
parents: 30226
diff changeset
   333
  the size-change principle, together with some other
14a658faadb6 mentioned method size_change in function tutorial
krauss
parents: 30226
diff changeset
   334
  techniques. While the details are discussed
58620
7435b6a3f72e more antiquotations;
wenzelm
parents: 58310
diff changeset
   335
  elsewhere @{cite krauss_phd},
33856
14a658faadb6 mentioned method size_change in function tutorial
krauss
parents: 30226
diff changeset
   336
  here are a few typical situations where
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67443
diff changeset
   337
  \<open>lexicographic_order\<close> has difficulties and \<open>size_change\<close>
33856
14a658faadb6 mentioned method size_change in function tutorial
krauss
parents: 30226
diff changeset
   338
  may be worth a try:
14a658faadb6 mentioned method size_change in function tutorial
krauss
parents: 30226
diff changeset
   339
  \begin{itemize}
14a658faadb6 mentioned method size_change in function tutorial
krauss
parents: 30226
diff changeset
   340
  \item Arguments are permuted in a recursive call.
14a658faadb6 mentioned method size_change in function tutorial
krauss
parents: 30226
diff changeset
   341
  \item Several mutually recursive functions with multiple arguments.
14a658faadb6 mentioned method size_change in function tutorial
krauss
parents: 30226
diff changeset
   342
  \item Unusual control flow (e.g., when some recursive calls cannot
14a658faadb6 mentioned method size_change in function tutorial
krauss
parents: 30226
diff changeset
   343
  occur in sequence).
14a658faadb6 mentioned method size_change in function tutorial
krauss
parents: 30226
diff changeset
   344
  \end{itemize}
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   345
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67443
diff changeset
   346
  Loading the theory \<open>Multiset\<close> makes the \<open>size_change\<close>
33856
14a658faadb6 mentioned method size_change in function tutorial
krauss
parents: 30226
diff changeset
   347
  method a bit stronger: it can then use multiset orders internally.
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   348
\<close>
21212
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
   349
70276
910dc065b869 documentation for termination_simp attribute
krauss
parents: 69597
diff changeset
   350
subsection \<open>Configuring simplification rules for termination proofs\<close>
910dc065b869 documentation for termination_simp attribute
krauss
parents: 69597
diff changeset
   351
910dc065b869 documentation for termination_simp attribute
krauss
parents: 69597
diff changeset
   352
text \<open>
910dc065b869 documentation for termination_simp attribute
krauss
parents: 69597
diff changeset
   353
  Since both \<open>lexicographic_order\<close> and \<open>size_change\<close> rely on the simplifier internally,
910dc065b869 documentation for termination_simp attribute
krauss
parents: 69597
diff changeset
   354
  there can sometimes be the need for adding additional simp rules to them.
910dc065b869 documentation for termination_simp attribute
krauss
parents: 69597
diff changeset
   355
  This can be done either as arguments to the methods themselves, or globally via the
910dc065b869 documentation for termination_simp attribute
krauss
parents: 69597
diff changeset
   356
  theorem attribute \<open>termination_simp\<close>, which is useful in rare cases.
910dc065b869 documentation for termination_simp attribute
krauss
parents: 69597
diff changeset
   357
\<close>
910dc065b869 documentation for termination_simp attribute
krauss
parents: 69597
diff changeset
   358
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   359
section \<open>Mutual Recursion\<close>
21212
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
   360
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   361
text \<open>
21212
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
   362
  If two or more functions call one another mutually, they have to be defined
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67443
diff changeset
   363
  in one step. Here are \<open>even\<close> and \<open>odd\<close>:
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   364
\<close>
21212
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
   365
22065
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   366
function even :: "nat \<Rightarrow> bool"
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   367
    and odd  :: "nat \<Rightarrow> bool"
21212
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
   368
where
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
   369
  "even 0 = True"
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
   370
| "odd 0 = False"
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
   371
| "even (Suc n) = odd n"
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
   372
| "odd (Suc n) = even n"
22065
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   373
by pat_completeness auto
21212
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
   374
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   375
text \<open>
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   376
  To eliminate the mutual dependencies, Isabelle internally
21212
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
   377
  creates a single function operating on the sum
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   378
  type \<^typ>\<open>nat + nat\<close>. Then, \<^const>\<open>even\<close> and \<^const>\<open>odd\<close> are
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   379
  defined as projections. Consequently, termination has to be proved
21212
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
   380
  simultaneously for both functions, by specifying a measure on the
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
   381
  sum type: 
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   382
\<close>
21212
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
   383
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
   384
termination 
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   385
by (relation "measure (\<lambda>x. case x of Inl n \<Rightarrow> n | Inr n \<Rightarrow> n)") auto
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   386
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   387
text \<open>
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67443
diff changeset
   388
  We could also have used \<open>lexicographic_order\<close>, which
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   389
  supports mutual recursive termination proofs to a certain extent.
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   390
\<close>
22065
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   391
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   392
subsection \<open>Induction for mutual recursion\<close>
22065
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   393
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   394
text \<open>
22065
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   395
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   396
  When functions are mutually recursive, proving properties about them
53107
57c7294eac0a more document antiquotations (for proper theorem names);
Christian Sternagel
parents: 50530
diff changeset
   397
  generally requires simultaneous induction. The induction rule @{thm [source] "even_odd.induct"}
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   398
  generated from the above definition reflects this.
22065
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   399
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   400
  Let us prove something about \<^const>\<open>even\<close> and \<^const>\<open>odd\<close>:
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   401
\<close>
22065
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   402
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   403
lemma even_odd_mod2:
22065
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   404
  "even n = (n mod 2 = 0)"
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   405
  "odd n = (n mod 2 = 1)"
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   406
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   407
text \<open>
22065
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   408
  We apply simultaneous induction, specifying the induction variable
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   409
  for both goals, separated by \cmd{and}:\<close>
22065
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   410
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   411
apply (induct n and n rule: even_odd.induct)
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   412
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   413
text \<open>
22065
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   414
  We get four subgoals, which correspond to the clauses in the
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   415
  definition of \<^const>\<open>even\<close> and \<^const>\<open>odd\<close>:
22065
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   416
  @{subgoals[display,indent=0]}
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   417
  Simplification solves the first two goals, leaving us with two
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67443
diff changeset
   418
  statements about the \<open>mod\<close> operation to prove:
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   419
\<close>
22065
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   420
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   421
apply simp_all
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   422
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   423
text \<open>
22065
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   424
  @{subgoals[display,indent=0]} 
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   425
23805
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
   426
  \noindent These can be handled by Isabelle's arithmetic decision procedures.
22065
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   427
  
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   428
\<close>
22065
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   429
23805
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
   430
apply arith
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
   431
apply arith
22065
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   432
done
21212
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
   433
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   434
text \<open>
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   435
  In proofs like this, the simultaneous induction is really essential:
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   436
  Even if we are just interested in one of the results, the other
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   437
  one is necessary to strengthen the induction hypothesis. If we leave
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   438
  out the statement about \<^const>\<open>odd\<close> and just write \<^term>\<open>True\<close> instead,
27026
3602b81665b5 Updated function tutorial.
krauss
parents: 26749
diff changeset
   439
  the same proof fails:
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   440
\<close>
22065
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   441
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   442
lemma failed_attempt:
22065
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   443
  "even n = (n mod 2 = 0)"
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   444
  "True"
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   445
apply (induct n rule: even_odd.induct)
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   446
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   447
text \<open>
22065
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   448
  \noindent Now the third subgoal is a dead end, since we have no
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   449
  useful induction hypothesis available:
22065
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   450
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   451
  @{subgoals[display,indent=0]} 
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   452
\<close>
22065
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   453
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   454
oops
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   455
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   456
section \<open>Elimination\<close>
54017
2a3c07f49615 basic documentation for function elimination rules and fun_cases
krauss
parents: 53107
diff changeset
   457
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   458
text \<open>
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67443
diff changeset
   459
  A definition of function \<open>f\<close> gives rise to two kinds of elimination rules. Rule \<open>f.cases\<close>
54017
2a3c07f49615 basic documentation for function elimination rules and fun_cases
krauss
parents: 53107
diff changeset
   460
  simply describes case analysis according to the patterns used in the definition:
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   461
\<close>
54017
2a3c07f49615 basic documentation for function elimination rules and fun_cases
krauss
parents: 53107
diff changeset
   462
2a3c07f49615 basic documentation for function elimination rules and fun_cases
krauss
parents: 53107
diff changeset
   463
fun list_to_option :: "'a list \<Rightarrow> 'a option"
2a3c07f49615 basic documentation for function elimination rules and fun_cases
krauss
parents: 53107
diff changeset
   464
where
2a3c07f49615 basic documentation for function elimination rules and fun_cases
krauss
parents: 53107
diff changeset
   465
  "list_to_option [x] = Some x"
2a3c07f49615 basic documentation for function elimination rules and fun_cases
krauss
parents: 53107
diff changeset
   466
| "list_to_option _ = None"
2a3c07f49615 basic documentation for function elimination rules and fun_cases
krauss
parents: 53107
diff changeset
   467
2a3c07f49615 basic documentation for function elimination rules and fun_cases
krauss
parents: 53107
diff changeset
   468
thm list_to_option.cases
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   469
text \<open>
54017
2a3c07f49615 basic documentation for function elimination rules and fun_cases
krauss
parents: 53107
diff changeset
   470
  @{thm[display] list_to_option.cases}
2a3c07f49615 basic documentation for function elimination rules and fun_cases
krauss
parents: 53107
diff changeset
   471
2a3c07f49615 basic documentation for function elimination rules and fun_cases
krauss
parents: 53107
diff changeset
   472
  Note that this rule does not mention the function at all, but only describes the cases used for
2a3c07f49615 basic documentation for function elimination rules and fun_cases
krauss
parents: 53107
diff changeset
   473
  defining it. In contrast, the rule @{thm[source] list_to_option.elims} also tell us what the function
2a3c07f49615 basic documentation for function elimination rules and fun_cases
krauss
parents: 53107
diff changeset
   474
  value will be in each case:
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   475
\<close>
54017
2a3c07f49615 basic documentation for function elimination rules and fun_cases
krauss
parents: 53107
diff changeset
   476
thm list_to_option.elims
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   477
text \<open>
54017
2a3c07f49615 basic documentation for function elimination rules and fun_cases
krauss
parents: 53107
diff changeset
   478
  @{thm[display] list_to_option.elims}
2a3c07f49615 basic documentation for function elimination rules and fun_cases
krauss
parents: 53107
diff changeset
   479
2a3c07f49615 basic documentation for function elimination rules and fun_cases
krauss
parents: 53107
diff changeset
   480
  \noindent
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   481
  This lets us eliminate an assumption of the form \<^prop>\<open>list_to_option xs = y\<close> and replace it
54017
2a3c07f49615 basic documentation for function elimination rules and fun_cases
krauss
parents: 53107
diff changeset
   482
  with the two cases, e.g.:
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   483
\<close>
54017
2a3c07f49615 basic documentation for function elimination rules and fun_cases
krauss
parents: 53107
diff changeset
   484
2a3c07f49615 basic documentation for function elimination rules and fun_cases
krauss
parents: 53107
diff changeset
   485
lemma "list_to_option xs = y \<Longrightarrow> P"
2a3c07f49615 basic documentation for function elimination rules and fun_cases
krauss
parents: 53107
diff changeset
   486
proof (erule list_to_option.elims)
2a3c07f49615 basic documentation for function elimination rules and fun_cases
krauss
parents: 53107
diff changeset
   487
  fix x assume "xs = [x]" "y = Some x" thus P sorry
2a3c07f49615 basic documentation for function elimination rules and fun_cases
krauss
parents: 53107
diff changeset
   488
next
2a3c07f49615 basic documentation for function elimination rules and fun_cases
krauss
parents: 53107
diff changeset
   489
  assume "xs = []" "y = None" thus P sorry
2a3c07f49615 basic documentation for function elimination rules and fun_cases
krauss
parents: 53107
diff changeset
   490
next
2a3c07f49615 basic documentation for function elimination rules and fun_cases
krauss
parents: 53107
diff changeset
   491
  fix a b xs' assume "xs = a # b # xs'" "y = None" thus P sorry
2a3c07f49615 basic documentation for function elimination rules and fun_cases
krauss
parents: 53107
diff changeset
   492
qed
2a3c07f49615 basic documentation for function elimination rules and fun_cases
krauss
parents: 53107
diff changeset
   493
2a3c07f49615 basic documentation for function elimination rules and fun_cases
krauss
parents: 53107
diff changeset
   494
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   495
text \<open>
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67443
diff changeset
   496
  Sometimes it is convenient to derive specialized versions of the \<open>elim\<close> rules above and
54017
2a3c07f49615 basic documentation for function elimination rules and fun_cases
krauss
parents: 53107
diff changeset
   497
  keep them around as facts explicitly. For example, it is natural to show that if 
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   498
  \<^prop>\<open>list_to_option xs = Some y\<close>, then \<^term>\<open>xs\<close> must be a singleton. The command 
54017
2a3c07f49615 basic documentation for function elimination rules and fun_cases
krauss
parents: 53107
diff changeset
   499
  \cmd{fun\_cases} derives such facts automatically, by instantiating and simplifying the general 
2a3c07f49615 basic documentation for function elimination rules and fun_cases
krauss
parents: 53107
diff changeset
   500
  elimination rules given some pattern:
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   501
\<close>
54017
2a3c07f49615 basic documentation for function elimination rules and fun_cases
krauss
parents: 53107
diff changeset
   502
2a3c07f49615 basic documentation for function elimination rules and fun_cases
krauss
parents: 53107
diff changeset
   503
fun_cases list_to_option_SomeE[elim]: "list_to_option xs = Some y"
2a3c07f49615 basic documentation for function elimination rules and fun_cases
krauss
parents: 53107
diff changeset
   504
2a3c07f49615 basic documentation for function elimination rules and fun_cases
krauss
parents: 53107
diff changeset
   505
thm list_to_option_SomeE
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   506
text \<open>
54017
2a3c07f49615 basic documentation for function elimination rules and fun_cases
krauss
parents: 53107
diff changeset
   507
  @{thm[display] list_to_option_SomeE}
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   508
\<close>
54017
2a3c07f49615 basic documentation for function elimination rules and fun_cases
krauss
parents: 53107
diff changeset
   509
2a3c07f49615 basic documentation for function elimination rules and fun_cases
krauss
parents: 53107
diff changeset
   510
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   511
section \<open>General pattern matching\<close>
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   512
text\<open>\label{genpats}\<close>
22065
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   513
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   514
subsection \<open>Avoiding automatic pattern splitting\<close>
22065
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   515
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   516
text \<open>
22065
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   517
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   518
  Up to now, we used pattern matching only on datatypes, and the
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   519
  patterns were always disjoint and complete, and if they weren't,
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   520
  they were made disjoint automatically like in the definition of
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   521
  \<^const>\<open>sep\<close> in \S\ref{patmatch}.
22065
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   522
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   523
  This automatic splitting can significantly increase the number of
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   524
  equations involved, and this is not always desirable. The following
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   525
  example shows the problem:
22065
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   526
  
23805
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
   527
  Suppose we are modeling incomplete knowledge about the world by a
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   528
  three-valued datatype, which has values \<^term>\<open>T\<close>, \<^term>\<open>F\<close>
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   529
  and \<^term>\<open>X\<close> for true, false and uncertain propositions, respectively. 
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   530
\<close>
22065
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   531
58310
91ea607a34d8 updated news
blanchet
parents: 58305
diff changeset
   532
datatype P3 = T | F | X
22065
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   533
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   534
text \<open>\noindent Then the conjunction of such values can be defined as follows:\<close>
22065
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   535
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   536
fun And :: "P3 \<Rightarrow> P3 \<Rightarrow> P3"
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   537
where
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   538
  "And T p = p"
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   539
| "And p T = p"
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   540
| "And p F = F"
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   541
| "And F p = F"
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   542
| "And X X = X"
21212
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
   543
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
   544
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   545
text \<open>
22065
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   546
  This definition is useful, because the equations can directly be used
28040
f47b4af3716a corrected some typos
krauss
parents: 27026
diff changeset
   547
  as simplification rules. But the patterns overlap: For example,
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   548
  the expression \<^term>\<open>And T T\<close> is matched by both the first and
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   549
  the second equation. By default, Isabelle makes the patterns disjoint by
22065
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   550
  splitting them up, producing instances:
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   551
\<close>
22065
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   552
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   553
thm And.simps
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   554
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   555
text \<open>
22065
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   556
  @{thm[indent=4] And.simps}
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   557
  
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   558
  \vspace*{1em}
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   559
  \noindent There are several problems with this:
22065
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   560
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   561
  \begin{enumerate}
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   562
  \item If the datatype has many constructors, there can be an
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   563
  explosion of equations. For \<^const>\<open>And\<close>, we get seven instead of
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   564
  five equations, which can be tolerated, but this is just a small
22065
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   565
  example.
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   566
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   567
  \item Since splitting makes the equations \qt{less general}, they
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   568
  do not always match in rewriting. While the term \<^term>\<open>And x F\<close>
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   569
  can be simplified to \<^term>\<open>F\<close> with the original equations, a
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   570
  (manual) case split on \<^term>\<open>x\<close> is now necessary.
22065
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   571
53107
57c7294eac0a more document antiquotations (for proper theorem names);
Christian Sternagel
parents: 50530
diff changeset
   572
  \item The splitting also concerns the induction rule @{thm [source]
22065
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   573
  "And.induct"}. Instead of five premises it now has seven, which
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   574
  means that our induction proofs will have more cases.
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   575
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   576
  \item In general, it increases clarity if we get the same definition
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   577
  back which we put in.
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   578
  \end{enumerate}
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   579
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   580
  If we do not want the automatic splitting, we can switch it off by
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   581
  leaving out the \cmd{sequential} option. However, we will have to
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   582
  prove that our pattern matching is consistent\footnote{This prevents
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   583
  us from defining something like \<^term>\<open>f x = True\<close> and \<^term>\<open>f x
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   584
  = False\<close> simultaneously.}:
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   585
\<close>
22065
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   586
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   587
function And2 :: "P3 \<Rightarrow> P3 \<Rightarrow> P3"
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   588
where
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   589
  "And2 T p = p"
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   590
| "And2 p T = p"
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   591
| "And2 p F = F"
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   592
| "And2 F p = F"
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   593
| "And2 X X = X"
22065
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   594
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   595
text \<open>
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   596
  \noindent Now let's look at the proof obligations generated by a
22065
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   597
  function definition. In this case, they are:
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   598
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   599
  @{subgoals[display,indent=0]}\vspace{-1.2em}\hspace{3cm}\vdots\vspace{1.2em}
22065
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   600
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   601
  The first subgoal expresses the completeness of the patterns. It has
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   602
  the form of an elimination rule and states that every \<^term>\<open>x\<close> of
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   603
  the function's input type must match at least one of the patterns\footnote{Completeness could
22065
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   604
  be equivalently stated as a disjunction of existential statements: 
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   605
\<^term>\<open>(\<exists>p. x = (T, p)) \<or> (\<exists>p. x = (p, T)) \<or> (\<exists>p. x = (p, F)) \<or>
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   606
  (\<exists>p. x = (F, p)) \<or> (x = (X, X))\<close>, and you can use the method \<open>atomize_elim\<close> to get that form instead.}. If the patterns just involve
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67443
diff changeset
   607
  datatypes, we can solve it with the \<open>pat_completeness\<close>
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   608
  method:
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   609
\<close>
22065
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   610
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   611
apply pat_completeness
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   612
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   613
text \<open>
22065
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   614
  The remaining subgoals express \emph{pattern compatibility}. We do
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   615
  allow that an input value matches multiple patterns, but in this
22065
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   616
  case, the result (i.e.~the right hand sides of the equations) must
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   617
  also be equal. For each pair of two patterns, there is one such
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   618
  subgoal. Usually this needs injectivity of the constructors, which
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67443
diff changeset
   619
  is used automatically by \<open>auto\<close>.
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   620
\<close>
22065
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   621
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   622
by auto
43042
0f9534b7ea75 function tutorial: do not omit termination proof, even when discussing other things
krauss
parents: 41847
diff changeset
   623
termination by (relation "{}") simp
21212
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
   624
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
   625
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   626
subsection \<open>Non-constructor patterns\<close>
21212
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
   627
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   628
text \<open>
23805
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
   629
  Most of Isabelle's basic types take the form of inductive datatypes,
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
   630
  and usually pattern matching works on the constructors of such types. 
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
   631
  However, this need not be always the case, and the \cmd{function}
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
   632
  command handles other kind of patterns, too.
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   633
23805
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
   634
  One well-known instance of non-constructor patterns are
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   635
  so-called \emph{$n+k$-patterns}, which are a little controversial in
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   636
  the functional programming world. Here is the initial fibonacci
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   637
  example with $n+k$-patterns:
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   638
\<close>
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   639
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   640
function fib2 :: "nat \<Rightarrow> nat"
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   641
where
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   642
  "fib2 0 = 1"
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   643
| "fib2 1 = 1"
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   644
| "fib2 (n + 2) = fib2 n + fib2 (Suc n)"
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   645
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   646
text \<open>
23805
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
   647
  This kind of matching is again justified by the proof of pattern
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
   648
  completeness and compatibility. 
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   649
  The proof obligation for pattern completeness states that every natural number is
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   650
  either \<^term>\<open>0::nat\<close>, \<^term>\<open>1::nat\<close> or \<^term>\<open>n +
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   651
  (2::nat)\<close>:
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   652
39752
06fc1a79b4bf removed unnecessary reference poking (cf. f45d332a90e3)
krauss
parents: 33856
diff changeset
   653
  @{subgoals[display,indent=0,goals_limit=1]}
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   654
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   655
  This is an arithmetic triviality, but unfortunately the
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67443
diff changeset
   656
  \<open>arith\<close> method cannot handle this specific form of an
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67443
diff changeset
   657
  elimination rule. However, we can use the method \<open>atomize_elim\<close> to do an ad-hoc conversion to a disjunction of
28040
f47b4af3716a corrected some typos
krauss
parents: 27026
diff changeset
   658
  existentials, which can then be solved by the arithmetic decision procedure.
23805
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
   659
  Pattern compatibility and termination are automatic as usual.
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   660
\<close>
26580
c3e597a476fd Generic conversion and tactic "atomize_elim" to convert elimination rules
krauss
parents: 25688
diff changeset
   661
apply atomize_elim
23805
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
   662
apply arith
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   663
apply auto
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   664
done
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   665
termination by lexicographic_order
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   666
text \<open>
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   667
  We can stretch the notion of pattern matching even more. The
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   668
  following function is not a sensible functional program, but a
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   669
  perfectly valid mathematical definition:
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   670
\<close>
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   671
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   672
function ev :: "nat \<Rightarrow> bool"
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   673
where
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   674
  "ev (2 * n) = True"
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   675
| "ev (2 * n + 1) = False"
26580
c3e597a476fd Generic conversion and tactic "atomize_elim" to convert elimination rules
krauss
parents: 25688
diff changeset
   676
apply atomize_elim
23805
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
   677
by arith+
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   678
termination by (relation "{}") simp
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   679
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   680
text \<open>
27026
3602b81665b5 Updated function tutorial.
krauss
parents: 26749
diff changeset
   681
  This general notion of pattern matching gives you a certain freedom
3602b81665b5 Updated function tutorial.
krauss
parents: 26749
diff changeset
   682
  in writing down specifications. However, as always, such freedom should
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   683
  be used with care:
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   684
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   685
  If we leave the area of constructor
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   686
  patterns, we have effectively departed from the world of functional
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   687
  programming. This means that it is no longer possible to use the
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   688
  code generator, and expect it to generate ML code for our
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   689
  definitions. Also, such a specification might not work very well together with
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   690
  simplification. Your mileage may vary.
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   691
\<close>
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   692
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   693
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   694
subsection \<open>Conditional equations\<close>
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   695
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   696
text \<open>
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   697
  The function package also supports conditional equations, which are
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   698
  similar to guards in a language like Haskell. Here is Euclid's
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   699
  algorithm written with conditional patterns\footnote{Note that the
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   700
  patterns are also overlapping in the base case}:
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   701
\<close>
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   702
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   703
function gcd :: "nat \<Rightarrow> nat \<Rightarrow> nat"
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   704
where
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   705
  "gcd x 0 = x"
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   706
| "gcd 0 y = y"
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   707
| "x < y \<Longrightarrow> gcd (Suc x) (Suc y) = gcd (Suc x) (y - x)"
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   708
| "\<not> x < y \<Longrightarrow> gcd (Suc x) (Suc y) = gcd (x - y) (Suc y)"
26580
c3e597a476fd Generic conversion and tactic "atomize_elim" to convert elimination rules
krauss
parents: 25688
diff changeset
   709
by (atomize_elim, auto, arith)
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   710
termination by lexicographic_order
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   711
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   712
text \<open>
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   713
  By now, you can probably guess what the proof obligations for the
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   714
  pattern completeness and compatibility look like. 
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   715
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   716
  Again, functions with conditional patterns are not supported by the
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   717
  code generator.
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   718
\<close>
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   719
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   720
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   721
subsection \<open>Pattern matching on strings\<close>
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   722
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   723
text \<open>
23805
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
   724
  As strings (as lists of characters) are normal datatypes, pattern
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   725
  matching on them is possible, but somewhat problematic. Consider the
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   726
  following definition:
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   727
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   728
\end{isamarkuptext}
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67443
diff changeset
   729
\noindent\cmd{fun} \<open>check :: "string \<Rightarrow> bool"\<close>\\%
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   730
\cmd{where}\\%
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67443
diff changeset
   731
\hspace*{2ex}\<open>"check (''good'') = True"\<close>\\%
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67443
diff changeset
   732
\<open>| "check s = False"\<close>
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   733
\begin{isamarkuptext}
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   734
23805
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
   735
  \noindent An invocation of the above \cmd{fun} command does not
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   736
  terminate. What is the problem? Strings are lists of characters, and
23805
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
   737
  characters are a datatype with a lot of constructors. Splitting the
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   738
  catch-all pattern thus leads to an explosion of cases, which cannot
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   739
  be handled by Isabelle.
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   740
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   741
  There are two things we can do here. Either we write an explicit
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67443
diff changeset
   742
  \<open>if\<close> on the right hand side, or we can use conditional patterns:
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   743
\<close>
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   744
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   745
function check :: "string \<Rightarrow> bool"
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   746
where
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   747
  "check (''good'') = True"
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   748
| "s \<noteq> ''good'' \<Longrightarrow> check s = False"
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   749
by auto
43042
0f9534b7ea75 function tutorial: do not omit termination proof, even when discussing other things
krauss
parents: 41847
diff changeset
   750
termination by (relation "{}") simp
22065
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   751
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   752
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   753
section \<open>Partiality\<close>
22065
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   754
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   755
text \<open>
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   756
  In HOL, all functions are total. A function \<^term>\<open>f\<close> applied to
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   757
  \<^term>\<open>x\<close> always has the value \<^term>\<open>f x\<close>, and there is no notion
22065
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   758
  of undefinedness. 
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   759
  This is why we have to do termination
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   760
  proofs when defining functions: The proof justifies that the
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   761
  function can be defined by wellfounded recursion.
22065
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
   762
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   763
  However, the \cmd{function} package does support partiality to a
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   764
  certain extent. Let's look at the following function which looks
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   765
  for a zero of a given function f. 
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   766
\<close>
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   767
41846
b368a7aee46a removed support for tail-recursion from function package (now implemented by partial_function)
krauss
parents: 39754
diff changeset
   768
function (*<*)(domintros)(*>*)findzero :: "(nat \<Rightarrow> nat) \<Rightarrow> nat \<Rightarrow> nat"
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   769
where
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   770
  "findzero f n = (if f n = 0 then n else findzero f (Suc n))"
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   771
by pat_completeness auto
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   772
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   773
text \<open>
23805
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
   774
  \noindent Clearly, any attempt of a termination proof must fail. And without
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67443
diff changeset
   775
  that, we do not get the usual rules \<open>findzero.simps\<close> and 
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67443
diff changeset
   776
  \<open>findzero.induct\<close>. So what was the definition good for at all?
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   777
\<close>
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   778
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   779
subsection \<open>Domain predicates\<close>
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   780
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   781
text \<open>
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   782
  The trick is that Isabelle has not only defined the function \<^const>\<open>findzero\<close>, but also
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   783
  a predicate \<^term>\<open>findzero_dom\<close> that characterizes the values where the function
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   784
  terminates: the \emph{domain} of the function. If we treat a
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   785
  partial function just as a total function with an additional domain
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   786
  predicate, we can derive simplification and
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   787
  induction rules as we do for total functions. They are guarded
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67443
diff changeset
   788
  by domain conditions and are called \<open>psimps\<close> and \<open>pinduct\<close>: 
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   789
\<close>
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   790
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   791
text \<open>
23805
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
   792
  \noindent\begin{minipage}{0.79\textwidth}@{thm[display,margin=85] findzero.psimps}\end{minipage}
53107
57c7294eac0a more document antiquotations (for proper theorem names);
Christian Sternagel
parents: 50530
diff changeset
   793
  \hfill(@{thm [source] "findzero.psimps"})
23805
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
   794
  \vspace{1em}
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   795
23805
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
   796
  \noindent\begin{minipage}{0.79\textwidth}@{thm[display,margin=85] findzero.pinduct}\end{minipage}
53107
57c7294eac0a more document antiquotations (for proper theorem names);
Christian Sternagel
parents: 50530
diff changeset
   797
  \hfill(@{thm [source] "findzero.pinduct"})
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   798
\<close>
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   799
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   800
text \<open>
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   801
  Remember that all we
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   802
  are doing here is use some tricks to make a total function appear
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   803
  as if it was partial. We can still write the term \<^term>\<open>findzero
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   804
  (\<lambda>x. 1) 0\<close> and like any other term of type \<^typ>\<open>nat\<close> it is equal
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   805
  to some natural number, although we might not be able to find out
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   806
  which one. The function is \emph{underdefined}.
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   807
23805
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
   808
  But it is defined enough to prove something interesting about it. We
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   809
  can prove that if \<^term>\<open>findzero f n\<close>
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   810
  terminates, it indeed returns a zero of \<^term>\<open>f\<close>:
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   811
\<close>
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   812
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   813
lemma findzero_zero: "findzero_dom (f, n) \<Longrightarrow> f (findzero f n) = 0"
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   814
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   815
text \<open>\noindent We apply induction as usual, but using the partial induction
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   816
  rule:\<close>
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   817
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   818
apply (induct f n rule: findzero.pinduct)
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   819
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   820
text \<open>\noindent This gives the following subgoals:
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   821
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   822
  @{subgoals[display,indent=0]}
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   823
23805
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
   824
  \noindent The hypothesis in our lemma was used to satisfy the first premise in
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   825
  the induction rule. However, we also get \<^term>\<open>findzero_dom (f, n)\<close> as a local assumption in the induction step. This
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   826
  allows unfolding \<^term>\<open>findzero f n\<close> using the \<open>psimps\<close>
39754
150f831ce4a3 no longer declare .psimps rules as [simp].
krauss
parents: 39752
diff changeset
   827
  rule, and the rest is trivial.
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   828
\<close>
39754
150f831ce4a3 no longer declare .psimps rules as [simp].
krauss
parents: 39752
diff changeset
   829
apply (simp add: findzero.psimps)
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   830
done
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   831
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   832
text \<open>
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   833
  Proofs about partial functions are often not harder than for total
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   834
  functions. Fig.~\ref{findzero_isar} shows a slightly more
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   835
  complicated proof written in Isar. It is verbose enough to show how
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   836
  partiality comes into play: From the partial induction, we get an
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   837
  additional domain condition hypothesis. Observe how this condition
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   838
  is applied when calls to \<^term>\<open>findzero\<close> are unfolded.
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   839
\<close>
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   840
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   841
text_raw \<open>
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   842
\begin{figure}
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   843
\hrule\vspace{6pt}
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   844
\begin{minipage}{0.8\textwidth}
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   845
\isabellestyle{it}
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   846
\isastyle\isamarkuptrue
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   847
\<close>
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   848
lemma "\<lbrakk>findzero_dom (f, n); x \<in> {n ..< findzero f n}\<rbrakk> \<Longrightarrow> f x \<noteq> 0"
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   849
proof (induct rule: findzero.pinduct)
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   850
  fix f n assume dom: "findzero_dom (f, n)"
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   851
               and IH: "\<lbrakk>f n \<noteq> 0; x \<in> {Suc n ..< findzero f (Suc n)}\<rbrakk> \<Longrightarrow> f x \<noteq> 0"
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   852
               and x_range: "x \<in> {n ..< findzero f n}"
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   853
  have "f n \<noteq> 0"
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   854
  proof 
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   855
    assume "f n = 0"
39754
150f831ce4a3 no longer declare .psimps rules as [simp].
krauss
parents: 39752
diff changeset
   856
    with dom have "findzero f n = n" by (simp add: findzero.psimps)
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   857
    with x_range show False by auto
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   858
  qed
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   859
  
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   860
  from x_range have "x = n \<or> x \<in> {Suc n ..< findzero f n}" by auto
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   861
  thus "f x \<noteq> 0"
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   862
  proof
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   863
    assume "x = n"
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   864
    with \<open>f n \<noteq> 0\<close> show ?thesis by simp
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   865
  next
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   866
    assume "x \<in> {Suc n ..< findzero f n}"
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   867
    with dom and \<open>f n \<noteq> 0\<close> have "x \<in> {Suc n ..< findzero f (Suc n)}" by (simp add: findzero.psimps)
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   868
    with IH and \<open>f n \<noteq> 0\<close>
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   869
    show ?thesis by simp
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   870
  qed
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   871
qed
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   872
text_raw \<open>
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   873
\isamarkupfalse\isabellestyle{tt}
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   874
\end{minipage}\vspace{6pt}\hrule
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   875
\caption{A proof about a partial function}\label{findzero_isar}
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   876
\end{figure}
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   877
\<close>
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   878
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   879
subsection \<open>Partial termination proofs\<close>
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   880
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   881
text \<open>
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   882
  Now that we have proved some interesting properties about our
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   883
  function, we should turn to the domain predicate and see if it is
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   884
  actually true for some values. Otherwise we would have just proved
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   885
  lemmas with \<^term>\<open>False\<close> as a premise.
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   886
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67443
diff changeset
   887
  Essentially, we need some introduction rules for \<open>findzero_dom\<close>. The function package can prove such domain
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   888
  introduction rules automatically. But since they are not used very
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   889
  often (they are almost never needed if the function is total), this
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   890
  functionality is disabled by default for efficiency reasons. So we have to go
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67443
diff changeset
   891
  back and ask for them explicitly by passing the \<open>(domintros)\<close> option to the function package:
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   892
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   893
\vspace{1ex}
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67443
diff changeset
   894
\noindent\cmd{function} \<open>(domintros) findzero :: "(nat \<Rightarrow> nat) \<Rightarrow> nat \<Rightarrow> nat"\<close>\\%
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   895
\cmd{where}\isanewline%
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   896
\ \ \ldots\\
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   897
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67443
diff changeset
   898
  \noindent Now the package has proved an introduction rule for \<open>findzero_dom\<close>:
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   899
\<close>
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   900
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   901
thm findzero.domintros
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   902
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   903
text \<open>
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   904
  @{thm[display] findzero.domintros}
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   905
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   906
  Domain introduction rules allow to show that a given value lies in the
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   907
  domain of a function, if the arguments of all recursive calls
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   908
  are in the domain as well. They allow to do a \qt{single step} in a
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   909
  termination proof. Usually, you want to combine them with a suitable
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   910
  induction principle.
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   911
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   912
  Since our function increases its argument at recursive calls, we
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   913
  need an induction principle which works \qt{backwards}. We will use
53107
57c7294eac0a more document antiquotations (for proper theorem names);
Christian Sternagel
parents: 50530
diff changeset
   914
  @{thm [source] inc_induct}, which allows to do induction from a fixed number
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   915
  \qt{downwards}:
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   916
53107
57c7294eac0a more document antiquotations (for proper theorem names);
Christian Sternagel
parents: 50530
diff changeset
   917
  \begin{center}@{thm inc_induct}\hfill(@{thm [source] "inc_induct"})\end{center}
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   918
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   919
  Figure \ref{findzero_term} gives a detailed Isar proof of the fact
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67443
diff changeset
   920
  that \<open>findzero\<close> terminates if there is a zero which is greater
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   921
  or equal to \<^term>\<open>n\<close>. First we derive two useful rules which will
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   922
  solve the base case and the step case of the induction. The
23805
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
   923
  induction is then straightforward, except for the unusual induction
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   924
  principle.
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   925
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   926
\<close>
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   927
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   928
text_raw \<open>
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   929
\begin{figure}
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   930
\hrule\vspace{6pt}
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   931
\begin{minipage}{0.8\textwidth}
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   932
\isabellestyle{it}
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   933
\isastyle\isamarkuptrue
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   934
\<close>
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   935
lemma findzero_termination:
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   936
  assumes "x \<ge> n" and "f x = 0"
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   937
  shows "findzero_dom (f, n)"
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   938
proof - 
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   939
  have base: "findzero_dom (f, x)"
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   940
    by (rule findzero.domintros) (simp add:\<open>f x = 0\<close>)
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   941
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   942
  have step: "\<And>i. findzero_dom (f, Suc i) 
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   943
    \<Longrightarrow> findzero_dom (f, i)"
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   944
    by (rule findzero.domintros) simp
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   945
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   946
  from \<open>x \<ge> n\<close> show ?thesis
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   947
  proof (induct rule:inc_induct)
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   948
    show "findzero_dom (f, x)" by (rule base)
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   949
  next
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   950
    fix i assume "findzero_dom (f, Suc i)"
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   951
    thus "findzero_dom (f, i)" by (rule step)
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   952
  qed
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   953
qed      
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   954
text_raw \<open>
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   955
\isamarkupfalse\isabellestyle{tt}
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   956
\end{minipage}\vspace{6pt}\hrule
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67443
diff changeset
   957
\caption{Termination proof for \<open>findzero\<close>}\label{findzero_term}
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   958
\end{figure}
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   959
\<close>
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   960
      
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   961
text \<open>
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   962
  Again, the proof given in Fig.~\ref{findzero_term} has a lot of
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   963
  detail in order to explain the principles. Using more automation, we
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   964
  can also have a short proof:
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   965
\<close>
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   966
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   967
lemma findzero_termination_short:
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   968
  assumes zero: "x >= n" 
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   969
  assumes [simp]: "f x = 0"
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   970
  shows "findzero_dom (f, n)"
23805
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
   971
using zero
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
   972
by (induct rule:inc_induct) (auto intro: findzero.domintros)
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   973
    
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   974
text \<open>
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   975
  \noindent It is simple to combine the partial correctness result with the
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   976
  termination lemma:
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   977
\<close>
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   978
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   979
lemma findzero_total_correctness:
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   980
  "f x = 0 \<Longrightarrow> f (findzero f 0) = 0"
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   981
by (blast intro: findzero_zero findzero_termination)
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   982
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   983
subsection \<open>Definition of the domain predicate\<close>
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   984
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
   985
text \<open>
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   986
  Sometimes it is useful to know what the definition of the domain
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67443
diff changeset
   987
  predicate looks like. Actually, \<open>findzero_dom\<close> is just an
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   988
  abbreviation:
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   989
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   990
  @{abbrev[display] findzero_dom}
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   991
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   992
  The domain predicate is the \emph{accessible part} of a relation \<^const>\<open>findzero_rel\<close>, which was also created internally by the function
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   993
  package. \<^const>\<open>findzero_rel\<close> is just a normal
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
   994
  inductive predicate, so we can inspect its definition by
53107
57c7294eac0a more document antiquotations (for proper theorem names);
Christian Sternagel
parents: 50530
diff changeset
   995
  looking at the introduction rules @{thm [source] findzero_rel.intros}.
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   996
  In our case there is just a single rule:
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   997
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   998
  @{thm[display] findzero_rel.intros}
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
   999
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
  1000
  The predicate \<^const>\<open>findzero_rel\<close>
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1001
  describes the \emph{recursion relation} of the function
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1002
  definition. The recursion relation is a binary relation on
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1003
  the arguments of the function that relates each argument to its
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1004
  recursive calls. In general, there is one introduction rule for each
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1005
  recursive call.
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1006
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
  1007
  The predicate \<^term>\<open>Wellfounded.accp findzero_rel\<close> is the accessible part of
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1008
  that relation. An argument belongs to the accessible part, if it can
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67443
diff changeset
  1009
  be reached in a finite number of steps (cf.~its definition in \<open>Wellfounded.thy\<close>).
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1010
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1011
  Since the domain predicate is just an abbreviation, you can use
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
  1012
  lemmas for \<^const>\<open>Wellfounded.accp\<close> and \<^const>\<open>findzero_rel\<close> directly. Some
53107
57c7294eac0a more document antiquotations (for proper theorem names);
Christian Sternagel
parents: 50530
diff changeset
  1013
  lemmas which are occasionally useful are @{thm [source] accpI}, @{thm [source]
23805
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
  1014
  accp_downward}, and of course the introduction and elimination rules
53107
57c7294eac0a more document antiquotations (for proper theorem names);
Christian Sternagel
parents: 50530
diff changeset
  1015
  for the recursion relation @{thm [source] "findzero_rel.intros"} and @{thm
57c7294eac0a more document antiquotations (for proper theorem names);
Christian Sternagel
parents: 50530
diff changeset
  1016
  [source] "findzero_rel.cases"}.
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
  1017
\<close>
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1018
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
  1019
section \<open>Nested recursion\<close>
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1020
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
  1021
text \<open>
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1022
  Recursive calls which are nested in one another frequently cause
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1023
  complications, since their termination proof can depend on a partial
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1024
  correctness property of the function itself. 
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1025
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1026
  As a small example, we define the \qt{nested zero} function:
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
  1027
\<close>
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1028
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1029
function nz :: "nat \<Rightarrow> nat"
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1030
where
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1031
  "nz 0 = 0"
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1032
| "nz (Suc n) = nz (nz n)"
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1033
by pat_completeness auto
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1034
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
  1035
text \<open>
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1036
  If we attempt to prove termination using the identity measure on
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1037
  naturals, this fails:
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
  1038
\<close>
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1039
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1040
termination
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1041
  apply (relation "measure (\<lambda>n. n)")
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1042
  apply auto
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1043
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
  1044
text \<open>
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1045
  We get stuck with the subgoal
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1046
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1047
  @{subgoals[display]}
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1048
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
  1049
  Of course this statement is true, since we know that \<^const>\<open>nz\<close> is
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1050
  the zero function. And in fact we have no problem proving this
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1051
  property by induction.
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
  1052
\<close>
23805
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
  1053
(*<*)oops(*>*)
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1054
lemma nz_is_zero: "nz_dom n \<Longrightarrow> nz n = 0"
39754
150f831ce4a3 no longer declare .psimps rules as [simp].
krauss
parents: 39752
diff changeset
  1055
  by (induct rule:nz.pinduct) (auto simp: nz.psimps)
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1056
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
  1057
text \<open>
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1058
  We formulate this as a partial correctness lemma with the condition
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
  1059
  \<^term>\<open>nz_dom n\<close>. This allows us to prove it with the \<open>pinduct\<close> rule before we have proved termination. With this lemma,
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1060
  the termination proof works as expected:
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
  1061
\<close>
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1062
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1063
termination
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1064
  by (relation "measure (\<lambda>n. n)") (auto simp: nz_is_zero)
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1065
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
  1066
text \<open>
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1067
  As a general strategy, one should prove the statements needed for
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1068
  termination as a partial property first. Then they can be used to do
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1069
  the termination proof. This also works for less trivial
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
  1070
  examples. Figure \ref{f91} defines the 91-function, a well-known
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
  1071
  challenge problem due to John McCarthy, and proves its termination.
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
  1072
\<close>
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1073
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
  1074
text_raw \<open>
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1075
\begin{figure}
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
  1076
\hrule\vspace{6pt}
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1077
\begin{minipage}{0.8\textwidth}
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1078
\isabellestyle{it}
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1079
\isastyle\isamarkuptrue
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
  1080
\<close>
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1081
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
  1082
function f91 :: "nat \<Rightarrow> nat"
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1083
where
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1084
  "f91 n = (if 100 < n then n - 10 else f91 (f91 (n + 11)))"
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1085
by pat_completeness auto
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1086
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1087
lemma f91_estimate: 
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1088
  assumes trm: "f91_dom n" 
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1089
  shows "n < f91 n + 11"
39754
150f831ce4a3 no longer declare .psimps rules as [simp].
krauss
parents: 39752
diff changeset
  1090
using trm by induct (auto simp: f91.psimps)
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1091
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1092
termination
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1093
proof
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1094
  let ?R = "measure (\<lambda>x. 101 - x)"
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1095
  show "wf ?R" ..
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1096
67443
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 67406
diff changeset
  1097
  fix n :: nat assume "\<not> 100 < n" \<comment> \<open>Assumptions for both calls\<close>
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1098
67443
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 67406
diff changeset
  1099
  thus "(n + 11, n) \<in> ?R" by simp \<comment> \<open>Inner call\<close>
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1100
67443
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 67406
diff changeset
  1101
  assume inner_trm: "f91_dom (n + 11)" \<comment> \<open>Outer call\<close>
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1102
  with f91_estimate have "n + 11 < f91 (n + 11) + 11" .
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
  1103
  with \<open>\<not> 100 < n\<close> show "(f91 (n + 11), n) \<in> ?R" by simp
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1104
qed
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1105
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
  1106
text_raw \<open>
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1107
\isamarkupfalse\isabellestyle{tt}
23188
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
  1108
\end{minipage}
595a0e24bd8e updated
krauss
parents: 23003
diff changeset
  1109
\vspace{6pt}\hrule
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1110
\caption{McCarthy's 91-function}\label{f91}
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1111
\end{figure}
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
  1112
\<close>
22065
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
  1113
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
  1114
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
  1115
section \<open>Higher-Order Recursion\<close>
22065
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
  1116
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
  1117
text \<open>
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1118
  Higher-order recursion occurs when recursive calls
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
  1119
  are passed as arguments to higher-order combinators such as \<^const>\<open>map\<close>, \<^term>\<open>filter\<close> etc.
23805
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
  1120
  As an example, imagine a datatype of n-ary trees:
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
  1121
\<close>
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1122
58310
91ea607a34d8 updated news
blanchet
parents: 58305
diff changeset
  1123
datatype 'a tree = 
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1124
  Leaf 'a 
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1125
| Branch "'a tree list"
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1126
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1127
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
  1128
text \<open>\noindent We can define a function which swaps the left and right subtrees recursively, using the 
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
  1129
  list functions \<^const>\<open>rev\<close> and \<^const>\<open>map\<close>:\<close>
25688
6c24a82a98f1 temporarily fixed documentation due to changed size functions
krauss
parents: 25278
diff changeset
  1130
27026
3602b81665b5 Updated function tutorial.
krauss
parents: 26749
diff changeset
  1131
fun mirror :: "'a tree \<Rightarrow> 'a tree"
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1132
where
25278
3026df96941d changed "treemap" example to "mirror"
krauss
parents: 25091
diff changeset
  1133
  "mirror (Leaf n) = Leaf n"
3026df96941d changed "treemap" example to "mirror"
krauss
parents: 25091
diff changeset
  1134
| "mirror (Branch l) = Branch (rev (map mirror l))"
22065
cdd077905eee added sections on mutual induction and patterns
krauss
parents: 21346
diff changeset
  1135
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
  1136
text \<open>
27026
3602b81665b5 Updated function tutorial.
krauss
parents: 26749
diff changeset
  1137
  Although the definition is accepted without problems, let us look at the termination proof:
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
  1138
\<close>
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1139
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1140
termination proof
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
  1141
  text \<open>
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1142
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1143
  As usual, we have to give a wellfounded relation, such that the
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1144
  arguments of the recursive calls get smaller. But what exactly are
27026
3602b81665b5 Updated function tutorial.
krauss
parents: 26749
diff changeset
  1145
  the arguments of the recursive calls when mirror is given as an
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
  1146
  argument to \<^const>\<open>map\<close>? Isabelle gives us the
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1147
  subgoals
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1148
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1149
  @{subgoals[display,indent=0]} 
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1150
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
  1151
  So the system seems to know that \<^const>\<open>map\<close> only
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
  1152
  applies the recursive call \<^term>\<open>mirror\<close> to elements
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
  1153
  of \<^term>\<open>l\<close>, which is essential for the termination proof.
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1154
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
  1155
  This knowledge about \<^const>\<open>map\<close> is encoded in so-called congruence rules,
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1156
  which are special theorems known to the \cmd{function} command. The
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
  1157
  rule for \<^const>\<open>map\<close> is
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1158
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1159
  @{thm[display] map_cong}
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1160
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
  1161
  You can read this in the following way: Two applications of \<^const>\<open>map\<close> are equal, if the list arguments are equal and the functions
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1162
  coincide on the elements of the list. This means that for the value 
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
  1163
  \<^term>\<open>map f l\<close> we only have to know how \<^term>\<open>f\<close> behaves on
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
  1164
  the elements of \<^term>\<open>l\<close>.
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1165
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1166
  Usually, one such congruence rule is
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1167
  needed for each higher-order construct that is used when defining
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
  1168
  new functions. In fact, even basic functions like \<^const>\<open>If\<close> and \<^const>\<open>Let\<close> are handled by this mechanism. The congruence
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
  1169
  rule for \<^const>\<open>If\<close> states that the \<open>then\<close> branch is only
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67443
diff changeset
  1170
  relevant if the condition is true, and the \<open>else\<close> branch only if it
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1171
  is false:
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1172
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1173
  @{thm[display] if_cong}
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1174
  
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1175
  Congruence rules can be added to the
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
  1176
  function package by giving them the \<^term>\<open>fundef_cong\<close> attribute.
23003
4b0bf04a4d68 updated
krauss
parents: 22065
diff changeset
  1177
23805
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
  1178
  The constructs that are predefined in Isabelle, usually
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
  1179
  come with the respective congruence rules.
27026
3602b81665b5 Updated function tutorial.
krauss
parents: 26749
diff changeset
  1180
  But if you define your own higher-order functions, you may have to
3602b81665b5 Updated function tutorial.
krauss
parents: 26749
diff changeset
  1181
  state and prove the required congruence rules yourself, if you want to use your
23805
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
  1182
  functions in recursive definitions. 
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
  1183
\<close>
27026
3602b81665b5 Updated function tutorial.
krauss
parents: 26749
diff changeset
  1184
(*<*)oops(*>*)
23805
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
  1185
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
  1186
subsection \<open>Congruence Rules and Evaluation Order\<close>
23805
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
  1187
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
  1188
text \<open>
23805
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
  1189
  Higher order logic differs from functional programming languages in
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
  1190
  that it has no built-in notion of evaluation order. A program is
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
  1191
  just a set of equations, and it is not specified how they must be
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
  1192
  evaluated. 
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
  1193
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
  1194
  However for the purpose of function definition, we must talk about
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
  1195
  evaluation order implicitly, when we reason about termination.
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
  1196
  Congruence rules express that a certain evaluation order is
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
  1197
  consistent with the logical definition. 
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
  1198
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
  1199
  Consider the following function.
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
  1200
\<close>
23805
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
  1201
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
  1202
function f :: "nat \<Rightarrow> bool"
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
  1203
where
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
  1204
  "f n = (n = 0 \<or> f (n - 1))"
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
  1205
(*<*)by pat_completeness auto(*>*)
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
  1206
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
  1207
text \<open>
27026
3602b81665b5 Updated function tutorial.
krauss
parents: 26749
diff changeset
  1208
  For this definition, the termination proof fails. The default configuration
23805
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
  1209
  specifies no congruence rule for disjunction. We have to add a
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
  1210
  congruence rule that specifies left-to-right evaluation order:
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
  1211
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
  1212
  \vspace{1ex}
53107
57c7294eac0a more document antiquotations (for proper theorem names);
Christian Sternagel
parents: 50530
diff changeset
  1213
  \noindent @{thm disj_cong}\hfill(@{thm [source] "disj_cong"})
23805
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
  1214
  \vspace{1ex}
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
  1215
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
  1216
  Now the definition works without problems. Note how the termination
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
  1217
  proof depends on the extra condition that we get from the congruence
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
  1218
  rule.
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
  1219
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
  1220
  However, as evaluation is not a hard-wired concept, we
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
  1221
  could just turn everything around by declaring a different
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
  1222
  congruence rule. Then we can make the reverse definition:
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
  1223
\<close>
23805
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
  1224
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
  1225
lemma disj_cong2[fundef_cong]: 
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
  1226
  "(\<not> Q' \<Longrightarrow> P = P') \<Longrightarrow> (Q = Q') \<Longrightarrow> (P \<or> Q) = (P' \<or> Q')"
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
  1227
  by blast
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
  1228
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
  1229
fun f' :: "nat \<Rightarrow> bool"
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
  1230
where
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
  1231
  "f' n = (f' (n - 1) \<or> n = 0)"
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
  1232
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
  1233
text \<open>
23805
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
  1234
  \noindent These examples show that, in general, there is no \qt{best} set of
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
  1235
  congruence rules.
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
  1236
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
  1237
  However, such tweaking should rarely be necessary in
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
  1238
  practice, as most of the time, the default set of congruence rules
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
  1239
  works well.
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 59005
diff changeset
  1240
\<close>
23805
953eb3c5f793 updated
krauss
parents: 23188
diff changeset
  1241
21212
547224bf9348 Added a (stub of a) function tutorial
krauss
parents:
diff changeset
  1242
end