| 
9722
 | 
     1  | 
%
  | 
| 
 | 
     2  | 
\begin{isabellebody}%
 | 
| 
9924
 | 
     3  | 
\def\isabellecontext{AdvancedInd}%
 | 
| 
17056
 | 
     4  | 
%
  | 
| 
 | 
     5  | 
\isadelimtheory
  | 
| 
 | 
     6  | 
%
  | 
| 
 | 
     7  | 
\endisadelimtheory
  | 
| 
 | 
     8  | 
%
  | 
| 
 | 
     9  | 
\isatagtheory
  | 
| 
 | 
    10  | 
%
  | 
| 
 | 
    11  | 
\endisatagtheory
  | 
| 
 | 
    12  | 
{\isafoldtheory}%
 | 
| 
 | 
    13  | 
%
  | 
| 
 | 
    14  | 
\isadelimtheory
  | 
| 
 | 
    15  | 
%
  | 
| 
 | 
    16  | 
\endisadelimtheory
  | 
| 
9670
 | 
    17  | 
%
  | 
| 
 | 
    18  | 
\begin{isamarkuptext}%
 | 
| 
 | 
    19  | 
\noindent
  | 
| 
 | 
    20  | 
Now that we have learned about rules and logic, we take another look at the
  | 
| 
11494
 | 
    21  | 
finer points of induction.  We consider two questions: what to do if the
  | 
| 
10396
 | 
    22  | 
proposition to be proved is not directly amenable to induction
  | 
| 
 | 
    23  | 
(\S\ref{sec:ind-var-in-prems}), and how to utilize (\S\ref{sec:complete-ind})
 | 
| 
 | 
    24  | 
and even derive (\S\ref{sec:derive-ind}) new induction schemas. We conclude
 | 
| 
 | 
    25  | 
with an extended example of induction (\S\ref{sec:CTL-revisited}).%
 | 
| 
9670
 | 
    26  | 
\end{isamarkuptext}%
 | 
| 
11866
 | 
    27  | 
\isamarkuptrue%
  | 
| 
9670
 | 
    28  | 
%
  | 
| 
10878
 | 
    29  | 
\isamarkupsubsection{Massaging the Proposition%
 | 
| 
10397
 | 
    30  | 
}
  | 
| 
11866
 | 
    31  | 
\isamarkuptrue%
  | 
| 
9670
 | 
    32  | 
%
  | 
| 
 | 
    33  | 
\begin{isamarkuptext}%
 | 
| 
10217
 | 
    34  | 
\label{sec:ind-var-in-prems}
 | 
| 
11494
 | 
    35  | 
Often we have assumed that the theorem to be proved is already in a form
  | 
| 
10878
 | 
    36  | 
that is amenable to induction, but sometimes it isn't.
  | 
| 
 | 
    37  | 
Here is an example.
  | 
| 
 | 
    38  | 
Since \isa{hd} and \isa{last} return the first and last element of a
 | 
| 
 | 
    39  | 
non-empty list, this lemma looks easy to prove:%
  | 
| 
9670
 | 
    40  | 
\end{isamarkuptext}%
 | 
| 
17175
 | 
    41  | 
\isamarkuptrue%
  | 
| 
 | 
    42  | 
\isacommand{lemma}\isamarkupfalse%
 | 
| 
 | 
    43  | 
\ {\isachardoublequoteopen}xs\ {\isasymnoteq}\ {\isacharbrackleft}{\isacharbrackright}\ {\isasymLongrightarrow}\ hd{\isacharparenleft}rev\ xs{\isacharparenright}\ {\isacharequal}\ last\ xs{\isachardoublequoteclose}\isanewline
 | 
| 
17056
 | 
    44  | 
%
  | 
| 
 | 
    45  | 
\isadelimproof
  | 
| 
 | 
    46  | 
%
  | 
| 
 | 
    47  | 
\endisadelimproof
  | 
| 
 | 
    48  | 
%
  | 
| 
 | 
    49  | 
\isatagproof
  | 
| 
17175
 | 
    50  | 
\isacommand{apply}\isamarkupfalse%
 | 
| 
 | 
    51  | 
{\isacharparenleft}induct{\isacharunderscore}tac\ xs{\isacharparenright}%
 | 
| 
16069
 | 
    52  | 
\begin{isamarkuptxt}%
 | 
| 
 | 
    53  | 
\noindent
  | 
| 
 | 
    54  | 
But induction produces the warning
  | 
| 
 | 
    55  | 
\begin{quote}\tt
 | 
| 
 | 
    56  | 
Induction variable occurs also among premises!
  | 
| 
 | 
    57  | 
\end{quote}
 | 
| 
 | 
    58  | 
and leads to the base case
  | 
| 
 | 
    59  | 
\begin{isabelle}%
 | 
| 
 | 
    60  | 
\ {\isadigit{1}}{\isachardot}\ xs\ {\isasymnoteq}\ {\isacharbrackleft}{\isacharbrackright}\ {\isasymLongrightarrow}\ hd\ {\isacharparenleft}rev\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\ {\isacharequal}\ last\ {\isacharbrackleft}{\isacharbrackright}%
 | 
| 
 | 
    61  | 
\end{isabelle}
 | 
| 
 | 
    62  | 
Simplification reduces the base case to this:
  | 
| 
 | 
    63  | 
\begin{isabelle}
 | 
| 
 | 
    64  | 
