doc-src/TutorialI/document/AdvancedInd.tex
changeset 48966 6e15de7dd871
parent 48965 1fead823c7c6
child 48967 389e44f9e47a
equal deleted inserted replaced
48965:1fead823c7c6 48966:6e15de7dd871
     1 %
       
     2 \begin{isabellebody}%
       
     3 \def\isabellecontext{AdvancedInd}%
       
     4 %
       
     5 \isadelimtheory
       
     6 %
       
     7 \endisadelimtheory
       
     8 %
       
     9 \isatagtheory
       
    10 %
       
    11 \endisatagtheory
       
    12 {\isafoldtheory}%
       
    13 %
       
    14 \isadelimtheory
       
    15 %
       
    16 \endisadelimtheory
       
    17 %
       
    18 \begin{isamarkuptext}%
       
    19 \noindent
       
    20 Now that we have learned about rules and logic, we take another look at the
       
    21 finer points of induction.  We consider two questions: what to do if the
       
    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}).%
       
    26 \end{isamarkuptext}%
       
    27 \isamarkuptrue%
       
    28 %
       
    29 \isamarkupsubsection{Massaging the Proposition%
       
    30 }
       
    31 \isamarkuptrue%
       
    32 %
       
    33 \begin{isamarkuptext}%
       
    34 \label{sec:ind-var-in-prems}
       
    35 Often we have assumed that the theorem to be proved is already in a form
       
    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:%
       
    40 \end{isamarkuptext}%
       
    41 \isamarkuptrue%
       
    42 \isacommand{lemma}\isamarkupfalse%
       
    43 \ {\isaliteral{22}{\isachardoublequoteopen}}xs\ {\isaliteral{5C3C6E6F7465713E}{\isasymnoteq}}\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ hd{\isaliteral{28}{\isacharparenleft}}rev\ xs{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ last\ xs{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
       
    44 %
       
    45 \isadelimproof
       
    46 %
       
    47 \endisadelimproof
       
    48 %
       
    49 \isatagproof
       
    50 \isacommand{apply}\isamarkupfalse%
       
    51 {\isaliteral{28}{\isacharparenleft}}induct{\isaliteral{5F}{\isacharunderscore}}tac\ xs{\isaliteral{29}{\isacharparenright}}%
       
    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}}{\isaliteral{2E}{\isachardot}}\ xs\ {\isaliteral{5C3C6E6F7465713E}{\isasymnoteq}}\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ hd\ {\isaliteral{28}{\isacharparenleft}}rev\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ last\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\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{{\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\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{{\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}}.}
       
    77 \end{quote}
       
    78 Thus we should state the lemma as an ordinary 
       
    79 implication~(\isa{{\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}}), letting
       
    80 \attrdx{rule_format} (\S\ref{sec:forward}) convert the
       
    81 result to the usual \isa{{\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}} form:%
       
    82 \end{isamarkuptxt}%
       
    83 \isamarkuptrue%
       
    84 %
       
    85 \endisatagproof
       
    86 {\isafoldproof}%
       
    87 %
       
    88 \isadelimproof
       
    89 %
       
    90 \endisadelimproof
       
    91 \isacommand{lemma}\isamarkupfalse%
       
    92 \ hd{\isaliteral{5F}{\isacharunderscore}}rev\ {\isaliteral{5B}{\isacharbrackleft}}rule{\isaliteral{5F}{\isacharunderscore}}format{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}xs\ {\isaliteral{5C3C6E6F7465713E}{\isasymnoteq}}\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ hd{\isaliteral{28}{\isacharparenleft}}rev\ xs{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ last\ xs{\isaliteral{22}{\isachardoublequoteclose}}%
       
    93 \isadelimproof
       
    94 %
       
    95 \endisadelimproof
       
    96 %
       
    97 \isatagproof
       
    98 %
       
    99 \begin{isamarkuptxt}%
       
   100 \noindent
       
   101 This time, induction leaves us with a trivial base case:
       
   102 \begin{isabelle}%
       
   103 \ {\isadigit{1}}{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{5C3C6E6F7465713E}{\isasymnoteq}}\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ hd\ {\isaliteral{28}{\isacharparenleft}}rev\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ last\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\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{\isaliteral{5F}{\isacharunderscore}}format} 
       
   112 can remove any number of occurrences of \isa{{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}} and
       
   113 \isa{{\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\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{{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}}, again using \isa{{\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\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}%
       
   154 \isamarkuptrue%
       
   155 %
       
   156 \endisatagproof
       
   157 {\isafoldproof}%
       
   158 %
       
   159 \isadelimproof
       
   160 %
       
   161 \endisadelimproof
       
   162 %
       
   163 \isamarkupsubsection{Beyond Structural and Recursion Induction%
       
   164 }
       
   165 \isamarkuptrue%
       
   166 %
       
   167 \begin{isamarkuptext}%
       
   168 \label{sec:complete-ind}
       
   169 So far, inductive proofs were by structural induction for
       
   170 primitive recursive functions and recursion induction for total recursive
       
   171 functions. But sometimes structural induction is awkward and there is no
       
   172 recursive function that could furnish a more appropriate
       
   173 induction schema. In such cases a general-purpose induction schema can
       
   174 be helpful. We show how to apply such induction schemas by an example.
       
   175 
       
   176 Structural induction on \isa{nat} is
       
   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}:
       
   181 \begin{isabelle}%
       
   182 \ \ \ \ \ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C416E643E}{\isasymAnd}}n{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}m{\isaliteral{3C}{\isacharless}}n{\isaliteral{2E}{\isachardot}}\ P\ m\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ P\ n{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ P\ n%
       
   183 \end{isabelle}
       
   184 As an application, we prove a property of the following
       
   185 function:%
       
   186 \end{isamarkuptext}%
       
   187 \isamarkuptrue%
       
   188 \isacommand{consts}\isamarkupfalse%
       
   189 \ f\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}nat\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ nat{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
       
   190 \isacommand{axioms}\isamarkupfalse%
       
   191 \ f{\isaliteral{5F}{\isacharunderscore}}ax{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}f{\isaliteral{28}{\isacharparenleft}}f{\isaliteral{28}{\isacharparenleft}}n{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3C}{\isacharless}}\ f{\isaliteral{28}{\isacharparenleft}}Suc{\isaliteral{28}{\isacharparenleft}}n{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}%
       
   192 \begin{isamarkuptext}%
       
   193 \begin{warn}
       
   194 We discourage the use of axioms because of the danger of
       
   195 inconsistencies.  Axiom \isa{f{\isaliteral{5F}{\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
       
   202 The axiom for \isa{f} implies \isa{n\ {\isaliteral{5C3C6C653E}{\isasymle}}\ f\ n}, which can
       
   203 be proved by induction on \mbox{\isa{f\ n}}. Following the recipe outlined
       
   204 above, we have to phrase the proposition as follows to allow induction:%
       
   205 \end{isamarkuptext}%
       
   206 \isamarkuptrue%
       
   207 \isacommand{lemma}\isamarkupfalse%
       
   208 \ f{\isaliteral{5F}{\isacharunderscore}}incr{\isaliteral{5F}{\isacharunderscore}}lem{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}i{\isaliteral{2E}{\isachardot}}\ k\ {\isaliteral{3D}{\isacharequal}}\ f\ i\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ i\ {\isaliteral{5C3C6C653E}{\isasymle}}\ f\ i{\isaliteral{22}{\isachardoublequoteclose}}%
       
   209 \isadelimproof
       
   210 %
       
   211 \endisadelimproof
       
   212 %
       
   213 \isatagproof
       
   214 %
       
   215 \begin{isamarkuptxt}%
       
   216 \noindent
       
   217 To perform induction on \isa{k} using \isa{nat{\isaliteral{5F}{\isacharunderscore}}less{\isaliteral{5F}{\isacharunderscore}}induct}, we use
       
   218 the same general induction method as for recursion induction (see
       
   219 \S\ref{sec:fun-induction}):%
       
   220 \end{isamarkuptxt}%
       
   221 \isamarkuptrue%
       
   222 \isacommand{apply}\isamarkupfalse%
       
   223 {\isaliteral{28}{\isacharparenleft}}induct{\isaliteral{5F}{\isacharunderscore}}tac\ k\ rule{\isaliteral{3A}{\isacharcolon}}\ nat{\isaliteral{5F}{\isacharunderscore}}less{\isaliteral{5F}{\isacharunderscore}}induct{\isaliteral{29}{\isacharparenright}}%
       
   224 \begin{isamarkuptxt}%
       
   225 \noindent
       
   226 We get the following proof state:
       
   227 \begin{isabelle}%
       
   228 \ {\isadigit{1}}{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C416E643E}{\isasymAnd}}n{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}m{\isaliteral{3C}{\isacharless}}n{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}i{\isaliteral{2E}{\isachardot}}\ m\ {\isaliteral{3D}{\isacharequal}}\ f\ i\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ i\ {\isaliteral{5C3C6C653E}{\isasymle}}\ f\ i\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}i{\isaliteral{2E}{\isachardot}}\ n\ {\isaliteral{3D}{\isacharequal}}\ f\ i\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ i\ {\isaliteral{5C3C6C653E}{\isasymle}}\ f\ i%
       
   229 \end{isabelle}
       
   230 After stripping the \isa{{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}i}, the proof continues with a case
       
   231 distinction on \isa{i}. The case \isa{i\ {\isaliteral{3D}{\isacharequal}}\ {\isadigit{0}}} is trivial and we focus on
       
   232 the other case:%
       
   233 \end{isamarkuptxt}%
       
   234 \isamarkuptrue%
       
   235 \isacommand{apply}\isamarkupfalse%
       
   236 {\isaliteral{28}{\isacharparenleft}}rule\ allI{\isaliteral{29}{\isacharparenright}}\isanewline
       
   237 \isacommand{apply}\isamarkupfalse%
       
   238 {\isaliteral{28}{\isacharparenleft}}case{\isaliteral{5F}{\isacharunderscore}}tac\ i{\isaliteral{29}{\isacharparenright}}\isanewline
       
   239 \ \isacommand{apply}\isamarkupfalse%
       
   240 {\isaliteral{28}{\isacharparenleft}}simp{\isaliteral{29}{\isacharparenright}}%
       
   241 \begin{isamarkuptxt}%
       
   242 \begin{isabelle}%
       
   243 \ {\isadigit{1}}{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C416E643E}{\isasymAnd}}n\ i\ nat{\isaliteral{2E}{\isachardot}}\isanewline
       
   244 \isaindent{\ {\isadigit{1}}{\isaliteral{2E}{\isachardot}}\ \ \ \ }{\isaliteral{5C3C6C6272616B6B3E}{\isasymlbrakk}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}m{\isaliteral{3C}{\isacharless}}n{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}i{\isaliteral{2E}{\isachardot}}\ m\ {\isaliteral{3D}{\isacharequal}}\ f\ i\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ i\ {\isaliteral{5C3C6C653E}{\isasymle}}\ f\ i{\isaliteral{3B}{\isacharsemicolon}}\ i\ {\isaliteral{3D}{\isacharequal}}\ Suc\ nat{\isaliteral{5C3C726272616B6B3E}{\isasymrbrakk}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ n\ {\isaliteral{3D}{\isacharequal}}\ f\ i\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ i\ {\isaliteral{5C3C6C653E}{\isasymle}}\ f\ i%
       
   245 \end{isabelle}%
       
   246 \end{isamarkuptxt}%
       
   247 \isamarkuptrue%
       
   248 \isacommand{by}\isamarkupfalse%
       
   249 {\isaliteral{28}{\isacharparenleft}}blast\ intro{\isaliteral{21}{\isacharbang}}{\isaliteral{3A}{\isacharcolon}}\ f{\isaliteral{5F}{\isacharunderscore}}ax\ Suc{\isaliteral{5F}{\isacharunderscore}}leI\ intro{\isaliteral{3A}{\isacharcolon}}\ le{\isaliteral{5F}{\isacharunderscore}}less{\isaliteral{5F}{\isacharunderscore}}trans{\isaliteral{29}{\isacharparenright}}%
       
   250 \endisatagproof
       
   251 {\isafoldproof}%
       
   252 %
       
   253 \isadelimproof
       
   254 %
       
   255 \endisadelimproof
       
   256 %
       
   257 \begin{isamarkuptext}%
       
   258 \noindent
       
   259 If you find the last step puzzling, here are the two lemmas it employs:
       
   260 \begin{isabelle}
       
   261 \isa{m\ {\isaliteral{3C}{\isacharless}}\ n\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ Suc\ m\ {\isaliteral{5C3C6C653E}{\isasymle}}\ n}
       
   262 \rulename{Suc_leI}\isanewline
       
   263 \isa{{\isaliteral{5C3C6C6272616B6B3E}{\isasymlbrakk}}x\ {\isaliteral{5C3C6C653E}{\isasymle}}\ y{\isaliteral{3B}{\isacharsemicolon}}\ y\ {\isaliteral{3C}{\isacharless}}\ z{\isaliteral{5C3C726272616B6B3E}{\isasymrbrakk}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ x\ {\isaliteral{3C}{\isacharless}}\ z}
       
   264 \rulename{le_less_trans}
       
   265 \end{isabelle}
       
   266 %
       
   267 The proof goes like this (writing \isa{j} instead of \isa{nat}).
       
   268 Since \isa{i\ {\isaliteral{3D}{\isacharequal}}\ Suc\ j} it suffices to show
       
   269 \hbox{\isa{j\ {\isaliteral{3C}{\isacharless}}\ f\ {\isaliteral{28}{\isacharparenleft}}Suc\ j{\isaliteral{29}{\isacharparenright}}}},
       
   270 by \isa{Suc{\isaliteral{5F}{\isacharunderscore}}leI}\@.  This is
       
   271 proved as follows. From \isa{f{\isaliteral{5F}{\isacharunderscore}}ax} we have \isa{f\ {\isaliteral{28}{\isacharparenleft}}f\ j{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3C}{\isacharless}}\ f\ {\isaliteral{28}{\isacharparenleft}}Suc\ j{\isaliteral{29}{\isacharparenright}}}
       
   272 (1) which implies \isa{f\ j\ {\isaliteral{5C3C6C653E}{\isasymle}}\ f\ {\isaliteral{28}{\isacharparenleft}}f\ j{\isaliteral{29}{\isacharparenright}}} by the induction hypothesis.
       
   273 Using (1) once more we obtain \isa{f\ j\ {\isaliteral{3C}{\isacharless}}\ f\ {\isaliteral{28}{\isacharparenleft}}Suc\ j{\isaliteral{29}{\isacharparenright}}} (2) by the transitivity
       
   274 rule \isa{le{\isaliteral{5F}{\isacharunderscore}}less{\isaliteral{5F}{\isacharunderscore}}trans}.
       
   275 Using the induction hypothesis once more we obtain \isa{j\ {\isaliteral{5C3C6C653E}{\isasymle}}\ f\ j}
       
   276 which, together with (2) yields \isa{j\ {\isaliteral{3C}{\isacharless}}\ f\ {\isaliteral{28}{\isacharparenleft}}Suc\ j{\isaliteral{29}{\isacharparenright}}} (again by
       
   277 \isa{le{\isaliteral{5F}{\isacharunderscore}}less{\isaliteral{5F}{\isacharunderscore}}trans}).
       
   278 
       
   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.
       
   283 
       
   284 The desired result, \isa{i\ {\isaliteral{5C3C6C653E}{\isasymle}}\ f\ i}, follows from \isa{f{\isaliteral{5F}{\isacharunderscore}}incr{\isaliteral{5F}{\isacharunderscore}}lem}:%
       
   285 \end{isamarkuptext}%
       
   286 \isamarkuptrue%
       
   287 \isacommand{lemmas}\isamarkupfalse%
       
   288 \ f{\isaliteral{5F}{\isacharunderscore}}incr\ {\isaliteral{3D}{\isacharequal}}\ f{\isaliteral{5F}{\isacharunderscore}}incr{\isaliteral{5F}{\isacharunderscore}}lem{\isaliteral{5B}{\isacharbrackleft}}rule{\isaliteral{5F}{\isacharunderscore}}format{\isaliteral{2C}{\isacharcomma}}\ OF\ refl{\isaliteral{5D}{\isacharbrackright}}%
       
   289 \begin{isamarkuptext}%
       
   290 \noindent
       
   291 The final \isa{refl} gets rid of the premise \isa{{\isaliteral{3F}{\isacharquery}}k\ {\isaliteral{3D}{\isacharequal}}\ f\ {\isaliteral{3F}{\isacharquery}}i}. 
       
   292 We could have included this derivation in the original statement of the lemma:%
       
   293 \end{isamarkuptext}%
       
   294 \isamarkuptrue%
       
   295 \isacommand{lemma}\isamarkupfalse%
       
   296 \ f{\isaliteral{5F}{\isacharunderscore}}incr{\isaliteral{5B}{\isacharbrackleft}}rule{\isaliteral{5F}{\isacharunderscore}}format{\isaliteral{2C}{\isacharcomma}}\ OF\ refl{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}i{\isaliteral{2E}{\isachardot}}\ k\ {\isaliteral{3D}{\isacharequal}}\ f\ i\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ i\ {\isaliteral{5C3C6C653E}{\isasymle}}\ f\ i{\isaliteral{22}{\isachardoublequoteclose}}%
       
   297 \isadelimproof
       
   298 %
       
   299 \endisadelimproof
       
   300 %
       
   301 \isatagproof
       
   302 %
       
   303 \endisatagproof
       
   304 {\isafoldproof}%
       
   305 %
       
   306 \isadelimproof
       
   307 %
       
   308 \endisadelimproof
       
   309 %
       
   310 \begin{isamarkuptext}%
       
   311 \begin{exercise}
       
   312 From the axiom and lemma for \isa{f}, show that \isa{f} is the
       
   313 identity function.
       
   314 \end{exercise}
       
   315 
       
   316 Method \methdx{induct_tac} can be applied with any rule $r$
       
   317 whose conclusion is of the form ${?}P~?x@1 \dots ?x@n$, in which case the
       
   318 format is
       
   319 \begin{quote}
       
   320 \isacommand{apply}\isa{{\isaliteral{28}{\isacharparenleft}}induct{\isaliteral{5F}{\isacharunderscore}}tac} $y@1 \dots y@n$ \isa{rule{\isaliteral{3A}{\isacharcolon}}} $r$\isa{{\isaliteral{29}{\isacharparenright}}}
       
   321 \end{quote}
       
   322 where $y@1, \dots, y@n$ are variables in the conclusion of the first subgoal.
       
   323 
       
   324 A further useful induction rule is \isa{length{\isaliteral{5F}{\isacharunderscore}}induct},
       
   325 induction on the length of a list\indexbold{*length_induct}
       
   326 \begin{isabelle}%
       
   327 \ \ \ \ \ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C416E643E}{\isasymAnd}}xs{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}ys{\isaliteral{2E}{\isachardot}}\ length\ ys\ {\isaliteral{3C}{\isacharless}}\ length\ xs\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ P\ ys\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ P\ xs{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ P\ xs%
       
   328 \end{isabelle}
       
   329 which is a special case of \isa{measure{\isaliteral{5F}{\isacharunderscore}}induct}
       
   330 \begin{isabelle}%
       
   331 \ \ \ \ \ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C416E643E}{\isasymAnd}}x{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}y{\isaliteral{2E}{\isachardot}}\ f\ y\ {\isaliteral{3C}{\isacharless}}\ f\ x\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ P\ y\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ P\ x{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ P\ a%
       
   332 \end{isabelle}
       
   333 where \isa{f} may be any function into type \isa{nat}.%
       
   334 \end{isamarkuptext}%
       
   335 \isamarkuptrue%
       
   336 %
       
   337 \isamarkupsubsection{Derivation of New Induction Schemas%
       
   338 }
       
   339 \isamarkuptrue%
       
   340 %
       
   341 \begin{isamarkuptext}%
       
   342 \label{sec:derive-ind}
       
   343 \index{induction!deriving new schemas}%
       
   344 Induction schemas are ordinary theorems and you can derive new ones
       
   345 whenever you wish.  This section shows you how, using the example
       
   346 of \isa{nat{\isaliteral{5F}{\isacharunderscore}}less{\isaliteral{5F}{\isacharunderscore}}induct}. Assume we only have structural induction
       
   347 available for \isa{nat} and want to derive complete induction.  We
       
   348 must generalize the statement as shown:%
       
   349 \end{isamarkuptext}%
       
   350 \isamarkuptrue%
       
   351 \isacommand{lemma}\isamarkupfalse%
       
   352 \ induct{\isaliteral{5F}{\isacharunderscore}}lem{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C416E643E}{\isasymAnd}}n{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}nat{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}m{\isaliteral{3C}{\isacharless}}n{\isaliteral{2E}{\isachardot}}\ P\ m\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ P\ n{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}m{\isaliteral{3C}{\isacharless}}n{\isaliteral{2E}{\isachardot}}\ P\ m{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
       
   353 %
       
   354 \isadelimproof
       
   355 %
       
   356 \endisadelimproof
       
   357 %
       
   358 \isatagproof
       
   359 \isacommand{apply}\isamarkupfalse%
       
   360 {\isaliteral{28}{\isacharparenleft}}induct{\isaliteral{5F}{\isacharunderscore}}tac\ n{\isaliteral{29}{\isacharparenright}}%
       
   361 \begin{isamarkuptxt}%
       
   362 \noindent
       
   363 The base case is vacuously true. For the induction step (\isa{m\ {\isaliteral{3C}{\isacharless}}\ Suc\ n}) we distinguish two cases: case \isa{m\ {\isaliteral{3C}{\isacharless}}\ n} is true by induction
       
   364 hypothesis and case \isa{m\ {\isaliteral{3D}{\isacharequal}}\ n} follows from the assumption, again using
       
   365 the induction hypothesis:%
       
   366 \end{isamarkuptxt}%
       
   367 \isamarkuptrue%
       
   368 \ \isacommand{apply}\isamarkupfalse%
       
   369 {\isaliteral{28}{\isacharparenleft}}blast{\isaliteral{29}{\isacharparenright}}\isanewline
       
   370 \isacommand{by}\isamarkupfalse%
       
   371 {\isaliteral{28}{\isacharparenleft}}blast\ elim{\isaliteral{3A}{\isacharcolon}}\ less{\isaliteral{5F}{\isacharunderscore}}SucE{\isaliteral{29}{\isacharparenright}}%
       
   372 \endisatagproof
       
   373 {\isafoldproof}%
       
   374 %
       
   375 \isadelimproof
       
   376 %
       
   377 \endisadelimproof
       
   378 %
       
   379 \begin{isamarkuptext}%
       
   380 \noindent
       
   381 The elimination rule \isa{less{\isaliteral{5F}{\isacharunderscore}}SucE} expresses the case distinction:
       
   382 \begin{isabelle}%
       
   383 \ \ \ \ \ {\isaliteral{5C3C6C6272616B6B3E}{\isasymlbrakk}}m\ {\isaliteral{3C}{\isacharless}}\ Suc\ n{\isaliteral{3B}{\isacharsemicolon}}\ m\ {\isaliteral{3C}{\isacharless}}\ n\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ P{\isaliteral{3B}{\isacharsemicolon}}\ m\ {\isaliteral{3D}{\isacharequal}}\ n\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ P{\isaliteral{5C3C726272616B6B3E}{\isasymrbrakk}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ P%
       
   384 \end{isabelle}
       
   385 
       
   386 Now it is straightforward to derive the original version of
       
   387 \isa{nat{\isaliteral{5F}{\isacharunderscore}}less{\isaliteral{5F}{\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\ {\isaliteral{3C}{\isacharless}}\ Suc\ n}. Fortunately, this
       
   390 happens automatically when we add the lemma as a new premise to the
       
   391 desired goal:%
       
   392 \end{isamarkuptext}%
       
   393 \isamarkuptrue%
       
   394 \isacommand{theorem}\isamarkupfalse%
       
   395 \ nat{\isaliteral{5F}{\isacharunderscore}}less{\isaliteral{5F}{\isacharunderscore}}induct{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C416E643E}{\isasymAnd}}n{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}nat{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}m{\isaliteral{3C}{\isacharless}}n{\isaliteral{2E}{\isachardot}}\ P\ m\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ P\ n{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ P\ n{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
       
   396 %
       
   397 \isadelimproof
       
   398 %
       
   399 \endisadelimproof
       
   400 %
       
   401 \isatagproof
       
   402 \isacommand{by}\isamarkupfalse%
       
   403 {\isaliteral{28}{\isacharparenleft}}insert\ induct{\isaliteral{5F}{\isacharunderscore}}lem{\isaliteral{2C}{\isacharcomma}}\ blast{\isaliteral{29}{\isacharparenright}}%
       
   404 \endisatagproof
       
   405 {\isafoldproof}%
       
   406 %
       
   407 \isadelimproof
       
   408 %
       
   409 \endisadelimproof
       
   410 %
       
   411 \begin{isamarkuptext}%
       
   412 HOL already provides the mother of
       
   413 all inductions, well-founded induction (see \S\ref{sec:Well-founded}).  For
       
   414 example theorem \isa{nat{\isaliteral{5F}{\isacharunderscore}}less{\isaliteral{5F}{\isacharunderscore}}induct} is
       
   415 a special case of \isa{wf{\isaliteral{5F}{\isacharunderscore}}induct} where \isa{r} is \isa{{\isaliteral{3C}{\isacharless}}} on
       
   416 \isa{nat}. The details can be found in theory \isa{Wellfounded_Recursion}.%
       
   417 \end{isamarkuptext}%
       
   418 \isamarkuptrue%
       
   419 %
       
   420 \isadelimtheory
       
   421 %
       
   422 \endisadelimtheory
       
   423 %
       
   424 \isatagtheory
       
   425 %
       
   426 \endisatagtheory
       
   427 {\isafoldtheory}%
       
   428 %
       
   429 \isadelimtheory
       
   430 %
       
   431 \endisadelimtheory
       
   432 \end{isabellebody}%
       
   433 %%% Local Variables:
       
   434 %%% mode: latex
       
   435 %%% TeX-master: "root"
       
   436 %%% End: