| 
10879
 | 
     1  | 
% $Id$
  | 
| 
 | 
     2  | 
\section{The Set of Even Numbers}
 | 
| 
 | 
     3  | 
  | 
| 
11411
 | 
     4  | 
\index{even numbers!defining inductively|(}%
 | 
| 
10879
 | 
     5  | 
The set of even numbers can be inductively defined as the least set
  | 
| 
11129
 | 
     6  | 
containing 0 and closed under the operation $+2$.  Obviously,
  | 
| 
10879
 | 
     7  | 
\emph{even} can also be expressed using the divides relation (\isa{dvd}). 
 | 
| 
 | 
     8  | 
We shall prove below that the two formulations coincide.  On the way we
  | 
| 
 | 
     9  | 
shall examine the primary means of reasoning about inductively defined
  | 
| 
 | 
    10  | 
sets: rule induction.
  | 
| 
 | 
    11  | 
  | 
| 
 | 
    12  | 
\subsection{Making an Inductive Definition}
 | 
| 
 | 
    13  | 
  | 
| 
 | 
    14  | 
Using \isacommand{consts}, we declare the constant \isa{even} to be a set
 | 
| 
11411
 | 
    15  | 
of natural numbers. The \commdx{inductive} declaration gives it the
 | 
| 
10879
 | 
    16  | 
desired properties.
  | 
| 
 | 
    17  | 
\begin{isabelle}
 | 
| 
 | 
    18  | 
\isacommand{consts}\ even\ ::\ "nat\ set"\isanewline
 | 
| 
 | 
    19  | 
\isacommand{inductive}\ even\isanewline
 | 
| 
 | 
    20  | 
\isakeyword{intros}\isanewline
 | 
| 
 | 
    21  | 
zero[intro!]:\ "0\ \isasymin \ even"\isanewline
  | 
| 
 | 
    22  | 
step[intro!]:\ "n\ \isasymin \ even\ \isasymLongrightarrow \ (Suc\ (Suc\
  | 
| 
 | 
    23  | 
n))\ \isasymin \ even"
  | 
| 
 | 
    24  | 
\end{isabelle}
 | 
| 
 | 
    25  | 
  | 
| 
 | 
    26  | 
An inductive definition consists of introduction rules.  The first one
  | 
| 
 | 
    27  | 
above states that 0 is even; the second states that if $n$ is even, then so
  | 
| 
11173
 | 
    28  | 
is~$n+2$.  Given this declaration, Isabelle generates a fixed point
  | 
| 
11201
 | 
    29  | 
definition for \isa{even} and proves theorems about it,
 | 
| 
11494
 | 
    30  | 
thus following the definitional approach (see {\S}\ref{sec:definitional}).
 | 
| 
11201
 | 
    31  | 
These theorems
  | 
| 
11173
 | 
    32  | 
include the introduction rules specified in the declaration, an elimination
  | 
| 
 | 
    33  | 
rule for case analysis and an induction rule.  We can refer to these
  | 
| 
 | 
    34  | 
theorems by automatically-generated names.  Here are two examples:
  | 
| 
10879
 | 
    35  | 
%
  | 
| 
 | 
    36  | 
\begin{isabelle}
 | 
| 
 | 
    37  | 
0\ \isasymin \ even
  | 
| 
 | 
    38  | 
\rulename{even.zero}
 | 
| 
 | 
    39  | 
\par\smallskip
  | 
| 
 | 
    40  | 
n\ \isasymin \ even\ \isasymLongrightarrow \ Suc\ (Suc\ n)\ \isasymin \
  | 
| 
 | 
    41  | 
even%
  | 
| 
 | 
    42  | 
\rulename{even.step}
 | 
| 
 | 
    43  | 
\end{isabelle}
 | 
| 
 | 
    44  | 
  | 
| 
11411
 | 
    45  | 
The introduction rules can be given attributes.  Here
  | 
| 
 | 
    46  | 
both rules are specified as \isa{intro!},%
 | 
| 
 | 
    47  | 
\index{intro"!@\isa {intro"!} (attribute)}
 | 
| 
 | 
    48  | 
directing the classical reasoner to 
  | 
| 
10879
 | 
    49  | 
apply them aggressively. Obviously, regarding 0 as even is safe.  The
  | 
| 
 | 
    50  | 