\ 1.\ xs\ {\isasymnoteq}\ []\ {\isasymLongrightarrow}\ hd\ []\ =\ last\ []
 | 
| 
 | 
    65  | 
\end{isabelle}
 | 
| 
 | 
    66  | 
We cannot prove this equality because we do not know what \isa{hd} and
 | 
| 
 | 
    67  | 
\isa{last} return when applied to \isa{{\isacharbrackleft}{\isacharbrackright}}.
 | 
| 
 | 
    68  | 
  | 
| 
 | 
    69  | 
We should not have ignored the warning. Because the induction
  | 
| 
 | 
    70  | 
formula is only the conclusion, induction does not affect the occurrence of \isa{xs} in the premises.  
 | 
| 
 | 
    71  | 
Thus the case that should have been trivial
  | 
| 
 | 
    72  | 
becomes unprovable. Fortunately, the solution is easy:\footnote{A similar
 | 
| 
 | 
    73  | 
heuristic applies to rule inductions; see \S\ref{sec:rtc}.}
 | 
| 
 | 
    74  | 
\begin{quote}
 | 
| 
 | 
    75  | 
\emph{Pull all occurrences of the induction variable into the conclusion
 | 
| 
 | 
    76  | 
using \isa{{\isasymlongrightarrow}}.}
 | 
| 
 | 
    77  | 
\end{quote}
 | 
| 
 | 
    78  | 
Thus we should state the lemma as an ordinary 
  | 
| 
 | 
    79  | 
implication~(\isa{{\isasymlongrightarrow}}), letting
 | 
| 
 | 
    80  | 
\attrdx{rule_format} (\S\ref{sec:forward}) convert the
 | 
| 
 | 
    81  | 
result to the usual \isa{{\isasymLongrightarrow}} form:%
 | 
| 
 | 
    82  | 
\end{isamarkuptxt}%
 | 
| 
17175
 | 
    83  | 
\isamarkuptrue%
  | 
| 
17056
 | 
    84  | 
%
  | 
| 
 | 
    85  | 
\endisatagproof
  | 
| 
 | 
    86  | 
{\isafoldproof}%
 | 
| 
 | 
    87  | 
%
  | 
| 
 | 
    88  | 
\isadelimproof
  | 
| 
 | 
    89  | 
%
  | 
| 
 | 
    90  | 
\endisadelimproof
  | 
| 
17175
 | 
    91  | 
\isacommand{lemma}\isamarkupfalse%
 | 
| 
 | 
    92  | 
\ hd{\isacharunderscore}rev\ {\isacharbrackleft}rule{\isacharunderscore}format{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}xs\ {\isasymnoteq}\ {\isacharbrackleft}{\isacharbrackright}\ {\isasymlongrightarrow}\ hd{\isacharparenleft}rev\ xs{\isacharparenright}\ {\isacharequal}\ last\ xs{\isachardoublequoteclose}%
 | 
| 
17056
 | 
    93  | 
\isadelimproof
  | 
| 
 | 
    94  | 
%
  | 
| 
 | 
    95  | 
\endisadelimproof
  | 
| 
 | 
    96  | 
%
  | 
| 
 | 
    97  | 
\isatagproof
  | 
| 
16069
 | 
    98  | 
%
  | 
| 
 | 
    99  | 
\begin{isamarkuptxt}%
 | 
| 
 | 
   100  | 
\noindent
  | 
| 
 | 
   101  | 
This time, induction leaves us with a trivial base case:
  | 
| 
 | 
   102  | 
\begin{isabelle}%
 | 
| 
 | 
   103  | 
\ {\isadigit{1}}{\isachardot}\ {\isacharbrackleft}{\isacharbrackright}\ {\isasymnoteq}\ {\isacharbrackleft}{\isacharbrackright}\ {\isasymlongrightarrow}\ hd\ {\isacharparenleft}rev\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\ {\isacharequal}\ last\ {\isacharbrackleft}{\isacharbrackright}%
 | 
| 
 | 
   104  | 
\end{isabelle}
 | 
| 
 | 
   105  | 
And \isa{auto} completes the proof.
 | 
| 
 | 
   106  | 
  | 
| 
 | 
   107  | 
If there are multiple premises $A@1$, \dots, $A@n$ containing the
  | 
| 
 | 
   108  | 
induction variable, you should turn the conclusion $C$ into
  | 
| 
 | 
   109  | 
\[ A@1 \longrightarrow \cdots A@n \longrightarrow C. \]
  | 
| 
 | 
   110  | 
Additionally, you may also have to universally quantify some other variables,
  | 
| 
 | 
   111  | 
which can yield a fairly complex conclusion.  However, \isa{rule{\isacharunderscore}format} 
 | 
| 
 | 
   112  | 
can remove any number of occurrences of \isa{{\isasymforall}} and
 | 
| 
 | 
   113  | 
\isa{{\isasymlongrightarrow}}.
 | 
| 
 | 
   114  | 
  | 
| 
 | 
   115  | 
\index{induction!on a term}%
 | 
| 
 | 
   116  | 
A second reason why your proposition may not be amenable to induction is that
  | 
| 
 | 
   117  | 
you want to induct on a complex term, rather than a variable. In
  | 
| 
 | 
   118  | 
general, induction on a term~$t$ requires rephrasing the conclusion~$C$
  | 
| 
 | 
   119  | 
as
  | 
| 
 | 
   120  | 
\begin{equation}\label{eqn:ind-over-term}
 | 
| 
 | 
   121  | 
