32
+ − 1
theory Ind_Examples
88
+ − 2
imports Main LaTeXsugar
32
+ − 3
begin
+ − 4
88
+ − 5
section{* Examples of Inductive Definitions \label{sec:ind-examples} *}
32
+ − 6
+ − 7
text {*
88
+ − 8
+ − 9
In this section, we will give three examples showing how to define inductive
+ − 10
predicates by hand and prove characteristic properties such as introduction
+ − 11
rules and an induction rule. From these examples, we will then figure out a
+ − 12
general method for defining inductive predicates. It should be noted that
+ − 13
our aim in this section is not to write proofs that are as beautiful as
+ − 14
possible, but as close as possible to the ML code producing the proofs that
+ − 15
we will develop later. As a first example, we consider the transitive
+ − 16
closure of a relation @{text R}. It is an inductive predicate characterized
+ − 17
by the two introduction rules
+ − 18
+ − 19
\begin{center}
+ − 20
@{term[mode=Axiom] "trcl R x x"} \hspace{5mm}
+ − 21
@{term[mode=Rule] "R x y \<Longrightarrow> trcl R y z \<Longrightarrow> trcl R x z"}
+ − 22
\end{center}
+ − 23
+ − 24
(FIXME first rule should be an ``axiom'')
+ − 25
+ − 26
Note that the @{text trcl} predicate has two different kinds of parameters: the
+ − 27
first parameter @{text R} stays \emph{fixed} throughout the definition, whereas
+ − 28
the second and third parameter changes in the ``recursive call''.
+ − 29
+ − 30
Since an inductively defined predicate is the least predicate closed under
+ − 31
a collection of introduction rules, we define the predicate @{text "trcl R x y"} in
+ − 32
such a way that it holds if and only if @{text "P x y"} holds for every predicate
+ − 33
@{text P} closed under the rules above. This gives rise to the definition
32
+ − 34
*}
+ − 35
88
+ − 36
definition "trcl R x y \<equiv>
+ − 37
\<forall>P. (\<forall>x. P x x) \<longrightarrow> (\<forall>x y z. R x y \<longrightarrow> P y z \<longrightarrow> P x z) \<longrightarrow> P x y"
32
+ − 38
+ − 39
text {*
88
+ − 40
where we quantify over the predicate @{text P}.
32
+ − 41
88
+ − 42
Since the predicate @{term "trcl R x y"} yields an element of the type of object
+ − 43
level truth values @{text bool}, the meta-level implications @{text "\<Longrightarrow>"} in the
+ − 44
above introduction rules have to be converted to object-level implications
+ − 45
@{text "\<longrightarrow>"}. Moreover, we use object-level universal quantifiers @{text "\<forall>"}
+ − 46
rather than meta-level universal quantifiers @{text "\<And>"} for quantifying over
+ − 47
the variable parameters of the introduction rules. Isabelle already offers some
+ − 48
infrastructure for converting between meta-level and object-level connectives,
+ − 49
which we will use later on.
+ − 50
+ − 51
With this definition of the transitive closure, the proof of the (weak) induction
+ − 52
theorem is almost immediate. It suffices to convert all the meta-level connectives
+ − 53
in the induction rule to object-level connectives using the @{text atomize} proof
+ − 54
method, expand the definition of @{text trcl}, eliminate the universal quantifier
+ − 55
contained in it, and then solve the goal by assumption.
32
+ − 56
*}
+ − 57
+ − 58
lemma trcl_induct:
+ − 59
assumes trcl: "trcl R x y"
+ − 60
shows "(\<And>x. P x x) \<Longrightarrow> (\<And>x y z. R x y \<Longrightarrow> P y z \<Longrightarrow> P x z) \<Longrightarrow> P x y"
+ − 61
apply (atomize (full))
+ − 62
apply (cut_tac trcl)
+ − 63
apply (unfold trcl_def)
+ − 64
apply (drule spec [where x=P])
+ − 65
apply assumption
+ − 66
done
+ − 67
(*<*)
+ − 68
lemma "trcl R x x"
+ − 69
apply (unfold trcl_def)
+ − 70
apply (rule allI impI)+
+ − 71
(*>*)
88
+ − 72
32
+ − 73
txt {*
88
+ − 74
+ − 75
The above induction rule is \emph{weak} in the sense that the induction step may
+ − 76
only be proved using the assumptions @{term "R x y"} and @{term "P y z"}, but not
+ − 77
using the additional assumption \mbox{@{term "trcl R y z"}}. A stronger induction rule
+ − 78
containing this additional assumption can be derived from the weaker one with the
+ − 79
help of the introduction rules for @{text trcl}.
32
+ − 80
88
+ − 81
We now turn to the proofs of the introduction rules, which are slightly more complicated.
+ − 82
In order to prove the first introduction rule, we again unfold the definition and
+ − 83
then apply the introdution rules for @{text "\<forall>"} and @{text "\<longrightarrow>"} as often as possible.
+ − 84
We then end up in a proof state of the following form:
+ − 85
+ − 86
@{subgoals [display]}
+ − 87
+ − 88
The two assumptions correspond to the introduction rules, where @{term "trcl R"} has been
+ − 89
replaced by @{term "P"}. Thus, all we have to do is to eliminate the universal quantifier
+ − 90
in front of the first assumption, and then solve the goal by assumption:
+ − 91
*}
32
+ − 92
(*<*)oops(*>*)
88
+ − 93
32
+ − 94
lemma trcl_base: "trcl R x x"
+ − 95
apply (unfold trcl_def)
+ − 96
apply (rule allI impI)+
+ − 97
apply (drule spec)
+ − 98
apply assumption
+ − 99
done
88
+ − 100
32
+ − 101
(*<*)
+ − 102
lemma "R x y \<Longrightarrow> trcl R y z \<Longrightarrow> trcl R x z"
+ − 103
apply (unfold trcl_def)
+ − 104
apply (rule allI impI)+
+ − 105
(*>*)
88
+ − 106
32
+ − 107
txt {*
88
+ − 108
+ − 109
Since the second introduction rule has premises, its proof is not as easy as the previous
+ − 110
one. After unfolding the definitions and applying the introduction rules for @{text "\<forall>"}
+ − 111
and @{text "\<longrightarrow>"}, we get the proof state
+ − 112
@{subgoals [display]}
+ − 113
The third and fourth assumption corresponds to the first and second introduction rule,
+ − 114
respectively, whereas the first and second assumption corresponds to the premises of
+ − 115
the introduction rule. Since we want to prove the second introduction rule, we apply
+ − 116
the fourth assumption to the goal @{term "P x z"}. In order for the assumption to
+ − 117
be applicable, we have to eliminate the universal quantifiers and turn the object-level
+ − 118
implications into meta-level ones. This can be accomplished using the @{text rule_format}
+ − 119
attribute. Applying the assumption produces two new subgoals, which can be solved using
+ − 120
the first and second assumption. The second assumption again involves a quantifier and
+ − 121
implications that have to be eliminated before it can be applied. To avoid problems
+ − 122
with higher order unification, it is advisable to provide an instantiation for the
+ − 123
universally quantified predicate variable in the assumption.
32
+ − 124
*}
+ − 125
(*<*)oops(*>*)
88
+ − 126
32
+ − 127
lemma trcl_step: "R x y \<Longrightarrow> trcl R y z \<Longrightarrow> trcl R x z"
+ − 128
apply (unfold trcl_def)
+ − 129
apply (rule allI impI)+
+ − 130
proof -
102
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 131
case (goal1 P)
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 132
have g1: "R x y" by fact
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 133
have g2: "\<forall>P. (\<forall>x. P x x) \<longrightarrow> (\<forall>x y z. R x y \<longrightarrow> P y z \<longrightarrow> P x z) \<longrightarrow> P y z" by fact
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 134
have g3: "\<forall>x. P x x" by fact
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 135
have g4: "\<forall>x y z. R x y \<longrightarrow> P y z \<longrightarrow> P x z" by fact
32
+ − 136
show ?case
102
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 137
apply (rule g4 [rule_format])
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 138
apply (rule g1)
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 139
apply (rule g2 [THEN spec [where x=P], THEN mp, THEN mp, OF g3, OF g4])
32
+ − 140
done
+ − 141
qed
+ − 142
+ − 143
text {*
88
+ − 144
+ − 145
This method of defining inductive predicates easily generalizes to mutually inductive
+ − 146
predicates, like the predicates @{text even} and @{text odd} characterized by the
+ − 147
following introduction rules:
+ − 148
+ − 149
\begin{center}
+ − 150
@{term[mode=Axiom] "even (0::nat)"} \hspace{5mm}
+ − 151
@{term[mode=Rule] "odd m \<Longrightarrow> even (Suc m)"} \hspace{5mm}
+ − 152
@{term[mode=Rule] "even m \<Longrightarrow> odd (Suc m)"}
+ − 153
\end{center}
+ − 154
+ − 155
Since the predicates are mutually inductive, each of the definitions contain two
+ − 156
quantifiers over the predicates @{text P} and @{text Q}.
32
+ − 157
*}
+ − 158
88
+ − 159
definition "even n \<equiv>
+ − 160
\<forall>P Q. P 0 \<longrightarrow> (\<forall>m. Q m \<longrightarrow> P (Suc m))
+ − 161
\<longrightarrow> (\<forall>m. P m \<longrightarrow> Q (Suc m)) \<longrightarrow> P n"
32
+ − 162
88
+ − 163
definition "odd n \<equiv>
+ − 164
\<forall>P Q. P 0 \<longrightarrow> (\<forall>m. Q m \<longrightarrow> P (Suc m))
+ − 165
\<longrightarrow> (\<forall>m. P m \<longrightarrow> Q (Suc m)) \<longrightarrow> Q n"
32
+ − 166
+ − 167
text {*
88
+ − 168
For proving the induction rule, we use exactly the same technique as in the transitive
+ − 169
closure example:
32
+ − 170
*}
+ − 171
+ − 172
lemma even_induct:
+ − 173
assumes even: "even n"
38
+ − 174
shows "P 0 \<Longrightarrow>
+ − 175
(\<And>m. Q m \<Longrightarrow> P (Suc m)) \<Longrightarrow> (\<And>m. P m \<Longrightarrow> Q (Suc m)) \<Longrightarrow> P n"
32
+ − 176
apply (atomize (full))
+ − 177
apply (cut_tac even)
+ − 178
apply (unfold even_def)
+ − 179
apply (drule spec [where x=P])
+ − 180
apply (drule spec [where x=Q])
+ − 181
apply assumption
+ − 182
done
+ − 183
+ − 184
text {*
88
+ − 185
A similar induction rule having @{term "Q n"} as a conclusion can be proved for
+ − 186
the @{text odd} predicate.
+ − 187
The proofs of the introduction rules are also very similar to the ones in the
+ − 188
previous example. We only show the proof of the second introduction rule,
+ − 189
since it is almost the same as the one for the third introduction rule,
+ − 190
and the proof of the first rule is trivial.
32
+ − 191
*}
+ − 192
+ − 193
lemma evenS: "odd m \<Longrightarrow> even (Suc m)"
+ − 194
apply (unfold odd_def even_def)
+ − 195
apply (rule allI impI)+
+ − 196
proof -
+ − 197
case goal1
+ − 198
show ?case
+ − 199
apply (rule goal1(3) [rule_format])
+ − 200
apply (rule goal1(1) [THEN spec [where x=P], THEN spec [where x=Q],
+ − 201
THEN mp, THEN mp, THEN mp, OF goal1(2-4)])
+ − 202
done
+ − 203
qed
+ − 204
(*<*)
+ − 205
lemma even0: "even 0"
+ − 206
apply (unfold even_def)
+ − 207
apply (rule allI impI)+
+ − 208
apply assumption
+ − 209
done
+ − 210
+ − 211
lemma oddS: "even m \<Longrightarrow> odd (Suc m)"
+ − 212
apply (unfold odd_def even_def)
+ − 213
apply (rule allI impI)+
+ − 214
proof -
+ − 215
case goal1
+ − 216
show ?case
+ − 217
apply (rule goal1(4) [rule_format])
+ − 218
apply (rule goal1(1) [THEN spec [where x=P], THEN spec [where x=Q],
+ − 219
THEN mp, THEN mp, THEN mp, OF goal1(2-4)])
+ − 220
done
+ − 221
qed
+ − 222
(*>*)
88
+ − 223
32
+ − 224
text {*
88
+ − 225
As a final example, we will consider the definition of the accessible
+ − 226
part of a relation @{text R} characterized by the introduction rule
+ − 227
+ − 228
\begin{center}
+ − 229
@{term[mode=Rule] "(\<forall>y. R y x \<longrightarrow> accpart R y) \<Longrightarrow> accpart R x"}
+ − 230
\end{center}
+ − 231
+ − 232
whose premise involves a universal quantifier and an implication. The
+ − 233
definition of @{text accpart} is as follows:
32
+ − 234
*}
+ − 235
88
+ − 236
definition "accpart R x \<equiv> \<forall>P. (\<forall>x. (\<forall>y. R y x \<longrightarrow> P y) \<longrightarrow> P x) \<longrightarrow> P x"
32
+ − 237
+ − 238
text {*
88
+ − 239
The proof of the induction theorem is again straightforward:
32
+ − 240
*}
+ − 241
+ − 242
lemma accpart_induct:
+ − 243
assumes acc: "accpart R x"
+ − 244
shows "(\<And>x. (\<And>y. R y x \<Longrightarrow> P y) \<Longrightarrow> P x) \<Longrightarrow> P x"
+ − 245
apply (atomize (full))
+ − 246
apply (cut_tac acc)
+ − 247
apply (unfold accpart_def)
+ − 248
apply (drule spec [where x=P])
+ − 249
apply assumption
+ − 250
done
+ − 251
(*<*)
+ − 252
lemma accpartI: "(\<And>y. R y x \<Longrightarrow> accpart R y) \<Longrightarrow> accpart R x"
+ − 253
apply (unfold accpart_def)
+ − 254
apply (rule allI impI)+
+ − 255
(*>*)
88
+ − 256
32
+ − 257
txt {*
88
+ − 258
+ − 259
Proving the introduction rule is a little more complicated, due to the quantifier
+ − 260
and the implication in the premise. We first convert the meta-level universal quantifier
+ − 261
and implication to their object-level counterparts. Unfolding the definition of
+ − 262
@{text accpart} and applying the introduction rules for @{text "\<forall>"} and @{text "\<longrightarrow>"}
+ − 263
yields the following proof state:
+ − 264
@{subgoals [display]}
+ − 265
Applying the second assumption produces a proof state with the new local assumption
+ − 266
@{term "R y x"}, which will then be used to solve the goal @{term "P y"} using the
+ − 267
first assumption.
32
+ − 268
*}
+ − 269
(*<*)oops(*>*)
88
+ − 270
32
+ − 271
lemma accpartI: "(\<And>y. R y x \<Longrightarrow> accpart R y) \<Longrightarrow> accpart R x"
+ − 272
apply (unfold accpart_def)
+ − 273
apply (rule allI impI)+
+ − 274
proof -
+ − 275
case goal1
+ − 276
note goal1' = this
+ − 277
show ?case
+ − 278
apply (rule goal1'(2) [rule_format])
+ − 279
proof -
+ − 280
case goal1
+ − 281
show ?case
+ − 282
apply (rule goal1'(1) [OF goal1, THEN spec [where x=P],
+ − 283
THEN mp, OF goal1'(2)])
+ − 284
done
+ − 285
qed
+ − 286
qed
+ − 287
+ − 288
end