\isa{step} rule is also safe because $n+2$ is even if and only if $n$ is
 | 
| 
 | 
    51  | 
even.  We prove this equivalence later.
  | 
| 
 | 
    52  | 
  | 
| 
 | 
    53  | 
\subsection{Using Introduction Rules}
 | 
| 
 | 
    54  | 
  | 
| 
 | 
    55  | 
Our first lemma states that numbers of the form $2\times k$ are even.
  | 
| 
 | 
    56  | 
Introduction rules are used to show that specific values belong to the
  | 
| 
 | 
    57  | 
inductive set.  Such proofs typically involve 
  | 
| 
 | 
    58  | 
induction, perhaps over some other inductive set.
  | 
| 
 | 
    59  | 
\begin{isabelle}
 | 
| 
12663
 | 
    60  | 
\isacommand{lemma}\ two_times_even[intro!]:\ "2*k\ \isasymin \ even"
 | 
| 
10879
 | 
    61  | 
\isanewline
  | 
| 
12328
 | 
    62  | 
\isacommand{apply}\ (induct_tac\ k)\isanewline
 | 
| 
10879
 | 
    63  | 
\ \isacommand{apply}\ auto\isanewline
 | 
| 
 | 
    64  | 
\isacommand{done}
 | 
| 
 | 
    65  | 
\end{isabelle}
 | 
| 
 | 
    66  | 
%
  | 
| 
 | 
    67  | 
The first step is induction on the natural number \isa{k}, which leaves
 | 
| 
 | 
    68  | 
two subgoals:
  | 
| 
 | 
    69  | 
\begin{isabelle}
 | 
| 
12663
 | 
    70  | 
\ 1.\ 2\ *\ 0\ \isasymin \ even\isanewline
  | 
| 
 | 
    71  | 
\ 2.\ \isasymAnd n.\ 2\ *\ n\ \isasymin \ even\ \isasymLongrightarrow \ 2\ *\ Suc\ n\ \isasymin \ even
  | 
| 
10879
 | 
    72  | 
\end{isabelle}
 | 
| 
 | 
    73  | 
%
  | 
| 
 | 
    74  | 
Here \isa{auto} simplifies both subgoals so that they match the introduction
 | 
| 
 | 
    75  | 
rules, which are then applied automatically.
  | 
| 
 | 
    76  | 
  | 
| 
 | 
    77  | 
Our ultimate goal is to prove the equivalence between the traditional
  | 
| 
 | 
    78  | 
definition of \isa{even} (using the divides relation) and our inductive
 | 
| 
 | 
    79  | 
definition.  One direction of this equivalence is immediate by the lemma
  | 
| 
11129
 | 
    80  | 
just proved, whose \isa{intro!} attribute ensures it is applied automatically.
 | 
| 
10879
 | 
    81  | 
\begin{isabelle}
 | 
| 
12663
 | 
    82  | 
\isacommand{lemma}\ dvd_imp_even:\ "2\ dvd\ n\ \isasymLongrightarrow \ n\ \isasymin \ even"\isanewline
 | 
| 
10879
 | 
    83  | 
\isacommand{by}\ (auto\ simp\ add:\ dvd_def)
 | 
| 
 | 
    84  | 
\end{isabelle}
 | 
| 
 | 
    85  | 
  | 
| 
 | 
    86  | 
\subsection{Rule Induction}
 | 
| 
 | 
    87  | 
\label{sec:rule-induction}
 | 
| 
 | 
    88  | 
  | 
| 
11411
 | 
    89  | 