\forall y@1 \dots y@n.~ x = t \longrightarrow C.
  | 
| 
 | 
   122  | 
\end{equation}
 | 
| 
 | 
   123  | 
where $y@1 \dots y@n$ are the free variables in $t$ and $x$ is a new variable.
  | 
| 
 | 
   124  | 
Now you can perform induction on~$x$. An example appears in
  | 
| 
 | 
   125  | 
\S\ref{sec:complete-ind} below.
 | 
| 
 | 
   126  | 
  | 
| 
 | 
   127  | 
The very same problem may occur in connection with rule induction. Remember
  | 
| 
 | 
   128  | 
that it requires a premise of the form $(x@1,\dots,x@k) \in R$, where $R$ is
  | 
| 
 | 
   129  | 
some inductively defined set and the $x@i$ are variables.  If instead we have
  | 
| 
 | 
   130  | 
a premise $t \in R$, where $t$ is not just an $n$-tuple of variables, we
  | 
| 
 | 
   131  | 
replace it with $(x@1,\dots,x@k) \in R$, and rephrase the conclusion $C$ as
  | 
| 
 | 
   132  | 
\[ \forall y@1 \dots y@n.~ (x@1,\dots,x@k) = t \longrightarrow C. \]
  | 
| 
 | 
   133  | 
For an example see \S\ref{sec:CTL-revisited} below.
 | 
| 
 | 
   134  | 
  | 
| 
 | 
   135  | 
Of course, all premises that share free variables with $t$ need to be pulled into
  | 
| 
 | 
   136  | 
the conclusion as well, under the \isa{{\isasymforall}}, again using \isa{{\isasymlongrightarrow}} as shown above.
 | 
| 
 | 
   137  | 
  | 
| 
 | 
   138  | 
Readers who are puzzled by the form of statement
  | 
| 
 | 
   139  | 
(\ref{eqn:ind-over-term}) above should remember that the
 | 
| 
 | 
   140  | 
transformation is only performed to permit induction. Once induction
  | 
| 
 | 
   141  | 
has been applied, the statement can be transformed back into something quite
  | 
| 
 | 
   142  | 
intuitive. For example, applying wellfounded induction on $x$ (w.r.t.\
  | 
| 
 | 
   143  | 
$\prec$) to (\ref{eqn:ind-over-term}) and transforming the result a
 | 
| 
 | 
   144  | 
little leads to the goal
  | 
| 
 | 
   145  | 
\[ \bigwedge\overline{y}.\ 
 | 
| 
 | 
   146  | 
   \forall \overline{z}.\ t\,\overline{z} \prec t\,\overline{y}\ \longrightarrow\ C\,\overline{z}
 | 
| 
 | 
   147  | 
    \ \Longrightarrow\ C\,\overline{y} \]
 | 
| 
 | 
   148  | 
where $\overline{y}$ stands for $y@1 \dots y@n$ and the dependence of $t$ and
 | 
| 
 | 
   149  | 
$C$ on the free variables of $t$ has been made explicit.
  | 
| 
 | 
   150  | 
Unfortunately, this induction schema cannot be expressed as a
  | 
| 
 | 
   151  | 
single theorem because it depends on the number of free variables in $t$ ---
  | 
| 
 | 
   152  | 
the notation $\overline{y}$ is merely an informal device.%
 | 
| 
 | 
   153  | 
\end{isamarkuptxt}%
 | 
| 
17175
 | 
   154  | 
\isamarkuptrue%
  | 
| 
17056
 | 
   155  | 
%
  | 
| 
 | 
   156  | 
\endisatagproof
  | 
| 
 | 
   157  | 
{\isafoldproof}%
 | 
| 
 | 
   158  | 
%
  | 
| 
 | 
   159  | 
\isadelimproof
  | 
| 
 | 
   160  | 
%
  | 
| 
 | 
   161  | 
\endisadelimproof
  | 
| 
9670
 | 
   162  | 
%
  | 
| 
10878
 | 
   163  | 
\isamarkupsubsection{Beyond Structural and Recursion Induction%
 | 
| 
10397
 | 
   164  | 
}
  | 
| 
11866
 | 
   165  | 
\isamarkuptrue%
  | 
| 
9670
 | 
   166  | 
%
  | 
| 
 | 
   167  | 
\begin{isamarkuptext}%
 | 
| 
10217
 | 
   168  | 
\label{sec:complete-ind}
 | 
| 
10878
 | 
   169  | 
So far, inductive proofs were by structural induction for
  | 
| 
9670
 | 
   170  | 
primitive recursive functions and recursion induction for total recursive
  | 
| 
 | 
   171  | 
functions. But sometimes structural induction is awkward and there is no
  | 
| 
10878
 | 
   172  | 
recursive function that could furnish a more appropriate
  | 
| 
 | 
   173  | 
induction schema. In such cases a general-purpose induction schema can
  | 
| 
9670
 | 
   174  | 
be helpful. We show how to apply such induction schemas by an example.
  | 
| 
 | 
   175  | 
  | 
| 
 | 
   176  | 
Structural induction on \isa{nat} is
 | 
| 
11494
 | 
   177  | 
usually known as mathematical induction. There is also \textbf{complete}
 | 
| 
 | 
   178  | 
\index{induction!complete}%
 | 
| 
 | 
   179  | 
induction, where you prove $P(n)$ under the assumption that $P(m)$
  | 
| 
 | 
   180  | 
holds for all $m<n$. In Isabelle, this is the theorem \tdx{nat_less_induct}:
 | 
| 
9670
 | 
   181  | 
\begin{isabelle}%
 | 
| 
19654
 | 
   182  | 
\ \ \ \ \ {\isacharparenleft}{\isasymAnd}n{\isachardot}\ {\isasymforall}m{\isacharless}n{\isachardot}\ P\ m\ {\isasymLongrightarrow}\ P\ n{\isacharparenright}\ {\isasymLongrightarrow}\ P\ n%
 | 
| 
9924
 | 
   183  | 
\end{isabelle}
 | 
| 
11494
 | 
   184  | 
As an application, we prove a property of the following
  | 
| 
11278
 | 
   185  | 
function:%
  | 
| 
9670
 | 
   186  | 
\end{isamarkuptext}%
 | 
| 
17175
 | 
   187  | 
\isamarkuptrue%
  | 
| 
 | 
   188  | 
\isacommand{consts}\isamarkupfalse%
 | 
| 
 | 
   189  | 
\ f\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\isanewline
 | 
| 
 | 
   190  | 
\isacommand{axioms}\isamarkupfalse%
 | 
| 
 | 
   191  | 
\ f{\isacharunderscore}ax{\isacharcolon}\ {\isachardoublequoteopen}f{\isacharparenleft}f{\isacharparenleft}n{\isacharparenright}{\isacharparenright}\ {\isacharless}\ f{\isacharparenleft}Suc{\isacharparenleft}n{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}%
 | 
| 
9670
 | 
   192  | 
\begin{isamarkuptext}%
 | 
| 
11256
 | 
   193  | 
\begin{warn}
 | 
| 
 | 
   194  | 
We discourage the use of axioms because of the danger of
  | 
| 
 | 
   195  | 
inconsistencies.  Axiom \isa{f{\isacharunderscore}ax} does
 | 
| 
 | 
   196  | 
not introduce an inconsistency because, for example, the identity function
  | 
| 
 | 
   197  | 
satisfies it.  Axioms can be useful in exploratory developments, say when 
  | 
| 
 | 
   198  | 
you assume some well-known theorems so that you can quickly demonstrate some
  | 
| 
 | 
   199  | 
point about methodology.  If your example turns into a substantial proof
  | 
| 
 | 
   200  | 
development, you should replace axioms by theorems.
  | 
| 
 | 
   201  | 
\end{warn}\noindent
 | 
| 
10878
 | 
   202  | 
The axiom for \isa{f} implies \isa{n\ {\isasymle}\ f\ n}, which can
 | 
| 
11196
 | 
   203  | 
be proved by induction on \mbox{\isa{f\ n}}. Following the recipe outlined
 | 
| 
9670
 | 
   204  | 
above, we have to phrase the proposition as follows to allow induction:%
  | 
| 
 | 
   205  | 
\end{isamarkuptext}%
 | 
| 
17175
 | 
   206  | 
\isamarkuptrue%
  | 
| 
 | 
   207  | 
\isacommand{lemma}\isamarkupfalse%
 | 
| 
 | 
   208  | 
\ f{\isacharunderscore}incr{\isacharunderscore}lem{\isacharcolon}\ {\isachardoublequoteopen}{\isasymforall}i{\isachardot}\ k\ {\isacharequal}\ f\ i\ {\isasymlongrightarrow}\ i\ {\isasymle}\ f\ i{\isachardoublequoteclose}%
 | 
| 
17056
 | 
   209  | 
\isadelimproof
  | 
| 
 | 
   210  | 
%
  | 
| 
 | 
   211  | 
\endisadelimproof
  | 
| 
 | 
   212  | 
%
  | 
| 
 | 
   213  | 
\isatagproof
  | 
| 
16069
 | 
   214  | 
%
  | 
| 
 | 
   215  | 
\begin{isamarkuptxt}%
 | 
| 
 | 
   216  | 
\noindent
  | 
| 
 | 
   217  | 
To perform induction on \isa{k} using \isa{nat{\isacharunderscore}less{\isacharunderscore}induct}, we use
 | 
| 
 | 
   218  | 
the same general induction method as for recursion induction (see
  | 
| 
25258
 | 
   219  | 
\S\ref{sec:fun-induction}):%
 | 
| 
16069
 | 
   220  | 
\end{isamarkuptxt}%
 | 
| 
17175
 | 
   221  | 
\isamarkuptrue%
  | 
| 
 | 
   222  | 
\isacommand{apply}\isamarkupfalse%
 | 
| 
 | 
   223  | 
{\isacharparenleft}induct{\isacharunderscore}tac\ k\ rule{\isacharcolon}\ nat{\isacharunderscore}less{\isacharunderscore}induct{\isacharparenright}%
 | 
| 
16069
 | 
   224  | 
\begin{isamarkuptxt}%
 | 
| 
 | 
   225  | 
\noindent
  | 
| 
 | 
   226  | 
We get the following proof state:
  | 
| 
 | 
   227  | 
\begin{isabelle}%
 | 
| 
19654
 | 
   228  | 
\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}n{\isachardot}\ {\isasymforall}m{\isacharless}n{\isachardot}\ {\isasymforall}i{\isachardot}\ m\ {\isacharequal}\ f\ i\ {\isasymlongrightarrow}\ i\ {\isasymle}\ f\ i\ {\isasymLongrightarrow}\ {\isasymforall}i{\isachardot}\ n\ {\isacharequal}\ f\ i\ {\isasymlongrightarrow}\ i\ {\isasymle}\ f\ i%
 | 
