| 8745 |      1 | (*<*)
 | 
| 17326 |      2 | theory Itrev
 | 
|  |      3 | imports Main
 | 
|  |      4 | begin
 | 
|  |      5 | ML"reset NameSpace.unique_names"
 | 
| 8745 |      6 | (*>*)
 | 
|  |      7 | 
 | 
| 10885 |      8 | section{*Induction Heuristics*}
 | 
| 9844 |      9 | 
 | 
|  |     10 | text{*\label{sec:InductionHeuristics}
 | 
| 11458 |     11 | \index{induction heuristics|(}%
 | 
| 9844 |     12 | The purpose of this section is to illustrate some simple heuristics for
 | 
|  |     13 | inductive proofs. The first one we have already mentioned in our initial
 | 
|  |     14 | example:
 | 
|  |     15 | \begin{quote}
 | 
|  |     16 | \emph{Theorems about recursive functions are proved by induction.}
 | 
|  |     17 | \end{quote}
 | 
|  |     18 | In case the function has more than one argument
 | 
|  |     19 | \begin{quote}
 | 
|  |     20 | \emph{Do induction on argument number $i$ if the function is defined by
 | 
|  |     21 | recursion in argument number $i$.}
 | 
|  |     22 | \end{quote}
 | 
| 11458 |     23 | When we look at the proof of @{text"(xs@ys) @ zs = xs @ (ys@zs)"}
 | 
|  |     24 | in \S\ref{sec:intro-proof} we find
 | 
|  |     25 | \begin{itemize}
 | 
|  |     26 | \item @{text"@"} is recursive in
 | 
|  |     27 | the first argument
 | 
|  |     28 | \item @{term xs}  occurs only as the first argument of
 | 
|  |     29 | @{text"@"}
 | 
|  |     30 | \item both @{term ys} and @{term zs} occur at least once as
 | 
|  |     31 | the second argument of @{text"@"}
 | 
|  |     32 | \end{itemize}
 | 
|  |     33 | Hence it is natural to perform induction on~@{term xs}.
 | 
| 9844 |     34 | 
 | 
|  |     35 | The key heuristic, and the main point of this section, is to
 | 
| 11458 |     36 | \emph{generalize the goal before induction}.
 | 
|  |     37 | The reason is simple: if the goal is
 | 
| 9844 |     38 | too specific, the induction hypothesis is too weak to allow the induction
 | 
| 10971 |     39 | step to go through. Let us illustrate the idea with an example.
 | 
| 9844 |     40 | 
 | 
| 11458 |     41 | Function \cdx{rev} has quadratic worst-case running time
 | 
| 9792 |     42 | because it calls function @{text"@"} for each element of the list and
 | 
|  |     43 | @{text"@"} is linear in its first argument.  A linear time version of
 | 
| 9754 |     44 | @{term"rev"} reqires an extra argument where the result is accumulated
 | 
| 11458 |     45 | gradually, using only~@{text"#"}:
 | 
| 9754 |     46 | *}
 | 
| 8745 |     47 | 
 | 
| 10362 |     48 | consts itrev :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list";
 | 
| 8745 |     49 | primrec
 | 
|  |     50 | "itrev []     ys = ys"
 | 
|  |     51 | "itrev (x#xs) ys = itrev xs (x#ys)";
 | 
|  |     52 | 
 | 
| 9754 |     53 | text{*\noindent
 | 
| 11458 |     54 | The behaviour of \cdx{itrev} is simple: it reverses
 | 
| 9493 |     55 | its first argument by stacking its elements onto the second argument,
 | 
|  |     56 | and returning that second argument when the first one becomes
 | 
| 11458 |     57 | empty. Note that @{term"itrev"} is tail-recursive: it can be
 | 
| 9493 |     58 | compiled into a loop.
 | 
|  |     59 | 
 | 
| 9754 |     60 | Naturally, we would like to show that @{term"itrev"} does indeed reverse
 | 
|  |     61 | its first argument provided the second one is empty:
 | 
|  |     62 | *};
 | 
| 8745 |     63 | 
 | 
|  |     64 | lemma "itrev xs [] = rev xs";
 | 
|  |     65 | 
 | 
|  |     66 | txt{*\noindent
 | 
|  |     67 | There is no choice as to the induction variable, and we immediately simplify:
 | 
| 9458 |     68 | *};
 | 
| 8745 |     69 | 
 | 
| 9689 |     70 | apply(induct_tac xs, simp_all);
 | 
| 8745 |     71 | 
 | 
|  |     72 | txt{*\noindent
 | 
| 11458 |     73 | Unfortunately, this attempt does not prove
 | 
|  |     74 | the induction step:
 | 
| 10971 |     75 | @{subgoals[display,indent=0,margin=70]}
 | 
| 11458 |     76 | The induction hypothesis is too weak.  The fixed
 | 
|  |     77 | argument,~@{term"[]"}, prevents it from rewriting the conclusion.  
 | 
|  |     78 | This example suggests a heuristic:
 | 
|  |     79 | \begin{quote}\index{generalizing induction formulae}%
 | 
| 9754 |     80 | \emph{Generalize goals for induction by replacing constants by variables.}
 | 
| 8745 |     81 | \end{quote}
 | 
| 9754 |     82 | Of course one cannot do this na\"{\i}vely: @{term"itrev xs ys = rev xs"} is
 | 
| 11458 |     83 | just not true.  The correct generalization is
 | 
| 9458 |     84 | *};
 | 
| 8745 |     85 | (*<*)oops;(*>*)
 | 
|  |     86 | lemma "itrev xs ys = rev xs @ ys";
 | 
| 10362 |     87 | (*<*)apply(induct_tac xs, simp_all)(*>*)
 | 
| 8745 |     88 | txt{*\noindent
 | 
| 9754 |     89 | If @{term"ys"} is replaced by @{term"[]"}, the right-hand side simplifies to
 | 
| 11458 |     90 | @{term"rev xs"}, as required.
 | 
| 8745 |     91 | 
 | 
| 11458 |     92 | In this instance it was easy to guess the right generalization.
 | 
|  |     93 | Other situations can require a good deal of creativity.  
 | 
| 8745 |     94 | 
 | 
| 9754 |     95 | Although we now have two variables, only @{term"xs"} is suitable for
 | 
| 11458 |     96 | induction, and we repeat our proof attempt. Unfortunately, we are still
 | 
| 8745 |     97 | not there:
 | 
| 10362 |     98 | @{subgoals[display,indent=0,goals_limit=1]}
 | 
| 8745 |     99 | The induction hypothesis is still too weak, but this time it takes no
 | 
| 9754 |    100 | intuition to generalize: the problem is that @{term"ys"} is fixed throughout
 | 
| 8745 |    101 | the subgoal, but the induction hypothesis needs to be applied with
 | 
| 9754 |    102 | @{term"a # ys"} instead of @{term"ys"}. Hence we prove the theorem
 | 
|  |    103 | for all @{term"ys"} instead of a fixed one:
 | 
| 9458 |    104 | *};
 | 
| 8745 |    105 | (*<*)oops;(*>*)
 | 
| 10362 |    106 | lemma "\<forall>ys. itrev xs ys = rev xs @ ys";
 | 
| 9844 |    107 | (*<*)
 | 
|  |    108 | by(induct_tac xs, simp_all);
 | 
|  |    109 | (*>*)
 | 
| 8745 |    110 | 
 | 
| 9844 |    111 | text{*\noindent
 | 
| 9754 |    112 | This time induction on @{term"xs"} followed by simplification succeeds. This
 | 
| 8745 |    113 | leads to another heuristic for generalization:
 | 
|  |    114 | \begin{quote}
 | 
| 9754 |    115 | \emph{Generalize goals for induction by universally quantifying all free
 | 
| 8745 |    116 | variables {\em(except the induction variable itself!)}.}
 | 
|  |    117 | \end{quote}
 | 
| 11458 |    118 | This prevents trivial failures like the one above and does not affect the
 | 
|  |    119 | validity of the goal.  However, this heuristic should not be applied blindly.
 | 
|  |    120 | It is not always required, and the additional quantifiers can complicate
 | 
| 13081 |    121 | matters in some cases. The variables that should be quantified are typically
 | 
| 11458 |    122 | those that change in recursive calls.
 | 
| 9644 |    123 | 
 | 
| 9844 |    124 | A final point worth mentioning is the orientation of the equation we just
 | 
| 15905 |    125 | proved: the more complex notion (@{const itrev}) is on the left-hand
 | 
| 9844 |    126 | side, the simpler one (@{term rev}) on the right-hand side. This constitutes
 | 
|  |    127 | another, albeit weak heuristic that is not restricted to induction:
 | 
|  |    128 | \begin{quote}
 | 
|  |    129 |   \emph{The right-hand side of an equation should (in some sense) be simpler
 | 
|  |    130 |     than the left-hand side.}
 | 
|  |    131 | \end{quote}
 | 
|  |    132 | This heuristic is tricky to apply because it is not obvious that
 | 
|  |    133 | @{term"rev xs @ ys"} is simpler than @{term"itrev xs ys"}. But see what
 | 
|  |    134 | happens if you try to prove @{prop"rev xs @ ys = itrev xs ys"}!
 | 
| 10971 |    135 | 
 | 
| 11458 |    136 | If you have tried these heuristics and still find your
 | 
| 10971 |    137 | induction does not go through, and no obvious lemma suggests itself, you may
 | 
|  |    138 | need to generalize your proposition even further. This requires insight into
 | 
| 11458 |    139 | the problem at hand and is beyond simple rules of thumb.  
 | 
|  |    140 | Additionally, you can read \S\ref{sec:advanced-ind}
 | 
|  |    141 | to learn about some advanced techniques for inductive proofs.%
 | 
|  |    142 | \index{induction heuristics|)}
 | 
| 9844 |    143 | *}
 | 
| 8745 |    144 | (*<*)
 | 
| 17326 |    145 | ML"set NameSpace.unique_names"
 | 
| 8745 |    146 | end
 | 
|  |    147 | (*>*)
 |