127
+ − 1
theory Ind_Prelims
346
+ − 2
imports Ind_Intro
32
+ − 3
begin
+ − 4
127
+ − 5
section{* Preliminaries *}
+ − 6
32
+ − 7
text {*
218
7ff7325e3b4e
started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 8
The user will just give a specification of inductive predicate(s) and
129
+ − 9
expects from the package to produce a convenient reasoning
+ − 10
infrastructure. This infrastructure needs to be derived from the definition
218
7ff7325e3b4e
started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 11
that correspond to the specified predicate(s). Before we start with
295
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 12
explaining all parts of the package, let us first give some examples showing
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 13
how to define inductive predicates and then also how to generate a reasoning
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 14
infrastructure for them. From the examples we will figure out a general
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 15
method for defining inductive predicates. This is usually the first step in
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 16
writing a package for Isabelle. The aim in this section is \emph{not} to
219
+ − 17
write proofs that are as beautiful as possible, but as close as possible to
244
+ − 18
the ML-implementation we will develop in later sections.
219
+ − 19
295
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 20
219
+ − 21
We first consider the transitive closure of a relation @{text R}. The
+ − 22
``pencil-and-paper'' specification for the transitive closure is:
+ − 23
127
+ − 24
\begin{center}\small
116
+ − 25
@{prop[mode=Axiom] "trcl R x x"} \hspace{5mm}
+ − 26
@{prop[mode=Rule] "R x y \<Longrightarrow> trcl R y z \<Longrightarrow> trcl R x z"}
88
+ − 27
\end{center}
+ − 28
295
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 29
As mentioned before, the package has to make an appropriate definition for
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 30
@{term "trcl"}. Since an inductively defined predicate is the least
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 31
predicate closed under a collection of introduction rules, the predicate
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 32
@{text "trcl R x y"} can be defined so that it holds if and only if @{text
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 33
"P x y"} holds for every predicate @{text P} closed under the rules
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 34
above. This gives rise to the definition
32
+ − 35
*}
+ − 36
127
+ − 37
definition "trcl \<equiv>
+ − 38
\<lambda>R x y. \<forall>P. (\<forall>x. P x x)
+ − 39
\<longrightarrow> (\<forall>x y z. R x y \<longrightarrow> P y z \<longrightarrow> P x z) \<longrightarrow> P x y"
32
+ − 40
+ − 41
text {*
295
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 42
Note we have to use the object implication @{text "\<longrightarrow>"} and object
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 43
quantification @{text "\<forall>"} for stating this definition (there is no other
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 44
way for definitions in HOL). However, the introduction rules and induction
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 45
principles associated with the transitive closure should use the
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 46
meta-connectives, since they simplify the reasoning for the user.
218
7ff7325e3b4e
started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 47
32
+ − 48
127
+ − 49
With this definition, the proof of the induction principle for @{term trcl}
129
+ − 50
is almost immediate. It suffices to convert all the meta-level
116
+ − 51
connectives in the lemma to object-level connectives using the
219
+ − 52
proof method @{text atomize} (Line 4 below), expand the definition of @{term trcl}
127
+ − 53
(Line 5 and 6), eliminate the universal quantifier contained in it (Line~7),
218
7ff7325e3b4e
started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 54
and then solve the goal by @{text assumption} (Line 8).
113
+ − 55
32
+ − 56
*}
+ − 57
114
+ − 58
lemma %linenos trcl_induct:
218
7ff7325e3b4e
started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 59
assumes "trcl R x y"
7ff7325e3b4e
started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 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"
113
+ − 61
apply(atomize (full))
165
+ − 62
apply(cut_tac prems)
113
+ − 63
apply(unfold trcl_def)
+ − 64
apply(drule spec[where x=P])
+ − 65
apply(assumption)
+ − 66
done
88
+ − 67
113
+ − 68
text {*
127
+ − 69
The proofs for the introduction rules are slightly more complicated.
+ − 70
For the first one, we need to prove the following lemma:
+ − 71
*}
+ − 72
+ − 73
lemma %linenos trcl_base:
218
7ff7325e3b4e
started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 74
shows "trcl R x x"
127
+ − 75
apply(unfold trcl_def)
+ − 76
apply(rule allI impI)+
+ − 77
apply(drule spec)
+ − 78
apply(assumption)
+ − 79
done
+ − 80
+ − 81
text {*
+ − 82
We again unfold first the definition and apply introduction rules
+ − 83
for @{text "\<forall>"} and @{text "\<longrightarrow>"} as often as possible (Lines 3 and 4).
116
+ − 84
We then end up in the goal state:
113
+ − 85
*}
88
+ − 86
113
+ − 87
(*<*)lemma "trcl R x x"
115
+ − 88
apply (unfold trcl_def)
+ − 89
apply (rule allI impI)+(*>*)
113
+ − 90
txt {* @{subgoals [display]} *}
32
+ − 91
(*<*)oops(*>*)
88
+ − 92
113
+ − 93
text {*
218
7ff7325e3b4e
started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 94
The two assumptions come from the definition of @{term trcl} and correspond
7ff7325e3b4e
started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 95
to the introduction rules. Thus, all we have to do is to eliminate the
7ff7325e3b4e
started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 96
universal quantifier in front of the first assumption (Line 5), and then
219
+ − 97
solve the goal by @{text assumption} (Line 6).
113
+ − 98
*}
+ − 99
+ − 100
text {*
127
+ − 101
Next we have to show that the second introduction rule also follows from the
+ − 102
definition. Since this rule has premises, the proof is a bit more
+ − 103
involved. After unfolding the definitions and applying the introduction
+ − 104
rules for @{text "\<forall>"} and @{text "\<longrightarrow>"}
113
+ − 105
*}
+ − 106
127
+ − 107
lemma trcl_step:
218
7ff7325e3b4e
started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 108
shows "R x y \<Longrightarrow> trcl R y z \<Longrightarrow> trcl R x z"
115
+ − 109
apply (unfold trcl_def)
127
+ − 110
apply (rule allI impI)+
113
+ − 111
127
+ − 112
txt {*
+ − 113
we obtain the goal state
+ − 114
+ − 115
@{subgoals [display]}
+ − 116
+ − 117
To see better where we are, let us explicitly name the assumptions
+ − 118
by starting a subproof.
115
+ − 119
*}
+ − 120
+ − 121
proof -
+ − 122
case (goal1 P)
127
+ − 123
have p1: "R x y" by fact
+ − 124
have p2: "\<forall>P. (\<forall>x. P x x)
+ − 125
\<longrightarrow> (\<forall>x y z. R x y \<longrightarrow> P y z \<longrightarrow> P x z) \<longrightarrow> P y z" by fact
+ − 126
have r1: "\<forall>x. P x x" by fact
+ − 127
have r2: "\<forall>x y z. R x y \<longrightarrow> P y z \<longrightarrow> P x z" by fact
+ − 128
show "P x z"
+ − 129
+ − 130
txt {*
+ − 131
The assumptions @{text "p1"} and @{text "p2"} correspond to the premises of
218
7ff7325e3b4e
started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 132
the second introduction rule (unfolded); the assumptions @{text "r1"} and @{text "r2"}
219
+ − 133
come from the definition of @{term trcl}. We apply @{text "r2"} to the goal
+ − 134
@{term "P x z"}. In order for this assumption to be applicable as a rule, we
127
+ − 135
have to eliminate the universal quantifier and turn the object-level
+ − 136
implications into meta-level ones. This can be accomplished using the @{text
+ − 137
rule_format} attribute. So we continue the proof with:
115
+ − 138
32
+ − 139
*}
113
+ − 140
127
+ − 141
apply (rule r2[rule_format])
88
+ − 142
127
+ − 143
txt {*
+ − 144
This gives us two new subgoals
+ − 145
+ − 146
@{subgoals [display]}
+ − 147
+ − 148
which can be solved using assumptions @{text p1} and @{text p2}. The latter
+ − 149
involves a quantifier and implications that have to be eliminated before it
+ − 150
can be applied. To avoid potential problems with higher-order unification,
+ − 151
we explicitly instantiate the quantifier to @{text "P"} and also match
+ − 152
explicitly the implications with @{text "r1"} and @{text "r2"}. This gives
+ − 153
the proof:
+ − 154
*}
+ − 155
+ − 156
apply(rule p1)
+ − 157
apply(rule p2[THEN spec[where x=P], THEN mp, THEN mp, OF r1, OF r2])
115
+ − 158
done
+ − 159
qed
32
+ − 160
+ − 161
text {*
127
+ − 162
Now we are done. It might be surprising that we are not using the automatic
244
+ − 163
tactics available in Isabelle/HOL for proving this lemmas. After all @{text
127
+ − 164
"blast"} would easily dispense of it.
115
+ − 165
*}
88
+ − 166
127
+ − 167
lemma trcl_step_blast:
218
7ff7325e3b4e
started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 168
shows "R x y \<Longrightarrow> trcl R y z \<Longrightarrow> trcl R x z"
115
+ − 169
apply(unfold trcl_def)
+ − 170
apply(blast)
+ − 171
done
+ − 172
+ − 173
text {*
116
+ − 174
Experience has shown that it is generally a bad idea to rely heavily on
+ − 175
@{text blast}, @{text auto} and the like in automated proofs. The reason is
+ − 176
that you do not have precise control over them (the user can, for example,
+ − 177
declare new intro- or simplification rules that can throw automatic tactics
+ − 178
off course) and also it is very hard to debug proofs involving automatic
+ − 179
tactics whenever something goes wrong. Therefore if possible, automatic
295
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 180
tactics in packages should be avoided or be constrained sufficiently.
115
+ − 181
116
+ − 182
The method of defining inductive predicates by impredicative quantification
+ − 183
also generalises to mutually inductive predicates. The next example defines
219
+ − 184
the predicates @{text even} and @{text odd} given by
32
+ − 185
219
+ − 186
\begin{center}\small
+ − 187
@{prop[mode=Axiom] "even (0::nat)"} \hspace{5mm}
+ − 188
@{prop[mode=Rule] "odd n \<Longrightarrow> even (Suc n)"} \hspace{5mm}
+ − 189
@{prop[mode=Rule] "even n \<Longrightarrow> odd (Suc n)"}
+ − 190
\end{center}
32
+ − 191
127
+ − 192
Since the predicates @{term even} and @{term odd} are mutually inductive, each
+ − 193
corresponding definition must quantify over both predicates (we name them
+ − 194
below @{text "P"} and @{text "Q"}).
+ − 195
*}
+ − 196
219
+ − 197
definition "even \<equiv>
127
+ − 198
\<lambda>n. \<forall>P Q. P 0 \<longrightarrow> (\<forall>m. Q m \<longrightarrow> P (Suc m))
+ − 199
\<longrightarrow> (\<forall>m. P m \<longrightarrow> Q (Suc m)) \<longrightarrow> P n"
+ − 200
219
+ − 201
definition "odd \<equiv>
127
+ − 202
\<lambda>n. \<forall>P Q. P 0 \<longrightarrow> (\<forall>m. Q m \<longrightarrow> P (Suc m))
+ − 203
\<longrightarrow> (\<forall>m. P m \<longrightarrow> Q (Suc m)) \<longrightarrow> Q n"
32
+ − 204
+ − 205
text {*
116
+ − 206
For proving the induction principles, we use exactly the same technique
+ − 207
as in the transitive closure example, namely:
32
+ − 208
*}
+ − 209
+ − 210
lemma even_induct:
218
7ff7325e3b4e
started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 211
assumes "even n"
7ff7325e3b4e
started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 212
shows "P 0 \<Longrightarrow> (\<And>m. Q m \<Longrightarrow> P (Suc m)) \<Longrightarrow> (\<And>m. P m \<Longrightarrow> Q (Suc m)) \<Longrightarrow> P n"
115
+ − 213
apply(atomize (full))
165
+ − 214
apply(cut_tac prems)
115
+ − 215
apply(unfold even_def)
+ − 216
apply(drule spec[where x=P])
+ − 217
apply(drule spec[where x=Q])
+ − 218
apply(assumption)
+ − 219
done
32
+ − 220
+ − 221
text {*
127
+ − 222
The only difference with the proof @{text "trcl_induct"} is that we have to
+ − 223
instantiate here two universal quantifiers. We omit the other induction
218
7ff7325e3b4e
started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 224
principle that has @{prop "even n"} as premise and @{term "Q n"} as conclusion.
7ff7325e3b4e
started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 225
The proofs of the introduction rules are also very similar to the ones in
7ff7325e3b4e
started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 226
the @{text "trcl"}-example. We only show the proof of the second introduction
7ff7325e3b4e
started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 227
rule.
32
+ − 228
*}
+ − 229
127
+ − 230
lemma %linenos evenS:
218
7ff7325e3b4e
started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 231
shows "odd m \<Longrightarrow> even (Suc m)"
115
+ − 232
apply (unfold odd_def even_def)
+ − 233
apply (rule allI impI)+
+ − 234
proof -
165
+ − 235
case (goal1 P Q)
127
+ − 236
have p1: "\<forall>P Q. P 0 \<longrightarrow> (\<forall>m. Q m \<longrightarrow> P (Suc m))
115
+ − 237
\<longrightarrow> (\<forall>m. P m \<longrightarrow> Q (Suc m)) \<longrightarrow> Q m" by fact
127
+ − 238
have r1: "P 0" by fact
+ − 239
have r2: "\<forall>m. Q m \<longrightarrow> P (Suc m)" by fact
+ − 240
have r3: "\<forall>m. P m \<longrightarrow> Q (Suc m)" by fact
115
+ − 241
show "P (Suc m)"
127
+ − 242
apply(rule r2[rule_format])
+ − 243
apply(rule p1[THEN spec[where x=P], THEN spec[where x=Q],
+ − 244
THEN mp, THEN mp, THEN mp, OF r1, OF r2, OF r3])
115
+ − 245
done
+ − 246
qed
88
+ − 247
32
+ − 248
text {*
295
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 249
The interesting lines are 7 to 15. Again, the assumptions fall into two categories:
218
7ff7325e3b4e
started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 250
@{text p1} corresponds to the premise of the introduction rule; @{text "r1"}
7ff7325e3b4e
started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 251
to @{text "r3"} come from the definition of @{text "even"}.
127
+ − 252
In Line 13, we apply the assumption @{text "r2"} (since we prove the second
+ − 253
introduction rule). In Lines 14 and 15 we apply assumption @{text "p1"} (if
+ − 254
the second introduction rule had more premises we have to do that for all
+ − 255
of them). In order for this assumption to be applicable, the quantifiers
+ − 256
need to be instantiated and then also the implications need to be resolved
+ − 257
with the other rules.
+ − 258
219
+ − 259
Next we define the accessible part of a relation @{text R} given by
+ − 260
the single rule:
+ − 261
+ − 262
\begin{center}\small
+ − 263
\mbox{\inferrule{@{term "\<And>y. R y x \<Longrightarrow> accpart R y"}}{@{term "accpart R x"}}}
+ − 264
\end{center}
+ − 265
295
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 266
The interesting point of this definition is that it contains a quantification
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 267
that ranges only over the premise and the premise has also a precondition.
219
+ − 268
The definition of @{text "accpart"} is:
32
+ − 269
*}
+ − 270
127
+ − 271
definition "accpart \<equiv> \<lambda>R x. \<forall>P. (\<forall>x. (\<forall>y. R y x \<longrightarrow> P y) \<longrightarrow> P x) \<longrightarrow> P x"
32
+ − 272
+ − 273
text {*
218
7ff7325e3b4e
started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 274
The proof of the induction principle is again straightforward and omitted.
7ff7325e3b4e
started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 275
Proving the introduction rule is a little more complicated, because the
7ff7325e3b4e
started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 276
quantifier and the implication in the premise. The proof is as follows.
115
+ − 277
*}
+ − 278
127
+ − 279
lemma %linenos accpartI:
218
7ff7325e3b4e
started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 280
shows "(\<And>y. R y x \<Longrightarrow> accpart R y) \<Longrightarrow> accpart R x"
115
+ − 281
apply (unfold accpart_def)
+ − 282
apply (rule allI impI)+
+ − 283
proof -
+ − 284
case (goal1 P)
127
+ − 285
have p1: "\<And>y. R y x \<Longrightarrow>
115
+ − 286
(\<forall>P. (\<forall>x. (\<forall>y. R y x \<longrightarrow> P y) \<longrightarrow> P x) \<longrightarrow> P y)" by fact
127
+ − 287
have r1: "\<forall>x. (\<forall>y. R y x \<longrightarrow> P y) \<longrightarrow> P x" by fact
115
+ − 288
show "P x"
127
+ − 289
apply(rule r1[rule_format])
115
+ − 290
proof -
+ − 291
case (goal1 y)
127
+ − 292
have r1_prem: "R y x" by fact
115
+ − 293
show "P y"
127
+ − 294
apply(rule p1[OF r1_prem, THEN spec[where x=P], THEN mp, OF r1])
115
+ − 295
done
+ − 296
qed
+ − 297
qed
+ − 298
+ − 299
text {*
295
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 300
As you can see, there are now two subproofs. The assumptions fall as usual into
219
+ − 301
two categories (Lines 7 to 9). In Line 11, applying the assumption @{text
+ − 302
"r1"} generates a goal state with the new local assumption @{term "R y x"},
+ − 303
named @{text "r1_prem"} in the second subproof (Line 14). This local assumption is
+ − 304
used to solve the goal @{term "P y"} with the help of assumption @{text
+ − 305
"p1"}.
+ − 306
124
+ − 307
219
+ − 308
\begin{exercise}
+ − 309
Give the definition for the freshness predicate for lambda-terms. The rules
+ − 310
for this predicate are:
+ − 311
+ − 312
\begin{center}\small
+ − 313
@{prop[mode=Rule] "a\<noteq>b \<Longrightarrow> fresh a (Var b)"}\hspace{5mm}
+ − 314
@{prop[mode=Rule] "\<lbrakk>fresh a t; fresh a s\<rbrakk> \<Longrightarrow> fresh a (App t s)"}\\[2mm]
+ − 315
@{prop[mode=Axiom] "fresh a (Lam a t)"}\hspace{5mm}
+ − 316
@{prop[mode=Rule] "\<lbrakk>a\<noteq>b; fresh a t\<rbrakk> \<Longrightarrow> fresh a (Lam b t)"}
+ − 317
\end{center}
+ − 318
+ − 319
From the definition derive the induction principle and the introduction
+ − 320
rules.
+ − 321
\end{exercise}
+ − 322
+ − 323
The point of all these examples is to get a feeling what the automatic
+ − 324
proofs should do in order to solve all inductive definitions we throw at
+ − 325
them. This is usually the first step in writing a package. We next explain
129
+ − 326
the parsing and typing part of the package.
127
+ − 327
115
+ − 328
*}
129
+ − 329
(*<*)end(*>*)