| 
16069
 | 
   229  | 
\end{isabelle}
 | 
| 
 | 
   230  | 
After stripping the \isa{{\isasymforall}i}, the proof continues with a case
 | 
| 
 | 
   231  | 
distinction on \isa{i}. The case \isa{i\ {\isacharequal}\ {\isadigit{0}}} is trivial and we focus on
 | 
| 
 | 
   232  | 
the other case:%
  | 
| 
 | 
   233  | 
\end{isamarkuptxt}%
 | 
| 
17175
 | 
   234  | 
\isamarkuptrue%
  | 
| 
 | 
   235  | 
\isacommand{apply}\isamarkupfalse%
 | 
| 
 | 
   236  | 
{\isacharparenleft}rule\ allI{\isacharparenright}\isanewline
 | 
| 
 | 
   237  | 
\isacommand{apply}\isamarkupfalse%
 | 
| 
 | 
   238  | 
{\isacharparenleft}case{\isacharunderscore}tac\ i{\isacharparenright}\isanewline
 | 
| 
 | 
   239  | 
\ \isacommand{apply}\isamarkupfalse%
 | 
| 
 | 
   240  | 
{\isacharparenleft}simp{\isacharparenright}%
 | 
| 
16069
 | 
   241  | 
\begin{isamarkuptxt}%
 | 
| 
 | 
   242  | 
\begin{isabelle}%
 | 
| 
 | 
   243  | 
\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}n\ i\ nat{\isachardot}\isanewline
 | 
| 
19654
 | 
   244  | 
\isaindent{\ {\isadigit{1}}{\isachardot}\ \ \ \ }{\isasymlbrakk}{\isasymforall}m{\isacharless}n{\isachardot}\ {\isasymforall}i{\isachardot}\ m\ {\isacharequal}\ f\ i\ {\isasymlongrightarrow}\ i\ {\isasymle}\ f\ i{\isacharsemicolon}\ i\ {\isacharequal}\ Suc\ nat{\isasymrbrakk}\ {\isasymLongrightarrow}\ n\ {\isacharequal}\ f\ i\ {\isasymlongrightarrow}\ i\ {\isasymle}\ f\ i%
 | 
| 
16069
 | 
   245  | 
\end{isabelle}%
 | 
| 
 | 
   246  | 
\end{isamarkuptxt}%
 | 
| 
17175
 | 
   247  | 
\isamarkuptrue%
  | 
| 
 | 
   248  | 
\isacommand{by}\isamarkupfalse%
 | 
| 
 | 
   249  | 
{\isacharparenleft}blast\ intro{\isacharbang}{\isacharcolon}\ f{\isacharunderscore}ax\ Suc{\isacharunderscore}leI\ intro{\isacharcolon}\ le{\isacharunderscore}less{\isacharunderscore}trans{\isacharparenright}%
 | 
| 
17056
 | 
   250  | 
\endisatagproof
  | 
| 
 | 
   251  | 
{\isafoldproof}%
 | 
| 
 | 
   252  | 
%
  | 
| 
 | 
   253  | 
\isadelimproof
  | 
| 
 | 
   254  | 
%
  | 
| 
 | 
   255  | 
\endisadelimproof
  | 
| 
11866
 | 
   256  | 
%
  | 
| 
9670
 | 
   257  | 
\begin{isamarkuptext}%
 | 
| 
 | 
   258  | 
\noindent
  | 
| 
11196
 | 
   259  | 
If you find the last step puzzling, here are the two lemmas it employs:
  | 
| 
10878
 | 
   260  | 
\begin{isabelle}
 | 
| 
 | 
   261  | 
\isa{m\ {\isacharless}\ n\ {\isasymLongrightarrow}\ Suc\ m\ {\isasymle}\ n}
 | 
| 
 | 
   262  | 
\rulename{Suc_leI}\isanewline
 | 
| 
25056
 | 
   263  | 
\isa{{\isasymlbrakk}x\ {\isasymle}\ y{\isacharsemicolon}\ y\ {\isacharless}\ z{\isasymrbrakk}\ {\isasymLongrightarrow}\ x\ {\isacharless}\ z}
 | 
| 
10878
 | 
   264  | 
\rulename{le_less_trans}
 | 
| 
 | 
   265  | 
\end{isabelle}
 | 
| 
 | 
   266  | 
%
  | 
| 
9670
 | 
   267  | 
The proof goes like this (writing \isa{j} instead of \isa{nat}).
 | 
| 
9792
 | 
   268  | 
Since \isa{i\ {\isacharequal}\ Suc\ j} it suffices to show
 | 
| 
10878
 | 
   269  | 
\hbox{\isa{j\ {\isacharless}\ f\ {\isacharparenleft}Suc\ j{\isacharparenright}}},
 | 
| 
 | 
   270  | 
by \isa{Suc{\isacharunderscore}leI}\@.  This is
 | 
| 
9792
 | 
   271  | 
proved as follows. From \isa{f{\isacharunderscore}ax} we have \isa{f\ {\isacharparenleft}f\ j{\isacharparenright}\ {\isacharless}\ f\ {\isacharparenleft}Suc\ j{\isacharparenright}}
 | 
| 
10878
 | 
   272  | 
(1) which implies \isa{f\ j\ {\isasymle}\ f\ {\isacharparenleft}f\ j{\isacharparenright}} by the induction hypothesis.
 | 
| 
 | 
   273  | 
