doc-src/TutorialI/Misc/AdvancedInd.thy
changeset 9645 20ae97cd2a16
child 9689 751fde5307e4
equal deleted inserted replaced
9644:6b0b6b471855 9645:20ae97cd2a16
       
     1 (*<*)
       
     2 theory AdvancedInd = Main:;
       
     3 (*>*)
       
     4 
       
     5 text{*\noindent
       
     6 Now that we have learned about rules and logic, we take another look at the
       
     7 finer points of induction. The two questions we answer are: what to do if the
       
     8 proposition to be proved is not directly amenable to induction, and how to
       
     9 utilize and even derive new induction schemas.
       
    10 *}
       
    11 
       
    12 subsection{*Massaging the proposition\label{sec:ind-var-in-prems}*}
       
    13 
       
    14 text{*
       
    15 \noindent
       
    16 So far we have assumed that the theorem we want to prove is already in a form
       
    17 that is amenable to induction, but this is not always the case:
       
    18 *}
       
    19 
       
    20 lemma "xs \\<noteq> [] \\<Longrightarrow> hd(rev xs) = last xs";
       
    21 apply(induct_tac xs);
       
    22 
       
    23 txt{*\noindent
       
    24 (where \isa{hd} and \isa{last} return the first and last element of a
       
    25 non-empty list)
       
    26 produces the warning
       
    27 \begin{quote}\tt
       
    28 Induction variable occurs also among premises!
       
    29 \end{quote}
       
    30 and leads to the base case
       
    31 \begin{isabellepar}%
       
    32 \ 1.\ xs\ {\isasymnoteq}\ []\ {\isasymLongrightarrow}\ hd\ (rev\ [])\ =\ last\ []
       
    33 \end{isabellepar}%
       
    34 which, after simplification, becomes
       
    35 \begin{isabellepar}%
       
    36 \ 1.\ xs\ {\isasymnoteq}\ []\ {\isasymLongrightarrow}\ hd\ []\ =\ last\ []
       
    37 \end{isabellepar}%
       
    38 We cannot prove this equality because we do not know what \isa{hd} and
       
    39 \isa{last} return when applied to \isa{[]}.
       
    40 
       
    41 The point is that we have violated the above warning. Because the induction
       
    42 formula is only the conclusion, the occurrence of \isa{xs} in the premises is
       
    43 not modified by induction. Thus the case that should have been trivial
       
    44 becomes unprovable. Fortunately, the solution is easy:
       
    45 \begin{quote}
       
    46 \emph{Pull all occurrences of the induction variable into the conclusion
       
    47 using \isa{\isasymlongrightarrow}.}
       
    48 \end{quote}
       
    49 This means we should prove
       
    50 *}
       
    51 (*<*)oops(*>*)
       
    52 lemma hd_rev: "xs \\<noteq> [] \\<longrightarrow> hd(rev xs) = last xs";
       
    53 (*<*)
       
    54 by(induct_tac xs, auto)
       
    55 (*>*)
       
    56 
       
    57 text{*\noindent
       
    58 This time, induction leaves us with the following base case
       
    59 \begin{isabellepar}%
       
    60 \ 1.\ []\ {\isasymnoteq}\ []\ {\isasymlongrightarrow}\ hd\ (rev\ [])\ =\ last\ []
       
    61 \end{isabellepar}%
       
    62 which is trivial, and \isa{auto} finishes the whole proof.
       
    63 
       
    64 If \isa{hd\_rev} is meant to be simplification rule, you are done. But if you
       
    65 really need the \isa{\isasymLongrightarrow}-version of \isa{hd\_rev}, for
       
    66 example because you want to apply it as an introduction rule, you need to
       
    67 derive it separately, by combining it with modus ponens:
       
    68 *}
       
    69 
       
    70 lemmas hd_revI = hd_rev[THEN mp]
       
    71  
       
    72 text{*\noindent
       
    73 which yields the lemma we originally set out to prove.
       
    74 
       
    75 In case there are multiple premises $A@1$, \dots, $A@n$ containing the
       
    76 induction variable, you should turn the conclusion $C$ into
       
    77 \[ A@1 \longrightarrow \cdots A@n \longrightarrow C \]
       
    78 (see the remark?? in \S\ref{??}).
       
    79 Additionally, you may also have to universally quantify some other variables,
       
    80 which can yield a fairly complex conclusion.
       
    81 Here is a simple example (which is proved by \isa{blast}):
       
    82 *}
       
    83 
       
    84 lemma simple: "\\<forall> y. A y \\<longrightarrow> B y \<longrightarrow> B y & A y"
       
    85 (*<*)by blast(*>*)
       
    86 
       
    87 text{*\noindent
       
    88 You can get the desired lemma by explicit
       
    89 application of modus ponens and \isa{spec}:
       
    90 *}
       
    91 
       
    92 lemmas myrule = simple[THEN spec, THEN mp, THEN mp]
       
    93 
       
    94 text{*\noindent
       
    95 or the wholesale stripping of \isa{\isasymforall} and
       
    96 \isa{\isasymlongrightarrow} in the conclusion via \isa{rulify} 
       
    97 *}
       
    98 
       
    99 lemmas myrule = simple[rulify]
       
   100 
       
   101 text{*\noindent
       
   102 yielding @{thm"myrule"}.
       
   103 You can go one step further and include these derivations already in the
       
   104 statement of your original lemma, thus avoiding the intermediate step:
       
   105 *}
       
   106 
       
   107 lemma myrule[rulify]:  "\\<forall> y. A y \\<longrightarrow> B y \<longrightarrow> B y & A y"
       
   108 (*<*)
       
   109 by blast
       
   110 (*>*)
       
   111 
       
   112 text{*
       
   113 \bigskip
       
   114 
       
   115 A second reason why your proposition may not be amenable to induction is that
       
   116 you want to induct on a whole term, rather than an individual variable. In
       
   117 general, when inducting on some term $t$ you must rephrase the conclusion as
       
   118 \[ \forall y@1 \dots y@n.~ x = t \longrightarrow C \] where $y@1 \dots y@n$
       
   119 are the free variables in $t$ and $x$ is new, and perform induction on $x$
       
   120 afterwards. An example appears below.
       
   121 
       
   122 *}
       
   123 
       
   124 subsection{*Beyond structural induction*}
       
   125 
       
   126 text{*
       
   127 So far, inductive proofs where by structural induction for
       
   128 primitive recursive functions and recursion induction for total recursive
       
   129 functions. But sometimes structural induction is awkward and there is no
       
   130 recursive function in sight either that could furnish a more appropriate
       
   131 induction schema. In such cases some existing standard induction schema can
       
   132 be helpful. We show how to apply such induction schemas by an example.
       
   133 
       
   134 Structural induction on \isa{nat} is
       
   135 usually known as ``mathematical induction''. There is also ``complete
       
   136 induction'', where you must prove $P(n)$ under the assumption that $P(m)$
       
   137 holds for all $m<n$. In Isabelle, this is the theorem \isa{less\_induct}:
       
   138 \begin{quote}
       
   139 @{thm[display]"less_induct"}
       
   140 \end{quote}
       
   141 Here is an example of its application.
       
   142 *}
       
   143 
       
   144 consts f :: "nat => nat"
       
   145 axioms f_ax: "f(f(n)) < f(Suc(n))"
       
   146 
       
   147 text{*\noindent
       
   148 From the above axiom\footnote{In general, the use of axioms is strongly
       
   149 discouraged, because of the danger of inconsistencies. The above axiom does
       
   150 not introduce an inconsistency because, for example, the identity function
       
   151 satisfies it.}
       
   152 for \isa{f} it follows that @{term"n <= f n"}, which can
       
   153 be proved by induction on @{term"f n"}. Following the recipy outlined
       
   154 above, we have to phrase the proposition as follows to allow induction:
       
   155 *}
       
   156 
       
   157 lemma f_incr_lem: "\\<forall>i. k = f i \\<longrightarrow> i \\<le> f i"
       
   158 
       
   159 txt{*\noindent
       
   160 To perform induction on \isa{k} using \isa{less\_induct}, we use the same
       
   161 general induction method as for recursion induction (see
       
   162 \S\ref{sec:recdef-induction}):
       
   163 *}
       
   164 
       
   165 apply(induct_tac k rule:less_induct)
       
   166 (*<*)
       
   167 apply(rule allI)
       
   168 apply(case_tac i);
       
   169  apply(simp);
       
   170 (*>*)
       
   171 txt{*\noindent
       
   172 which leaves us with the following proof state:
       
   173 \begin{isabellepar}%
       
   174 \ 1.\ {\isasymAnd}\mbox{n}.\ {\isasymforall}\mbox{m}.\ \mbox{m}\ <\ \mbox{n}\ {\isasymlongrightarrow}\ ({\isasymforall}\mbox{i}.\ \mbox{m}\ =\ f\ \mbox{i}\ {\isasymlongrightarrow}\ \mbox{i}\ {\isasymle}\ f\ \mbox{i})\isanewline
       
   175 \ \ \ \ \ \ \ {\isasymLongrightarrow}\ {\isasymforall}\mbox{i}.\ \mbox{n}\ =\ f\ \mbox{i}\ {\isasymlongrightarrow}\ \mbox{i}\ {\isasymle}\ f\ \mbox{i}
       
   176 \end{isabellepar}%
       
   177 After stripping the \isa{\isasymforall i}, the proof continues with a case
       
   178 distinction on \isa{i}. The case \isa{i = 0} is trivial and we focus on the
       
   179 other case:
       
   180 \begin{isabellepar}%
       
   181 \ 1.\ {\isasymAnd}\mbox{n}\ \mbox{i}\ \mbox{nat}.\isanewline
       
   182 \ \ \ \ \ \ \ {\isasymlbrakk}{\isasymforall}\mbox{m}.\ \mbox{m}\ <\ \mbox{n}\ {\isasymlongrightarrow}\ ({\isasymforall}\mbox{i}.\ \mbox{m}\ =\ f\ \mbox{i}\ {\isasymlongrightarrow}\ \mbox{i}\ {\isasymle}\ f\ \mbox{i});\ \mbox{i}\ =\ Suc\ \mbox{nat}{\isasymrbrakk}\isanewline
       
   183 \ \ \ \ \ \ \ {\isasymLongrightarrow}\ \mbox{n}\ =\ f\ \mbox{i}\ {\isasymlongrightarrow}\ \mbox{i}\ {\isasymle}\ f\ \mbox{i}
       
   184 \end{isabellepar}%
       
   185 *}
       
   186 
       
   187 by(blast intro!: f_ax Suc_leI intro:le_less_trans);
       
   188 
       
   189 text{*\noindent
       
   190 It is not surprising if you find the last step puzzling.
       
   191 The proof goes like this (writing \isa{j} instead of \isa{nat}).
       
   192 Since @{term"i = Suc j"} it suffices to show
       
   193 @{term"j < f(Suc j)"} (by \isa{Suc\_leI}: @{thm"Suc_leI"}). This is
       
   194 proved as follows. From \isa{f\_ax} we have @{term"f (f j) < f (Suc j)"}
       
   195 (1) which implies @{term"f j <= f (f j)"} (by the induction hypothesis).
       
   196 Using (1) once more we obtain @{term"f j < f(Suc j)"} (2) by transitivity
       
   197 (\isa{le_less_trans}: @{thm"le_less_trans"}).
       
   198 Using the induction hypothesis once more we obtain @{term"j <= f j"}
       
   199 which, together with (2) yields @{term"j < f (Suc j)"} (again by
       
   200 \isa{le_less_trans}).
       
   201 
       
   202 This last step shows both the power and the danger of automatic proofs: they
       
   203 will usually not tell you how the proof goes, because it can be very hard to
       
   204 translate the internal proof into a human-readable format. Therefore
       
   205 \S\ref{sec:part2?} introduces a language for writing readable yet concise
       
   206 proofs.
       
   207 
       
   208 We can now derive the desired @{term"i <= f i"} from \isa{f\_incr}:
       
   209 *}
       
   210 
       
   211 lemmas f_incr = f_incr_lem[rulify, OF refl];
       
   212 
       
   213 text{*
       
   214 The final \isa{refl} gets rid of the premise \isa{?k = f ?i}. Again, we could
       
   215 have included this derivation in the original statement of the lemma:
       
   216 *}
       
   217 
       
   218 lemma f_incr[rulify, OF refl]: "\\<forall>i. k = f i \\<longrightarrow> i \\<le> f i"
       
   219 (*<*)oops(*>*)
       
   220 
       
   221 text{*
       
   222 \begin{exercise}
       
   223 From the above axiom and lemma for \isa{f} show that \isa{f} is the identity.
       
   224 \end{exercise}
       
   225 
       
   226 In general, \isa{induct\_tac} can be applied with any rule \isa{r}
       
   227 whose conclusion is of the form \isa{?P ?x1 \dots ?xn}, in which case the
       
   228 format is
       
   229 \begin{ttbox}
       
   230 apply(induct_tac y1 ... yn rule: r)
       
   231 \end{ttbox}\index{*induct_tac}%
       
   232 where \isa{y1}, \dots, \isa{yn} are variables in the first subgoal.
       
   233 In fact, \isa{induct\_tac} even allows the conclusion of
       
   234 \isa{r} to be an (iterated) conjunction of formulae of the above form, in
       
   235 which case the application is
       
   236 \begin{ttbox}
       
   237 apply(induct_tac y1 ... yn and ... and z1 ... zm rule: r)
       
   238 \end{ttbox}
       
   239 
       
   240 Finally we should mention that HOL already provides the mother of all
       
   241 inductions, \emph{wellfounded induction} (\isa{wf\_induct}):
       
   242 \begin{quote}
       
   243 @{thm[display]"wf_induct"}
       
   244 \end{quote}
       
   245 For details see the library.
       
   246 *}
       
   247 
       
   248 (*<*)
       
   249 end
       
   250 (*>*)