\index{rule induction|(}%
 | 
| 
10879
 | 
    90  | 
From the definition of the set
  | 
| 
 | 
    91  | 
\isa{even}, Isabelle has
 | 
| 
 | 
    92  | 
generated an induction rule:
  | 
| 
 | 
    93  | 
\begin{isabelle}
 | 
| 
 | 
    94  | 
\isasymlbrakk xa\ \isasymin \ even;\isanewline
  | 
| 
 | 
    95  | 
\ P\ 0;\isanewline
  | 
| 
 | 
    96  | 
\ \isasymAnd n.\ \isasymlbrakk n\ \isasymin \ even;\ P\ n\isasymrbrakk \
  | 
| 
 | 
    97  | 
\isasymLongrightarrow \ P\ (Suc\ (Suc\ n))\isasymrbrakk\isanewline
  | 
| 
 | 
    98  | 
\ \isasymLongrightarrow \ P\ xa%
  | 
| 
 | 
    99  | 
\rulename{even.induct}
 | 
| 
 | 
   100  | 
\end{isabelle}
 | 
| 
 | 
   101  | 
A property \isa{P} holds for every even number provided it
 | 
| 
 | 
   102  | 
holds for~\isa{0} and is closed under the operation
 | 
| 
11129
 | 
   103  | 
\isa{Suc(Suc \(\cdot\))}.  Then \isa{P} is closed under the introduction
 | 
| 
10879
 | 
   104  | 
rules for \isa{even}, which is the least set closed under those rules. 
 | 
| 
 | 
   105  | 
This type of inductive argument is called \textbf{rule induction}. 
 | 
| 
 | 
   106  | 
  | 
| 
 | 
   107  | 
Apart from the double application of \isa{Suc}, the induction rule above
 | 
| 
 | 
   108  | 
resembles the familiar mathematical induction, which indeed is an instance
  | 
| 
 | 
   109  | 
of rule induction; the natural numbers can be defined inductively to be
  | 
| 
 | 
   110  | 
the least set containing \isa{0} and closed under~\isa{Suc}.
 | 
| 
 | 
   111  | 
  | 
| 
 | 
   112  | 
Induction is the usual way of proving a property of the elements of an
  | 
| 
 | 
   113  | 
inductively defined set.  Let us prove that all members of the set
  | 
| 
 | 
   114  | 
\isa{even} are multiples of two.  
 | 
| 
 | 
   115  | 
\begin{isabelle}
 | 
| 
12663
 | 
   116  | 
\isacommand{lemma}\ even_imp_dvd:\ "n\ \isasymin \ even\ \isasymLongrightarrow \ 2\ dvd\ n"
 | 
| 
10879
 | 
   117  | 
\end{isabelle}
 | 
| 
 | 
   118  | 
%
  | 
| 
 | 
   119  | 
We begin by applying induction.  Note that \isa{even.induct} has the form
 | 
| 
 | 
   120  | 
of an elimination rule, so we use the method \isa{erule}.  We get two
 | 
| 
 | 
   121  | 
subgoals:
  | 
| 
 | 
   122  | 
\begin{isabelle}
 | 
| 
 | 
   123  | 
\isacommand{apply}\ (erule\ even.induct)
 | 
| 
 | 
   124  | 
\isanewline\isanewline
  | 
| 
12663
 | 
   125  | 
\ 1.\ 2\ dvd\ 0\isanewline
  | 
| 
 | 
   126  | 
\ 2.\ \isasymAnd n.\ \isasymlbrakk n\ \isasymin \ even;\ 2\ dvd\ n\isasymrbrakk \ \isasymLongrightarrow \ 2\ dvd\ Suc\ (Suc\ n)
  | 
| 
10879
 | 
   127  | 
\end{isabelle}
 | 
| 
 | 
   128  | 
%
  | 
| 
 | 
   129  | 
We unfold the definition of \isa{dvd} in both subgoals, proving the first
 | 
| 
 | 
   130  | 
one and simplifying the second:
  | 
| 
 | 
   131  | 
\begin{isabelle}
 | 
| 
 | 
   132  | 
\isacommand{apply}\ (simp_all\ add:\ dvd_def)
 | 
| 
 | 
   133  | 
\isanewline\isanewline
  | 
| 
 | 
   134  | 
\ 1.\ \isasymAnd n.\ \isasymlbrakk n\ \isasymin \ even;\ \isasymexists k.\
  | 
| 
12663
 | 
   135  | 
n\ =\ 2\ *\ k\isasymrbrakk \ \isasymLongrightarrow \ \isasymexists k.\
  | 
| 
 | 
   136  | 
Suc\ (Suc\ n)\ =\ 2\ *\ k
  | 
| 
10879
 | 
   137  | 
\end{isabelle}
 | 
| 
 | 
   138  | 
%
  | 
| 
 | 
   139  | 
The next command eliminates the existential quantifier from the assumption
  | 
| 
12663
 | 
   140  | 
and replaces \isa{n} by \isa{2\ *\ k}.
 | 
| 
10879
 | 
   141  | 
\begin{isabelle}
 | 
| 
 | 
   142  | 
\isacommand{apply}\ clarify
 | 
| 
 | 
   143  | 
\isanewline\isanewline
  | 
| 
12663
 | 
   144  | 
\ 1.\ \isasymAnd n\ k.\ 2\ *\ k\ \isasymin \ even\ \isasymLongrightarrow \ \isasymexists ka.\ Suc\ (Suc\ (2\ *\ k))\ =\ 2\ *\ ka%
  | 
| 
10879
 | 
   145  | 
\end{isabelle}
 | 
| 
 | 
   146  | 
%
  | 
| 
 | 
   147  | 
To conclude, we tell Isabelle that the desired value is
  | 
| 
 | 
   148  | 
\isa{Suc\ k}.  With this hint, the subgoal falls to \isa{simp}.
 | 
| 
 | 
   149  | 
\begin{isabelle}
 | 
| 
11156
 | 
   150  | 
\isacommand{apply}\ (rule_tac\ x\ =\ "Suc\ k"\ \isakeyword{in}\ exI, simp)
 | 
| 
10879
 | 
   151  | 
\end{isabelle}
 | 
| 
 | 
   152  | 
  | 
| 
 | 
   153  | 
  | 
| 
 | 
   154  | 
\medskip
  | 
| 
 | 
   155  | 
Combining the previous two results yields our objective, the
  | 
| 
 | 
   156  | 
equivalence relating \isa{even} and \isa{dvd}. 
 | 
| 
 | 
   157  | 
%
  | 
| 
 | 
   158  | 
%we don't want [iff]: discuss?
  | 
| 
 | 
   159  | 
\begin{isabelle}
 | 
| 
12663
 | 
   160  | 
\isacommand{theorem}\ even_iff_dvd:\ "(n\ \isasymin \ even)\ =\ (2\ dvd\ n)"\isanewline
 | 
| 
10879
 | 
   161  | 
\isacommand{by}\ (blast\ intro:\ dvd_imp_even\ even_imp_dvd)
 | 
| 
 | 
   162  | 
\end{isabelle}
 | 
| 
 | 
   163  | 
  | 
| 
 | 
   164  | 
  | 
| 
 | 
   165  | 
\subsection{Generalization and Rule Induction}
 | 
| 
 | 
   166  | 
\label{sec:gen-rule-induction}
 | 
| 
 | 
   167  | 
  | 
| 
11411
 | 
   168  | 
\index{generalizing for induction}%
 | 
| 
10879
 | 
   169  | 
Before applying induction, we typically must generalize
  | 
| 
 | 
   170  | 
the induction formula.  With rule induction, the required generalization
  | 
| 
 | 
   171  | 
can be hard to find and sometimes requires a complete reformulation of the
  | 
| 
11156
 | 
   172  | 
problem.  In this  example, our first attempt uses the obvious statement of
  | 
| 
 | 
   173  | 
the result.  It fails:
  | 
| 
10879
 | 
   174  | 
%
  | 
| 
 | 
   175  | 
\begin{isabelle}
 | 
| 
 | 
   176  | 
\isacommand{lemma}\ "Suc\ (Suc\ n)\ \isasymin \ even\
 | 
| 
 | 
   177  | 
\isasymLongrightarrow \ n\ \isasymin \ even"\isanewline
  | 
| 
 | 
   178  | 
\isacommand{apply}\ (erule\ even.induct)\isanewline
 | 
| 
 | 
   179  | 
\isacommand{oops}
 | 
| 
 | 
   180  | 
\end{isabelle}
 | 
| 
 | 
   181  | 
%
  | 
| 
 | 
   182  | 
Rule induction finds no occurrences of \isa{Suc(Suc\ n)} in the
 | 
| 
 | 
   183  | 
conclusion, which it therefore leaves unchanged.  (Look at
  | 
| 
 | 
   184  | 
\isa{even.induct} to see why this happens.)  We have these subgoals:
 | 
| 
 | 
   185  | 
\begin{isabelle}
 | 
| 
 | 
   186  | 
\ 1.\ n\ \isasymin \ even\isanewline
  | 
| 
 | 
   187  | 
\ 2.\ \isasymAnd na.\ \isasymlbrakk na\ \isasymin \ even;\ n\ \isasymin \ even\isasymrbrakk \ \isasymLongrightarrow \ n\ \isasymin \ even%
  | 
| 
 | 
   188  | 
\end{isabelle}
 | 
| 
 | 
   189  | 
The first one is hopeless.  Rule inductions involving
  | 
| 
 | 
   190  | 
non-trivial terms usually fail.  How to deal with such situations
  | 
| 
 | 
   191  | 
in general is described in {\S}\ref{sec:ind-var-in-prems} below.
 | 
| 
 | 
   192  | 
In the current case the solution is easy because
  | 
| 
 | 
   193  | 
we have the necessary inverse, subtraction:
  | 
| 
 | 
   194  | 
\begin{isabelle}
 | 
| 
12663
 | 
   195  | 
\isacommand{lemma}\ even_imp_even_minus_2:\ "n\ \isasymin \ even\ \isasymLongrightarrow \ n-2\ \isasymin \ even"\isanewline
 | 
| 
10879
 | 
   196  | 
\isacommand{apply}\ (erule\ even.induct)\isanewline
 | 
| 
 | 
   197  | 
\ \isacommand{apply}\ auto\isanewline
 | 
| 
 | 
   198  | 
\isacommand{done}
 | 
| 
 | 
   199  | 
\end{isabelle}
 | 
| 
 | 
   200  | 
%
  | 
| 
 | 
   201  | 
This lemma is trivially inductive.  Here are the subgoals:
  | 
| 
 | 
   202  | 
\begin{isabelle}
 | 
| 
12663
 | 
   203  | 
\ 1.\ 0\ -\ 2\ \isasymin \ even\isanewline
  | 
| 
 | 
   204  | 
\ 2.\ \isasymAnd n.\ \isasymlbrakk n\ \isasymin \ even;\ n\ -\ 2\ \isasymin \ even\isasymrbrakk \ \isasymLongrightarrow \ Suc\ (Suc\ n)\ -\ 2\ \isasymin \ even%
  | 
| 
10879
 | 
   205  | 
\end{isabelle}
 | 
| 
12663
 | 
   206  | 
The first is trivial because \isa{0\ -\ 2} simplifies to \isa{0}, which is
 | 
| 
 | 
   207  | 
even.  The second is trivial too: \isa{Suc\ (Suc\ n)\ -\ 2} simplifies to
 | 
| 
11411
 | 
   208  | 
\isa{n}, matching the assumption.%
 | 
| 
 | 
   209  | 
\index{rule induction|)}  %the sequel isn't really about induction
 | 