Using (1) once more we obtain \isa{f\ j\ {\isacharless}\ f\ {\isacharparenleft}Suc\ j{\isacharparenright}} (2) by the transitivity
 | 
| 
 | 
   274  | 
rule \isa{le{\isacharunderscore}less{\isacharunderscore}trans}.
 | 
| 
9792
 | 
   275  | 
Using the induction hypothesis once more we obtain \isa{j\ {\isasymle}\ f\ j}
 | 
| 
 | 
   276  | 
which, together with (2) yields \isa{j\ {\isacharless}\ f\ {\isacharparenleft}Suc\ j{\isacharparenright}} (again by
 | 
| 
 | 
   277  | 
\isa{le{\isacharunderscore}less{\isacharunderscore}trans}).
 | 
| 
9670
 | 
   278  | 
  | 
| 
11494
 | 
   279  | 
This last step shows both the power and the danger of automatic proofs.  They
  | 
| 
 | 
   280  | 
will usually not tell you how the proof goes, because it can be hard to
  | 
| 
 | 
   281  | 
translate the internal proof into a human-readable format.  Automatic
  | 
| 
 | 
   282  | 
proofs are easy to write but hard to read and understand.
  | 
| 
9670
 | 
   283  | 
  | 
| 
11494
 | 
   284  | 
The desired result, \isa{i\ {\isasymle}\ f\ i}, follows from \isa{f{\isacharunderscore}incr{\isacharunderscore}lem}:%
 | 
| 
9670
 | 
   285  | 
\end{isamarkuptext}%
 | 
| 
17175
 | 
   286  | 
\isamarkuptrue%
  | 
| 
 | 
   287  | 
\isacommand{lemmas}\isamarkupfalse%
 | 
| 
 | 
   288  | 
\ f{\isacharunderscore}incr\ {\isacharequal}\ f{\isacharunderscore}incr{\isacharunderscore}lem{\isacharbrackleft}rule{\isacharunderscore}format{\isacharcomma}\ OF\ refl{\isacharbrackright}%
 | 
| 
9670
 | 
   289  | 
\begin{isamarkuptext}%
 | 
| 
9698
 | 
   290  | 
\noindent
  | 
| 
10878
 | 
   291  | 
The final \isa{refl} gets rid of the premise \isa{{\isacharquery}k\ {\isacharequal}\ f\ {\isacharquery}i}. 
 | 
| 
 | 
   292  | 
We could have included this derivation in the original statement of the lemma:%
  | 
| 
9670
 | 
   293  | 
\end{isamarkuptext}%
 | 
| 
17175
 | 
   294  | 
\isamarkuptrue%
  | 
| 
 | 
   295  | 
\isacommand{lemma}\isamarkupfalse%
 | 
| 
 | 
   296  | 
\ f{\isacharunderscore}incr{\isacharbrackleft}rule{\isacharunderscore}format{\isacharcomma}\ OF\ refl{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymforall}i{\isachardot}\ k\ {\isacharequal}\ f\ i\ {\isasymlongrightarrow}\ i\ {\isasymle}\ f\ i{\isachardoublequoteclose}%
 | 
| 
17056
 | 
   297  | 
\isadelimproof
  | 
| 
 | 
   298  | 
%
  | 
| 
 | 
   299  | 
\endisadelimproof
  | 
| 
 | 
   300  | 
%
  | 
| 
 | 
   301  | 
\isatagproof
  | 
| 
 | 
   302  | 
%
  | 
| 
 | 
   303  | 
\endisatagproof
  | 
| 
 | 
   304  | 
{\isafoldproof}%
 | 
| 
 | 
   305  | 
%
  | 
| 
 | 
   306  | 
\isadelimproof
  | 
| 
 | 
   307  | 
%
  | 
| 
 | 
   308  | 
\endisadelimproof
  | 
| 
11866
 | 
   309  | 
%
  | 
| 
9670
 | 
   310  | 
\begin{isamarkuptext}%
 | 
| 
11256
 | 
   311  | 
\begin{exercise}
 | 
| 
 | 
   312  | 
From the axiom and lemma for \isa{f}, show that \isa{f} is the
 | 
| 
 | 
   313  | 
identity function.
  | 
| 
 | 
   314  | 
\end{exercise}
 | 
| 
9670
 | 
   315  | 
  | 
| 
11428
 | 
   316  | 
Method \methdx{induct_tac} can be applied with any rule $r$
 | 
| 
9792
 | 
   317  | 
whose conclusion is of the form ${?}P~?x@1 \dots ?x@n$, in which case the
 | 
| 
9670
 | 
   318  | 
format is
  | 
| 
9792
 | 
   319  | 
\begin{quote}
 | 
| 
 | 
   320  | 
\isacommand{apply}\isa{{\isacharparenleft}induct{\isacharunderscore}tac} $y@1 \dots y@n$ \isa{rule{\isacharcolon}} $r$\isa{{\isacharparenright}}
 | 
| 
11428
 | 
   321  | 
\end{quote}
 | 
| 
27319
 | 
   322  | 
where $y@1, \dots, y@n$ are variables in the conclusion of the first subgoal.
  | 
| 
10878
 | 
   323  | 
  | 
| 
11256
 | 
   324  | 
A further useful induction rule is \isa{length{\isacharunderscore}induct},
 | 
| 
 | 
   325  | 
induction on the length of a list\indexbold{*length_induct}
 | 
| 
 | 
   326  | 
\begin{isabelle}%
 | 
