author | wenzelm |
Sat, 22 May 2010 20:02:26 +0200 | |
changeset 37060 | 6f2731bdba11 |
parent 25330 | 15bf0f47a87d |
child 42637 | 381fdcab0f36 |
permissions | -rw-r--r-- |
10341 | 1 |
(* ID: $Id$ *) |
23842
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
2 |
(*<*)theory Even imports Main uses "../../antiquote_setup.ML" begin(*>*) |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
3 |
|
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
4 |
section{* The Set of Even Numbers *} |
10314 | 5 |
|
23842
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
6 |
text {* |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
7 |
\index{even numbers!defining inductively|(}% |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
8 |
The set of even numbers can be inductively defined as the least set |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
9 |
containing 0 and closed under the operation $+2$. Obviously, |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
10 |
\emph{even} can also be expressed using the divides relation (@{text dvd}). |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
11 |
We shall prove below that the two formulations coincide. On the way we |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
12 |
shall examine the primary means of reasoning about inductively defined |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
13 |
sets: rule induction. |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
14 |
*} |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
15 |
|
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
16 |
subsection{* Making an Inductive Definition *} |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
17 |
|
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
18 |
text {* |
23928 | 19 |
Using \commdx{inductive\protect\_set}, we declare the constant @{text even} to be |
23842
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
20 |
a set of natural numbers with the desired properties. |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
21 |
*} |
10314 | 22 |
|
25330 | 23 |
inductive_set even :: "nat set" where |
24 |
zero[intro!]: "0 \<in> even" | |
|
25 |
step[intro!]: "n \<in> even \<Longrightarrow> (Suc (Suc n)) \<in> even" |
|
10314 | 26 |
|
23842
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
27 |
text {* |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
28 |
An inductive definition consists of introduction rules. The first one |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
29 |
above states that 0 is even; the second states that if $n$ is even, then so |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
30 |
is~$n+2$. Given this declaration, Isabelle generates a fixed point |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
31 |
definition for @{term even} and proves theorems about it, |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
32 |
thus following the definitional approach (see {\S}\ref{sec:definitional}). |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
33 |
These theorems |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
34 |
include the introduction rules specified in the declaration, an elimination |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
35 |
rule for case analysis and an induction rule. We can refer to these |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
36 |
theorems by automatically-generated names. Here are two examples: |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
37 |
@{named_thms[display,indent=0] even.zero[no_vars] (even.zero) even.step[no_vars] (even.step)} |
10314 | 38 |
|
23842
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
39 |
The introduction rules can be given attributes. Here |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
40 |
both rules are specified as \isa{intro!},% |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
41 |
\index{intro"!@\isa {intro"!} (attribute)} |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
42 |
directing the classical reasoner to |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
43 |
apply them aggressively. Obviously, regarding 0 as even is safe. The |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
44 |
@{text step} rule is also safe because $n+2$ is even if and only if $n$ is |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
45 |
even. We prove this equivalence later. |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
46 |
*} |
10314 | 47 |
|
23842
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
48 |
subsection{*Using Introduction Rules*} |
10314 | 49 |
|
23842
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
50 |
text {* |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
51 |
Our first lemma states that numbers of the form $2\times k$ are even. |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
52 |
Introduction rules are used to show that specific values belong to the |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
53 |
inductive set. Such proofs typically involve |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
54 |
induction, perhaps over some other inductive set. |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
55 |
*} |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
56 |
|
11705 | 57 |
lemma two_times_even[intro!]: "2*k \<in> even" |
12328 | 58 |
apply (induct_tac k) |
23842
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
59 |
apply auto |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
60 |
done |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
61 |
(*<*) |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
62 |
lemma "2*k \<in> even" |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
63 |
apply (induct_tac k) |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
64 |
(*>*) |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
65 |
txt {* |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
66 |
\noindent |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
67 |
The first step is induction on the natural number @{text k}, which leaves |
10883 | 68 |
two subgoals: |
69 |
@{subgoals[display,indent=0,margin=65]} |
|
23842
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
70 |
Here @{text auto} simplifies both subgoals so that they match the introduction |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
71 |
rules, which are then applied automatically. |
10314 | 72 |
|
23842
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
73 |
Our ultimate goal is to prove the equivalence between the traditional |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
74 |
definition of @{text even} (using the divides relation) and our inductive |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
75 |
definition. One direction of this equivalence is immediate by the lemma |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
76 |
just proved, whose @{text "intro!"} attribute ensures it is applied automatically. |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
77 |
*} |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
78 |
(*<*)oops(*>*) |
11705 | 79 |
lemma dvd_imp_even: "2 dvd n \<Longrightarrow> n \<in> even" |
10883 | 80 |
by (auto simp add: dvd_def) |
10314 | 81 |
|
23842
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
82 |
subsection{* Rule Induction \label{sec:rule-induction} *} |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
83 |
|
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
84 |
text {* |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
85 |
\index{rule induction|(}% |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
86 |
From the definition of the set |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
87 |
@{term even}, Isabelle has |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
88 |
generated an induction rule: |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
89 |
@{named_thms [display,indent=0,margin=40] even.induct [no_vars] (even.induct)} |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
90 |
A property @{term P} holds for every even number provided it |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
91 |
holds for~@{text 0} and is closed under the operation |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
92 |
\isa{Suc(Suc \(\cdot\))}. Then @{term P} is closed under the introduction |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
93 |
rules for @{term even}, which is the least set closed under those rules. |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
94 |
This type of inductive argument is called \textbf{rule induction}. |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
95 |
|
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
96 |
Apart from the double application of @{term Suc}, the induction rule above |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
97 |
resembles the familiar mathematical induction, which indeed is an instance |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
98 |
of rule induction; the natural numbers can be defined inductively to be |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
99 |
the least set containing @{text 0} and closed under~@{term Suc}. |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
100 |
|
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
101 |
Induction is the usual way of proving a property of the elements of an |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
102 |
inductively defined set. Let us prove that all members of the set |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
103 |
@{term even} are multiples of two. |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
104 |
*} |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
105 |
|
11705 | 106 |
lemma even_imp_dvd: "n \<in> even \<Longrightarrow> 2 dvd n" |
23842
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
107 |
txt {* |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
108 |
We begin by applying induction. Note that @{text even.induct} has the form |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
109 |
of an elimination rule, so we use the method @{text erule}. We get two |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
110 |
subgoals: |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
111 |
*} |
10314 | 112 |
apply (erule even.induct) |
23842
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
113 |
txt {* |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
114 |
@{subgoals[display,indent=0]} |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
115 |
We unfold the definition of @{text dvd} in both subgoals, proving the first |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
116 |
one and simplifying the second: |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
117 |
*} |
10883 | 118 |
apply (simp_all add: dvd_def) |
23842
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
119 |
txt {* |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
120 |
@{subgoals[display,indent=0]} |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
121 |
The next command eliminates the existential quantifier from the assumption |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
122 |
and replaces @{text n} by @{text "2 * k"}. |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
123 |
*} |
10314 | 124 |
apply clarify |
23842
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
125 |
txt {* |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
126 |
@{subgoals[display,indent=0]} |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
127 |
To conclude, we tell Isabelle that the desired value is |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
128 |
@{term "Suc k"}. With this hint, the subgoal falls to @{text simp}. |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
129 |
*} |
10883 | 130 |
apply (rule_tac x = "Suc k" in exI, simp) |
23842
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
131 |
(*<*)done(*>*) |
10314 | 132 |
|
23842
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
133 |
text {* |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
134 |
Combining the previous two results yields our objective, the |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
135 |
equivalence relating @{term even} and @{text dvd}. |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
136 |
% |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
137 |
%we don't want [iff]: discuss? |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
138 |
*} |
10314 | 139 |
|
11705 | 140 |
theorem even_iff_dvd: "(n \<in> even) = (2 dvd n)" |
10883 | 141 |
by (blast intro: dvd_imp_even even_imp_dvd) |
10314 | 142 |
|
23842
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
143 |
|
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
144 |
subsection{* Generalization and Rule Induction \label{sec:gen-rule-induction} *} |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
145 |
|
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
146 |
text {* |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
147 |
\index{generalizing for induction}% |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
148 |
Before applying induction, we typically must generalize |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
149 |
the induction formula. With rule induction, the required generalization |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
150 |
can be hard to find and sometimes requires a complete reformulation of the |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
151 |
problem. In this example, our first attempt uses the obvious statement of |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
152 |
the result. It fails: |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
153 |
*} |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
154 |
|
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
155 |
lemma "Suc (Suc n) \<in> even \<Longrightarrow> n \<in> even" |
10314 | 156 |
apply (erule even.induct) |
157 |
oops |
|
23842
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
158 |
(*<*) |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
159 |
lemma "Suc (Suc n) \<in> even \<Longrightarrow> n \<in> even" |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
160 |
apply (erule even.induct) |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
161 |
(*>*) |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
162 |
txt {* |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
163 |
Rule induction finds no occurrences of @{term "Suc(Suc n)"} in the |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
164 |
conclusion, which it therefore leaves unchanged. (Look at |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
165 |
@{text even.induct} to see why this happens.) We have these subgoals: |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
166 |
@{subgoals[display,indent=0]} |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
167 |
The first one is hopeless. Rule induction on |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
168 |
a non-variable term discards information, and usually fails. |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
169 |
How to deal with such situations |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
170 |
in general is described in {\S}\ref{sec:ind-var-in-prems} below. |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
171 |
In the current case the solution is easy because |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
172 |
we have the necessary inverse, subtraction: |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
173 |
*} |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
174 |
(*<*)oops(*>*) |
11705 | 175 |
lemma even_imp_even_minus_2: "n \<in> even \<Longrightarrow> n - 2 \<in> even" |
10314 | 176 |
apply (erule even.induct) |
23842
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
177 |
apply auto |
10314 | 178 |
done |
23842
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
179 |
(*<*) |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
180 |
lemma "n \<in> even \<Longrightarrow> n - 2 \<in> even" |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
181 |
apply (erule even.induct) |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
182 |
(*>*) |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
183 |
txt {* |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
184 |
This lemma is trivially inductive. Here are the subgoals: |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
185 |
@{subgoals[display,indent=0]} |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
186 |
The first is trivial because @{text "0 - 2"} simplifies to @{text 0}, which is |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
187 |
even. The second is trivial too: @{term "Suc (Suc n) - 2"} simplifies to |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
188 |
@{term n}, matching the assumption.% |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
189 |
\index{rule induction|)} %the sequel isn't really about induction |
10314 | 190 |
|
23842
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
191 |
\medskip |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
192 |
Using our lemma, we can easily prove the result we originally wanted: |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
193 |
*} |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
194 |
(*<*)oops(*>*) |
10883 | 195 |
lemma Suc_Suc_even_imp_even: "Suc (Suc n) \<in> even \<Longrightarrow> n \<in> even" |
196 |
by (drule even_imp_even_minus_2, simp) |
|
10326 | 197 |
|
23842
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
198 |
text {* |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
199 |
We have just proved the converse of the introduction rule @{text even.step}. |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
200 |
This suggests proving the following equivalence. We give it the |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
201 |
\attrdx{iff} attribute because of its obvious value for simplification. |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
202 |
*} |
10326 | 203 |
|
204 |
lemma [iff]: "((Suc (Suc n)) \<in> even) = (n \<in> even)" |
|
10883 | 205 |
by (blast dest: Suc_Suc_even_imp_even) |
10314 | 206 |
|
23842
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
207 |
|
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
208 |
subsection{* Rule Inversion \label{sec:rule-inversion} *} |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
209 |
|
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
210 |
text {* |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
211 |
\index{rule inversion|(}% |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
212 |
Case analysis on an inductive definition is called \textbf{rule |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
213 |
inversion}. It is frequently used in proofs about operational |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
214 |
semantics. It can be highly effective when it is applied |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
215 |
automatically. Let us look at how rule inversion is done in |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
216 |
Isabelle/HOL\@. |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
217 |
|
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
218 |
Recall that @{term even} is the minimal set closed under these two rules: |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
219 |
@{thm [display,indent=0] even.intros [no_vars]} |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
220 |
Minimality means that @{term even} contains only the elements that these |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
221 |
rules force it to contain. If we are told that @{term a} |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
222 |
belongs to |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
223 |
@{term even} then there are only two possibilities. Either @{term a} is @{text 0} |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
224 |
or else @{term a} has the form @{term "Suc(Suc n)"}, for some suitable @{term n} |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
225 |
that belongs to |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
226 |
@{term even}. That is the gist of the @{term cases} rule, which Isabelle proves |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
227 |
for us when it accepts an inductive definition: |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
228 |
@{named_thms [display,indent=0,margin=40] even.cases [no_vars] (even.cases)} |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
229 |
This general rule is less useful than instances of it for |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
230 |
specific patterns. For example, if @{term a} has the form |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
231 |
@{term "Suc(Suc n)"} then the first case becomes irrelevant, while the second |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
232 |
case tells us that @{term n} belongs to @{term even}. Isabelle will generate |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
233 |
this instance for us: |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
234 |
*} |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
235 |
|
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
236 |
inductive_cases Suc_Suc_cases [elim!]: "Suc(Suc n) \<in> even" |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
237 |
|
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
238 |
text {* |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
239 |
The \commdx{inductive\protect\_cases} command generates an instance of |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
240 |
the @{text cases} rule for the supplied pattern and gives it the supplied name: |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
241 |
@{named_thms [display,indent=0] Suc_Suc_cases [no_vars] (Suc_Suc_cases)} |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
242 |
Applying this as an elimination rule yields one case where @{text even.cases} |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
243 |
would yield two. Rule inversion works well when the conclusions of the |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
244 |
introduction rules involve datatype constructors like @{term Suc} and @{text "#"} |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
245 |
(list ``cons''); freeness reasoning discards all but one or two cases. |
10314 | 246 |
|
23842
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
247 |
In the \isacommand{inductive\_cases} command we supplied an |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
248 |
attribute, @{text "elim!"}, |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
249 |
\index{elim"!@\isa {elim"!} (attribute)}% |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
250 |
indicating that this elimination rule can be |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
251 |
applied aggressively. The original |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
252 |
@{term cases} rule would loop if used in that manner because the |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
253 |
pattern~@{term a} matches everything. |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
254 |
|
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
255 |
The rule @{text Suc_Suc_cases} is equivalent to the following implication: |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
256 |
@{term [display,indent=0] "Suc (Suc n) \<in> even \<Longrightarrow> n \<in> even"} |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
257 |
Just above we devoted some effort to reaching precisely |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
258 |
this result. Yet we could have obtained it by a one-line declaration, |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
259 |
dispensing with the lemma @{text even_imp_even_minus_2}. |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
260 |
This example also justifies the terminology |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
261 |
\textbf{rule inversion}: the new rule inverts the introduction rule |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
262 |
@{text even.step}. In general, a rule can be inverted when the set of elements |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
263 |
it introduces is disjoint from those of the other introduction rules. |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
264 |
|
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
265 |
For one-off applications of rule inversion, use the \methdx{ind_cases} method. |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
266 |
Here is an example: |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
267 |
*} |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
268 |
|
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
269 |
(*<*)lemma "Suc(Suc n) \<in> even \<Longrightarrow> P"(*>*) |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
270 |
apply (ind_cases "Suc(Suc n) \<in> even") |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
271 |
(*<*)oops(*>*) |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
272 |
|
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
273 |
text {* |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
274 |
The specified instance of the @{text cases} rule is generated, then applied |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
275 |
as an elimination rule. |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
276 |
|
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
277 |
To summarize, every inductive definition produces a @{text cases} rule. The |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
278 |
\commdx{inductive\protect\_cases} command stores an instance of the |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
279 |
@{text cases} rule for a given pattern. Within a proof, the |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
280 |
@{text ind_cases} method applies an instance of the @{text cases} |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
281 |
rule. |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
282 |
|
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
283 |
The even numbers example has shown how inductive definitions can be |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
284 |
used. Later examples will show that they are actually worth using.% |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
285 |
\index{rule inversion|)}% |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
286 |
\index{even numbers!defining inductively|)} |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
287 |
*} |
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
288 |
|
9d87177f1f89
LaTeX code is now generated directly from theory file.
berghofe
parents:
23733
diff
changeset
|
289 |
(*<*)end(*>*) |