| 9742 |      1 | (*<*)
 | 
| 16417 |      2 | theory case_exprs imports Main begin
 | 
| 9742 |      3 | (*>*)
 | 
|  |      4 | 
 | 
| 67406 |      5 | text\<open>
 | 
| 25340 |      6 | \subsection{Case Expressions}
 | 
|  |      7 | \label{sec:case-expressions}\index{*case expressions}%
 | 
| 11456 |      8 | HOL also features \isa{case}-expressions for analyzing
 | 
| 9742 |      9 | elements of a datatype. For example,
 | 
| 12699 |     10 | @{term[display]"case xs of [] => [] | y#ys => y"}
 | 
| 69597 |     11 | evaluates to \<^term>\<open>[]\<close> if \<^term>\<open>xs\<close> is \<^term>\<open>[]\<close> and to \<^term>\<open>y\<close> if 
 | 
|  |     12 | \<^term>\<open>xs\<close> is \<^term>\<open>y#ys\<close>. (Since the result in both branches must be of
 | 
|  |     13 | the same type, it follows that \<^term>\<open>y\<close> is of type \<^typ>\<open>'a list\<close> and hence
 | 
|  |     14 | that \<^term>\<open>xs\<close> is of type \<^typ>\<open>'a list list\<close>.)
 | 
| 9742 |     15 | 
 | 
| 25340 |     16 | In general, case expressions are of the form
 | 
| 9742 |     17 | \[
 | 
| 25340 |     18 | \begin{array}{c}
 | 
| 69505 |     19 | \<open>case\<close>~e~\<open>of\<close>\ pattern@1~\<open>\<Rightarrow>\<close>~e@1\ \<open>|\<close>\ \dots\
 | 
|  |     20 |  \<open>|\<close>~pattern@m~\<open>\<Rightarrow>\<close>~e@m
 | 
| 9742 |     21 | \end{array}
 | 
|  |     22 | \]
 | 
| 25340 |     23 | Like in functional programming, patterns are expressions consisting of
 | 
| 69597 |     24 | datatype constructors (e.g. \<^term>\<open>[]\<close> and \<open>#\<close>)
 | 
| 25340 |     25 | and variables, including the wildcard ``\verb$_$''.
 | 
|  |     26 | Not all cases need to be covered and the order of cases matters.
 | 
|  |     27 | However, one is well-advised not to wallow in complex patterns because
 | 
|  |     28 | complex case distinctions tend to induce complex proofs.
 | 
| 9742 |     29 | 
 | 
|  |     30 | \begin{warn}
 | 
| 25340 |     31 | Internally Isabelle only knows about exhaustive case expressions with
 | 
|  |     32 | non-nested patterns: $pattern@i$ must be of the form
 | 
|  |     33 | $C@i~x@ {i1}~\dots~x@ {ik@i}$ and $C@1, \dots, C@m$ must be exactly the
 | 
|  |     34 | constructors of the type of $e$.
 | 
|  |     35 | %
 | 
|  |     36 | More complex case expressions are automatically
 | 
|  |     37 | translated into the simpler form upon parsing but are not translated
 | 
|  |     38 | back for printing. This may lead to surprising output.
 | 
| 9742 |     39 | \end{warn}
 | 
|  |     40 | 
 | 
| 25340 |     41 | \begin{warn}
 | 
| 69505 |     42 | Like \<open>if\<close>, \<open>case\<close>-expressions may need to be enclosed in
 | 
| 25340 |     43 | parentheses to indicate their scope.
 | 
|  |     44 | \end{warn}
 | 
| 9742 |     45 | 
 | 
| 25340 |     46 | \subsection{Structural Induction and Case Distinction}
 | 
|  |     47 | \label{sec:struct-ind-case}
 | 
| 11456 |     48 | \index{case distinctions}\index{induction!structural}%
 | 
|  |     49 | Induction is invoked by \methdx{induct_tac}, as we have seen above; 
 | 
|  |     50 | it works for any datatype.  In some cases, induction is overkill and a case
 | 
|  |     51 | distinction over all constructors of the datatype suffices.  This is performed
 | 
|  |     52 | by \methdx{case_tac}.  Here is a trivial example:
 | 
| 67406 |     53 | \<close>
 | 
| 9742 |     54 | 
 | 
| 58860 |     55 | lemma "(case xs of [] \<Rightarrow> [] | y#ys \<Rightarrow> xs) = xs"
 | 
|  |     56 | apply(case_tac xs)
 | 
| 9742 |     57 | 
 | 
| 67406 |     58 | txt\<open>\noindent
 | 
| 9742 |     59 | results in the proof state
 | 
| 10420 |     60 | @{subgoals[display,indent=0,margin=65]}
 | 
| 9742 |     61 | which is solved automatically:
 | 
| 67406 |     62 | \<close>
 | 
| 9742 |     63 | 
 | 
| 10171 |     64 | apply(auto)
 | 
|  |     65 | (*<*)done(*>*)
 | 
| 67406 |     66 | text\<open>
 | 
| 9742 |     67 | Note that we do not need to give a lemma a name if we do not intend to refer
 | 
|  |     68 | to it explicitly in the future.
 | 
| 11456 |     69 | Other basic laws about a datatype are applied automatically during
 | 
|  |     70 | simplification, so no special methods are provided for them.
 | 
| 10824 |     71 | 
 | 
|  |     72 | \begin{warn}
 | 
|  |     73 |   Induction is only allowed on free (or \isasymAnd-bound) variables that
 | 
|  |     74 |   should not occur among the assumptions of the subgoal; see
 | 
|  |     75 |   \S\ref{sec:ind-var-in-prems} for details. Case distinction
 | 
| 69505 |     76 |   (\<open>case_tac\<close>) works for arbitrary terms, which need to be
 | 
|  |     77 |   quoted if they are non-atomic. However, apart from \<open>\<And>\<close>-bound
 | 
| 10824 |     78 |   variables, the terms must not contain variables that are bound outside.
 | 
| 69597 |     79 |   For example, given the goal \<^prop>\<open>\<forall>xs. xs = [] \<or> (\<exists>y ys. xs = y#ys)\<close>,
 | 
| 69505 |     80 |   \<open>case_tac xs\<close> will not work as expected because Isabelle interprets
 | 
| 69597 |     81 |   the \<^term>\<open>xs\<close> as a new free variable distinct from the bound
 | 
|  |     82 |   \<^term>\<open>xs\<close> in the goal.
 | 
| 10824 |     83 | \end{warn}
 | 
| 67406 |     84 | \<close>
 | 
| 9742 |     85 | 
 | 
|  |     86 | (*<*)
 | 
|  |     87 | end
 | 
|  |     88 | (*>*)
 |