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