| author | wenzelm |
| Tue, 28 Nov 2000 01:22:56 +0100 | |
| changeset 10531 | a9e7786db49e |
| parent 10420 | ef006735bee8 |
| child 10654 | 458068404143 |
| permissions | -rw-r--r-- |
| 9645 | 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 |
|
| 10396 | 8 |
proposition to be proved is not directly amenable to induction |
9 |
(\S\ref{sec:ind-var-in-prems}), and how to utilize (\S\ref{sec:complete-ind})
|
|
10 |
and even derive (\S\ref{sec:derive-ind}) new induction schemas. We conclude
|
|
11 |
with an extended example of induction (\S\ref{sec:CTL-revisited}).
|
|
| 9689 | 12 |
*}; |
| 9645 | 13 |
|
| 10217 | 14 |
subsection{*Massaging the proposition*};
|
| 9645 | 15 |
|
| 10217 | 16 |
text{*\label{sec:ind-var-in-prems}
|
| 9645 | 17 |
So far we have assumed that the theorem we want to prove is already in a form |
18 |
that is amenable to induction, but this is not always the case: |
|
| 9689 | 19 |
*}; |
| 9645 | 20 |
|
| 9933 | 21 |
lemma "xs \<noteq> [] \<Longrightarrow> hd(rev xs) = last xs"; |
| 9645 | 22 |
apply(induct_tac xs); |
23 |
||
24 |
txt{*\noindent
|
|
| 9792 | 25 |
(where @{term"hd"} and @{term"last"} return the first and last element of a
|
| 9645 | 26 |
non-empty list) |
27 |
produces the warning |
|
28 |
\begin{quote}\tt
|
|
29 |
Induction variable occurs also among premises! |
|
30 |
\end{quote}
|
|
31 |
and leads to the base case |
|
| 10363 | 32 |
@{subgoals[display,indent=0,goals_limit=1]}
|
| 9645 | 33 |
which, after simplification, becomes |
| 9723 | 34 |
\begin{isabelle}
|
| 9645 | 35 |
\ 1.\ xs\ {\isasymnoteq}\ []\ {\isasymLongrightarrow}\ hd\ []\ =\ last\ []
|
| 9723 | 36 |
\end{isabelle}
|
| 10242 | 37 |
We cannot prove this equality because we do not know what @{term hd} and
|
38 |
@{term last} return when applied to @{term"[]"}.
|
|
| 9645 | 39 |
|
40 |
The point is that we have violated the above warning. Because the induction |
|
| 10242 | 41 |
formula is only the conclusion, the occurrence of @{term xs} in the premises is
|
| 9645 | 42 |
not modified by induction. Thus the case that should have been trivial |
| 10242 | 43 |
becomes unprovable. Fortunately, the solution is easy:\footnote{A very similar
|
44 |
heuristic applies to rule inductions; see \S\ref{sec:rtc}.}
|
|
| 9645 | 45 |
\begin{quote}
|
46 |
\emph{Pull all occurrences of the induction variable into the conclusion
|
|
| 9792 | 47 |
using @{text"\<longrightarrow>"}.}
|
| 9645 | 48 |
\end{quote}
|
49 |
This means we should prove |
|
| 9689 | 50 |
*}; |
51 |
(*<*)oops;(*>*) |
|
| 9933 | 52 |
lemma hd_rev: "xs \<noteq> [] \<longrightarrow> hd(rev xs) = last xs"; |
| 9645 | 53 |
(*<*) |
| 10420 | 54 |
apply(induct_tac xs); |
| 9645 | 55 |
(*>*) |
56 |
||
| 10420 | 57 |
txt{*\noindent
|
| 9645 | 58 |
This time, induction leaves us with the following base case |
| 10420 | 59 |
@{subgoals[display,indent=0,goals_limit=1]}
|
| 9792 | 60 |
which is trivial, and @{text"auto"} finishes the whole proof.
|
| 9645 | 61 |
|
| 10420 | 62 |
If @{text hd_rev} is meant to be a simplification rule, you are
|
| 9792 | 63 |
done. But if you really need the @{text"\<Longrightarrow>"}-version of
|
| 10420 | 64 |
@{text hd_rev}, for example because you want to apply it as an
|
| 9792 | 65 |
introduction rule, you need to derive it separately, by combining it with |
66 |
modus ponens: |
|
| 10420 | 67 |
*}(*<*)by(auto);(*>*) |
| 9689 | 68 |
lemmas hd_revI = hd_rev[THEN mp]; |
| 9645 | 69 |
|
70 |
text{*\noindent
|
|
71 |
which yields the lemma we originally set out to prove. |
|
72 |
||
73 |
In case there are multiple premises $A@1$, \dots, $A@n$ containing the |
|
74 |
induction variable, you should turn the conclusion $C$ into |
|
75 |
\[ A@1 \longrightarrow \cdots A@n \longrightarrow C \] |
|
76 |
(see the remark?? in \S\ref{??}).
|
|
77 |
Additionally, you may also have to universally quantify some other variables, |
|
78 |
which can yield a fairly complex conclusion. |
|
| 9792 | 79 |
Here is a simple example (which is proved by @{text"blast"}):
|
| 9689 | 80 |
*}; |
| 9645 | 81 |
|
| 10281 | 82 |
lemma simple: "\<forall>y. A y \<longrightarrow> B y \<longrightarrow> B y \<and> A y"; |
| 9689 | 83 |
(*<*)by blast;(*>*) |
| 9645 | 84 |
|
85 |
text{*\noindent
|
|
86 |
You can get the desired lemma by explicit |
|
| 9792 | 87 |
application of modus ponens and @{thm[source]spec}:
|
| 9689 | 88 |
*}; |
| 9645 | 89 |
|
| 9689 | 90 |
lemmas myrule = simple[THEN spec, THEN mp, THEN mp]; |
| 9645 | 91 |
|
92 |
text{*\noindent
|
|
| 9792 | 93 |
or the wholesale stripping of @{text"\<forall>"} and
|
|
9941
fe05af7ec816
renamed atts: rulify to rule_format, elimify to elim_format;
wenzelm
parents:
9933
diff
changeset
|
94 |
@{text"\<longrightarrow>"} in the conclusion via @{text"rule_format"}
|
| 9689 | 95 |
*}; |
| 9645 | 96 |
|
|
9941
fe05af7ec816
renamed atts: rulify to rule_format, elimify to elim_format;
wenzelm
parents:
9933
diff
changeset
|
97 |
lemmas myrule = simple[rule_format]; |
| 9645 | 98 |
|
99 |
text{*\noindent
|
|
| 9689 | 100 |
yielding @{thm"myrule"[no_vars]}.
|
| 9645 | 101 |
You can go one step further and include these derivations already in the |
102 |
statement of your original lemma, thus avoiding the intermediate step: |
|
| 9689 | 103 |
*}; |
| 9645 | 104 |
|
| 10281 | 105 |
lemma myrule[rule_format]: "\<forall>y. A y \<longrightarrow> B y \<longrightarrow> B y \<and> A y"; |
| 9645 | 106 |
(*<*) |
| 9689 | 107 |
by blast; |
| 9645 | 108 |
(*>*) |
109 |
||
110 |
text{*
|
|
111 |
\bigskip |
|
112 |
||
113 |
A second reason why your proposition may not be amenable to induction is that |
|
114 |
you want to induct on a whole term, rather than an individual variable. In |
|
| 10217 | 115 |
general, when inducting on some term $t$ you must rephrase the conclusion $C$ |
116 |
as |
|
117 |
\[ \forall y@1 \dots y@n.~ x = t \longrightarrow C \] |
|
118 |
where $y@1 \dots y@n$ are the free variables in $t$ and $x$ is new, and |
|
119 |
perform induction on $x$ afterwards. An example appears in |
|
120 |
\S\ref{sec:complete-ind} below.
|
|
121 |
||
122 |
The very same problem may occur in connection with rule induction. Remember |
|
123 |
that it requires a premise of the form $(x@1,\dots,x@k) \in R$, where $R$ is |
|
124 |
some inductively defined set and the $x@i$ are variables. If instead we have |
|
125 |
a premise $t \in R$, where $t$ is not just an $n$-tuple of variables, we |
|
126 |
replace it with $(x@1,\dots,x@k) \in R$, and rephrase the conclusion $C$ as |
|
127 |
\[ \forall y@1 \dots y@n.~ (x@1,\dots,x@k) = t \longrightarrow C \] |
|
128 |
For an example see \S\ref{sec:CTL-revisited} below.
|
|
| 10281 | 129 |
|
130 |
Of course, all premises that share free variables with $t$ need to be pulled into |
|
131 |
the conclusion as well, under the @{text"\<forall>"}, again using @{text"\<longrightarrow>"} as shown above.
|
|
| 9689 | 132 |
*}; |
| 9645 | 133 |
|
| 9689 | 134 |
subsection{*Beyond structural and recursion induction*};
|
| 9645 | 135 |
|
| 10217 | 136 |
text{*\label{sec:complete-ind}
|
| 9645 | 137 |
So far, inductive proofs where by structural induction for |
138 |
primitive recursive functions and recursion induction for total recursive |
|
139 |
functions. But sometimes structural induction is awkward and there is no |
|
140 |
recursive function in sight either that could furnish a more appropriate |
|
141 |
induction schema. In such cases some existing standard induction schema can |
|
142 |
be helpful. We show how to apply such induction schemas by an example. |
|
143 |
||
| 9792 | 144 |
Structural induction on @{typ"nat"} is
|
| 9645 | 145 |
usually known as ``mathematical induction''. There is also ``complete |
146 |
induction'', where you must prove $P(n)$ under the assumption that $P(m)$ |
|
| 9933 | 147 |
holds for all $m<n$. In Isabelle, this is the theorem @{thm[source]nat_less_induct}:
|
148 |
@{thm[display]"nat_less_induct"[no_vars]}
|
|
| 9645 | 149 |
Here is an example of its application. |
| 9689 | 150 |
*}; |
| 9645 | 151 |
|
| 10281 | 152 |
consts f :: "nat \<Rightarrow> nat"; |
| 9689 | 153 |
axioms f_ax: "f(f(n)) < f(Suc(n))"; |
| 9645 | 154 |
|
155 |
text{*\noindent
|
|
156 |
From the above axiom\footnote{In general, the use of axioms is strongly
|
|
157 |
discouraged, because of the danger of inconsistencies. The above axiom does |
|
158 |
not introduce an inconsistency because, for example, the identity function |
|
159 |
satisfies it.} |
|
| 9792 | 160 |
for @{term"f"} it follows that @{prop"n <= f n"}, which can
|
| 10396 | 161 |
be proved by induction on @{term"f n"}. Following the recipe outlined
|
| 9645 | 162 |
above, we have to phrase the proposition as follows to allow induction: |
| 9689 | 163 |
*}; |
| 9645 | 164 |
|
| 9933 | 165 |
lemma f_incr_lem: "\<forall>i. k = f i \<longrightarrow> i \<le> f i"; |
| 9645 | 166 |
|
167 |
txt{*\noindent
|
|
| 10363 | 168 |
To perform induction on @{term k} using @{thm[source]nat_less_induct}, we use
|
169 |
the same general induction method as for recursion induction (see |
|
| 9645 | 170 |
\S\ref{sec:recdef-induction}):
|
| 9689 | 171 |
*}; |
| 9645 | 172 |
|
| 9923 | 173 |
apply(induct_tac k rule: nat_less_induct); |
| 10363 | 174 |
|
175 |
txt{*\noindent
|
|
176 |
which leaves us with the following proof state: |
|
177 |
@{subgoals[display,indent=0,margin=65]}
|
|
178 |
After stripping the @{text"\<forall>i"}, the proof continues with a case
|
|
179 |
distinction on @{term"i"}. The case @{prop"i = 0"} is trivial and we focus on
|
|
180 |
the other case: |
|
181 |
*} |
|
182 |
||
| 9689 | 183 |
apply(rule allI); |
| 9645 | 184 |
apply(case_tac i); |
185 |
apply(simp); |
|
| 10363 | 186 |
|
187 |
txt{*
|
|
188 |
@{subgoals[display,indent=0]}
|
|
| 9689 | 189 |
*}; |
| 9645 | 190 |
|
| 9923 | 191 |
by(blast intro!: f_ax Suc_leI intro: le_less_trans); |
| 9645 | 192 |
|
193 |
text{*\noindent
|
|
194 |
It is not surprising if you find the last step puzzling. |
|
| 9792 | 195 |
The proof goes like this (writing @{term"j"} instead of @{typ"nat"}).
|
196 |
Since @{prop"i = Suc j"} it suffices to show
|
|
197 |
@{prop"j < f(Suc j)"} (by @{thm[source]Suc_leI}: @{thm"Suc_leI"[no_vars]}). This is
|
|
198 |
proved as follows. From @{thm[source]f_ax} we have @{prop"f (f j) < f (Suc j)"}
|
|
199 |
(1) which implies @{prop"f j <= f (f j)"} (by the induction hypothesis).
|
|
200 |
Using (1) once more we obtain @{prop"f j < f(Suc j)"} (2) by transitivity
|
|
201 |
(@{thm[source]le_less_trans}: @{thm"le_less_trans"[no_vars]}).
|
|
202 |
Using the induction hypothesis once more we obtain @{prop"j <= f j"}
|
|
203 |
which, together with (2) yields @{prop"j < f (Suc j)"} (again by
|
|
204 |
@{thm[source]le_less_trans}).
|
|
| 9645 | 205 |
|
206 |
This last step shows both the power and the danger of automatic proofs: they |
|
207 |
will usually not tell you how the proof goes, because it can be very hard to |
|
208 |
translate the internal proof into a human-readable format. Therefore |
|
209 |
\S\ref{sec:part2?} introduces a language for writing readable yet concise
|
|
210 |
proofs. |
|
211 |
||
| 9792 | 212 |
We can now derive the desired @{prop"i <= f i"} from @{text"f_incr"}:
|
| 9689 | 213 |
*}; |
| 9645 | 214 |
|
|
9941
fe05af7ec816
renamed atts: rulify to rule_format, elimify to elim_format;
wenzelm
parents:
9933
diff
changeset
|
215 |
lemmas f_incr = f_incr_lem[rule_format, OF refl]; |
| 9645 | 216 |
|
| 9689 | 217 |
text{*\noindent
|
| 9792 | 218 |
The final @{thm[source]refl} gets rid of the premise @{text"?k = f ?i"}. Again,
|
219 |
we could have included this derivation in the original statement of the lemma: |
|
| 9689 | 220 |
*}; |
| 9645 | 221 |
|
|
9941
fe05af7ec816
renamed atts: rulify to rule_format, elimify to elim_format;
wenzelm
parents:
9933
diff
changeset
|
222 |
lemma f_incr[rule_format, OF refl]: "\<forall>i. k = f i \<longrightarrow> i \<le> f i"; |
| 9689 | 223 |
(*<*)oops;(*>*) |
| 9645 | 224 |
|
225 |
text{*
|
|
226 |
\begin{exercise}
|
|
| 9792 | 227 |
From the above axiom and lemma for @{term"f"} show that @{term"f"} is the
|
228 |
identity. |
|
| 9645 | 229 |
\end{exercise}
|
230 |
||
| 10236 | 231 |
In general, @{text induct_tac} can be applied with any rule $r$
|
| 9792 | 232 |
whose conclusion is of the form ${?}P~?x@1 \dots ?x@n$, in which case the
|
| 9645 | 233 |
format is |
| 9792 | 234 |
\begin{quote}
|
235 |
\isacommand{apply}@{text"(induct_tac"} $y@1 \dots y@n$ @{text"rule:"} $r$@{text")"}
|
|
236 |
\end{quote}\index{*induct_tac}%
|
|
237 |
where $y@1, \dots, y@n$ are variables in the first subgoal. |
|
| 10236 | 238 |
A further example of a useful induction rule is @{thm[source]length_induct},
|
239 |
induction on the length of a list:\indexbold{*length_induct}
|
|
240 |
@{thm[display]length_induct[no_vars]}
|
|
241 |
||
| 9792 | 242 |
In fact, @{text"induct_tac"} even allows the conclusion of
|
243 |
$r$ to be an (iterated) conjunction of formulae of the above form, in |
|
| 9645 | 244 |
which case the application is |
| 9792 | 245 |
\begin{quote}
|
246 |
\isacommand{apply}@{text"(induct_tac"} $y@1 \dots y@n$ @{text"and"} \dots\ @{text"and"} $z@1 \dots z@m$ @{text"rule:"} $r$@{text")"}
|
|
247 |
\end{quote}
|
|
| 9689 | 248 |
*}; |
| 9645 | 249 |
|
| 9689 | 250 |
subsection{*Derivation of new induction schemas*};
|
251 |
||
252 |
text{*\label{sec:derive-ind}
|
|
253 |
Induction schemas are ordinary theorems and you can derive new ones |
|
254 |
whenever you wish. This section shows you how to, using the example |
|
| 9933 | 255 |
of @{thm[source]nat_less_induct}. Assume we only have structural induction
|
| 9689 | 256 |
available for @{typ"nat"} and want to derive complete induction. This
|
257 |
requires us to generalize the statement first: |
|
258 |
*}; |
|
259 |
||
| 9933 | 260 |
lemma induct_lem: "(\<And>n::nat. \<forall>m<n. P m \<Longrightarrow> P n) \<Longrightarrow> \<forall>m<n. P m"; |
| 9689 | 261 |
apply(induct_tac n); |
262 |
||
263 |
txt{*\noindent
|
|
| 9792 | 264 |
The base case is trivially true. For the induction step (@{prop"m <
|
| 9933 | 265 |
Suc n"}) we distinguish two cases: case @{prop"m < n"} is true by induction
|
266 |
hypothesis and case @{prop"m = n"} follows from the assumption, again using
|
|
| 9689 | 267 |
the induction hypothesis: |
268 |
*}; |
|
269 |
apply(blast); |
|
| 9933 | 270 |
by(blast elim:less_SucE) |
| 9689 | 271 |
|
272 |
text{*\noindent
|
|
| 9792 | 273 |
The elimination rule @{thm[source]less_SucE} expresses the case distinction:
|
| 9689 | 274 |
@{thm[display]"less_SucE"[no_vars]}
|
275 |
||
276 |
Now it is straightforward to derive the original version of |
|
| 9933 | 277 |
@{thm[source]nat_less_induct} by manipulting the conclusion of the above lemma:
|
| 9792 | 278 |
instantiate @{term"n"} by @{term"Suc n"} and @{term"m"} by @{term"n"} and
|
| 10396 | 279 |
remove the trivial condition @{prop"n < Suc n"}. Fortunately, this
|
| 9689 | 280 |
happens automatically when we add the lemma as a new premise to the |
281 |
desired goal: |
|
282 |
*}; |
|
283 |
||
| 9933 | 284 |
theorem nat_less_induct: "(\<And>n::nat. \<forall>m<n. P m \<Longrightarrow> P n) \<Longrightarrow> P n"; |
| 9689 | 285 |
by(insert induct_lem, blast); |
286 |
||
| 9933 | 287 |
text{*
|
| 10396 | 288 |
Finally we should remind the reader that HOL already provides the mother of |
289 |
all inductions, well-founded induction (see \S\ref{sec:Well-founded}). For
|
|
290 |
example theorem @{thm[source]nat_less_induct} can be viewed (and derived) as
|
|
291 |
a special case of @{thm[source]wf_induct} where @{term r} is @{text"<"} on
|
|
292 |
@{typ nat}. The details can be found in the HOL library.
|
|
| 9689 | 293 |
*}; |
| 9645 | 294 |
|
295 |
(*<*) |
|
296 |
end |
|
297 |
(*>*) |