127
+ − 1
theory Ind_Prelims
346
+ − 2
imports Ind_Intro
32
+ − 3
begin
+ − 4
565
+ − 5
section\<open>Preliminaries\<close>
127
+ − 6
565
+ − 7
text \<open>
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
565
+ − 21
We first consider the transitive closure of a relation \<open>R\<close>. The
219
+ − 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
565
+ − 32
\<open>trcl R x y\<close> can be defined so that it holds if and only if \<open>P x y\<close> holds for every predicate \<open>P\<close> closed under the rules
295
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 33
above. This gives rise to the definition
565
+ − 34
\<close>
32
+ − 35
127
+ − 36
definition "trcl \<equiv>
+ − 37
\<lambda>R x y. \<forall>P. (\<forall>x. P x x)
+ − 38
\<longrightarrow> (\<forall>x y z. R x y \<longrightarrow> P y z \<longrightarrow> P x z) \<longrightarrow> P x y"
32
+ − 39
565
+ − 40
text \<open>
+ − 41
Note we have to use the object implication \<open>\<longrightarrow>\<close> and object
+ − 42
quantification \<open>\<forall>\<close> for stating this definition (there is no other
295
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 43
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
+ − 44
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
+ − 45
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
+ − 46
32
+ − 47
127
+ − 48
With this definition, the proof of the induction principle for @{term trcl}
129
+ − 49
is almost immediate. It suffices to convert all the meta-level
116
+ − 50
connectives in the lemma to object-level connectives using the
565
+ − 51
proof method \<open>atomize\<close> (Line 4 below), expand the definition of @{term trcl}
127
+ − 52
(Line 5 and 6), eliminate the universal quantifier contained in it (Line~7),
565
+ − 53
and then solve the goal by \<open>assumption\<close> (Line 8).
113
+ − 54
565
+ − 55
\<close>
32
+ − 56
114
+ − 57
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
+ − 58
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
+ − 59
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
+ − 60
apply(atomize (full))
513
+ − 61
apply(cut_tac assms)
113
+ − 62
apply(unfold trcl_def)
+ − 63
apply(drule spec[where x=P])
+ − 64
apply(assumption)
+ − 65
done
88
+ − 66
565
+ − 67
text \<open>
127
+ − 68
The proofs for the introduction rules are slightly more complicated.
+ − 69
For the first one, we need to prove the following lemma:
565
+ − 70
\<close>
127
+ − 71
+ − 72
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
+ − 73
shows "trcl R x x"
127
+ − 74
apply(unfold trcl_def)
+ − 75
apply(rule allI impI)+
+ − 76
apply(drule spec)
+ − 77
apply(assumption)
+ − 78
done
+ − 79
565
+ − 80
text \<open>
127
+ − 81
We again unfold first the definition and apply introduction rules
565
+ − 82
for \<open>\<forall>\<close> and \<open>\<longrightarrow>\<close> as often as possible (Lines 3 and 4).
116
+ − 83
We then end up in the goal state:
565
+ − 84
\<close>
88
+ − 85
113
+ − 86
(*<*)lemma "trcl R x x"
115
+ − 87
apply (unfold trcl_def)
+ − 88
apply (rule allI impI)+(*>*)
565
+ − 89
txt \<open>@{subgoals [display]}\<close>
32
+ − 90
(*<*)oops(*>*)
88
+ − 91
565
+ − 92
text \<open>
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
+ − 93
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
+ − 94
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
+ − 95
universal quantifier in front of the first assumption (Line 5), and then
565
+ − 96
solve the goal by \<open>assumption\<close> (Line 6).
+ − 97
\<close>
113
+ − 98
565
+ − 99
text \<open>
127
+ − 100
Next we have to show that the second introduction rule also follows from the
+ − 101
definition. Since this rule has premises, the proof is a bit more
+ − 102
involved. After unfolding the definitions and applying the introduction
565
+ − 103
rules for \<open>\<forall>\<close> and \<open>\<longrightarrow>\<close>
+ − 104
\<close>
113
+ − 105
127
+ − 106
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
+ − 107
shows "R x y \<Longrightarrow> trcl R y z \<Longrightarrow> trcl R x z"
115
+ − 108
apply (unfold trcl_def)
127
+ − 109
apply (rule allI impI)+
113
+ − 110
565
+ − 111
txt \<open>
127
+ − 112
we obtain the goal state
+ − 113
+ − 114
@{subgoals [display]}
+ − 115
+ − 116
To see better where we are, let us explicitly name the assumptions
+ − 117
by starting a subproof.
565
+ − 118
\<close>
115
+ − 119
+ − 120
proof -
+ − 121
case (goal1 P)
127
+ − 122
have p1: "R x y" by fact
+ − 123
have p2: "\<forall>P. (\<forall>x. P x x)
+ − 124
\<longrightarrow> (\<forall>x y z. R x y \<longrightarrow> P y z \<longrightarrow> P x z) \<longrightarrow> P y z" by fact
+ − 125
have r1: "\<forall>x. P x x" by fact
+ − 126
have r2: "\<forall>x y z. R x y \<longrightarrow> P y z \<longrightarrow> P x z" by fact
+ − 127
show "P x z"
+ − 128
565
+ − 129
txt \<open>
+ − 130
The assumptions \<open>p1\<close> and \<open>p2\<close> correspond to the premises of
+ − 131
the second introduction rule (unfolded); the assumptions \<open>r1\<close> and \<open>r2\<close>
+ − 132
come from the definition of @{term trcl}. We apply \<open>r2\<close> to the goal
219
+ − 133
@{term "P x z"}. In order for this assumption to be applicable as a rule, we
127
+ − 134
have to eliminate the universal quantifier and turn the object-level
565
+ − 135
implications into meta-level ones. This can be accomplished using the \<open>rule_format\<close> attribute. So we continue the proof with:
115
+ − 136
565
+ − 137
\<close>
113
+ − 138
127
+ − 139
apply (rule r2[rule_format])
88
+ − 140
565
+ − 141
txt \<open>
127
+ − 142
This gives us two new subgoals
+ − 143
+ − 144
@{subgoals [display]}
+ − 145
565
+ − 146
which can be solved using assumptions \<open>p1\<close> and \<open>p2\<close>. The latter
127
+ − 147
involves a quantifier and implications that have to be eliminated before it
+ − 148
can be applied. To avoid potential problems with higher-order unification,
565
+ − 149
we explicitly instantiate the quantifier to \<open>P\<close> and also match
+ − 150
explicitly the implications with \<open>r1\<close> and \<open>r2\<close>. This gives
127
+ − 151
the proof:
565
+ − 152
\<close>
127
+ − 153
+ − 154
apply(rule p1)
+ − 155
apply(rule p2[THEN spec[where x=P], THEN mp, THEN mp, OF r1, OF r2])
115
+ − 156
done
+ − 157
qed
32
+ − 158
565
+ − 159
text \<open>
127
+ − 160
Now we are done. It might be surprising that we are not using the automatic
565
+ − 161
tactics available in Isabelle/HOL for proving this lemmas. After all \<open>blast\<close> would easily dispense of it.
+ − 162
\<close>
88
+ − 163
127
+ − 164
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
+ − 165
shows "R x y \<Longrightarrow> trcl R y z \<Longrightarrow> trcl R x z"
115
+ − 166
apply(unfold trcl_def)
+ − 167
apply(blast)
+ − 168
done
562
+ − 169
term "even"
565
+ − 170
text \<open>
116
+ − 171
Experience has shown that it is generally a bad idea to rely heavily on
565
+ − 172
\<open>blast\<close>, \<open>auto\<close> and the like in automated proofs. The reason is
116
+ − 173
that you do not have precise control over them (the user can, for example,
+ − 174
declare new intro- or simplification rules that can throw automatic tactics
+ − 175
off course) and also it is very hard to debug proofs involving automatic
+ − 176
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
+ − 177
tactics in packages should be avoided or be constrained sufficiently.
115
+ − 178
116
+ − 179
The method of defining inductive predicates by impredicative quantification
+ − 180
also generalises to mutually inductive predicates. The next example defines
565
+ − 181
the predicates \<open>even\<close> and \<open>odd\<close> given by
32
+ − 182
219
+ − 183
\begin{center}\small
+ − 184
@{prop[mode=Axiom] "even (0::nat)"} \hspace{5mm}
+ − 185
@{prop[mode=Rule] "odd n \<Longrightarrow> even (Suc n)"} \hspace{5mm}
+ − 186
@{prop[mode=Rule] "even n \<Longrightarrow> odd (Suc n)"}
+ − 187
\end{center}
32
+ − 188
127
+ − 189
Since the predicates @{term even} and @{term odd} are mutually inductive, each
+ − 190
corresponding definition must quantify over both predicates (we name them
565
+ − 191
below \<open>P\<close> and \<open>Q\<close>).
+ − 192
\<close>
127
+ − 193
562
+ − 194
hide_const even
+ − 195
hide_const odd
+ − 196
+ − 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
565
+ − 205
text \<open>
116
+ − 206
For proving the induction principles, we use exactly the same technique
+ − 207
as in the transitive closure example, namely:
565
+ − 208
\<close>
32
+ − 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))
513
+ − 214
apply(cut_tac assms)
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
565
+ − 221
text \<open>
+ − 222
The only difference with the proof \<open>trcl_induct\<close> is that we have to
127
+ − 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
565
+ − 226
the \<open>trcl\<close>-example. We only show the proof of the second introduction
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
+ − 227
rule.
565
+ − 228
\<close>
32
+ − 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
565
+ − 248
text \<open>
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:
565
+ − 250
\<open>p1\<close> corresponds to the premise of the introduction rule; \<open>r1\<close>
+ − 251
to \<open>r3\<close> come from the definition of \<open>even\<close>.
+ − 252
In Line 13, we apply the assumption \<open>r2\<close> (since we prove the second
+ − 253
introduction rule). In Lines 14 and 15 we apply assumption \<open>p1\<close> (if
127
+ − 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
565
+ − 259
Next we define the accessible part of a relation \<open>R\<close> given by
219
+ − 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.
565
+ − 268
The definition of \<open>accpart\<close> is:
+ − 269
\<close>
32
+ − 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
565
+ − 273
text \<open>
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.
565
+ − 277
\<close>
115
+ − 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
565
+ − 299
text \<open>
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
565
+ − 301
two categories (Lines 7 to 9). In Line 11, applying the assumption \<open>r1\<close> generates a goal state with the new local assumption @{term "R y x"},
+ − 302
named \<open>r1_prem\<close> in the second subproof (Line 14). This local assumption is
+ − 303
used to solve the goal @{term "P y"} with the help of assumption \<open>p1\<close>.
219
+ − 304
124
+ − 305
219
+ − 306
\begin{exercise}
+ − 307
Give the definition for the freshness predicate for lambda-terms. The rules
+ − 308
for this predicate are:
+ − 309
+ − 310
\begin{center}\small
+ − 311
@{prop[mode=Rule] "a\<noteq>b \<Longrightarrow> fresh a (Var b)"}\hspace{5mm}
+ − 312
@{prop[mode=Rule] "\<lbrakk>fresh a t; fresh a s\<rbrakk> \<Longrightarrow> fresh a (App t s)"}\\[2mm]
+ − 313
@{prop[mode=Axiom] "fresh a (Lam a t)"}\hspace{5mm}
+ − 314
@{prop[mode=Rule] "\<lbrakk>a\<noteq>b; fresh a t\<rbrakk> \<Longrightarrow> fresh a (Lam b t)"}
+ − 315
\end{center}
+ − 316
+ − 317
From the definition derive the induction principle and the introduction
+ − 318
rules.
+ − 319
\end{exercise}
+ − 320
+ − 321
The point of all these examples is to get a feeling what the automatic
+ − 322
proofs should do in order to solve all inductive definitions we throw at
+ − 323
them. This is usually the first step in writing a package. We next explain
129
+ − 324
the parsing and typing part of the package.
127
+ − 325
565
+ − 326
\<close>
129
+ − 327
(*<*)end(*>*)