| 
10879
 | 
   210  | 
  | 
| 
 | 
   211  | 
\medskip
  | 
| 
 | 
   212  | 
Using our lemma, we can easily prove the result we originally wanted:
  | 
| 
 | 
   213  | 
\begin{isabelle}
 | 
| 
 | 
   214  | 
\isacommand{lemma}\ Suc_Suc_even_imp_even:\ "Suc\ (Suc\ n)\ \isasymin \ even\ \isasymLongrightarrow \ n\ \isasymin \ even"\isanewline
 | 
| 
11156
 | 
   215  | 
\isacommand{by}\ (drule\ even_imp_even_minus_2, simp)
 | 
| 
10879
 | 
   216  | 
\end{isabelle}
 | 
| 
 | 
   217  | 
  | 
| 
 | 
   218  | 
We have just proved the converse of the introduction rule \isa{even.step}. 
 | 
| 
11411
 | 
   219  | 
This suggests proving the following equivalence.  We give it the
  | 
| 
 | 
   220  | 
\attrdx{iff} attribute because of its obvious value for simplification.
 | 
| 
10879
 | 
   221  | 
\begin{isabelle}
 | 
| 
 | 
   222  | 
\isacommand{lemma}\ [iff]:\ "((Suc\ (Suc\ n))\ \isasymin \ even)\ =\ (n\
 | 
