10341
|
1 |
(* ID: $Id$ *)
|
10314
|
2 |
theory Even = Main:
|
|
3 |
|
|
4 |
|
|
5 |
consts even :: "nat set"
|
|
6 |
inductive even
|
|
7 |
intros
|
|
8 |
zero[intro!]: "0 \<in> even"
|
|
9 |
step[intro!]: "n \<in> even \<Longrightarrow> (Suc (Suc n)) \<in> even"
|
|
10 |
|
10883
|
11 |
text{*An inductive definition consists of introduction rules.
|
10314
|
12 |
|
10326
|
13 |
@{thm[display] even.step[no_vars]}
|
10314
|
14 |
\rulename{even.step}
|
|
15 |
|
10326
|
16 |
@{thm[display] even.induct[no_vars]}
|
10314
|
17 |
\rulename{even.induct}
|
|
18 |
|
|
19 |
Attributes can be given to the introduction rules. Here both rules are
|
10883
|
20 |
specified as \isa{intro!}
|
10314
|
21 |
|
10883
|
22 |
Our first lemma states that numbers of the form $2\times k$ are even. *}
|
11705
|
23 |
lemma two_times_even[intro!]: "2*k \<in> even"
|
12328
|
24 |
apply (induct_tac k)
|
10883
|
25 |
txt{*
|
|
26 |
The first step is induction on the natural number \isa{k}, which leaves
|
|
27 |
two subgoals:
|
|
28 |
@{subgoals[display,indent=0,margin=65]}
|
|
29 |
Here \isa{auto} simplifies both subgoals so that they match the introduction
|
|
30 |
rules, which then are applied automatically.*};
|
10314
|
31 |
apply auto
|
|
32 |
done
|
|
33 |
|
|
34 |
text{*Our goal is to prove the equivalence between the traditional definition
|
|
35 |
of even (using the divides relation) and our inductive definition. Half of
|
|
36 |
this equivalence is trivial using the lemma just proved, whose \isa{intro!}
|
|
37 |
attribute ensures it will be applied automatically. *}
|
|
38 |
|
11705
|
39 |
lemma dvd_imp_even: "2 dvd n \<Longrightarrow> n \<in> even"
|
10883
|
40 |
by (auto simp add: dvd_def)
|
10314
|
41 |
|
|
42 |
text{*our first rule induction!*}
|
11705
|
43 |
lemma even_imp_dvd: "n \<in> even \<Longrightarrow> 2 dvd n"
|
10314
|
44 |
apply (erule even.induct)
|
10883
|
45 |
txt{*
|
|
46 |
@{subgoals[display,indent=0,margin=65]}
|
|
47 |
*};
|
|
48 |
apply (simp_all add: dvd_def)
|
|
49 |
txt{*
|
|
50 |
@{subgoals[display,indent=0,margin=65]}
|
|
51 |
*};
|
10314
|
52 |
apply clarify
|
10883
|
53 |
txt{*
|
|
54 |
@{subgoals[display,indent=0,margin=65]}
|
|
55 |
*};
|
|
56 |
apply (rule_tac x = "Suc k" in exI, simp)
|
10314
|
57 |
done
|
|
58 |
|
|
59 |
|
|
60 |
text{*no iff-attribute because we don't always want to use it*}
|
11705
|
61 |
theorem even_iff_dvd: "(n \<in> even) = (2 dvd n)"
|
10883
|
62 |
by (blast intro: dvd_imp_even even_imp_dvd)
|
10314
|
63 |
|
|
64 |
text{*this result ISN'T inductive...*}
|
|
65 |
lemma Suc_Suc_even_imp_even: "Suc (Suc n) \<in> even \<Longrightarrow> n \<in> even"
|
|
66 |
apply (erule even.induct)
|
10883
|
67 |
txt{*
|
|
68 |
@{subgoals[display,indent=0,margin=65]}
|
|
69 |
*};
|
10314
|
70 |
oops
|
|
71 |
|
|
72 |
text{*...so we need an inductive lemma...*}
|
11705
|
73 |
lemma even_imp_even_minus_2: "n \<in> even \<Longrightarrow> n - 2 \<in> even"
|
10314
|
74 |
apply (erule even.induct)
|
10883
|
75 |
txt{*
|
|
76 |
@{subgoals[display,indent=0,margin=65]}
|
|
77 |
*};
|
10314
|
78 |
apply auto
|
|
79 |
done
|
|
80 |
|
10883
|
81 |
text{*...and prove it in a separate step*}
|
|
82 |
lemma Suc_Suc_even_imp_even: "Suc (Suc n) \<in> even \<Longrightarrow> n \<in> even"
|
|
83 |
by (drule even_imp_even_minus_2, simp)
|
10326
|
84 |
|
|
85 |
|
|
86 |
lemma [iff]: "((Suc (Suc n)) \<in> even) = (n \<in> even)"
|
10883
|
87 |
by (blast dest: Suc_Suc_even_imp_even)
|
10314
|
88 |
|
|
89 |
end
|
|
90 |
|