| 
 | 
   327  | 
\ \ \ \ \ {\isacharparenleft}{\isasymAnd}xs{\isachardot}\ {\isasymforall}ys{\isachardot}\ length\ ys\ {\isacharless}\ length\ xs\ {\isasymlongrightarrow}\ P\ ys\ {\isasymLongrightarrow}\ P\ xs{\isacharparenright}\ {\isasymLongrightarrow}\ P\ xs%
 | 
| 
 | 
   328  | 
\end{isabelle}
 | 
| 
 | 
   329  | 
which is a special case of \isa{measure{\isacharunderscore}induct}
 | 
| 
 | 
   330  | 
\begin{isabelle}%
 | 
| 
 | 
   331  | 
\ \ \ \ \ {\isacharparenleft}{\isasymAnd}x{\isachardot}\ {\isasymforall}y{\isachardot}\ f\ y\ {\isacharless}\ f\ x\ {\isasymlongrightarrow}\ P\ y\ {\isasymLongrightarrow}\ P\ x{\isacharparenright}\ {\isasymLongrightarrow}\ P\ a%
 | 
| 
 | 
   332  | 
\end{isabelle}
 | 
| 
 | 
   333  | 
where \isa{f} may be any function into type \isa{nat}.%
 | 
| 
9698
 | 
   334  | 
\end{isamarkuptext}%
 | 
| 
11866
 | 
   335  | 
\isamarkuptrue%
  | 
| 
9698
 | 
   336  | 
%
  | 
| 
10878
 | 
   337  | 
\isamarkupsubsection{Derivation of New Induction Schemas%
 | 
| 
10397
 | 
   338  | 
}
  | 
| 
11866
 | 
   339  | 
\isamarkuptrue%
  | 
| 
9698
 | 
   340  | 
%
  | 
| 
 | 
   341  | 
\begin{isamarkuptext}%
 | 
| 
 | 
   342  | 
\label{sec:derive-ind}
 | 
| 
11494
 | 
   343  | 
\index{induction!deriving new schemas}%
 | 
| 
9698
 | 
   344  | 
Induction schemas are ordinary theorems and you can derive new ones
  | 
| 
11494
 | 
   345  | 
whenever you wish.  This section shows you how, using the example
  | 
| 
9924
 | 
   346  | 
of \isa{nat{\isacharunderscore}less{\isacharunderscore}induct}. Assume we only have structural induction
 | 
| 
11494
 | 
   347  | 
available for \isa{nat} and want to derive complete induction.  We
 | 
| 
 | 
   348  | 
must generalize the statement as shown:%
  | 
| 
9698
 | 
   349  | 
\end{isamarkuptext}%
 | 
| 
17175
 | 
   350  | 
\isamarkuptrue%
  | 
| 
 | 
   351  | 
\isacommand{lemma}\isamarkupfalse%
 | 
| 
 | 
   352  | 
\ induct{\isacharunderscore}lem{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}{\isasymAnd}n{\isacharcolon}{\isacharcolon}nat{\isachardot}\ {\isasymforall}m{\isacharless}n{\isachardot}\ P\ m\ {\isasymLongrightarrow}\ P\ n{\isacharparenright}\ {\isasymLongrightarrow}\ {\isasymforall}m{\isacharless}n{\isachardot}\ P\ m{\isachardoublequoteclose}\isanewline
 | 
| 
17056
 | 
   353  | 
%
  | 
| 
 | 
   354  | 
\isadelimproof
  | 
| 
 | 
   355  | 
%
  | 
| 
 | 
   356  | 
\endisadelimproof
  | 
| 
 | 
   357  | 
%
  | 
| 
 | 
   358  | 
\isatagproof
  | 
| 
17175
 | 
   359  | 
\isacommand{apply}\isamarkupfalse%
 | 
| 
 | 
   360  | 
{\isacharparenleft}induct{\isacharunderscore}tac\ n{\isacharparenright}%
 | 
| 
16069
 | 
   361  | 
\begin{isamarkuptxt}%
 | 
| 
 | 
   362  | 
\noindent
  | 
| 
 | 
   363  | 
The base case is vacuously true. For the induction step (\isa{m\ {\isacharless}\ Suc\ n}) we distinguish two cases: case \isa{m\ {\isacharless}\ n} is true by induction
 | 
| 
 | 
   364  | 
hypothesis and case \isa{m\ {\isacharequal}\ n} follows from the assumption, again using
 | 
| 
 | 
   365  | 
the induction hypothesis:%
  | 
| 
 | 
   366  | 
\end{isamarkuptxt}%
 | 
| 
17175
 | 
   367  | 
\isamarkuptrue%
  | 
| 
 | 
   368  | 
\ \isacommand{apply}\isamarkupfalse%
 | 
| 
 | 
   369  | 
{\isacharparenleft}blast{\isacharparenright}\isanewline
 | 
| 
 | 
   370  | 
\isacommand{by}\isamarkupfalse%
 | 
| 
 | 
   371  | 
{\isacharparenleft}blast\ elim{\isacharcolon}\ less{\isacharunderscore}SucE{\isacharparenright}%
 | 
| 
17056
 | 
   372  | 
\endisatagproof
  | 
| 
 | 
   373  | 
{\isafoldproof}%
 | 
| 
 | 
   374  | 
%
  | 
| 
 | 
   375  | 
\isadelimproof
  | 
| 
 | 
   376  | 
%
  | 
| 
 | 
   377  | 
