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