| 9742 |      1 | (*<*)
 | 
| 16417 |      2 | theory case_exprs imports Main begin
 | 
| 9742 |      3 | (*>*)
 | 
|  |      4 | 
 | 
| 10885 |      5 | subsection{*Case Expressions*}
 | 
| 9742 |      6 | 
 | 
| 11456 |      7 | text{*\label{sec:case-expressions}\index{*case expressions}%
 | 
|  |      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"}
 | 
|  |     11 | evaluates to @{term"[]"} if @{term"xs"} is @{term"[]"} and to @{term"y"} if 
 | 
| 9742 |     12 | @{term"xs"} is @{term"y#ys"}. (Since the result in both branches must be of
 | 
| 12699 |     13 | the same type, it follows that @{term y} is of type @{typ"'a list"} and hence
 | 
|  |     14 | that @{term xs} is of type @{typ"'a list list"}.)
 | 
| 9742 |     15 | 
 | 
|  |     16 | In general, if $e$ is a term of the datatype $t$ defined in
 | 
|  |     17 | \S\ref{sec:general-datatype} above, the corresponding
 | 
| 9792 |     18 | @{text"case"}-expression analyzing $e$ is
 | 
| 9742 |     19 | \[
 | 
|  |     20 | \begin{array}{rrcl}
 | 
| 9792 |     21 | @{text"case"}~e~@{text"of"} & C@1~x@ {11}~\dots~x@ {1k@1} & \To & e@1 \\
 | 
| 9742 |     22 |                            \vdots \\
 | 
|  |     23 |                            \mid & C@m~x@ {m1}~\dots~x@ {mk@m} & \To & e@m
 | 
|  |     24 | \end{array}
 | 
|  |     25 | \]
 | 
|  |     26 | 
 | 
|  |     27 | \begin{warn}
 | 
|  |     28 | \emph{All} constructors must be present, their order is fixed, and nested
 | 
|  |     29 | patterns are not supported.  Violating these restrictions results in strange
 | 
|  |     30 | error messages.
 | 
|  |     31 | \end{warn}
 | 
|  |     32 | \noindent
 | 
| 9792 |     33 | Nested patterns can be simulated by nested @{text"case"}-expressions: instead
 | 
| 9742 |     34 | of
 | 
| 12699 |     35 | @{text[display]"case xs of [] => [] | [x] => x | x # (y # zs) => y"}
 | 
| 9742 |     36 | write
 | 
| 12699 |     37 | @{term[display,eta_contract=false,margin=50]"case xs of [] => [] | x#ys => (case ys of [] => x | y#zs => y)"}
 | 
| 9742 |     38 | 
 | 
| 9792 |     39 | Note that @{text"case"}-expressions may need to be enclosed in parentheses to
 | 
| 9742 |     40 | indicate their scope
 | 
|  |     41 | *}
 | 
|  |     42 | 
 | 
| 10885 |     43 | subsection{*Structural Induction and Case Distinction*}
 | 
| 9742 |     44 | 
 | 
| 10824 |     45 | text{*\label{sec:struct-ind-case}
 | 
| 11456 |     46 | \index{case distinctions}\index{induction!structural}%
 | 
|  |     47 | Induction is invoked by \methdx{induct_tac}, as we have seen above; 
 | 
|  |     48 | it works for any datatype.  In some cases, induction is overkill and a case
 | 
|  |     49 | distinction over all constructors of the datatype suffices.  This is performed
 | 
|  |     50 | by \methdx{case_tac}.  Here is a trivial example:
 | 
| 9742 |     51 | *}
 | 
|  |     52 | 
 | 
|  |     53 | lemma "(case xs of [] \<Rightarrow> [] | y#ys \<Rightarrow> xs) = xs";
 | 
|  |     54 | apply(case_tac xs);
 | 
|  |     55 | 
 | 
|  |     56 | txt{*\noindent
 | 
|  |     57 | results in the proof state
 | 
| 10420 |     58 | @{subgoals[display,indent=0,margin=65]}
 | 
| 9742 |     59 | which is solved automatically:
 | 
|  |     60 | *}
 | 
|  |     61 | 
 | 
| 10171 |     62 | apply(auto)
 | 
|  |     63 | (*<*)done(*>*)
 | 
| 9742 |     64 | text{*
 | 
|  |     65 | Note that we do not need to give a lemma a name if we do not intend to refer
 | 
|  |     66 | to it explicitly in the future.
 | 
| 11456 |     67 | Other basic laws about a datatype are applied automatically during
 | 
|  |     68 | simplification, so no special methods are provided for them.
 | 
| 10824 |     69 | 
 | 
|  |     70 | \begin{warn}
 | 
|  |     71 |   Induction is only allowed on free (or \isasymAnd-bound) variables that
 | 
|  |     72 |   should not occur among the assumptions of the subgoal; see
 | 
|  |     73 |   \S\ref{sec:ind-var-in-prems} for details. Case distinction
 | 
|  |     74 |   (@{text"case_tac"}) works for arbitrary terms, which need to be
 | 
|  |     75 |   quoted if they are non-atomic. However, apart from @{text"\<And>"}-bound
 | 
|  |     76 |   variables, the terms must not contain variables that are bound outside.
 | 
|  |     77 |   For example, given the goal @{prop"\<forall>xs. xs = [] \<or> (\<exists>y ys. xs = y#ys)"},
 | 
|  |     78 |   @{text"case_tac xs"} will not work as expected because Isabelle interprets
 | 
|  |     79 |   the @{term xs} as a new free variable distinct from the bound
 | 
|  |     80 |   @{term xs} in the goal.
 | 
|  |     81 | \end{warn}
 | 
| 9742 |     82 | *}
 | 
|  |     83 | 
 | 
|  |     84 | (*<*)
 | 
|  |     85 | end
 | 
|  |     86 | (*>*)
 |