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