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