src/Doc/Functions/Functions.thy
changeset 61419 3c3f8b182e4b
parent 59005 1c54ebc68394
child 67406 23307fd33906
equal deleted inserted replaced
61418:44d9d55b9ef4 61419:3c3f8b182e4b
     6 
     6 
     7 theory Functions
     7 theory Functions
     8 imports Main
     8 imports Main
     9 begin
     9 begin
    10 
    10 
    11 section {* Function Definitions for Dummies *}
    11 section \<open>Function Definitions for Dummies\<close>
    12 
    12 
    13 text {*
    13 text \<open>
    14   In most cases, defining a recursive function is just as simple as other definitions:
    14   In most cases, defining a recursive function is just as simple as other definitions:
    15 *}
    15 \<close>
    16 
    16 
    17 fun fib :: "nat \<Rightarrow> nat"
    17 fun fib :: "nat \<Rightarrow> nat"
    18 where
    18 where
    19   "fib 0 = 1"
    19   "fib 0 = 1"
    20 | "fib (Suc 0) = 1"
    20 | "fib (Suc 0) = 1"
    21 | "fib (Suc (Suc n)) = fib n + fib (Suc n)"
    21 | "fib (Suc (Suc n)) = fib n + fib (Suc n)"
    22 
    22 
    23 text {*
    23 text \<open>
    24   The syntax is rather self-explanatory: We introduce a function by
    24   The syntax is rather self-explanatory: We introduce a function by
    25   giving its name, its type, 
    25   giving its name, its type, 
    26   and a set of defining recursive equations.
    26   and a set of defining recursive equations.
    27   If we leave out the type, the most general type will be
    27   If we leave out the type, the most general type will be
    28   inferred, which can sometimes lead to surprises: Since both @{term
    28   inferred, which can sometimes lead to surprises: Since both @{term
    29   "1::nat"} and @{text "+"} are overloaded, we would end up
    29   "1::nat"} and @{text "+"} are overloaded, we would end up
    30   with @{text "fib :: nat \<Rightarrow> 'a::{one,plus}"}.
    30   with @{text "fib :: nat \<Rightarrow> 'a::{one,plus}"}.
    31 *}
    31 \<close>
    32 
    32 
    33 text {*
    33 text \<open>
    34   The function always terminates, since its argument gets smaller in
    34   The function always terminates, since its argument gets smaller in
    35   every recursive call. 
    35   every recursive call. 
    36   Since HOL is a logic of total functions, termination is a
    36   Since HOL is a logic of total functions, termination is a
    37   fundamental requirement to prevent inconsistencies\footnote{From the
    37   fundamental requirement to prevent inconsistencies\footnote{From the
    38   \qt{definition} @{text "f(n) = f(n) + 1"} we could prove 
    38   \qt{definition} @{text "f(n) = f(n) + 1"} we could prove 
    39   @{text "0 = 1"} by subtracting @{text "f(n)"} on both sides.}.
    39   @{text "0 = 1"} by subtracting @{text "f(n)"} on both sides.}.
    40   Isabelle tries to prove termination automatically when a definition
    40   Isabelle tries to prove termination automatically when a definition
    41   is made. In \S\ref{termination}, we will look at cases where this
    41   is made. In \S\ref{termination}, we will look at cases where this
    42   fails and see what to do then.
    42   fails and see what to do then.
    43 *}
    43 \<close>
    44 
    44 
    45 subsection {* Pattern matching *}
    45 subsection \<open>Pattern matching\<close>
    46 
    46 
    47 text {* \label{patmatch}
    47 text \<open>\label{patmatch}
    48   Like in functional programming, we can use pattern matching to
    48   Like in functional programming, we can use pattern matching to
    49   define functions. At the moment we will only consider \emph{constructor
    49   define functions. At the moment we will only consider \emph{constructor
    50   patterns}, which only consist of datatype constructors and
    50   patterns}, which only consist of datatype constructors and
    51   variables. Furthermore, patterns must be linear, i.e.\ all variables
    51   variables. Furthermore, patterns must be linear, i.e.\ all variables
    52   on the left hand side of an equation must be distinct. In
    52   on the left hand side of an equation must be distinct. In
    53   \S\ref{genpats} we discuss more general pattern matching.
    53   \S\ref{genpats} we discuss more general pattern matching.
    54 
    54 
    55   If patterns overlap, the order of the equations is taken into
    55   If patterns overlap, the order of the equations is taken into
    56   account. The following function inserts a fixed element between any
    56   account. The following function inserts a fixed element between any
    57   two elements of a list:
    57   two elements of a list:
    58 *}
    58 \<close>
    59 
    59 
    60 fun sep :: "'a \<Rightarrow> 'a list \<Rightarrow> 'a list"
    60 fun sep :: "'a \<Rightarrow> 'a list \<Rightarrow> 'a list"
    61 where
    61 where
    62   "sep a (x#y#xs) = x # a # sep a (y # xs)"
    62   "sep a (x#y#xs) = x # a # sep a (y # xs)"
    63 | "sep a xs       = xs"
    63 | "sep a xs       = xs"
    64 
    64 
    65 text {* 
    65 text \<open>
    66   Overlapping patterns are interpreted as \qt{increments} to what is
    66   Overlapping patterns are interpreted as \qt{increments} to what is
    67   already there: The second equation is only meant for the cases where
    67   already there: The second equation is only meant for the cases where
    68   the first one does not match. Consequently, Isabelle replaces it
    68   the first one does not match. Consequently, Isabelle replaces it
    69   internally by the remaining cases, making the patterns disjoint:
    69   internally by the remaining cases, making the patterns disjoint:
    70 *}
    70 \<close>
    71 
    71 
    72 thm sep.simps
    72 thm sep.simps
    73 
    73 
    74 text {* @{thm [display] sep.simps[no_vars]} *}
    74 text \<open>@{thm [display] sep.simps[no_vars]}\<close>
    75 
    75 
    76 text {* 
    76 text \<open>
    77   \noindent The equations from function definitions are automatically used in
    77   \noindent The equations from function definitions are automatically used in
    78   simplification:
    78   simplification:
    79 *}
    79 \<close>
    80 
    80 
    81 lemma "sep 0 [1, 2, 3] = [1, 0, 2, 0, 3]"
    81 lemma "sep 0 [1, 2, 3] = [1, 0, 2, 0, 3]"
    82 by simp
    82 by simp
    83 
    83 
    84 subsection {* Induction *}
    84 subsection \<open>Induction\<close>
    85 
    85 
    86 text {*
    86 text \<open>
    87 
    87 
    88   Isabelle provides customized induction rules for recursive
    88   Isabelle provides customized induction rules for recursive
    89   functions. These rules follow the recursive structure of the
    89   functions. These rules follow the recursive structure of the
    90   definition. Here is the rule @{thm [source] sep.induct} arising from the
    90   definition. Here is the rule @{thm [source] sep.induct} arising from the
    91   above definition of @{const sep}:
    91   above definition of @{const sep}:
    93   @{thm [display] sep.induct}
    93   @{thm [display] sep.induct}
    94   
    94   
    95   We have a step case for list with at least two elements, and two
    95   We have a step case for list with at least two elements, and two
    96   base cases for the zero- and the one-element list. Here is a simple
    96   base cases for the zero- and the one-element list. Here is a simple
    97   proof about @{const sep} and @{const map}
    97   proof about @{const sep} and @{const map}
    98 *}
    98 \<close>
    99 
    99 
   100 lemma "map f (sep x ys) = sep (f x) (map f ys)"
   100 lemma "map f (sep x ys) = sep (f x) (map f ys)"
   101 apply (induct x ys rule: sep.induct)
   101 apply (induct x ys rule: sep.induct)
   102 
   102 
   103 text {*
   103 text \<open>
   104   We get three cases, like in the definition.
   104   We get three cases, like in the definition.
   105 
   105 
   106   @{subgoals [display]}
   106   @{subgoals [display]}
   107 *}
   107 \<close>
   108 
   108 
   109 apply auto 
   109 apply auto 
   110 done
   110 done
   111 text {*
   111 text \<open>
   112 
   112 
   113   With the \cmd{fun} command, you can define about 80\% of the
   113   With the \cmd{fun} command, you can define about 80\% of the
   114   functions that occur in practice. The rest of this tutorial explains
   114   functions that occur in practice. The rest of this tutorial explains
   115   the remaining 20\%.
   115   the remaining 20\%.
   116 *}
   116 \<close>
   117 
   117 
   118 
   118 
   119 section {* fun vs.\ function *}
   119 section \<open>fun vs.\ function\<close>
   120 
   120 
   121 text {* 
   121 text \<open>
   122   The \cmd{fun} command provides a
   122   The \cmd{fun} command provides a
   123   convenient shorthand notation for simple function definitions. In
   123   convenient shorthand notation for simple function definitions. In
   124   this mode, Isabelle tries to solve all the necessary proof obligations
   124   this mode, Isabelle tries to solve all the necessary proof obligations
   125   automatically. If any proof fails, the definition is
   125   automatically. If any proof fails, the definition is
   126   rejected. This can either mean that the definition is indeed faulty,
   126   rejected. This can either mean that the definition is indeed faulty,
   169   \cmd{termination} command. This will be explained in \S\ref{termination}.
   169   \cmd{termination} command. This will be explained in \S\ref{termination}.
   170  \end{enumerate}
   170  \end{enumerate}
   171   Whenever a \cmd{fun} command fails, it is usually a good idea to
   171   Whenever a \cmd{fun} command fails, it is usually a good idea to
   172   expand the syntax to the more verbose \cmd{function} form, to see
   172   expand the syntax to the more verbose \cmd{function} form, to see
   173   what is actually going on.
   173   what is actually going on.
   174  *}
   174 \<close>
   175 
   175 
   176 
   176 
   177 section {* Termination *}
   177 section \<open>Termination\<close>
   178 
   178 
   179 text {*\label{termination}
   179 text \<open>\label{termination}
   180   The method @{text "lexicographic_order"} is the default method for
   180   The method @{text "lexicographic_order"} is the default method for
   181   termination proofs. It can prove termination of a
   181   termination proofs. It can prove termination of a
   182   certain class of functions by searching for a suitable lexicographic
   182   certain class of functions by searching for a suitable lexicographic
   183   combination of size measures. Of course, not all functions have such
   183   combination of size measures. Of course, not all functions have such
   184   a simple termination argument. For them, we can specify the termination
   184   a simple termination argument. For them, we can specify the termination
   185   relation manually.
   185   relation manually.
   186 *}
   186 \<close>
   187 
   187 
   188 subsection {* The {\tt relation} method *}
   188 subsection \<open>The {\tt relation} method\<close>
   189 text{*
   189 text\<open>
   190   Consider the following function, which sums up natural numbers up to
   190   Consider the following function, which sums up natural numbers up to
   191   @{text "N"}, using a counter @{text "i"}:
   191   @{text "N"}, using a counter @{text "i"}:
   192 *}
   192 \<close>
   193 
   193 
   194 function sum :: "nat \<Rightarrow> nat \<Rightarrow> nat"
   194 function sum :: "nat \<Rightarrow> nat \<Rightarrow> nat"
   195 where
   195 where
   196   "sum i N = (if i > N then 0 else i + sum (Suc i) N)"
   196   "sum i N = (if i > N then 0 else i + sum (Suc i) N)"
   197 by pat_completeness auto
   197 by pat_completeness auto
   198 
   198 
   199 text {*
   199 text \<open>
   200   \noindent The @{text "lexicographic_order"} method fails on this example, because none of the
   200   \noindent The @{text "lexicographic_order"} method fails on this example, because none of the
   201   arguments decreases in the recursive call, with respect to the standard size ordering.
   201   arguments decreases in the recursive call, with respect to the standard size ordering.
   202   To prove termination manually, we must provide a custom wellfounded relation.
   202   To prove termination manually, we must provide a custom wellfounded relation.
   203 
   203 
   204   The termination argument for @{text "sum"} is based on the fact that
   204   The termination argument for @{text "sum"} is based on the fact that
   206   smaller in every step, and that the recursion stops when @{text "i"}
   206   smaller in every step, and that the recursion stops when @{text "i"}
   207   is greater than @{text "N"}. Phrased differently, the expression 
   207   is greater than @{text "N"}. Phrased differently, the expression 
   208   @{text "N + 1 - i"} always decreases.
   208   @{text "N + 1 - i"} always decreases.
   209 
   209 
   210   We can use this expression as a measure function suitable to prove termination.
   210   We can use this expression as a measure function suitable to prove termination.
   211 *}
   211 \<close>
   212 
   212 
   213 termination sum
   213 termination sum
   214 apply (relation "measure (\<lambda>(i,N). N + 1 - i)")
   214 apply (relation "measure (\<lambda>(i,N). N + 1 - i)")
   215 
   215 
   216 text {*
   216 text \<open>
   217   The \cmd{termination} command sets up the termination goal for the
   217   The \cmd{termination} command sets up the termination goal for the
   218   specified function @{text "sum"}. If the function name is omitted, it
   218   specified function @{text "sum"}. If the function name is omitted, it
   219   implicitly refers to the last function definition.
   219   implicitly refers to the last function definition.
   220 
   220 
   221   The @{text relation} method takes a relation of
   221   The @{text relation} method takes a relation of
   234   relation:
   234   relation:
   235 
   235 
   236   @{subgoals[display,indent=0]}
   236   @{subgoals[display,indent=0]}
   237 
   237 
   238   These goals are all solved by @{text "auto"}:
   238   These goals are all solved by @{text "auto"}:
   239 *}
   239 \<close>
   240 
   240 
   241 apply auto
   241 apply auto
   242 done
   242 done
   243 
   243 
   244 text {*
   244 text \<open>
   245   Let us complicate the function a little, by adding some more
   245   Let us complicate the function a little, by adding some more
   246   recursive calls: 
   246   recursive calls: 
   247 *}
   247 \<close>
   248 
   248 
   249 function foo :: "nat \<Rightarrow> nat \<Rightarrow> nat"
   249 function foo :: "nat \<Rightarrow> nat \<Rightarrow> nat"
   250 where
   250 where
   251   "foo i N = (if i > N 
   251   "foo i N = (if i > N 
   252               then (if N = 0 then 0 else foo 0 (N - 1))
   252               then (if N = 0 then 0 else foo 0 (N - 1))
   253               else i + foo (Suc i) N)"
   253               else i + foo (Suc i) N)"
   254 by pat_completeness auto
   254 by pat_completeness auto
   255 
   255 
   256 text {*
   256 text \<open>
   257   When @{text "i"} has reached @{text "N"}, it starts at zero again
   257   When @{text "i"} has reached @{text "N"}, it starts at zero again
   258   and @{text "N"} is decremented.
   258   and @{text "N"} is decremented.
   259   This corresponds to a nested
   259   This corresponds to a nested
   260   loop where one index counts up and the other down. Termination can
   260   loop where one index counts up and the other down. Termination can
   261   be proved using a lexicographic combination of two measures, namely
   261   be proved using a lexicographic combination of two measures, namely
   262   the value of @{text "N"} and the above difference. The @{const
   262   the value of @{text "N"} and the above difference. The @{const
   263   "measures"} combinator generalizes @{text "measure"} by taking a
   263   "measures"} combinator generalizes @{text "measure"} by taking a
   264   list of measure functions.  
   264   list of measure functions.  
   265 *}
   265 \<close>
   266 
   266 
   267 termination 
   267 termination 
   268 by (relation "measures [\<lambda>(i, N). N, \<lambda>(i,N). N + 1 - i]") auto
   268 by (relation "measures [\<lambda>(i, N). N, \<lambda>(i,N). N + 1 - i]") auto
   269 
   269 
   270 subsection {* How @{text "lexicographic_order"} works *}
   270 subsection \<open>How @{text "lexicographic_order"} works\<close>
   271 
   271 
   272 (*fun fails :: "nat \<Rightarrow> nat list \<Rightarrow> nat"
   272 (*fun fails :: "nat \<Rightarrow> nat list \<Rightarrow> nat"
   273 where
   273 where
   274   "fails a [] = a"
   274   "fails a [] = a"
   275 | "fails a (x#xs) = fails (x + a) (x # xs)"
   275 | "fails a (x#xs) = fails (x + a) (x # xs)"
   276 *)
   276 *)
   277 
   277 
   278 text {*
   278 text \<open>
   279   To see how the automatic termination proofs work, let's look at an
   279   To see how the automatic termination proofs work, let's look at an
   280   example where it fails\footnote{For a detailed discussion of the
   280   example where it fails\footnote{For a detailed discussion of the
   281   termination prover, see @{cite bulwahnKN07}}:
   281   termination prover, see @{cite bulwahnKN07}}:
   282 
   282 
   283 \end{isamarkuptext}  
   283 \end{isamarkuptext}  
   306 *** \ \ \ \ 1\ \ 2  \newline
   306 *** \ \ \ \ 1\ \ 2  \newline
   307 *** a:  ?   <= \newline
   307 *** a:  ?   <= \newline
   308 *** Could not find lexicographic termination order.\newline
   308 *** Could not find lexicographic termination order.\newline
   309 *** At command "fun".\newline
   309 *** At command "fun".\newline
   310 \end{isabelle}
   310 \end{isabelle}
   311 *}
   311 \<close>
   312 text {*
   312 text \<open>
   313   The key to this error message is the matrix at the bottom. The rows
   313   The key to this error message is the matrix at the bottom. The rows
   314   of that matrix correspond to the different recursive calls (In our
   314   of that matrix correspond to the different recursive calls (In our
   315   case, there is just one). The columns are the function's arguments 
   315   case, there is just one). The columns are the function's arguments 
   316   (expressed through different measure functions, which map the
   316   (expressed through different measure functions, which map the
   317   argument tuple to a natural number). 
   317   argument tuple to a natural number). 
   322   which is expressed by @{text "?"}. In general, there are the values
   322   which is expressed by @{text "?"}. In general, there are the values
   323   @{text "<"}, @{text "<="} and @{text "?"}.
   323   @{text "<"}, @{text "<="} and @{text "?"}.
   324 
   324 
   325   For the failed proof attempts, the unfinished subgoals are also
   325   For the failed proof attempts, the unfinished subgoals are also
   326   printed. Looking at these will often point to a missing lemma.
   326   printed. Looking at these will often point to a missing lemma.
   327 *}
   327 \<close>
   328 
   328 
   329 subsection {* The @{text size_change} method *}
   329 subsection \<open>The @{text size_change} method\<close>
   330 
   330 
   331 text {*
   331 text \<open>
   332   Some termination goals that are beyond the powers of
   332   Some termination goals that are beyond the powers of
   333   @{text lexicographic_order} can be solved automatically by the
   333   @{text lexicographic_order} can be solved automatically by the
   334   more powerful @{text size_change} method, which uses a variant of
   334   more powerful @{text size_change} method, which uses a variant of
   335   the size-change principle, together with some other
   335   the size-change principle, together with some other
   336   techniques. While the details are discussed
   336   techniques. While the details are discussed
   345   occur in sequence).
   345   occur in sequence).
   346   \end{itemize}
   346   \end{itemize}
   347 
   347 
   348   Loading the theory @{text Multiset} makes the @{text size_change}
   348   Loading the theory @{text Multiset} makes the @{text size_change}
   349   method a bit stronger: it can then use multiset orders internally.
   349   method a bit stronger: it can then use multiset orders internally.
   350 *}
   350 \<close>
   351 
   351 
   352 section {* Mutual Recursion *}
   352 section \<open>Mutual Recursion\<close>
   353 
   353 
   354 text {*
   354 text \<open>
   355   If two or more functions call one another mutually, they have to be defined
   355   If two or more functions call one another mutually, they have to be defined
   356   in one step. Here are @{text "even"} and @{text "odd"}:
   356   in one step. Here are @{text "even"} and @{text "odd"}:
   357 *}
   357 \<close>
   358 
   358 
   359 function even :: "nat \<Rightarrow> bool"
   359 function even :: "nat \<Rightarrow> bool"
   360     and odd  :: "nat \<Rightarrow> bool"
   360     and odd  :: "nat \<Rightarrow> bool"
   361 where
   361 where
   362   "even 0 = True"
   362   "even 0 = True"
   363 | "odd 0 = False"
   363 | "odd 0 = False"
   364 | "even (Suc n) = odd n"
   364 | "even (Suc n) = odd n"
   365 | "odd (Suc n) = even n"
   365 | "odd (Suc n) = even n"
   366 by pat_completeness auto
   366 by pat_completeness auto
   367 
   367 
   368 text {*
   368 text \<open>
   369   To eliminate the mutual dependencies, Isabelle internally
   369   To eliminate the mutual dependencies, Isabelle internally
   370   creates a single function operating on the sum
   370   creates a single function operating on the sum
   371   type @{typ "nat + nat"}. Then, @{const even} and @{const odd} are
   371   type @{typ "nat + nat"}. Then, @{const even} and @{const odd} are
   372   defined as projections. Consequently, termination has to be proved
   372   defined as projections. Consequently, termination has to be proved
   373   simultaneously for both functions, by specifying a measure on the
   373   simultaneously for both functions, by specifying a measure on the
   374   sum type: 
   374   sum type: 
   375 *}
   375 \<close>
   376 
   376 
   377 termination 
   377 termination 
   378 by (relation "measure (\<lambda>x. case x of Inl n \<Rightarrow> n | Inr n \<Rightarrow> n)") auto
   378 by (relation "measure (\<lambda>x. case x of Inl n \<Rightarrow> n | Inr n \<Rightarrow> n)") auto
   379 
   379 
   380 text {* 
   380 text \<open>
   381   We could also have used @{text lexicographic_order}, which
   381   We could also have used @{text lexicographic_order}, which
   382   supports mutual recursive termination proofs to a certain extent.
   382   supports mutual recursive termination proofs to a certain extent.
   383 *}
   383 \<close>
   384 
   384 
   385 subsection {* Induction for mutual recursion *}
   385 subsection \<open>Induction for mutual recursion\<close>
   386 
   386 
   387 text {*
   387 text \<open>
   388 
   388 
   389   When functions are mutually recursive, proving properties about them
   389   When functions are mutually recursive, proving properties about them
   390   generally requires simultaneous induction. The induction rule @{thm [source] "even_odd.induct"}
   390   generally requires simultaneous induction. The induction rule @{thm [source] "even_odd.induct"}
   391   generated from the above definition reflects this.
   391   generated from the above definition reflects this.
   392 
   392 
   393   Let us prove something about @{const even} and @{const odd}:
   393   Let us prove something about @{const even} and @{const odd}:
   394 *}
   394 \<close>
   395 
   395 
   396 lemma even_odd_mod2:
   396 lemma even_odd_mod2:
   397   "even n = (n mod 2 = 0)"
   397   "even n = (n mod 2 = 0)"
   398   "odd n = (n mod 2 = 1)"
   398   "odd n = (n mod 2 = 1)"
   399 
   399 
   400 text {* 
   400 text \<open>
   401   We apply simultaneous induction, specifying the induction variable
   401   We apply simultaneous induction, specifying the induction variable
   402   for both goals, separated by \cmd{and}:  *}
   402   for both goals, separated by \cmd{and}:\<close>
   403 
   403 
   404 apply (induct n and n rule: even_odd.induct)
   404 apply (induct n and n rule: even_odd.induct)
   405 
   405 
   406 text {* 
   406 text \<open>
   407   We get four subgoals, which correspond to the clauses in the
   407   We get four subgoals, which correspond to the clauses in the
   408   definition of @{const even} and @{const odd}:
   408   definition of @{const even} and @{const odd}:
   409   @{subgoals[display,indent=0]}
   409   @{subgoals[display,indent=0]}
   410   Simplification solves the first two goals, leaving us with two
   410   Simplification solves the first two goals, leaving us with two
   411   statements about the @{text "mod"} operation to prove:
   411   statements about the @{text "mod"} operation to prove:
   412 *}
   412 \<close>
   413 
   413 
   414 apply simp_all
   414 apply simp_all
   415 
   415 
   416 text {* 
   416 text \<open>
   417   @{subgoals[display,indent=0]} 
   417   @{subgoals[display,indent=0]} 
   418 
   418 
   419   \noindent These can be handled by Isabelle's arithmetic decision procedures.
   419   \noindent These can be handled by Isabelle's arithmetic decision procedures.
   420   
   420   
   421 *}
   421 \<close>
   422 
   422 
   423 apply arith
   423 apply arith
   424 apply arith
   424 apply arith
   425 done
   425 done
   426 
   426 
   427 text {*
   427 text \<open>
   428   In proofs like this, the simultaneous induction is really essential:
   428   In proofs like this, the simultaneous induction is really essential:
   429   Even if we are just interested in one of the results, the other
   429   Even if we are just interested in one of the results, the other
   430   one is necessary to strengthen the induction hypothesis. If we leave
   430   one is necessary to strengthen the induction hypothesis. If we leave
   431   out the statement about @{const odd} and just write @{term True} instead,
   431   out the statement about @{const odd} and just write @{term True} instead,
   432   the same proof fails:
   432   the same proof fails:
   433 *}
   433 \<close>
   434 
   434 
   435 lemma failed_attempt:
   435 lemma failed_attempt:
   436   "even n = (n mod 2 = 0)"
   436   "even n = (n mod 2 = 0)"
   437   "True"
   437   "True"
   438 apply (induct n rule: even_odd.induct)
   438 apply (induct n rule: even_odd.induct)
   439 
   439 
   440 text {*
   440 text \<open>
   441   \noindent Now the third subgoal is a dead end, since we have no
   441   \noindent Now the third subgoal is a dead end, since we have no
   442   useful induction hypothesis available:
   442   useful induction hypothesis available:
   443 
   443 
   444   @{subgoals[display,indent=0]} 
   444   @{subgoals[display,indent=0]} 
   445 *}
   445 \<close>
   446 
   446 
   447 oops
   447 oops
   448 
   448 
   449 section {* Elimination *}
   449 section \<open>Elimination\<close>
   450 
   450 
   451 text {* 
   451 text \<open>
   452   A definition of function @{text f} gives rise to two kinds of elimination rules. Rule @{text f.cases}
   452   A definition of function @{text f} gives rise to two kinds of elimination rules. Rule @{text f.cases}
   453   simply describes case analysis according to the patterns used in the definition:
   453   simply describes case analysis according to the patterns used in the definition:
   454 *}
   454 \<close>
   455 
   455 
   456 fun list_to_option :: "'a list \<Rightarrow> 'a option"
   456 fun list_to_option :: "'a list \<Rightarrow> 'a option"
   457 where
   457 where
   458   "list_to_option [x] = Some x"
   458   "list_to_option [x] = Some x"
   459 | "list_to_option _ = None"
   459 | "list_to_option _ = None"
   460 
   460 
   461 thm list_to_option.cases
   461 thm list_to_option.cases
   462 text {*
   462 text \<open>
   463   @{thm[display] list_to_option.cases}
   463   @{thm[display] list_to_option.cases}
   464 
   464 
   465   Note that this rule does not mention the function at all, but only describes the cases used for
   465   Note that this rule does not mention the function at all, but only describes the cases used for
   466   defining it. In contrast, the rule @{thm[source] list_to_option.elims} also tell us what the function
   466   defining it. In contrast, the rule @{thm[source] list_to_option.elims} also tell us what the function
   467   value will be in each case:
   467   value will be in each case:
   468 *}
   468 \<close>
   469 thm list_to_option.elims
   469 thm list_to_option.elims
   470 text {*
   470 text \<open>
   471   @{thm[display] list_to_option.elims}
   471   @{thm[display] list_to_option.elims}
   472 
   472 
   473   \noindent
   473   \noindent
   474   This lets us eliminate an assumption of the form @{prop "list_to_option xs = y"} and replace it
   474   This lets us eliminate an assumption of the form @{prop "list_to_option xs = y"} and replace it
   475   with the two cases, e.g.:
   475   with the two cases, e.g.:
   476 *}
   476 \<close>
   477 
   477 
   478 lemma "list_to_option xs = y \<Longrightarrow> P"
   478 lemma "list_to_option xs = y \<Longrightarrow> P"
   479 proof (erule list_to_option.elims)
   479 proof (erule list_to_option.elims)
   480   fix x assume "xs = [x]" "y = Some x" thus P sorry
   480   fix x assume "xs = [x]" "y = Some x" thus P sorry
   481 next
   481 next
   483 next
   483 next
   484   fix a b xs' assume "xs = a # b # xs'" "y = None" thus P sorry
   484   fix a b xs' assume "xs = a # b # xs'" "y = None" thus P sorry
   485 qed
   485 qed
   486 
   486 
   487 
   487 
   488 text {*
   488 text \<open>
   489   Sometimes it is convenient to derive specialized versions of the @{text elim} rules above and
   489   Sometimes it is convenient to derive specialized versions of the @{text elim} rules above and
   490   keep them around as facts explicitly. For example, it is natural to show that if 
   490   keep them around as facts explicitly. For example, it is natural to show that if 
   491   @{prop "list_to_option xs = Some y"}, then @{term xs} must be a singleton. The command 
   491   @{prop "list_to_option xs = Some y"}, then @{term xs} must be a singleton. The command 
   492   \cmd{fun\_cases} derives such facts automatically, by instantiating and simplifying the general 
   492   \cmd{fun\_cases} derives such facts automatically, by instantiating and simplifying the general 
   493   elimination rules given some pattern:
   493   elimination rules given some pattern:
   494 *}
   494 \<close>
   495 
   495 
   496 fun_cases list_to_option_SomeE[elim]: "list_to_option xs = Some y"
   496 fun_cases list_to_option_SomeE[elim]: "list_to_option xs = Some y"
   497 
   497 
   498 thm list_to_option_SomeE
   498 thm list_to_option_SomeE
   499 text {*
   499 text \<open>
   500   @{thm[display] list_to_option_SomeE}
   500   @{thm[display] list_to_option_SomeE}
   501 *}
   501 \<close>
   502 
   502 
   503 
   503 
   504 section {* General pattern matching *}
   504 section \<open>General pattern matching\<close>
   505 text{*\label{genpats} *}
   505 text\<open>\label{genpats}\<close>
   506 
   506 
   507 subsection {* Avoiding automatic pattern splitting *}
   507 subsection \<open>Avoiding automatic pattern splitting\<close>
   508 
   508 
   509 text {*
   509 text \<open>
   510 
   510 
   511   Up to now, we used pattern matching only on datatypes, and the
   511   Up to now, we used pattern matching only on datatypes, and the
   512   patterns were always disjoint and complete, and if they weren't,
   512   patterns were always disjoint and complete, and if they weren't,
   513   they were made disjoint automatically like in the definition of
   513   they were made disjoint automatically like in the definition of
   514   @{const "sep"} in \S\ref{patmatch}.
   514   @{const "sep"} in \S\ref{patmatch}.
   518   example shows the problem:
   518   example shows the problem:
   519   
   519   
   520   Suppose we are modeling incomplete knowledge about the world by a
   520   Suppose we are modeling incomplete knowledge about the world by a
   521   three-valued datatype, which has values @{term "T"}, @{term "F"}
   521   three-valued datatype, which has values @{term "T"}, @{term "F"}
   522   and @{term "X"} for true, false and uncertain propositions, respectively. 
   522   and @{term "X"} for true, false and uncertain propositions, respectively. 
   523 *}
   523 \<close>
   524 
   524 
   525 datatype P3 = T | F | X
   525 datatype P3 = T | F | X
   526 
   526 
   527 text {* \noindent Then the conjunction of such values can be defined as follows: *}
   527 text \<open>\noindent Then the conjunction of such values can be defined as follows:\<close>
   528 
   528 
   529 fun And :: "P3 \<Rightarrow> P3 \<Rightarrow> P3"
   529 fun And :: "P3 \<Rightarrow> P3 \<Rightarrow> P3"
   530 where
   530 where
   531   "And T p = p"
   531   "And T p = p"
   532 | "And p T = p"
   532 | "And p T = p"
   533 | "And p F = F"
   533 | "And p F = F"
   534 | "And F p = F"
   534 | "And F p = F"
   535 | "And X X = X"
   535 | "And X X = X"
   536 
   536 
   537 
   537 
   538 text {* 
   538 text \<open>
   539   This definition is useful, because the equations can directly be used
   539   This definition is useful, because the equations can directly be used
   540   as simplification rules. But the patterns overlap: For example,
   540   as simplification rules. But the patterns overlap: For example,
   541   the expression @{term "And T T"} is matched by both the first and
   541   the expression @{term "And T T"} is matched by both the first and
   542   the second equation. By default, Isabelle makes the patterns disjoint by
   542   the second equation. By default, Isabelle makes the patterns disjoint by
   543   splitting them up, producing instances:
   543   splitting them up, producing instances:
   544 *}
   544 \<close>
   545 
   545 
   546 thm And.simps
   546 thm And.simps
   547 
   547 
   548 text {*
   548 text \<open>
   549   @{thm[indent=4] And.simps}
   549   @{thm[indent=4] And.simps}
   550   
   550   
   551   \vspace*{1em}
   551   \vspace*{1em}
   552   \noindent There are several problems with this:
   552   \noindent There are several problems with this:
   553 
   553 
   573   If we do not want the automatic splitting, we can switch it off by
   573   If we do not want the automatic splitting, we can switch it off by
   574   leaving out the \cmd{sequential} option. However, we will have to
   574   leaving out the \cmd{sequential} option. However, we will have to
   575   prove that our pattern matching is consistent\footnote{This prevents
   575   prove that our pattern matching is consistent\footnote{This prevents
   576   us from defining something like @{term "f x = True"} and @{term "f x
   576   us from defining something like @{term "f x = True"} and @{term "f x
   577   = False"} simultaneously.}:
   577   = False"} simultaneously.}:
   578 *}
   578 \<close>
   579 
   579 
   580 function And2 :: "P3 \<Rightarrow> P3 \<Rightarrow> P3"
   580 function And2 :: "P3 \<Rightarrow> P3 \<Rightarrow> P3"
   581 where
   581 where
   582   "And2 T p = p"
   582   "And2 T p = p"
   583 | "And2 p T = p"
   583 | "And2 p T = p"
   584 | "And2 p F = F"
   584 | "And2 p F = F"
   585 | "And2 F p = F"
   585 | "And2 F p = F"
   586 | "And2 X X = X"
   586 | "And2 X X = X"
   587 
   587 
   588 text {*
   588 text \<open>
   589   \noindent Now let's look at the proof obligations generated by a
   589   \noindent Now let's look at the proof obligations generated by a
   590   function definition. In this case, they are:
   590   function definition. In this case, they are:
   591 
   591 
   592   @{subgoals[display,indent=0]}\vspace{-1.2em}\hspace{3cm}\vdots\vspace{1.2em}
   592   @{subgoals[display,indent=0]}\vspace{-1.2em}\hspace{3cm}\vdots\vspace{1.2em}
   593 
   593 
   597   be equivalently stated as a disjunction of existential statements: 
   597   be equivalently stated as a disjunction of existential statements: 
   598 @{term "(\<exists>p. x = (T, p)) \<or> (\<exists>p. x = (p, T)) \<or> (\<exists>p. x = (p, F)) \<or>
   598 @{term "(\<exists>p. x = (T, p)) \<or> (\<exists>p. x = (p, T)) \<or> (\<exists>p. x = (p, F)) \<or>
   599   (\<exists>p. x = (F, p)) \<or> (x = (X, X))"}, and you can use the method @{text atomize_elim} to get that form instead.}. If the patterns just involve
   599   (\<exists>p. x = (F, p)) \<or> (x = (X, X))"}, and you can use the method @{text atomize_elim} to get that form instead.}. If the patterns just involve
   600   datatypes, we can solve it with the @{text "pat_completeness"}
   600   datatypes, we can solve it with the @{text "pat_completeness"}
   601   method:
   601   method:
   602 *}
   602 \<close>
   603 
   603 
   604 apply pat_completeness
   604 apply pat_completeness
   605 
   605 
   606 text {*
   606 text \<open>
   607   The remaining subgoals express \emph{pattern compatibility}. We do
   607   The remaining subgoals express \emph{pattern compatibility}. We do
   608   allow that an input value matches multiple patterns, but in this
   608   allow that an input value matches multiple patterns, but in this
   609   case, the result (i.e.~the right hand sides of the equations) must
   609   case, the result (i.e.~the right hand sides of the equations) must
   610   also be equal. For each pair of two patterns, there is one such
   610   also be equal. For each pair of two patterns, there is one such
   611   subgoal. Usually this needs injectivity of the constructors, which
   611   subgoal. Usually this needs injectivity of the constructors, which
   612   is used automatically by @{text "auto"}.
   612   is used automatically by @{text "auto"}.
   613 *}
   613 \<close>
   614 
   614 
   615 by auto
   615 by auto
   616 termination by (relation "{}") simp
   616 termination by (relation "{}") simp
   617 
   617 
   618 
   618 
   619 subsection {* Non-constructor patterns *}
   619 subsection \<open>Non-constructor patterns\<close>
   620 
   620 
   621 text {*
   621 text \<open>
   622   Most of Isabelle's basic types take the form of inductive datatypes,
   622   Most of Isabelle's basic types take the form of inductive datatypes,
   623   and usually pattern matching works on the constructors of such types. 
   623   and usually pattern matching works on the constructors of such types. 
   624   However, this need not be always the case, and the \cmd{function}
   624   However, this need not be always the case, and the \cmd{function}
   625   command handles other kind of patterns, too.
   625   command handles other kind of patterns, too.
   626 
   626 
   627   One well-known instance of non-constructor patterns are
   627   One well-known instance of non-constructor patterns are
   628   so-called \emph{$n+k$-patterns}, which are a little controversial in
   628   so-called \emph{$n+k$-patterns}, which are a little controversial in
   629   the functional programming world. Here is the initial fibonacci
   629   the functional programming world. Here is the initial fibonacci
   630   example with $n+k$-patterns:
   630   example with $n+k$-patterns:
   631 *}
   631 \<close>
   632 
   632 
   633 function fib2 :: "nat \<Rightarrow> nat"
   633 function fib2 :: "nat \<Rightarrow> nat"
   634 where
   634 where
   635   "fib2 0 = 1"
   635   "fib2 0 = 1"
   636 | "fib2 1 = 1"
   636 | "fib2 1 = 1"
   637 | "fib2 (n + 2) = fib2 n + fib2 (Suc n)"
   637 | "fib2 (n + 2) = fib2 n + fib2 (Suc n)"
   638 
   638 
   639 text {*
   639 text \<open>
   640   This kind of matching is again justified by the proof of pattern
   640   This kind of matching is again justified by the proof of pattern
   641   completeness and compatibility. 
   641   completeness and compatibility. 
   642   The proof obligation for pattern completeness states that every natural number is
   642   The proof obligation for pattern completeness states that every natural number is
   643   either @{term "0::nat"}, @{term "1::nat"} or @{term "n +
   643   either @{term "0::nat"}, @{term "1::nat"} or @{term "n +
   644   (2::nat)"}:
   644   (2::nat)"}:
   649   @{text arith} method cannot handle this specific form of an
   649   @{text arith} method cannot handle this specific form of an
   650   elimination rule. However, we can use the method @{text
   650   elimination rule. However, we can use the method @{text
   651   "atomize_elim"} to do an ad-hoc conversion to a disjunction of
   651   "atomize_elim"} to do an ad-hoc conversion to a disjunction of
   652   existentials, which can then be solved by the arithmetic decision procedure.
   652   existentials, which can then be solved by the arithmetic decision procedure.
   653   Pattern compatibility and termination are automatic as usual.
   653   Pattern compatibility and termination are automatic as usual.
   654 *}
   654 \<close>
   655 apply atomize_elim
   655 apply atomize_elim
   656 apply arith
   656 apply arith
   657 apply auto
   657 apply auto
   658 done
   658 done
   659 termination by lexicographic_order
   659 termination by lexicographic_order
   660 text {*
   660 text \<open>
   661   We can stretch the notion of pattern matching even more. The
   661   We can stretch the notion of pattern matching even more. The
   662   following function is not a sensible functional program, but a
   662   following function is not a sensible functional program, but a
   663   perfectly valid mathematical definition:
   663   perfectly valid mathematical definition:
   664 *}
   664 \<close>
   665 
   665 
   666 function ev :: "nat \<Rightarrow> bool"
   666 function ev :: "nat \<Rightarrow> bool"
   667 where
   667 where
   668   "ev (2 * n) = True"
   668   "ev (2 * n) = True"
   669 | "ev (2 * n + 1) = False"
   669 | "ev (2 * n + 1) = False"
   670 apply atomize_elim
   670 apply atomize_elim
   671 by arith+
   671 by arith+
   672 termination by (relation "{}") simp
   672 termination by (relation "{}") simp
   673 
   673 
   674 text {*
   674 text \<open>
   675   This general notion of pattern matching gives you a certain freedom
   675   This general notion of pattern matching gives you a certain freedom
   676   in writing down specifications. However, as always, such freedom should
   676   in writing down specifications. However, as always, such freedom should
   677   be used with care:
   677   be used with care:
   678 
   678 
   679   If we leave the area of constructor
   679   If we leave the area of constructor
   680   patterns, we have effectively departed from the world of functional
   680   patterns, we have effectively departed from the world of functional
   681   programming. This means that it is no longer possible to use the
   681   programming. This means that it is no longer possible to use the
   682   code generator, and expect it to generate ML code for our
   682   code generator, and expect it to generate ML code for our
   683   definitions. Also, such a specification might not work very well together with
   683   definitions. Also, such a specification might not work very well together with
   684   simplification. Your mileage may vary.
   684   simplification. Your mileage may vary.
   685 *}
   685 \<close>
   686 
   686 
   687 
   687 
   688 subsection {* Conditional equations *}
   688 subsection \<open>Conditional equations\<close>
   689 
   689 
   690 text {* 
   690 text \<open>
   691   The function package also supports conditional equations, which are
   691   The function package also supports conditional equations, which are
   692   similar to guards in a language like Haskell. Here is Euclid's
   692   similar to guards in a language like Haskell. Here is Euclid's
   693   algorithm written with conditional patterns\footnote{Note that the
   693   algorithm written with conditional patterns\footnote{Note that the
   694   patterns are also overlapping in the base case}:
   694   patterns are also overlapping in the base case}:
   695 *}
   695 \<close>
   696 
   696 
   697 function gcd :: "nat \<Rightarrow> nat \<Rightarrow> nat"
   697 function gcd :: "nat \<Rightarrow> nat \<Rightarrow> nat"
   698 where
   698 where
   699   "gcd x 0 = x"
   699   "gcd x 0 = x"
   700 | "gcd 0 y = y"
   700 | "gcd 0 y = y"
   701 | "x < y \<Longrightarrow> gcd (Suc x) (Suc y) = gcd (Suc x) (y - x)"
   701 | "x < y \<Longrightarrow> gcd (Suc x) (Suc y) = gcd (Suc x) (y - x)"
   702 | "\<not> x < y \<Longrightarrow> gcd (Suc x) (Suc y) = gcd (x - y) (Suc y)"
   702 | "\<not> x < y \<Longrightarrow> gcd (Suc x) (Suc y) = gcd (x - y) (Suc y)"
   703 by (atomize_elim, auto, arith)
   703 by (atomize_elim, auto, arith)
   704 termination by lexicographic_order
   704 termination by lexicographic_order
   705 
   705 
   706 text {*
   706 text \<open>
   707   By now, you can probably guess what the proof obligations for the
   707   By now, you can probably guess what the proof obligations for the
   708   pattern completeness and compatibility look like. 
   708   pattern completeness and compatibility look like. 
   709 
   709 
   710   Again, functions with conditional patterns are not supported by the
   710   Again, functions with conditional patterns are not supported by the
   711   code generator.
   711   code generator.
   712 *}
   712 \<close>
   713 
   713 
   714 
   714 
   715 subsection {* Pattern matching on strings *}
   715 subsection \<open>Pattern matching on strings\<close>
   716 
   716 
   717 text {*
   717 text \<open>
   718   As strings (as lists of characters) are normal datatypes, pattern
   718   As strings (as lists of characters) are normal datatypes, pattern
   719   matching on them is possible, but somewhat problematic. Consider the
   719   matching on them is possible, but somewhat problematic. Consider the
   720   following definition:
   720   following definition:
   721 
   721 
   722 \end{isamarkuptext}
   722 \end{isamarkuptext}
   732   catch-all pattern thus leads to an explosion of cases, which cannot
   732   catch-all pattern thus leads to an explosion of cases, which cannot
   733   be handled by Isabelle.
   733   be handled by Isabelle.
   734 
   734 
   735   There are two things we can do here. Either we write an explicit
   735   There are two things we can do here. Either we write an explicit
   736   @{text "if"} on the right hand side, or we can use conditional patterns:
   736   @{text "if"} on the right hand side, or we can use conditional patterns:
   737 *}
   737 \<close>
   738 
   738 
   739 function check :: "string \<Rightarrow> bool"
   739 function check :: "string \<Rightarrow> bool"
   740 where
   740 where
   741   "check (''good'') = True"
   741   "check (''good'') = True"
   742 | "s \<noteq> ''good'' \<Longrightarrow> check s = False"
   742 | "s \<noteq> ''good'' \<Longrightarrow> check s = False"
   743 by auto
   743 by auto
   744 termination by (relation "{}") simp
   744 termination by (relation "{}") simp
   745 
   745 
   746 
   746 
   747 section {* Partiality *}
   747 section \<open>Partiality\<close>
   748 
   748 
   749 text {* 
   749 text \<open>
   750   In HOL, all functions are total. A function @{term "f"} applied to
   750   In HOL, all functions are total. A function @{term "f"} applied to
   751   @{term "x"} always has the value @{term "f x"}, and there is no notion
   751   @{term "x"} always has the value @{term "f x"}, and there is no notion
   752   of undefinedness. 
   752   of undefinedness. 
   753   This is why we have to do termination
   753   This is why we have to do termination
   754   proofs when defining functions: The proof justifies that the
   754   proofs when defining functions: The proof justifies that the
   755   function can be defined by wellfounded recursion.
   755   function can be defined by wellfounded recursion.
   756 
   756 
   757   However, the \cmd{function} package does support partiality to a
   757   However, the \cmd{function} package does support partiality to a
   758   certain extent. Let's look at the following function which looks
   758   certain extent. Let's look at the following function which looks
   759   for a zero of a given function f. 
   759   for a zero of a given function f. 
   760 *}
   760 \<close>
   761 
   761 
   762 function (*<*)(domintros)(*>*)findzero :: "(nat \<Rightarrow> nat) \<Rightarrow> nat \<Rightarrow> nat"
   762 function (*<*)(domintros)(*>*)findzero :: "(nat \<Rightarrow> nat) \<Rightarrow> nat \<Rightarrow> nat"
   763 where
   763 where
   764   "findzero f n = (if f n = 0 then n else findzero f (Suc n))"
   764   "findzero f n = (if f n = 0 then n else findzero f (Suc n))"
   765 by pat_completeness auto
   765 by pat_completeness auto
   766 
   766 
   767 text {*
   767 text \<open>
   768   \noindent Clearly, any attempt of a termination proof must fail. And without
   768   \noindent Clearly, any attempt of a termination proof must fail. And without
   769   that, we do not get the usual rules @{text "findzero.simps"} and 
   769   that, we do not get the usual rules @{text "findzero.simps"} and 
   770   @{text "findzero.induct"}. So what was the definition good for at all?
   770   @{text "findzero.induct"}. So what was the definition good for at all?
   771 *}
   771 \<close>
   772 
   772 
   773 subsection {* Domain predicates *}
   773 subsection \<open>Domain predicates\<close>
   774 
   774 
   775 text {*
   775 text \<open>
   776   The trick is that Isabelle has not only defined the function @{const findzero}, but also
   776   The trick is that Isabelle has not only defined the function @{const findzero}, but also
   777   a predicate @{term "findzero_dom"} that characterizes the values where the function
   777   a predicate @{term "findzero_dom"} that characterizes the values where the function
   778   terminates: the \emph{domain} of the function. If we treat a
   778   terminates: the \emph{domain} of the function. If we treat a
   779   partial function just as a total function with an additional domain
   779   partial function just as a total function with an additional domain
   780   predicate, we can derive simplification and
   780   predicate, we can derive simplification and
   781   induction rules as we do for total functions. They are guarded
   781   induction rules as we do for total functions. They are guarded
   782   by domain conditions and are called @{text psimps} and @{text
   782   by domain conditions and are called @{text psimps} and @{text
   783   pinduct}: 
   783   pinduct}: 
   784 *}
   784 \<close>
   785 
   785 
   786 text {*
   786 text \<open>
   787   \noindent\begin{minipage}{0.79\textwidth}@{thm[display,margin=85] findzero.psimps}\end{minipage}
   787   \noindent\begin{minipage}{0.79\textwidth}@{thm[display,margin=85] findzero.psimps}\end{minipage}
   788   \hfill(@{thm [source] "findzero.psimps"})
   788   \hfill(@{thm [source] "findzero.psimps"})
   789   \vspace{1em}
   789   \vspace{1em}
   790 
   790 
   791   \noindent\begin{minipage}{0.79\textwidth}@{thm[display,margin=85] findzero.pinduct}\end{minipage}
   791   \noindent\begin{minipage}{0.79\textwidth}@{thm[display,margin=85] findzero.pinduct}\end{minipage}
   792   \hfill(@{thm [source] "findzero.pinduct"})
   792   \hfill(@{thm [source] "findzero.pinduct"})
   793 *}
   793 \<close>
   794 
   794 
   795 text {*
   795 text \<open>
   796   Remember that all we
   796   Remember that all we
   797   are doing here is use some tricks to make a total function appear
   797   are doing here is use some tricks to make a total function appear
   798   as if it was partial. We can still write the term @{term "findzero
   798   as if it was partial. We can still write the term @{term "findzero
   799   (\<lambda>x. 1) 0"} and like any other term of type @{typ nat} it is equal
   799   (\<lambda>x. 1) 0"} and like any other term of type @{typ nat} it is equal
   800   to some natural number, although we might not be able to find out
   800   to some natural number, although we might not be able to find out
   801   which one. The function is \emph{underdefined}.
   801   which one. The function is \emph{underdefined}.
   802 
   802 
   803   But it is defined enough to prove something interesting about it. We
   803   But it is defined enough to prove something interesting about it. We
   804   can prove that if @{term "findzero f n"}
   804   can prove that if @{term "findzero f n"}
   805   terminates, it indeed returns a zero of @{term f}:
   805   terminates, it indeed returns a zero of @{term f}:
   806 *}
   806 \<close>
   807 
   807 
   808 lemma findzero_zero: "findzero_dom (f, n) \<Longrightarrow> f (findzero f n) = 0"
   808 lemma findzero_zero: "findzero_dom (f, n) \<Longrightarrow> f (findzero f n) = 0"
   809 
   809 
   810 text {* \noindent We apply induction as usual, but using the partial induction
   810 text \<open>\noindent We apply induction as usual, but using the partial induction
   811   rule: *}
   811   rule:\<close>
   812 
   812 
   813 apply (induct f n rule: findzero.pinduct)
   813 apply (induct f n rule: findzero.pinduct)
   814 
   814 
   815 text {* \noindent This gives the following subgoals:
   815 text \<open>\noindent This gives the following subgoals:
   816 
   816 
   817   @{subgoals[display,indent=0]}
   817   @{subgoals[display,indent=0]}
   818 
   818 
   819   \noindent The hypothesis in our lemma was used to satisfy the first premise in
   819   \noindent The hypothesis in our lemma was used to satisfy the first premise in
   820   the induction rule. However, we also get @{term
   820   the induction rule. However, we also get @{term
   821   "findzero_dom (f, n)"} as a local assumption in the induction step. This
   821   "findzero_dom (f, n)"} as a local assumption in the induction step. This
   822   allows unfolding @{term "findzero f n"} using the @{text psimps}
   822   allows unfolding @{term "findzero f n"} using the @{text psimps}
   823   rule, and the rest is trivial.
   823   rule, and the rest is trivial.
   824  *}
   824 \<close>
   825 apply (simp add: findzero.psimps)
   825 apply (simp add: findzero.psimps)
   826 done
   826 done
   827 
   827 
   828 text {*
   828 text \<open>
   829   Proofs about partial functions are often not harder than for total
   829   Proofs about partial functions are often not harder than for total
   830   functions. Fig.~\ref{findzero_isar} shows a slightly more
   830   functions. Fig.~\ref{findzero_isar} shows a slightly more
   831   complicated proof written in Isar. It is verbose enough to show how
   831   complicated proof written in Isar. It is verbose enough to show how
   832   partiality comes into play: From the partial induction, we get an
   832   partiality comes into play: From the partial induction, we get an
   833   additional domain condition hypothesis. Observe how this condition
   833   additional domain condition hypothesis. Observe how this condition
   834   is applied when calls to @{term findzero} are unfolded.
   834   is applied when calls to @{term findzero} are unfolded.
   835 *}
   835 \<close>
   836 
   836 
   837 text_raw {*
   837 text_raw \<open>
   838 \begin{figure}
   838 \begin{figure}
   839 \hrule\vspace{6pt}
   839 \hrule\vspace{6pt}
   840 \begin{minipage}{0.8\textwidth}
   840 \begin{minipage}{0.8\textwidth}
   841 \isabellestyle{it}
   841 \isabellestyle{it}
   842 \isastyle\isamarkuptrue
   842 \isastyle\isamarkuptrue
   843 *}
   843 \<close>
   844 lemma "\<lbrakk>findzero_dom (f, n); x \<in> {n ..< findzero f n}\<rbrakk> \<Longrightarrow> f x \<noteq> 0"
   844 lemma "\<lbrakk>findzero_dom (f, n); x \<in> {n ..< findzero f n}\<rbrakk> \<Longrightarrow> f x \<noteq> 0"
   845 proof (induct rule: findzero.pinduct)
   845 proof (induct rule: findzero.pinduct)
   846   fix f n assume dom: "findzero_dom (f, n)"
   846   fix f n assume dom: "findzero_dom (f, n)"
   847                and IH: "\<lbrakk>f n \<noteq> 0; x \<in> {Suc n ..< findzero f (Suc n)}\<rbrakk> \<Longrightarrow> f x \<noteq> 0"
   847                and IH: "\<lbrakk>f n \<noteq> 0; x \<in> {Suc n ..< findzero f (Suc n)}\<rbrakk> \<Longrightarrow> f x \<noteq> 0"
   848                and x_range: "x \<in> {n ..< findzero f n}"
   848                and x_range: "x \<in> {n ..< findzero f n}"
   855   
   855   
   856   from x_range have "x = n \<or> x \<in> {Suc n ..< findzero f n}" by auto
   856   from x_range have "x = n \<or> x \<in> {Suc n ..< findzero f n}" by auto
   857   thus "f x \<noteq> 0"
   857   thus "f x \<noteq> 0"
   858   proof
   858   proof
   859     assume "x = n"
   859     assume "x = n"
   860     with `f n \<noteq> 0` show ?thesis by simp
   860     with \<open>f n \<noteq> 0\<close> show ?thesis by simp
   861   next
   861   next
   862     assume "x \<in> {Suc n ..< findzero f n}"
   862     assume "x \<in> {Suc n ..< findzero f n}"
   863     with dom and `f n \<noteq> 0` have "x \<in> {Suc n ..< findzero f (Suc n)}" by (simp add: findzero.psimps)
   863     with dom and \<open>f n \<noteq> 0\<close> have "x \<in> {Suc n ..< findzero f (Suc n)}" by (simp add: findzero.psimps)
   864     with IH and `f n \<noteq> 0`
   864     with IH and \<open>f n \<noteq> 0\<close>
   865     show ?thesis by simp
   865     show ?thesis by simp
   866   qed
   866   qed
   867 qed
   867 qed
   868 text_raw {*
   868 text_raw \<open>
   869 \isamarkupfalse\isabellestyle{tt}
   869 \isamarkupfalse\isabellestyle{tt}
   870 \end{minipage}\vspace{6pt}\hrule
   870 \end{minipage}\vspace{6pt}\hrule
   871 \caption{A proof about a partial function}\label{findzero_isar}
   871 \caption{A proof about a partial function}\label{findzero_isar}
   872 \end{figure}
   872 \end{figure}
   873 *}
   873 \<close>
   874 
   874 
   875 subsection {* Partial termination proofs *}
   875 subsection \<open>Partial termination proofs\<close>
   876 
   876 
   877 text {*
   877 text \<open>
   878   Now that we have proved some interesting properties about our
   878   Now that we have proved some interesting properties about our
   879   function, we should turn to the domain predicate and see if it is
   879   function, we should turn to the domain predicate and see if it is
   880   actually true for some values. Otherwise we would have just proved
   880   actually true for some values. Otherwise we would have just proved
   881   lemmas with @{term False} as a premise.
   881   lemmas with @{term False} as a premise.
   882 
   882 
   892 \noindent\cmd{function} @{text "(domintros) findzero :: \"(nat \<Rightarrow> nat) \<Rightarrow> nat \<Rightarrow> nat\""}\\%
   892 \noindent\cmd{function} @{text "(domintros) findzero :: \"(nat \<Rightarrow> nat) \<Rightarrow> nat \<Rightarrow> nat\""}\\%
   893 \cmd{where}\isanewline%
   893 \cmd{where}\isanewline%
   894 \ \ \ldots\\
   894 \ \ \ldots\\
   895 
   895 
   896   \noindent Now the package has proved an introduction rule for @{text findzero_dom}:
   896   \noindent Now the package has proved an introduction rule for @{text findzero_dom}:
   897 *}
   897 \<close>
   898 
   898 
   899 thm findzero.domintros
   899 thm findzero.domintros
   900 
   900 
   901 text {*
   901 text \<open>
   902   @{thm[display] findzero.domintros}
   902   @{thm[display] findzero.domintros}
   903 
   903 
   904   Domain introduction rules allow to show that a given value lies in the
   904   Domain introduction rules allow to show that a given value lies in the
   905   domain of a function, if the arguments of all recursive calls
   905   domain of a function, if the arguments of all recursive calls
   906   are in the domain as well. They allow to do a \qt{single step} in a
   906   are in the domain as well. They allow to do a \qt{single step} in a
   919   or equal to @{term n}. First we derive two useful rules which will
   919   or equal to @{term n}. First we derive two useful rules which will
   920   solve the base case and the step case of the induction. The
   920   solve the base case and the step case of the induction. The
   921   induction is then straightforward, except for the unusual induction
   921   induction is then straightforward, except for the unusual induction
   922   principle.
   922   principle.
   923 
   923 
   924 *}
   924 \<close>
   925 
   925 
   926 text_raw {*
   926 text_raw \<open>
   927 \begin{figure}
   927 \begin{figure}
   928 \hrule\vspace{6pt}
   928 \hrule\vspace{6pt}
   929 \begin{minipage}{0.8\textwidth}
   929 \begin{minipage}{0.8\textwidth}
   930 \isabellestyle{it}
   930 \isabellestyle{it}
   931 \isastyle\isamarkuptrue
   931 \isastyle\isamarkuptrue
   932 *}
   932 \<close>
   933 lemma findzero_termination:
   933 lemma findzero_termination:
   934   assumes "x \<ge> n" and "f x = 0"
   934   assumes "x \<ge> n" and "f x = 0"
   935   shows "findzero_dom (f, n)"
   935   shows "findzero_dom (f, n)"
   936 proof - 
   936 proof - 
   937   have base: "findzero_dom (f, x)"
   937   have base: "findzero_dom (f, x)"
   938     by (rule findzero.domintros) (simp add:`f x = 0`)
   938     by (rule findzero.domintros) (simp add:\<open>f x = 0\<close>)
   939 
   939 
   940   have step: "\<And>i. findzero_dom (f, Suc i) 
   940   have step: "\<And>i. findzero_dom (f, Suc i) 
   941     \<Longrightarrow> findzero_dom (f, i)"
   941     \<Longrightarrow> findzero_dom (f, i)"
   942     by (rule findzero.domintros) simp
   942     by (rule findzero.domintros) simp
   943 
   943 
   944   from `x \<ge> n` show ?thesis
   944   from \<open>x \<ge> n\<close> show ?thesis
   945   proof (induct rule:inc_induct)
   945   proof (induct rule:inc_induct)
   946     show "findzero_dom (f, x)" by (rule base)
   946     show "findzero_dom (f, x)" by (rule base)
   947   next
   947   next
   948     fix i assume "findzero_dom (f, Suc i)"
   948     fix i assume "findzero_dom (f, Suc i)"
   949     thus "findzero_dom (f, i)" by (rule step)
   949     thus "findzero_dom (f, i)" by (rule step)
   950   qed
   950   qed
   951 qed      
   951 qed      
   952 text_raw {*
   952 text_raw \<open>
   953 \isamarkupfalse\isabellestyle{tt}
   953 \isamarkupfalse\isabellestyle{tt}
   954 \end{minipage}\vspace{6pt}\hrule
   954 \end{minipage}\vspace{6pt}\hrule
   955 \caption{Termination proof for @{text findzero}}\label{findzero_term}
   955 \caption{Termination proof for @{text findzero}}\label{findzero_term}
   956 \end{figure}
   956 \end{figure}
   957 *}
   957 \<close>
   958       
   958       
   959 text {*
   959 text \<open>
   960   Again, the proof given in Fig.~\ref{findzero_term} has a lot of
   960   Again, the proof given in Fig.~\ref{findzero_term} has a lot of
   961   detail in order to explain the principles. Using more automation, we
   961   detail in order to explain the principles. Using more automation, we
   962   can also have a short proof:
   962   can also have a short proof:
   963 *}
   963 \<close>
   964 
   964 
   965 lemma findzero_termination_short:
   965 lemma findzero_termination_short:
   966   assumes zero: "x >= n" 
   966   assumes zero: "x >= n" 
   967   assumes [simp]: "f x = 0"
   967   assumes [simp]: "f x = 0"
   968   shows "findzero_dom (f, n)"
   968   shows "findzero_dom (f, n)"
   969 using zero
   969 using zero
   970 by (induct rule:inc_induct) (auto intro: findzero.domintros)
   970 by (induct rule:inc_induct) (auto intro: findzero.domintros)
   971     
   971     
   972 text {*
   972 text \<open>
   973   \noindent It is simple to combine the partial correctness result with the
   973   \noindent It is simple to combine the partial correctness result with the
   974   termination lemma:
   974   termination lemma:
   975 *}
   975 \<close>
   976 
   976 
   977 lemma findzero_total_correctness:
   977 lemma findzero_total_correctness:
   978   "f x = 0 \<Longrightarrow> f (findzero f 0) = 0"
   978   "f x = 0 \<Longrightarrow> f (findzero f 0) = 0"
   979 by (blast intro: findzero_zero findzero_termination)
   979 by (blast intro: findzero_zero findzero_termination)
   980 
   980 
   981 subsection {* Definition of the domain predicate *}
   981 subsection \<open>Definition of the domain predicate\<close>
   982 
   982 
   983 text {*
   983 text \<open>
   984   Sometimes it is useful to know what the definition of the domain
   984   Sometimes it is useful to know what the definition of the domain
   985   predicate looks like. Actually, @{text findzero_dom} is just an
   985   predicate looks like. Actually, @{text findzero_dom} is just an
   986   abbreviation:
   986   abbreviation:
   987 
   987 
   988   @{abbrev[display] findzero_dom}
   988   @{abbrev[display] findzero_dom}
  1012   lemmas for @{const Wellfounded.accp} and @{const findzero_rel} directly. Some
  1012   lemmas for @{const Wellfounded.accp} and @{const findzero_rel} directly. Some
  1013   lemmas which are occasionally useful are @{thm [source] accpI}, @{thm [source]
  1013   lemmas which are occasionally useful are @{thm [source] accpI}, @{thm [source]
  1014   accp_downward}, and of course the introduction and elimination rules
  1014   accp_downward}, and of course the introduction and elimination rules
  1015   for the recursion relation @{thm [source] "findzero_rel.intros"} and @{thm
  1015   for the recursion relation @{thm [source] "findzero_rel.intros"} and @{thm
  1016   [source] "findzero_rel.cases"}.
  1016   [source] "findzero_rel.cases"}.
  1017 *}
  1017 \<close>
  1018 
  1018 
  1019 section {* Nested recursion *}
  1019 section \<open>Nested recursion\<close>
  1020 
  1020 
  1021 text {*
  1021 text \<open>
  1022   Recursive calls which are nested in one another frequently cause
  1022   Recursive calls which are nested in one another frequently cause
  1023   complications, since their termination proof can depend on a partial
  1023   complications, since their termination proof can depend on a partial
  1024   correctness property of the function itself. 
  1024   correctness property of the function itself. 
  1025 
  1025 
  1026   As a small example, we define the \qt{nested zero} function:
  1026   As a small example, we define the \qt{nested zero} function:
  1027 *}
  1027 \<close>
  1028 
  1028 
  1029 function nz :: "nat \<Rightarrow> nat"
  1029 function nz :: "nat \<Rightarrow> nat"
  1030 where
  1030 where
  1031   "nz 0 = 0"
  1031   "nz 0 = 0"
  1032 | "nz (Suc n) = nz (nz n)"
  1032 | "nz (Suc n) = nz (nz n)"
  1033 by pat_completeness auto
  1033 by pat_completeness auto
  1034 
  1034 
  1035 text {*
  1035 text \<open>
  1036   If we attempt to prove termination using the identity measure on
  1036   If we attempt to prove termination using the identity measure on
  1037   naturals, this fails:
  1037   naturals, this fails:
  1038 *}
  1038 \<close>
  1039 
  1039 
  1040 termination
  1040 termination
  1041   apply (relation "measure (\<lambda>n. n)")
  1041   apply (relation "measure (\<lambda>n. n)")
  1042   apply auto
  1042   apply auto
  1043 
  1043 
  1044 text {*
  1044 text \<open>
  1045   We get stuck with the subgoal
  1045   We get stuck with the subgoal
  1046 
  1046 
  1047   @{subgoals[display]}
  1047   @{subgoals[display]}
  1048 
  1048 
  1049   Of course this statement is true, since we know that @{const nz} is
  1049   Of course this statement is true, since we know that @{const nz} is
  1050   the zero function. And in fact we have no problem proving this
  1050   the zero function. And in fact we have no problem proving this
  1051   property by induction.
  1051   property by induction.
  1052 *}
  1052 \<close>
  1053 (*<*)oops(*>*)
  1053 (*<*)oops(*>*)
  1054 lemma nz_is_zero: "nz_dom n \<Longrightarrow> nz n = 0"
  1054 lemma nz_is_zero: "nz_dom n \<Longrightarrow> nz n = 0"
  1055   by (induct rule:nz.pinduct) (auto simp: nz.psimps)
  1055   by (induct rule:nz.pinduct) (auto simp: nz.psimps)
  1056 
  1056 
  1057 text {*
  1057 text \<open>
  1058   We formulate this as a partial correctness lemma with the condition
  1058   We formulate this as a partial correctness lemma with the condition
  1059   @{term "nz_dom n"}. This allows us to prove it with the @{text
  1059   @{term "nz_dom n"}. This allows us to prove it with the @{text
  1060   pinduct} rule before we have proved termination. With this lemma,
  1060   pinduct} rule before we have proved termination. With this lemma,
  1061   the termination proof works as expected:
  1061   the termination proof works as expected:
  1062 *}
  1062 \<close>
  1063 
  1063 
  1064 termination
  1064 termination
  1065   by (relation "measure (\<lambda>n. n)") (auto simp: nz_is_zero)
  1065   by (relation "measure (\<lambda>n. n)") (auto simp: nz_is_zero)
  1066 
  1066 
  1067 text {*
  1067 text \<open>
  1068   As a general strategy, one should prove the statements needed for
  1068   As a general strategy, one should prove the statements needed for
  1069   termination as a partial property first. Then they can be used to do
  1069   termination as a partial property first. Then they can be used to do
  1070   the termination proof. This also works for less trivial
  1070   the termination proof. This also works for less trivial
  1071   examples. Figure \ref{f91} defines the 91-function, a well-known
  1071   examples. Figure \ref{f91} defines the 91-function, a well-known
  1072   challenge problem due to John McCarthy, and proves its termination.
  1072   challenge problem due to John McCarthy, and proves its termination.
  1073 *}
  1073 \<close>
  1074 
  1074 
  1075 text_raw {*
  1075 text_raw \<open>
  1076 \begin{figure}
  1076 \begin{figure}
  1077 \hrule\vspace{6pt}
  1077 \hrule\vspace{6pt}
  1078 \begin{minipage}{0.8\textwidth}
  1078 \begin{minipage}{0.8\textwidth}
  1079 \isabellestyle{it}
  1079 \isabellestyle{it}
  1080 \isastyle\isamarkuptrue
  1080 \isastyle\isamarkuptrue
  1081 *}
  1081 \<close>
  1082 
  1082 
  1083 function f91 :: "nat \<Rightarrow> nat"
  1083 function f91 :: "nat \<Rightarrow> nat"
  1084 where
  1084 where
  1085   "f91 n = (if 100 < n then n - 10 else f91 (f91 (n + 11)))"
  1085   "f91 n = (if 100 < n then n - 10 else f91 (f91 (n + 11)))"
  1086 by pat_completeness auto
  1086 by pat_completeness auto
  1099 
  1099 
  1100   thus "(n + 11, n) \<in> ?R" by simp -- "Inner call"
  1100   thus "(n + 11, n) \<in> ?R" by simp -- "Inner call"
  1101 
  1101 
  1102   assume inner_trm: "f91_dom (n + 11)" -- "Outer call"
  1102   assume inner_trm: "f91_dom (n + 11)" -- "Outer call"
  1103   with f91_estimate have "n + 11 < f91 (n + 11) + 11" .
  1103   with f91_estimate have "n + 11 < f91 (n + 11) + 11" .
  1104   with `\<not> 100 < n` show "(f91 (n + 11), n) \<in> ?R" by simp
  1104   with \<open>\<not> 100 < n\<close> show "(f91 (n + 11), n) \<in> ?R" by simp
  1105 qed
  1105 qed
  1106 
  1106 
  1107 text_raw {*
  1107 text_raw \<open>
  1108 \isamarkupfalse\isabellestyle{tt}
  1108 \isamarkupfalse\isabellestyle{tt}
  1109 \end{minipage}
  1109 \end{minipage}
  1110 \vspace{6pt}\hrule
  1110 \vspace{6pt}\hrule
  1111 \caption{McCarthy's 91-function}\label{f91}
  1111 \caption{McCarthy's 91-function}\label{f91}
  1112 \end{figure}
  1112 \end{figure}
  1113 *}
  1113 \<close>
  1114 
  1114 
  1115 
  1115 
  1116 section {* Higher-Order Recursion *}
  1116 section \<open>Higher-Order Recursion\<close>
  1117 
  1117 
  1118 text {*
  1118 text \<open>
  1119   Higher-order recursion occurs when recursive calls
  1119   Higher-order recursion occurs when recursive calls
  1120   are passed as arguments to higher-order combinators such as @{const
  1120   are passed as arguments to higher-order combinators such as @{const
  1121   map}, @{term filter} etc.
  1121   map}, @{term filter} etc.
  1122   As an example, imagine a datatype of n-ary trees:
  1122   As an example, imagine a datatype of n-ary trees:
  1123 *}
  1123 \<close>
  1124 
  1124 
  1125 datatype 'a tree = 
  1125 datatype 'a tree = 
  1126   Leaf 'a 
  1126   Leaf 'a 
  1127 | Branch "'a tree list"
  1127 | Branch "'a tree list"
  1128 
  1128 
  1129 
  1129 
  1130 text {* \noindent We can define a function which swaps the left and right subtrees recursively, using the 
  1130 text \<open>\noindent We can define a function which swaps the left and right subtrees recursively, using the 
  1131   list functions @{const rev} and @{const map}: *}
  1131   list functions @{const rev} and @{const map}:\<close>
  1132 
  1132 
  1133 fun mirror :: "'a tree \<Rightarrow> 'a tree"
  1133 fun mirror :: "'a tree \<Rightarrow> 'a tree"
  1134 where
  1134 where
  1135   "mirror (Leaf n) = Leaf n"
  1135   "mirror (Leaf n) = Leaf n"
  1136 | "mirror (Branch l) = Branch (rev (map mirror l))"
  1136 | "mirror (Branch l) = Branch (rev (map mirror l))"
  1137 
  1137 
  1138 text {*
  1138 text \<open>
  1139   Although the definition is accepted without problems, let us look at the termination proof:
  1139   Although the definition is accepted without problems, let us look at the termination proof:
  1140 *}
  1140 \<close>
  1141 
  1141 
  1142 termination proof
  1142 termination proof
  1143   text {*
  1143   text \<open>
  1144 
  1144 
  1145   As usual, we have to give a wellfounded relation, such that the
  1145   As usual, we have to give a wellfounded relation, such that the
  1146   arguments of the recursive calls get smaller. But what exactly are
  1146   arguments of the recursive calls get smaller. But what exactly are
  1147   the arguments of the recursive calls when mirror is given as an
  1147   the arguments of the recursive calls when mirror is given as an
  1148   argument to @{const map}? Isabelle gives us the
  1148   argument to @{const map}? Isabelle gives us the
  1182   The constructs that are predefined in Isabelle, usually
  1182   The constructs that are predefined in Isabelle, usually
  1183   come with the respective congruence rules.
  1183   come with the respective congruence rules.
  1184   But if you define your own higher-order functions, you may have to
  1184   But if you define your own higher-order functions, you may have to
  1185   state and prove the required congruence rules yourself, if you want to use your
  1185   state and prove the required congruence rules yourself, if you want to use your
  1186   functions in recursive definitions. 
  1186   functions in recursive definitions. 
  1187 *}
  1187 \<close>
  1188 (*<*)oops(*>*)
  1188 (*<*)oops(*>*)
  1189 
  1189 
  1190 subsection {* Congruence Rules and Evaluation Order *}
  1190 subsection \<open>Congruence Rules and Evaluation Order\<close>
  1191 
  1191 
  1192 text {* 
  1192 text \<open>
  1193   Higher order logic differs from functional programming languages in
  1193   Higher order logic differs from functional programming languages in
  1194   that it has no built-in notion of evaluation order. A program is
  1194   that it has no built-in notion of evaluation order. A program is
  1195   just a set of equations, and it is not specified how they must be
  1195   just a set of equations, and it is not specified how they must be
  1196   evaluated. 
  1196   evaluated. 
  1197 
  1197 
  1199   evaluation order implicitly, when we reason about termination.
  1199   evaluation order implicitly, when we reason about termination.
  1200   Congruence rules express that a certain evaluation order is
  1200   Congruence rules express that a certain evaluation order is
  1201   consistent with the logical definition. 
  1201   consistent with the logical definition. 
  1202 
  1202 
  1203   Consider the following function.
  1203   Consider the following function.
  1204 *}
  1204 \<close>
  1205 
  1205 
  1206 function f :: "nat \<Rightarrow> bool"
  1206 function f :: "nat \<Rightarrow> bool"
  1207 where
  1207 where
  1208   "f n = (n = 0 \<or> f (n - 1))"
  1208   "f n = (n = 0 \<or> f (n - 1))"
  1209 (*<*)by pat_completeness auto(*>*)
  1209 (*<*)by pat_completeness auto(*>*)
  1210 
  1210 
  1211 text {*
  1211 text \<open>
  1212   For this definition, the termination proof fails. The default configuration
  1212   For this definition, the termination proof fails. The default configuration
  1213   specifies no congruence rule for disjunction. We have to add a
  1213   specifies no congruence rule for disjunction. We have to add a
  1214   congruence rule that specifies left-to-right evaluation order:
  1214   congruence rule that specifies left-to-right evaluation order:
  1215 
  1215 
  1216   \vspace{1ex}
  1216   \vspace{1ex}
  1222   rule.
  1222   rule.
  1223 
  1223 
  1224   However, as evaluation is not a hard-wired concept, we
  1224   However, as evaluation is not a hard-wired concept, we
  1225   could just turn everything around by declaring a different
  1225   could just turn everything around by declaring a different
  1226   congruence rule. Then we can make the reverse definition:
  1226   congruence rule. Then we can make the reverse definition:
  1227 *}
  1227 \<close>
  1228 
  1228 
  1229 lemma disj_cong2[fundef_cong]: 
  1229 lemma disj_cong2[fundef_cong]: 
  1230   "(\<not> Q' \<Longrightarrow> P = P') \<Longrightarrow> (Q = Q') \<Longrightarrow> (P \<or> Q) = (P' \<or> Q')"
  1230   "(\<not> Q' \<Longrightarrow> P = P') \<Longrightarrow> (Q = Q') \<Longrightarrow> (P \<or> Q) = (P' \<or> Q')"
  1231   by blast
  1231   by blast
  1232 
  1232 
  1233 fun f' :: "nat \<Rightarrow> bool"
  1233 fun f' :: "nat \<Rightarrow> bool"
  1234 where
  1234 where
  1235   "f' n = (f' (n - 1) \<or> n = 0)"
  1235   "f' n = (f' (n - 1) \<or> n = 0)"
  1236 
  1236 
  1237 text {*
  1237 text \<open>
  1238   \noindent These examples show that, in general, there is no \qt{best} set of
  1238   \noindent These examples show that, in general, there is no \qt{best} set of
  1239   congruence rules.
  1239   congruence rules.
  1240 
  1240 
  1241   However, such tweaking should rarely be necessary in
  1241   However, such tweaking should rarely be necessary in
  1242   practice, as most of the time, the default set of congruence rules
  1242   practice, as most of the time, the default set of congruence rules
  1243   works well.
  1243   works well.
  1244 *}
  1244 \<close>
  1245 
  1245 
  1246 end
  1246 end