| 
 | 
   223  | 
\isasymin \ even)"\isanewline
  | 
| 
 | 
   224  | 
\isacommand{by}\ (blast\ dest:\ Suc_Suc_even_imp_even)
 | 
| 
 | 
   225  | 
\end{isabelle}
 | 
| 
 | 
   226  | 
  | 
| 
11173
 | 
   227  | 
  | 
| 
 | 
   228  | 
\subsection{Rule Inversion}\label{sec:rule-inversion}
 | 
| 
 | 
   229  | 
  | 
| 
11411
 | 
   230  | 
\index{rule inversion|(}%
 | 
| 
 | 
   231  | 
Case analysis on an inductive definition is called \textbf{rule
 | 
| 
 | 
   232  | 
inversion}.  It is frequently used in proofs about operational
  | 
| 
 | 
   233  | 
semantics.  It can be highly effective when it is applied
  | 
| 
 | 
   234  | 
automatically.  Let us look at how rule inversion is done in
  | 
| 
 | 
   235  | 
Isabelle/HOL\@.
  | 
| 
11173
 | 
   236  | 
  | 
| 
 | 
   237  | 
Recall that \isa{even} is the minimal set closed under these two rules:
 | 
| 
 | 
   238  | 
\begin{isabelle}
 | 
| 
 | 
   239  | 
0\ \isasymin \ even\isanewline
  | 
| 
 | 
   240  | 
n\ \isasymin \ even\ \isasymLongrightarrow \ Suc\ (Suc\ n)\ \isasymin
  | 
| 
 | 
   241  | 
\ even
  | 
| 
 | 
   242  | 
\end{isabelle}
 | 
| 
 | 
   243  | 
Minimality means that \isa{even} contains only the elements that these
 | 
| 
 | 
   244  | 
rules force it to contain.  If we are told that \isa{a}
 | 
| 
 | 
   245  | 
belongs to
  | 
| 
 | 
   246  | 
\isa{even} then there are only two possibilities.  Either \isa{a} is \isa{0}
 | 
| 
 | 
   247  | 
or else \isa{a} has the form \isa{Suc(Suc~n)}, for some suitable \isa{n}
 | 
| 
 | 
   248  | 
that belongs to
  | 
| 
 | 
   249  | 
\isa{even}.  That is the gist of the \isa{cases} rule, which Isabelle proves
 | 
| 
 | 
   250  | 
for us when it accepts an inductive definition:
  | 
| 
 | 
   251  | 
\begin{isabelle}
 | 
| 
 | 
   252  | 
\isasymlbrakk a\ \isasymin \ even;\isanewline
  | 
| 
 | 
   253  | 
\ a\ =\ 0\ \isasymLongrightarrow \ P;\isanewline
  | 
| 
 | 
   254  | 
\ \isasymAnd n.\ \isasymlbrakk a\ =\ Suc(Suc\ n);\ n\ \isasymin \
  | 
| 
 | 
   255  | 
even\isasymrbrakk \ \isasymLongrightarrow \ P\isasymrbrakk \
  | 
| 
 | 
   256  | 
\isasymLongrightarrow \ P
  | 
| 
 | 
   257  | 
\rulename{even.cases}
 | 
| 
 | 
   258  | 
\end{isabelle}
 | 
| 
 | 
   259  | 
  | 
| 
 | 
   260  | 
This general rule is less useful than instances of it for
  | 
| 
 | 
   261  | 
specific patterns.  For example, if \isa{a} has the form
 | 
| 
 | 
   262  | 
\isa{Suc(Suc~n)} then the first case becomes irrelevant, while the second
 | 
| 
 | 
   263  | 
case tells us that \isa{n} belongs to \isa{even}.  Isabelle will generate
 | 
| 
 | 
   264  | 
this instance for us:
  | 
| 
 | 
   265  | 
\begin{isabelle}
 | 
| 
 | 
   266  | 
\isacommand{inductive\_cases}\ Suc_Suc_cases\ [elim!]:
 | 
| 
 | 
   267  | 
\ "Suc(Suc\ n)\ \isasymin \ even"
  | 
| 
 | 
   268  | 
\end{isabelle}
 | 
| 
11411
 | 
   269  | 
The \commdx{inductive\protect\_cases} command generates an instance of
 | 
| 
 | 
   270  | 
the
  | 
| 
11173
 | 
   271  | 
\isa{cases} rule for the supplied pattern and gives it the supplied name:
 | 
| 
 | 
   272  | 
%
  | 
| 
 | 
   273  | 
\begin{isabelle}
 | 
| 
 | 
   274  | 
\isasymlbrakk Suc(Suc\ n)\ \isasymin \ even;\ n\ \isasymin \ even\
  | 
| 
 | 
   275  | 
\isasymLongrightarrow \ P\isasymrbrakk \ \isasymLongrightarrow \ P%
  | 
| 
 | 
   276  | 
\rulename{Suc_Suc_cases}
 | 
| 
 | 
   277  | 
\end{isabelle}
 | 
| 
 | 
   278  | 
%
  | 
| 
 | 
   279  | 
Applying this as an elimination rule yields one case where \isa{even.cases}
 | 
| 
 | 
   280  | 
would yield two.  Rule inversion works well when the conclusions of the
  | 
| 
 | 
   281  | 
introduction rules involve datatype constructors like \isa{Suc} and \isa{\#}
 | 
| 
 | 
   282  | 
(list ``cons''); freeness reasoning discards all but one or two cases.
  | 
| 
 | 
   283  | 
  | 
| 
 | 
   284  | 
In the \isacommand{inductive\_cases} command we supplied an
 | 
| 
11411
 | 
   285  | 
attribute, \isa{elim!},
 | 
| 
 | 
   286  | 
\index{elim"!@\isa {elim"!} (attribute)}%
 | 
| 
 | 
   287  | 
indicating that this elimination rule can be
  | 
| 
 | 
   288  | 
applied aggressively.  The original
  | 
| 
11173
 | 
   289  | 
\isa{cases} rule would loop if used in that manner because the
 | 
| 
 | 
   290  | 
pattern~\isa{a} matches everything.
 | 
| 
 | 
   291  | 
  | 
| 
 | 
   292  | 
The rule \isa{Suc_Suc_cases} is equivalent to the following implication:
 | 
| 
 | 
   293  | 
\begin{isabelle}
 | 
| 
 | 
   294  | 
Suc (Suc\ n)\ \isasymin \ even\ \isasymLongrightarrow \ n\ \isasymin \
  | 
| 
 | 
   295  | 
even
  | 
| 
 | 
   296  | 
\end{isabelle}
 | 
| 
 | 
   297  | 
%
  | 
| 
 | 
   298  | 
Just above we devoted some effort to reaching precisely
  | 
| 
 | 
   299  | 
this result.  Yet we could have obtained it by a one-line declaration,
  | 
| 
 | 
   300  | 
dispensing with the lemma \isa{even_imp_even_minus_2}. 
 | 
| 
 | 
   301  | 
This example also justifies the terminology
  | 
| 
 | 
   302  | 
\textbf{rule inversion}: the new rule inverts the introduction rule
 | 
| 
11494
 | 
   303  | 
\isa{even.step}.  In general, a rule can be inverted when the set of elements
 | 
| 
 | 
   304  | 
it introduces is disjoint from those of the other introduction rules.
  | 
| 
11173
 | 
   305  | 
  | 
| 
11494
 | 
   306  | 
For one-off applications of rule inversion, use the \methdx{ind_cases} method. 
 | 
| 
11173
 | 
   307  | 
Here is an example:
  | 
| 
 | 
   308  | 
\begin{isabelle}
 | 
| 
 | 
   309  | 
\isacommand{apply}\ (ind_cases\ "Suc(Suc\ n)\ \isasymin \ even")
 | 
| 
 | 
   310  | 
\end{isabelle}
 | 
| 
 | 
   311  | 
The specified instance of the \isa{cases} rule is generated, then applied
 | 
| 
 | 
   312  | 
as an elimination rule.
  | 
| 
 | 
   313  | 
  | 
| 
 | 
   314  | 
To summarize, every inductive definition produces a \isa{cases} rule.  The
 | 
| 
11411
 | 
   315  | 
\commdx{inductive\protect\_cases} command stores an instance of the
 | 
| 
 | 
   316  | 
\isa{cases} rule for a given pattern.  Within a proof, the
 | 
| 
 | 
   317  | 
\isa{ind_cases} method applies an instance of the \isa{cases}
 | 
| 
11173
 | 
   318  | 
rule.
  | 
| 
 | 
   319  | 
  | 
| 
 | 
   320  | 
The even numbers example has shown how inductive definitions can be
  | 
| 
11411
 | 
   321  | 
used.  Later examples will show that they are actually worth using.%
  | 
| 
 | 
   322  | 
\index{rule inversion|)}%
 | 
| 
 | 
   323  | 
\index{even numbers!defining inductively|)}
 |