\endisadelimproof
  | 
| 
11866
 | 
   378  | 
%
  | 
| 
9698
 | 
   379  | 
\begin{isamarkuptext}%
 | 
| 
 | 
   380  | 
\noindent
  | 
| 
11196
 | 
   381  | 
The elimination rule \isa{less{\isacharunderscore}SucE} expresses the case distinction:
 | 
| 
9698
 | 
   382  | 
\begin{isabelle}%
 | 
| 
10696
 | 
   383  | 
\ \ \ \ \ {\isasymlbrakk}m\ {\isacharless}\ Suc\ n{\isacharsemicolon}\ m\ {\isacharless}\ n\ {\isasymLongrightarrow}\ P{\isacharsemicolon}\ m\ {\isacharequal}\ n\ {\isasymLongrightarrow}\ P{\isasymrbrakk}\ {\isasymLongrightarrow}\ P%
 | 
| 
9924
 | 
   384  | 
\end{isabelle}
 | 
| 
9698
 | 
   385  | 
  | 
| 
 | 
   386  | 
Now it is straightforward to derive the original version of
  | 
| 
11256
 | 
   387  | 
\isa{nat{\isacharunderscore}less{\isacharunderscore}induct} by manipulating the conclusion of the above
 | 
| 
 | 
   388  | 
lemma: instantiate \isa{n} by \isa{Suc\ n} and \isa{m} by \isa{n}
 | 
| 
 | 
   389  | 
and remove the trivial condition \isa{n\ {\isacharless}\ Suc\ n}. Fortunately, this
 | 
| 
9698
 | 
   390  | 
happens automatically when we add the lemma as a new premise to the
  | 
| 
 | 
   391  | 
desired goal:%
  | 
| 
 | 
   392  | 
\end{isamarkuptext}%
 | 
| 
17175
 | 
   393  | 
\isamarkuptrue%
  | 
| 
 | 
   394  | 
\isacommand{theorem}\isamarkupfalse%
 | 
| 
 | 
   395  | 
\ nat{\isacharunderscore}less{\isacharunderscore}induct{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}{\isasymAnd}n{\isacharcolon}{\isacharcolon}nat{\isachardot}\ {\isasymforall}m{\isacharless}n{\isachardot}\ P\ m\ {\isasymLongrightarrow}\ P\ n{\isacharparenright}\ {\isasymLongrightarrow}\ P\ n{\isachardoublequoteclose}\isanewline
 | 
| 
17056
 | 
   396  | 
%
  | 
| 
 | 
   397  | 
\isadelimproof
  | 
| 
 | 
   398  | 
%
  | 
| 
 | 
   399  | 
\endisadelimproof
  | 
| 
 | 
   400  | 
%
  | 
| 
 | 
   401  | 
\isatagproof
  | 
| 
17175
 | 
   402  | 
\isacommand{by}\isamarkupfalse%
 | 
| 
 | 
   403  | 
{\isacharparenleft}insert\ induct{\isacharunderscore}lem{\isacharcomma}\ blast{\isacharparenright}%
 | 
| 
17056
 | 
   404  | 
\endisatagproof
  | 
| 
 | 
   405  | 
{\isafoldproof}%
 | 
| 
 | 
   406  | 
%
  | 
| 
 | 
   407  | 
\isadelimproof
  | 
| 
 | 
   408  | 
%
  | 
| 
 | 
   409  | 
\endisadelimproof
  | 
| 
11866
 | 
   410  | 
%
  | 
| 
9698
 | 
   411  | 
\begin{isamarkuptext}%
 | 
| 
11494
 | 
   412  | 
HOL already provides the mother of
  | 
| 
10396
 | 
   413  | 
all inductions, well-founded induction (see \S\ref{sec:Well-founded}).  For
 | 
| 
10878
 | 
   414  | 
example theorem \isa{nat{\isacharunderscore}less{\isacharunderscore}induct} is
 | 
| 
10396
 | 
   415  | 
a special case of \isa{wf{\isacharunderscore}induct} where \isa{r} is \isa{{\isacharless}} on
 | 
| 
10878
 | 
   416  | 
\isa{nat}. The details can be found in theory \isa{Wellfounded_Recursion}.%
 | 
| 
9670
 | 
   417  | 
\end{isamarkuptext}%
 | 
| 
17175
 | 
   418  | 
\isamarkuptrue%
  | 
| 
17056
 | 
   419  | 
%
  | 
| 
 | 
   420  | 
\isadelimtheory
  | 
| 
 | 
   421  | 
%
  | 
| 
 | 
   422  | 
\endisadelimtheory
  | 
| 
 | 
   423  | 
%
  | 
| 
 | 
   424  | 
\isatagtheory
  | 
| 
 | 
   425  | 
%
  | 
| 
 | 
   426  | 
\endisatagtheory
  | 
| 
 | 
   427  | 
{\isafoldtheory}%
 | 
| 
 | 
   428  | 
%
  | 
| 
 | 
   429  | 
\isadelimtheory
  | 
| 
 | 
   430  | 
%
  | 
| 
 | 
   431  | 
\endisadelimtheory
  | 
| 
9722
 | 
   432  | 
\end{isabellebody}%
 | 
| 
9670
 | 
   433  | 
%%% Local Variables:
  | 
| 
 | 
   434  | 
%%% mode: latex
  | 
| 
 | 
   435  | 
%%% TeX-master: "root"
  | 
| 
 | 
   436  | 
%%% End:
  |