annotate src/InductiveTypes.v @ 35:6d05ee182b65

Nested Inductive Types
author Adam Chlipala <adamc@hcoop.net>
date Wed, 10 Sep 2008 15:47:22 -0400
parents d44274542f8b
children 9e46ade5af21
rev   line source
adamc@26 1 (* Copyright (c) 2008, Adam Chlipala
adamc@26 2 *
adamc@26 3 * This work is licensed under a
adamc@26 4 * Creative Commons Attribution-Noncommercial-No Derivative Works 3.0
adamc@26 5 * Unported License.
adamc@26 6 * The license text is available at:
adamc@26 7 * http://creativecommons.org/licenses/by-nc-nd/3.0/
adamc@26 8 *)
adamc@26 9
adamc@26 10 (* begin hide *)
adamc@26 11 Require Import List.
adamc@26 12
adamc@26 13 Require Import Tactics.
adamc@26 14
adamc@26 15 Set Implicit Arguments.
adamc@26 16 (* end hide *)
adamc@26 17
adamc@26 18
adamc@26 19 (** %\chapter{Inductive Types}% *)
adamc@26 20
adamc@26 21 (** In a sense, CIC is built from just two relatively straightforward features: function types and inductive types. From this modest foundation, we can prove effectively all of the theorems of math and carry out effectively all program verifications, with enough effort expended. This chapter introduces induction and recursion in Coq and shares some "design patterns" for overcoming common pitfalls with them. *)
adamc@26 22
adamc@26 23
adamc@26 24 (** * Enumerations *)
adamc@26 25
adamc@26 26 (** Coq inductive types generalize the algebraic datatypes found in Haskell and ML. Confusingly enough, inductive types also generalize generalized algebraic datatypes (GADTs), by adding the possibility for type dependency. Even so, it is worth backing up from the examples of the last chapter and going over basic, algebraic datatype uses of inductive datatypes, because the chance to prove things about the values of these types adds new wrinkles beyond usual practice in Haskell and ML.
adamc@26 27
adamc@26 28 The singleton type [unit] is an inductive type: *)
adamc@26 29
adamc@26 30 Inductive unit : Set :=
adamc@26 31 | tt.
adamc@26 32
adamc@26 33 (** This vernacular command defines a new inductive type [unit] whose only value is [tt], as we can see by checking the types of the two identifiers: *)
adamc@26 34
adamc@26 35 Check unit.
adamc@26 36 (** [[
adamc@26 37
adamc@26 38 unit : Set
adamc@26 39 ]] *)
adamc@26 40 Check tt.
adamc@26 41 (** [[
adamc@26 42
adamc@26 43 tt : unit
adamc@26 44 ]] *)
adamc@26 45
adamc@26 46 (** We can prove that [unit] is a genuine singleton type. *)
adamc@26 47
adamc@26 48 Theorem unit_singleton : forall x : unit, x = tt.
adamc@26 49 (** The important thing about an inductive type is, unsurprisingly, that you can do induction over its values, and induction is the key to proving this theorem. We ask to proceed by induction on the variable [x]. *)
adamc@26 50 induction x.
adamc@26 51 (** The goal changes to: [[
adamc@26 52
adamc@26 53 tt = tt
adamc@26 54 ]] *)
adamc@26 55 (** ...which we can discharge trivially. *)
adamc@26 56 reflexivity.
adamc@26 57 Qed.
adamc@26 58
adamc@26 59 (** It seems kind of odd to write a proof by induction with no inductive hypotheses. We could have arrived at the same result by beginning the proof with: [[
adamc@26 60
adamc@26 61 destruct x.
adamc@26 62 ...which corresponds to "proof by case analysis" in classical math. For non-recursive inductive types, the two tactics will always have identical behavior. Often case analysis is sufficient, even in proofs about recursive types, and it is nice to avoid introducing unneeded induction hypotheses.
adamc@26 63
adamc@26 64 What exactly %\textit{%#<i>#is#</i>#%}% the induction principle for [unit]? We can ask Coq: *)
adamc@26 65
adamc@26 66 Check unit_ind.
adamc@26 67 (** [[
adamc@26 68
adamc@26 69 unit_ind : forall P : unit -> Prop, P tt -> forall u : unit, P u
adamc@26 70 ]]
adamc@26 71
adamc@26 72 Every [Inductive] command defining a type [T] also defines an induction principle named [T_ind]. Coq follows the Curry-Howard correspondence and includes the ingredients of programming and proving in the same single syntactic class. Thus, our type, operations over it, and principles for reasoning about it all live in the same language and are described by the same type system. The key to telling what is a program and what is a proof lies in the distinction between the type [Prop], which appears in our induction principle; and the type [Set], which we have seen a few times already.
adamc@26 73
adamc@26 74 The convention goes like this: [Set] is the type of normal types, and the values of such types are programs. [Prop] is the type of logical propositions, and the values of such types are proofs. Thus, an induction principle has a type that shows us that it is a function for building proofs.
adamc@26 75
adamc@26 76 Specifically, [unit_ind] quantifies over a predicate [P] over [unit] values. If we can present a proof that [P] holds of [tt], then we are rewarded with a proof that [P] holds for any value [u] of type [unit]. In our last proof, the predicate was [(fun u : unit => u = tt)].
adamc@26 77
adamc@26 78 %\medskip%
adamc@26 79
adamc@26 80 We can define an inductive type even simpler than [unit]: *)
adamc@26 81
adamc@26 82 Inductive Empty_set : Set := .
adamc@26 83
adamc@26 84 (** [Empty_set] has no elements. We can prove fun theorems about it: *)
adamc@26 85
adamc@26 86 Theorem the_sky_is_falling : forall x : Empty_set, 2 + 2 = 5.
adamc@26 87 destruct 1.
adamc@26 88 Qed.
adamc@26 89
adamc@32 90 (** Because [Empty_set] has no elements, the fact of having an element of this type implies anything. We use [destruct 1] instead of [destruct x] in the proof because unused quantified variables are relegated to being referred to by number. (There is a good reason for this, related to the unity of quantifiers and implication. An implication is just a quantification over a proof, where the quantified variable is never used. It generally makes more sense to refer to implication hypotheses by number than by name, and Coq treats our quantifier over an unused variable as an implication in determining the proper behavior.)
adamc@26 91
adamc@26 92 We can see the induction principle that made this proof so easy: *)
adamc@26 93
adamc@26 94 Check Empty_set_ind.
adamc@26 95 (** [[
adamc@26 96
adamc@26 97 Empty_set_ind : forall (P : Empty_set -> Prop) (e : Empty_set), P e
adamc@26 98 ]]
adamc@26 99
adamc@26 100 In other words, any predicate over values from the empty set holds vacuously of every such element. In the last proof, we chose the predicate [(fun _ : Empty_set => 2 + 2 = 5)].
adamc@26 101
adamc@26 102 We can also apply this get-out-of-jail-free card programmatically. Here is a lazy way of converting values of [Empty_set] to values of [unit]: *)
adamc@26 103
adamc@26 104 Definition e2u (e : Empty_set) : unit := match e with end.
adamc@26 105
adamc@26 106 (** We employ [match] pattern matching as in the last chapter. Since we match on a value whose type has no constructors, there is no need to provide any branches.
adamc@26 107
adamc@26 108 %\medskip%
adamc@26 109
adamc@26 110 Moving up the ladder of complexity, we can define the booleans: *)
adamc@26 111
adamc@26 112 Inductive bool : Set :=
adamc@26 113 | true
adamc@26 114 | false.
adamc@26 115
adamc@26 116 (** We can use less vacuous pattern matching to define boolean negation. *)
adamc@26 117
adamc@26 118 Definition not (b : bool) : bool :=
adamc@26 119 match b with
adamc@26 120 | true => false
adamc@26 121 | false => true
adamc@26 122 end.
adamc@26 123
adamc@27 124 (** An alternative definition desugars to the above: *)
adamc@27 125
adamc@27 126 Definition not' (b : bool) : bool :=
adamc@27 127 if b then false else true.
adamc@27 128
adamc@26 129 (** We might want to prove that [not] is its own inverse operation. *)
adamc@26 130
adamc@26 131 Theorem not_inverse : forall b : bool, not (not b) = b.
adamc@26 132 destruct b.
adamc@26 133
adamc@26 134 (** After we case analyze on [b], we are left with one subgoal for each constructor of [bool].
adamc@26 135
adamc@26 136 [[
adamc@26 137
adamc@26 138 2 subgoals
adamc@26 139
adamc@26 140 ============================
adamc@26 141 not (not true) = true
adamc@26 142 ]]
adamc@26 143
adamc@26 144 [[
adamc@26 145 subgoal 2 is:
adamc@26 146 not (not false) = false
adamc@26 147 ]]
adamc@26 148
adamc@26 149 The first subgoal follows by Coq's rules of computation, so we can dispatch it easily: *)
adamc@26 150
adamc@26 151 reflexivity.
adamc@26 152
adamc@26 153 (** Likewise for the second subgoal, so we can restart the proof and give a very compact justification. *)
adamc@26 154
adamc@26 155 Restart.
adamc@26 156 destruct b; reflexivity.
adamc@26 157 Qed.
adamc@27 158
adamc@27 159 (** Another theorem about booleans illustrates another useful tactic. *)
adamc@27 160
adamc@27 161 Theorem not_ineq : forall b : bool, not b <> b.
adamc@27 162 destruct b; discriminate.
adamc@27 163 Qed.
adamc@27 164
adamc@27 165 (** [discriminate] is used to prove that two values of an inductive type are not equal, whenever the values are formed with different constructors. In this case, the different constructors are [true] and [false].
adamc@27 166
adamc@27 167 At this point, it is probably not hard to guess what the underlying induction principle for [bool] is. *)
adamc@27 168
adamc@27 169 Check bool_ind.
adamc@27 170 (** [[
adamc@27 171
adamc@27 172 bool_ind : forall P : bool -> Prop, P true -> P false -> forall b : bool, P b
adamc@27 173 ]] *)
adamc@28 174
adamc@28 175
adamc@28 176 (** * Simple Recursive Types *)
adamc@28 177
adamc@28 178 (** The natural numbers are the simplest common example of an inductive type that actually deserves the name. *)
adamc@28 179
adamc@28 180 Inductive nat : Set :=
adamc@28 181 | O : nat
adamc@28 182 | S : nat -> nat.
adamc@28 183
adamc@28 184 (** [O] is zero, and [S] is the successor function, so that [0] is syntactic sugar for [O], [1] for [S O], [2] for [S (S O)], and so on.
adamc@28 185
adamc@28 186 Pattern matching works as we demonstrated in the last chapter: *)
adamc@28 187
adamc@28 188 Definition isZero (n : nat) : bool :=
adamc@28 189 match n with
adamc@28 190 | O => true
adamc@28 191 | S _ => false
adamc@28 192 end.
adamc@28 193
adamc@28 194 Definition pred (n : nat) : nat :=
adamc@28 195 match n with
adamc@28 196 | O => O
adamc@28 197 | S n' => n'
adamc@28 198 end.
adamc@28 199
adamc@28 200 (** We can prove theorems by case analysis: *)
adamc@28 201
adamc@28 202 Theorem S_isZero : forall n : nat, isZero (pred (S (S n))) = false.
adamc@28 203 destruct n; reflexivity.
adamc@28 204 Qed.
adamc@28 205
adamc@28 206 (** We can also now get into genuine inductive theorems. First, we will need a recursive function, to make things interesting. *)
adamc@28 207
adamc@28 208 Fixpoint plus (n m : nat) {struct n} : nat :=
adamc@28 209 match n with
adamc@28 210 | O => m
adamc@28 211 | S n' => S (plus n' m)
adamc@28 212 end.
adamc@28 213
adamc@28 214 (** Recall that [Fixpoint] is Coq's mechanism for recursive function definitions, and that the [{struct n}] annotation is noting which function argument decreases structurally at recursive calls.
adamc@28 215
adamc@28 216 Some theorems about [plus] can be proved without induction. *)
adamc@28 217
adamc@28 218 Theorem O_plus_n : forall n : nat, plus O n = n.
adamc@28 219 intro; reflexivity.
adamc@28 220 Qed.
adamc@28 221
adamc@28 222 (** Coq's computation rules automatically simplify the application of [plus]. If we just reverse the order of the arguments, though, this no longer works, and we need induction. *)
adamc@28 223
adamc@28 224 Theorem n_plus_O : forall n : nat, plus n O = n.
adamc@28 225 induction n.
adamc@28 226
adamc@28 227 (** Our first subgoal is [plus O O = O], which %\textit{%#<i>#is#</i>#%}% trivial by computation. *)
adamc@28 228
adamc@28 229 reflexivity.
adamc@28 230
adamc@28 231 (** Our second subgoal is more work and also demonstrates our first inductive hypothesis.
adamc@28 232
adamc@28 233 [[
adamc@28 234
adamc@28 235 n : nat
adamc@28 236 IHn : plus n O = n
adamc@28 237 ============================
adamc@28 238 plus (S n) O = S n
adamc@28 239 ]]
adamc@28 240
adamc@28 241 We can start out by using computation to simplify the goal as far as we can. *)
adamc@28 242
adamc@28 243 simpl.
adamc@28 244
adamc@28 245 (** Now the conclusion is [S (plus n O) = S n]. Using our inductive hypothesis: *)
adamc@28 246
adamc@28 247 rewrite IHn.
adamc@28 248
adamc@28 249 (** ...we get a trivial conclusion [S n = S n]. *)
adamc@28 250
adamc@28 251 reflexivity.
adamc@28 252
adamc@28 253 (** Not much really went on in this proof, so the [crush] tactic from the [Tactics] module can prove this theorem automatically. *)
adamc@28 254
adamc@28 255 Restart.
adamc@28 256 induction n; crush.
adamc@28 257 Qed.
adamc@28 258
adamc@28 259 (** We can check out the induction principle at work here: *)
adamc@28 260
adamc@28 261 Check nat_ind.
adamc@28 262 (** [[
adamc@28 263
adamc@28 264 nat_ind : forall P : nat -> Prop,
adamc@28 265 P O -> (forall n : nat, P n -> P (S n)) -> forall n : nat, P n
adamc@28 266 ]]
adamc@28 267
adamc@28 268 Each of the two cases of our last proof came from the type of one of the arguments to [nat_ind]. We chose [P] to be [(fun n : nat => plus n O = n)]. The first proof case corresponded to [P O], and the second case to [(forall n : nat, P n -> P (S n))]. The free variable [n] and inductive hypothesis [IHn] came from the argument types given here.
adamc@28 269
adamc@28 270 Since [nat] has a constructor that takes an argument, we may sometimes need to know that that constructor is injective. *)
adamc@28 271
adamc@28 272 Theorem S_inj : forall n m : nat, S n = S m -> n = m.
adamc@28 273 injection 1; trivial.
adamc@28 274 Qed.
adamc@28 275
adamc@28 276 (** [injection] refers to a premise by number, adding new equalities between the corresponding arguments of equated terms that are formed with the same constructor. We end up needing to prove [n = m -> n = m], so it is unsurprising that a tactic named [trivial] is able to finish the proof.
adamc@28 277
adamc@29 278 There is also a very useful tactic called [congruence] that can prove this theorem immediately. [congruence] generalizes [discriminate] and [injection], and it also adds reasoning about the general properties of equality, such as that a function returns equal results on equal arguments. That is, [congruence] is a %\textit{%#<i>#complete decision procedure for the theory of equality and uninterpreted functions#</i>#%}%, plus some smarts about inductive types.
adamc@29 279
adamc@29 280 %\medskip%
adamc@29 281
adamc@29 282 We can define a type of lists of natural numbers. *)
adamc@29 283
adamc@29 284 Inductive nat_list : Set :=
adamc@29 285 | NNil : nat_list
adamc@29 286 | NCons : nat -> nat_list -> nat_list.
adamc@29 287
adamc@29 288 (** Recursive definitions are straightforward extensions of what we have seen before. *)
adamc@29 289
adamc@29 290 Fixpoint nlength (ls : nat_list) : nat :=
adamc@29 291 match ls with
adamc@29 292 | NNil => O
adamc@29 293 | NCons _ ls' => S (nlength ls')
adamc@29 294 end.
adamc@29 295
adamc@29 296 Fixpoint napp (ls1 ls2 : nat_list) {struct ls1} : nat_list :=
adamc@29 297 match ls1 with
adamc@29 298 | NNil => ls2
adamc@29 299 | NCons n ls1' => NCons n (napp ls1' ls2)
adamc@29 300 end.
adamc@29 301
adamc@29 302 (** Inductive theorem proving can again be automated quite effectively. *)
adamc@29 303
adamc@29 304 Theorem nlength_napp : forall ls1 ls2 : nat_list, nlength (napp ls1 ls2)
adamc@29 305 = plus (nlength ls1) (nlength ls2).
adamc@29 306 induction ls1; crush.
adamc@29 307 Qed.
adamc@29 308
adamc@29 309 Check nat_list_ind.
adamc@29 310 (** [[
adamc@29 311
adamc@29 312 nat_list_ind
adamc@29 313 : forall P : nat_list -> Prop,
adamc@29 314 P NNil ->
adamc@29 315 (forall (n : nat) (n0 : nat_list), P n0 -> P (NCons n n0)) ->
adamc@29 316 forall n : nat_list, P n
adamc@29 317 ]]
adamc@29 318
adamc@29 319 %\medskip%
adamc@29 320
adamc@29 321 In general, we can implement any "tree" types as inductive types. For example, here are binary trees of naturals. *)
adamc@29 322
adamc@29 323 Inductive nat_btree : Set :=
adamc@29 324 | NLeaf : nat_btree
adamc@29 325 | NNode : nat_btree -> nat -> nat_btree -> nat_btree.
adamc@29 326
adamc@29 327 Fixpoint nsize (tr : nat_btree) : nat :=
adamc@29 328 match tr with
adamc@35 329 | NLeaf => S O
adamc@29 330 | NNode tr1 _ tr2 => plus (nsize tr1) (nsize tr2)
adamc@29 331 end.
adamc@29 332
adamc@29 333 Fixpoint nsplice (tr1 tr2 : nat_btree) {struct tr1} : nat_btree :=
adamc@29 334 match tr1 with
adamc@35 335 | NLeaf => NNode tr2 O NLeaf
adamc@29 336 | NNode tr1' n tr2' => NNode (nsplice tr1' tr2) n tr2'
adamc@29 337 end.
adamc@29 338
adamc@29 339 Theorem plus_assoc : forall n1 n2 n3 : nat, plus (plus n1 n2) n3 = plus n1 (plus n2 n3).
adamc@29 340 induction n1; crush.
adamc@29 341 Qed.
adamc@29 342
adamc@29 343 Theorem nsize_nsplice : forall tr1 tr2 : nat_btree, nsize (nsplice tr1 tr2)
adamc@29 344 = plus (nsize tr2) (nsize tr1).
adamc@29 345 Hint Rewrite n_plus_O plus_assoc : cpdt.
adamc@29 346
adamc@29 347 induction tr1; crush.
adamc@29 348 Qed.
adamc@29 349
adamc@29 350 Check nat_btree_ind.
adamc@29 351 (** [[
adamc@29 352
adamc@29 353 nat_btree_ind
adamc@29 354 : forall P : nat_btree -> Prop,
adamc@29 355 P NLeaf ->
adamc@29 356 (forall n : nat_btree,
adamc@29 357 P n -> forall (n0 : nat) (n1 : nat_btree), P n1 -> P (NNode n n0 n1)) ->
adamc@29 358 forall n : nat_btree, P n
adamc@29 359 ]] *)
adamc@30 360
adamc@30 361
adamc@30 362 (** * Parameterized Types *)
adamc@30 363
adamc@30 364 (** We can also define polymorphic inductive types, as with algebraic datatypes in Haskell and ML. *)
adamc@30 365
adamc@30 366 Inductive list (T : Set) : Set :=
adamc@30 367 | Nil : list T
adamc@30 368 | Cons : T -> list T -> list T.
adamc@30 369
adamc@30 370 Fixpoint length T (ls : list T) : nat :=
adamc@30 371 match ls with
adamc@30 372 | Nil => O
adamc@30 373 | Cons _ ls' => S (length ls')
adamc@30 374 end.
adamc@30 375
adamc@30 376 Fixpoint app T (ls1 ls2 : list T) {struct ls1} : list T :=
adamc@30 377 match ls1 with
adamc@30 378 | Nil => ls2
adamc@30 379 | Cons x ls1' => Cons x (app ls1' ls2)
adamc@30 380 end.
adamc@30 381
adamc@30 382 Theorem length_app : forall T (ls1 ls2 : list T), length (app ls1 ls2)
adamc@30 383 = plus (length ls1) (length ls2).
adamc@30 384 induction ls1; crush.
adamc@30 385 Qed.
adamc@30 386
adamc@30 387 (** There is a useful shorthand for writing many definitions that share the same parameter, based on Coq's %\textit{%#<i>#section#</i>#%}% mechanism. The following block of code is equivalent to the above: *)
adamc@30 388
adamc@30 389 (* begin hide *)
adamc@30 390 Reset list.
adamc@30 391 (* end hide *)
adamc@30 392
adamc@30 393 Section list.
adamc@30 394 Variable T : Set.
adamc@30 395
adamc@30 396 Inductive list : Set :=
adamc@30 397 | Nil : list
adamc@30 398 | Cons : T -> list -> list.
adamc@30 399
adamc@30 400 Fixpoint length (ls : list) : nat :=
adamc@30 401 match ls with
adamc@30 402 | Nil => O
adamc@30 403 | Cons _ ls' => S (length ls')
adamc@30 404 end.
adamc@30 405
adamc@30 406 Fixpoint app (ls1 ls2 : list) {struct ls1} : list :=
adamc@30 407 match ls1 with
adamc@30 408 | Nil => ls2
adamc@30 409 | Cons x ls1' => Cons x (app ls1' ls2)
adamc@30 410 end.
adamc@30 411
adamc@30 412 Theorem length_app : forall ls1 ls2 : list, length (app ls1 ls2)
adamc@30 413 = plus (length ls1) (length ls2).
adamc@30 414 induction ls1; crush.
adamc@30 415 Qed.
adamc@30 416 End list.
adamc@30 417
adamc@35 418 (* begin hide *)
adamc@35 419 Implicit Arguments Nil [T].
adamc@35 420 (* end hide *)
adamc@35 421
adamc@30 422 (** After we end the section, the [Variable]s we used are added as extra function parameters for each defined identifier, as needed. *)
adamc@30 423
adamc@30 424 Check list.
adamc@30 425 (** [[
adamc@30 426
adamc@30 427 list
adamc@30 428 : Set -> Set
adamc@30 429 ]] *)
adamc@30 430
adamc@30 431 Check Cons.
adamc@30 432 (** [[
adamc@30 433
adamc@30 434 Cons
adamc@30 435 : forall T : Set, T -> list T -> list T
adamc@30 436 ]] *)
adamc@30 437
adamc@30 438 Check length.
adamc@30 439 (** [[
adamc@30 440
adamc@30 441 length
adamc@30 442 : forall T : Set, list T -> nat
adamc@30 443 ]]
adamc@30 444
adamc@30 445 The extra parameter [T] is treated as a new argument to the induction principle, too. *)
adamc@30 446
adamc@30 447 Check list_ind.
adamc@30 448 (** [[
adamc@30 449
adamc@30 450 list_ind
adamc@30 451 : forall (T : Set) (P : list T -> Prop),
adamc@30 452 P (Nil T) ->
adamc@30 453 (forall (t : T) (l : list T), P l -> P (Cons t l)) ->
adamc@30 454 forall l : list T, P l
adamc@30 455 ]]
adamc@30 456
adamc@30 457 Thus, even though we just saw that [T] is added as an extra argument to the constructor [Cons], there is no quantifier for [T] in the type of the inductive case like there is for each of the other arguments. *)
adamc@31 458
adamc@31 459
adamc@31 460 (** * Mutually Inductive Types *)
adamc@31 461
adamc@31 462 (** We can define inductive types that refer to each other: *)
adamc@31 463
adamc@31 464 Inductive even_list : Set :=
adamc@31 465 | ENil : even_list
adamc@31 466 | ECons : nat -> odd_list -> even_list
adamc@31 467
adamc@31 468 with odd_list : Set :=
adamc@31 469 | OCons : nat -> even_list -> odd_list.
adamc@31 470
adamc@31 471 Fixpoint elength (el : even_list) : nat :=
adamc@31 472 match el with
adamc@31 473 | ENil => O
adamc@31 474 | ECons _ ol => S (olength ol)
adamc@31 475 end
adamc@31 476
adamc@31 477 with olength (ol : odd_list) : nat :=
adamc@31 478 match ol with
adamc@31 479 | OCons _ el => S (elength el)
adamc@31 480 end.
adamc@31 481
adamc@31 482 Fixpoint eapp (el1 el2 : even_list) {struct el1} : even_list :=
adamc@31 483 match el1 with
adamc@31 484 | ENil => el2
adamc@31 485 | ECons n ol => ECons n (oapp ol el2)
adamc@31 486 end
adamc@31 487
adamc@31 488 with oapp (ol : odd_list) (el : even_list) {struct ol} : odd_list :=
adamc@31 489 match ol with
adamc@31 490 | OCons n el' => OCons n (eapp el' el)
adamc@31 491 end.
adamc@31 492
adamc@31 493 (** Everything is going roughly the same as in past examples, until we try to prove a theorem similar to those that came before. *)
adamc@31 494
adamc@31 495 Theorem elength_eapp : forall el1 el2 : even_list,
adamc@31 496 elength (eapp el1 el2) = plus (elength el1) (elength el2).
adamc@31 497 induction el1; crush.
adamc@31 498
adamc@31 499 (** One goal remains: [[
adamc@31 500
adamc@31 501 n : nat
adamc@31 502 o : odd_list
adamc@31 503 el2 : even_list
adamc@31 504 ============================
adamc@31 505 S (olength (oapp o el2)) = S (plus (olength o) (elength el2))
adamc@31 506 ]]
adamc@31 507
adamc@31 508 We have no induction hypothesis, so we cannot prove this goal without starting another induction, which would reach a similar point, sending us into a futile infinite chain of inductions. The problem is that Coq's generation of [T_ind] principles is incomplete. We only get non-mutual induction principles generated by default. *)
adamc@31 509
adamc@31 510 Abort.
adamc@31 511 Check even_list_ind.
adamc@31 512 (** [[
adamc@31 513
adamc@31 514 even_list_ind
adamc@31 515 : forall P : even_list -> Prop,
adamc@31 516 P ENil ->
adamc@31 517 (forall (n : nat) (o : odd_list), P (ECons n o)) ->
adamc@31 518 forall e : even_list, P e
adamc@31 519 ]]
adamc@31 520
adamc@31 521 We see that no inductive hypotheses are included anywhere in the type. To get them, we must ask for mutual principles as we need them, using the [Scheme] command. *)
adamc@31 522
adamc@31 523 Scheme even_list_mut := Induction for even_list Sort Prop
adamc@31 524 with odd_list_mut := Induction for odd_list Sort Prop.
adamc@31 525
adamc@31 526 Check even_list_mut.
adamc@31 527 (** [[
adamc@31 528
adamc@31 529 even_list_mut
adamc@31 530 : forall (P : even_list -> Prop) (P0 : odd_list -> Prop),
adamc@31 531 P ENil ->
adamc@31 532 (forall (n : nat) (o : odd_list), P0 o -> P (ECons n o)) ->
adamc@31 533 (forall (n : nat) (e : even_list), P e -> P0 (OCons n e)) ->
adamc@31 534 forall e : even_list, P e
adamc@31 535 ]]
adamc@31 536
adamc@31 537 This is the principle we wanted in the first place. There is one more wrinkle left in using it: the [induction] tactic will not apply it for us automatically. It will be helpful to look at how to prove one of our past examples without using [induction], so that we can then generalize the technique to mutual inductive types. *)
adamc@31 538
adamc@31 539 Theorem n_plus_O' : forall n : nat, plus n O = n.
adamc@31 540 apply (nat_ind (fun n => plus n O = n)); crush.
adamc@31 541 Qed.
adamc@31 542
adamc@31 543 (** From this example, we can see that [induction] is not magic. It only does some bookkeeping for us to make it easy to apply a theorem, which we can do directly with the [apply] tactic. We apply not just an identifier but a partial application of it, specifying the predicate we mean to prove holds for all naturals.
adamc@31 544
adamc@31 545 This technique generalizes to our mutual example: *)
adamc@31 546
adamc@31 547 Theorem elength_eapp : forall el1 el2 : even_list,
adamc@31 548 elength (eapp el1 el2) = plus (elength el1) (elength el2).
adamc@31 549 apply (even_list_mut
adamc@31 550 (fun el1 : even_list => forall el2 : even_list,
adamc@31 551 elength (eapp el1 el2) = plus (elength el1) (elength el2))
adamc@31 552 (fun ol : odd_list => forall el : even_list,
adamc@31 553 olength (oapp ol el) = plus (olength ol) (elength el))); crush.
adamc@31 554 Qed.
adamc@31 555
adamc@31 556 (** We simply need to specify two predicates, one for each of the mutually inductive types. In general, it would not be a good idea to assume that a proof assistant could infer extra predicates, so this way of applying mutual induction is about as straightforward as we could hope for. *)
adamc@33 557
adamc@33 558
adamc@33 559 (** * Reflexive Types *)
adamc@33 560
adamc@33 561 (** A kind of inductive type called a %\textit{%#<i>#reflexive type#</i>#%}% is defined in terms of functions that have the type being defined as their range. One very useful class of examples is in modeling variable binders. For instance, here is a type for encoding the syntax of a subset of first-order logic: *)
adamc@33 562
adamc@33 563 Inductive formula : Set :=
adamc@33 564 | Eq : nat -> nat -> formula
adamc@33 565 | And : formula -> formula -> formula
adamc@33 566 | Forall : (nat -> formula) -> formula.
adamc@33 567
adamc@33 568 (** Our kinds of formulas are equalities between naturals, conjunction, and universal quantification over natural numbers. We avoid needing to include a notion of "variables" in our type, by using Coq functions to encode quantification. For instance, here is the encoding of [forall x : nat, x = x]: *)
adamc@33 569
adamc@33 570 Example forall_refl : formula := Forall (fun x => Eq x x).
adamc@33 571
adamc@33 572 (** We can write recursive functions over reflexive types quite naturally. Here is one translating our formulas into native Coq propositions. *)
adamc@33 573
adamc@33 574 Fixpoint formulaDenote (f : formula) : Prop :=
adamc@33 575 match f with
adamc@33 576 | Eq n1 n2 => n1 = n2
adamc@33 577 | And f1 f2 => formulaDenote f1 /\ formulaDenote f2
adamc@33 578 | Forall f' => forall n : nat, formulaDenote (f' n)
adamc@33 579 end.
adamc@33 580
adamc@33 581 (** We can also encode a trivial formula transformation that swaps the order of equality and conjunction operands. *)
adamc@33 582
adamc@33 583 Fixpoint swapper (f : formula) : formula :=
adamc@33 584 match f with
adamc@33 585 | Eq n1 n2 => Eq n2 n1
adamc@33 586 | And f1 f2 => And (swapper f2) (swapper f1)
adamc@33 587 | Forall f' => Forall (fun n => swapper (f' n))
adamc@33 588 end.
adamc@33 589
adamc@33 590 (** It is helpful to prove that this transformation does not make true formulas false. *)
adamc@33 591
adamc@33 592 Theorem swapper_preserves_truth : forall f, formulaDenote f -> formulaDenote (swapper f).
adamc@33 593 induction f; crush.
adamc@33 594 Qed.
adamc@33 595
adamc@33 596 (** We can take a look at the induction principle behind this proof. *)
adamc@33 597
adamc@33 598 Check formula_ind.
adamc@33 599 (** [[
adamc@33 600
adamc@33 601 formula_ind
adamc@33 602 : forall P : formula -> Prop,
adamc@33 603 (forall n n0 : nat, P (Eq n n0)) ->
adamc@33 604 (forall f0 : formula,
adamc@33 605 P f0 -> forall f1 : formula, P f1 -> P (And f0 f1)) ->
adamc@33 606 (forall f1 : nat -> formula,
adamc@33 607 (forall n : nat, P (f1 n)) -> P (Forall f1)) ->
adamc@33 608 forall f2 : formula, P f2
adamc@33 609 ]] *)
adamc@33 610
adamc@33 611 (** Focusing on the [Forall] case, which comes third, we see that we are allowed to assume that the theorem holds %\textit{%#<i>#for any application of the argument function [f1]#</i>#%}%. That is, Coq induction principles do not follow a simple rule that the textual representations of induction variables must get shorter in appeals to induction hypotheses. Luckily for us, the people behind the metatheory of Coq have verified that this flexibility does not introduce unsoundness.
adamc@33 612
adamc@33 613 %\medskip%
adamc@33 614
adamc@33 615 Up to this point, we have seen how to encode in Coq more and more of what is possible with algebraic datatypes in Haskell and ML. This may have given the inaccurate impression that inductive types are a strict extension of algebraic datatypes. In fact, Coq must rule out some types allowed by Haskell and ML, for reasons of soundness. Reflexive types provide our first good example of such a case.
adamc@33 616
adamc@33 617 Given our last example of an inductive type, many readers are probably eager to try encoding the syntax of lambda calculus. Indeed, the function-based representation technique that we just used, called %\textit{%#<i>#higher-order abstract syntax (HOAS)#</i>#%}%, is the representation of choice for lambda calculi in Twelf and in many applications implemented in Haskell and ML. Let us try to import that choice to Coq: *)
adamc@33 618
adamc@33 619 (** [[
adamc@33 620
adamc@33 621 Inductive term : Set :=
adamc@33 622 | App : term -> term -> term
adamc@33 623 | Abs : (term -> term) -> term.
adamc@33 624
adamc@33 625 [[
adamc@33 626 Error: Non strictly positive occurrence of "term" in "(term -> term) -> term"
adamc@33 627 ]]
adamc@33 628
adamc@33 629 We have run afoul of the %\textit{%#<i>#strict positivity requirement#</i>#%}% for inductive definitions, which says that the type being defined may not occur to the left of an arrow in the type of a constructor argument. It is important that the type of a constructor is viewed in terms of a series of arguments and a result, since obviously we need recursive occurrences to the lefts of the outermost arrows if we are to have recursive occurrences at all.
adamc@33 630
adamc@33 631 Why must Coq enforce this restriction? Imagine that our last definition had been accepted, allowing us to write this function:
adamc@33 632
adamc@33 633 [[
adamc@33 634 Definition uhoh (t : term) : term :=
adamc@33 635 match t with
adamc@33 636 | Abs f => f t
adamc@33 637 | _ => t
adamc@33 638 end.
adamc@33 639
adamc@33 640 Using an informal idea of Coq's semantics, it is easy to verify that the application [uhoh (Abs uhoh)] will run forever. This would be a mere curiosity in OCaml and Haskell, where non-termination is commonplace, though the fact that we have a non-terminating program without explicit recursive function definitions is unusual.
adamc@33 641
adamc@33 642 For Coq, however, this would be a disaster. The possibility of writing such a function would destroy all our confidence that proving a theorem means anything. Since Coq combines programs and proofs in one language, we would be able to prove every theorem with an infinite loop.
adamc@33 643
adamc@33 644 Nonetheless, the basic insight of HOAS is a very useful one, and there are ways to realize most benefits of HOAS in Coq. We will study a particular technique of this kind in the later chapters on programming language syntax and semantics. *)
adamc@34 645
adamc@34 646
adamc@34 647 (** * An Interlude on Proof Terms *)
adamc@34 648
adamc@34 649 (** As we have emphasized a few times already, Coq proofs are actually programs, written in the same language we have been using in our examples all along. We can get a first sense of what this means by taking a look at the definitions of some of the induction principles we have used. *)
adamc@34 650
adamc@34 651 Print unit_ind.
adamc@34 652 (** [[
adamc@34 653
adamc@34 654 unit_ind =
adamc@34 655 fun P : unit -> Prop => unit_rect P
adamc@34 656 : forall P : unit -> Prop, P tt -> forall u : unit, P u
adamc@34 657 ]]
adamc@34 658
adamc@34 659 We see that this induction principle is defined in terms of a more general principle, [unit_rect]. *)
adamc@34 660
adamc@34 661 Check unit_rect.
adamc@34 662 (** [[
adamc@34 663
adamc@34 664 unit_rect
adamc@34 665 : forall P : unit -> Type, P tt -> forall u : unit, P u
adamc@34 666 ]]
adamc@34 667
adamc@34 668 [unit_rect] gives [P] type [unit -> Type] instead of [unit -> Prop]. [Type] is another universe, like [Set] and [Prop]. In fact, it is a common supertype of both. Later on, we will discuss exactly what the significances of the different universes are. For now, it is just important that we can use [Type] as a sort of meta-universe that may turn out to be either [Set] or [Prop]. We can see the symmetry inherent in the subtyping relationship by printing the definition of another principle that was generated for [unit] automatically: *)
adamc@34 669
adamc@34 670 Print unit_rec.
adamc@34 671 (** [[
adamc@34 672
adamc@34 673 unit_rec =
adamc@34 674 fun P : unit -> Set => unit_rect P
adamc@34 675 : forall P : unit -> Set, P tt -> forall u : unit, P u
adamc@34 676 ]]
adamc@34 677
adamc@34 678 This is identical to the definition for [unit_ind], except that we have substituted [Set] for [Prop]. For most inductive types [T], then, we get not just induction principles [T_ind], but also recursion principles [T_rec]. We can use [T_rec] to write recursive definitions without explicit [Fixpoint] recursion. For instance, the following two definitions are equivalent: *)
adamc@34 679
adamc@34 680 Definition always_O (u : unit) : nat :=
adamc@34 681 match u with
adamc@34 682 | tt => O
adamc@34 683 end.
adamc@34 684
adamc@34 685 Definition always_O' (u : unit) : nat :=
adamc@34 686 unit_rec (fun _ : unit => nat) O u.
adamc@34 687
adamc@34 688 (** Going even further down the rabbit hole, [unit_rect] itself is not even a primitive. It is a functional program that we can write manually. *)
adamc@34 689
adamc@34 690 Print unit_rect.
adamc@34 691
adamc@34 692 (** [[
adamc@34 693
adamc@34 694 unit_rect =
adamc@34 695 fun (P : unit -> Type) (f : P tt) (u : unit) =>
adamc@34 696 match u as u0 return (P u0) with
adamc@34 697 | tt => f
adamc@34 698 end
adamc@34 699 : forall P : unit -> Type, P tt -> forall u : unit, P u
adamc@34 700 ]]
adamc@34 701
adamc@34 702 The only new feature we see is an [as] clause for a [match], which is used in concert with the [return] clause that we saw in the introduction. Since the type of the [match] is dependent on the value of the object being analyzed, we must give that object a name so that we can refer to it in the [return] clause.
adamc@34 703
adamc@34 704 To prove that [unit_rect] is nothing special, we can reimplement it manually. *)
adamc@34 705
adamc@34 706 Definition unit_rect' (P : unit -> Type) (f : P tt) (u : unit) :=
adamc@34 707 match u return (P u) with
adamc@34 708 | tt => f
adamc@34 709 end.
adamc@34 710
adamc@34 711 (** We use the handy shorthand that lets us omit an [as] annotation when matching on a variable, simply using that variable directly in the [return] clause.
adamc@34 712
adamc@34 713 We can check the implement of [nat_rect] as well: *)
adamc@34 714
adamc@34 715 Print nat_rect.
adamc@34 716 (** [[
adamc@34 717
adamc@34 718 nat_rect =
adamc@34 719 fun (P : nat -> Type) (f : P O) (f0 : forall n : nat, P n -> P (S n)) =>
adamc@34 720 fix F (n : nat) : P n :=
adamc@34 721 match n as n0 return (P n0) with
adamc@34 722 | O => f
adamc@34 723 | S n0 => f0 n0 (F n0)
adamc@34 724 end
adamc@34 725 : forall P : nat -> Type,
adamc@34 726 P O -> (forall n : nat, P n -> P (S n)) -> forall n : nat, P n
adamc@34 727 ]]
adamc@34 728
adamc@34 729 Now we have an actual recursive definition. [fix] expressions are an anonymous form of [Fixpoint], just as [fun] expressions stand for anonymous non-recursive functions. Beyond that, the syntax of [fix] mirrors that of [Fixpoint]. We can understand the definition of [nat_rect] better by reimplementing [nat_ind] using sections. *)
adamc@34 730
adamc@34 731 Section nat_ind'.
adamc@34 732 (** First, we have the property of natural numbers that we aim to prove. *)
adamc@34 733 Variable P : nat -> Prop.
adamc@34 734
adamc@34 735 (** Then we require a proof of the [O] case. *)
adamc@34 736 Variable O_case : P O.
adamc@34 737
adamc@34 738 (** Next is a proof of the [S] case, which may assume an inductive hypothesis. *)
adamc@34 739 Variable S_case : forall n : nat, P n -> P (S n).
adamc@34 740
adamc@34 741 (** Finally, we define a recursive function to tie the pieces together. *)
adamc@34 742 Fixpoint nat_ind' (n : nat) : P n :=
adamc@34 743 match n return (P n) with
adamc@34 744 | O => O_case
adamc@34 745 | S n' => S_case (nat_ind' n')
adamc@34 746 end.
adamc@34 747 End nat_ind'.
adamc@34 748
adamc@34 749 (** Closing the section adds the [Variable]s as new [fun]-bound arguments to [nat_ind'], and, modulo the use of [Prop] instead of [Type], we end up with the exact same definition that was generated automatically for [nat_rect].
adamc@34 750
adamc@34 751 %\medskip%
adamc@34 752
adamc@34 753 We can also examine the definition of [even_list_mut], which we generated with [Scheme] for a mutually-recursive type. *)
adamc@34 754
adamc@34 755 Print even_list_mut.
adamc@34 756 (** [[
adamc@34 757
adamc@34 758 even_list_mut =
adamc@34 759 fun (P : even_list -> Prop) (P0 : odd_list -> Prop)
adamc@34 760 (f : P ENil) (f0 : forall (n : nat) (o : odd_list), P0 o -> P (ECons n o))
adamc@34 761 (f1 : forall (n : nat) (e : even_list), P e -> P0 (OCons n e)) =>
adamc@34 762 fix F (e : even_list) : P e :=
adamc@34 763 match e as e0 return (P e0) with
adamc@34 764 | ENil => f
adamc@34 765 | ECons n o => f0 n o (F0 o)
adamc@34 766 end
adamc@34 767 with F0 (o : odd_list) : P0 o :=
adamc@34 768 match o as o0 return (P0 o0) with
adamc@34 769 | OCons n e => f1 n e (F e)
adamc@34 770 end
adamc@34 771 for F
adamc@34 772 : forall (P : even_list -> Prop) (P0 : odd_list -> Prop),
adamc@34 773 P ENil ->
adamc@34 774 (forall (n : nat) (o : odd_list), P0 o -> P (ECons n o)) ->
adamc@34 775 (forall (n : nat) (e : even_list), P e -> P0 (OCons n e)) ->
adamc@34 776 forall e : even_list, P e
adamc@34 777 ]]
adamc@34 778
adamc@34 779 We see a mutually-recursive [fix], with the different functions separated by [with] in the same way that they would be separated by [and] in ML. A final [for] clause identifies which of the mutually-recursive functions should be the final value of the [fix] expression. Using this definition as a template, we can reimplement [even_list_mut] directly. *)
adamc@34 780
adamc@34 781 Section even_list_mut'.
adamc@34 782 (** First, we need the properties that we are proving. *)
adamc@34 783 Variable Peven : even_list -> Prop.
adamc@34 784 Variable Podd : odd_list -> Prop.
adamc@34 785
adamc@34 786 (** Next, we need proofs of the three cases. *)
adamc@34 787 Variable ENil_case : Peven ENil.
adamc@34 788 Variable ECons_case : forall (n : nat) (o : odd_list), Podd o -> Peven (ECons n o).
adamc@34 789 Variable OCons_case : forall (n : nat) (e : even_list), Peven e -> Podd (OCons n e).
adamc@34 790
adamc@34 791 (** Finally, we define the recursive functions. *)
adamc@34 792 Fixpoint even_list_mut' (e : even_list) : Peven e :=
adamc@34 793 match e return (Peven e) with
adamc@34 794 | ENil => ENil_case
adamc@34 795 | ECons n o => ECons_case n (odd_list_mut' o)
adamc@34 796 end
adamc@34 797 with odd_list_mut' (o : odd_list) : Podd o :=
adamc@34 798 match o return (Podd o) with
adamc@34 799 | OCons n e => OCons_case n (even_list_mut' e)
adamc@34 800 end.
adamc@34 801 End even_list_mut'.
adamc@34 802
adamc@34 803 (** Even induction principles for reflexive types are easy to implement directly. For our [formula] type, we can use a recursive definition much like those we wrote above. *)
adamc@34 804
adamc@34 805 Section formula_ind'.
adamc@34 806 Variable P : formula -> Prop.
adamc@34 807 Variable Eq_case : forall n1 n2 : nat, P (Eq n1 n2).
adamc@34 808 Variable And_case : forall f1 f2 : formula,
adamc@34 809 P f1 -> P f2 -> P (And f1 f2).
adamc@34 810 Variable Forall_case : forall f : nat -> formula,
adamc@34 811 (forall n : nat, P (f n)) -> P (Forall f).
adamc@34 812
adamc@34 813 Fixpoint formula_ind' (f : formula) : P f :=
adamc@34 814 match f return (P f) with
adamc@34 815 | Eq n1 n2 => Eq_case n1 n2
adamc@34 816 | And f1 f2 => And_case (formula_ind' f1) (formula_ind' f2)
adamc@34 817 | Forall f' => Forall_case f' (fun n => formula_ind' (f' n))
adamc@34 818 end.
adamc@34 819 End formula_ind'.
adamc@34 820
adamc@35 821
adamc@35 822 (** * Nested Inductive Types *)
adamc@35 823
adamc@35 824 (** Suppose we want to extend our earlier type of binary trees to trees with arbitrary finite branching. We can use lists to give a simple definition. *)
adamc@35 825
adamc@35 826 Inductive nat_tree : Set :=
adamc@35 827 | NLeaf' : nat_tree
adamc@35 828 | NNode' : nat -> list nat_tree -> nat_tree.
adamc@35 829
adamc@35 830 (** This is an example of a %\textit{%#<i>#nested#</i>#%}% inductive type definition, because we use the type we are defining as an argument to a parametrized type family. Coq will not allow all such definitions; it effectively pretends that we are defining [nat_tree] mutually with a version of [list] specialized to [nat_tree], checking that the resulting expanded definition satisfies the usual rules. For instance, if we replaced [list] with a type family that used its parameter as a function argument, then the definition would be rejected as violating the positivity restriction.
adamc@35 831
adamc@35 832 Like we encountered for mutual inductive types, we find that the automatically-generated induction principle for [nat_tree] is too weak. *)
adamc@35 833
adamc@35 834 Check nat_tree_ind.
adamc@35 835 (** [[
adamc@35 836
adamc@35 837 nat_tree_ind
adamc@35 838 : forall P : nat_tree -> Prop,
adamc@35 839 P NLeaf' ->
adamc@35 840 (forall (n : nat) (l : list nat_tree), P (NNode' n l)) ->
adamc@35 841 forall n : nat_tree, P n
adamc@35 842 ]]
adamc@35 843
adamc@35 844 There is no command like [Scheme] that will implement an improved principle for us. In general, it takes creativity to figure out how to incorporate nested uses to different type families. Now that we know how to implement induction principles manually, we are in a position to apply just such creativity to this problem.
adamc@35 845
adamc@35 846 First, we will need an auxiliary definition, characterizing what it means for a property to hold of every element of a list. *)
adamc@35 847
adamc@35 848 Section All.
adamc@35 849 Variable T : Set.
adamc@35 850 Variable P : T -> Prop.
adamc@35 851
adamc@35 852 Fixpoint All (ls : list T) : Prop :=
adamc@35 853 match ls with
adamc@35 854 | Nil => True
adamc@35 855 | Cons h t => P h /\ All t
adamc@35 856 end.
adamc@35 857 End All.
adamc@35 858
adamc@35 859 (** It will be useful to look at the definitions of [True] and [/\], since we will want to write manual proofs of them below. *)
adamc@35 860
adamc@35 861 Print True.
adamc@35 862 (** [[
adamc@35 863
adamc@35 864 Inductive True : Prop := I : True
adamc@35 865 ]]
adamc@35 866
adamc@35 867 That is, [True] is a proposition with exactly one proof, [I], which we may always supply trivially.
adamc@35 868
adamc@35 869 Finding the definition of [/\] takes a little more work. Coq supports user registration of arbitrary parsing rules, and it is such a rule that is letting us write [/\] instead of an application of some inductive type family. We can find the underlying inductive type with the [Locate] command. *)
adamc@35 870
adamc@35 871 Locate "/\".
adamc@35 872 (** [[
adamc@35 873
adamc@35 874 Notation Scope
adamc@35 875 "A /\ B" := and A B : type_scope
adamc@35 876 (default interpretation)
adamc@35 877 ]] *)
adamc@35 878
adamc@35 879 Print and.
adamc@35 880 (** [[
adamc@35 881
adamc@35 882 Inductive and (A : Prop) (B : Prop) : Prop := conj : A -> B -> A /\ B
adamc@35 883 For conj: Arguments A, B are implicit
adamc@35 884 For and: Argument scopes are [type_scope type_scope]
adamc@35 885 For conj: Argument scopes are [type_scope type_scope _ _]
adamc@35 886 ]]
adamc@35 887
adamc@35 888 In addition to the definition of [and] itself, we get information on implicit arguments and parsing rules for [and] and its constructor [conj]. We will ignore the parsing information for now. The implicit argument information tells us that we build a proof of a conjunction by calling the constructor [conj] on proofs of the conjuncts, with no need to include the types of those proofs as explicit arguments.
adamc@35 889
adamc@35 890 %\medskip%
adamc@35 891
adamc@35 892 Now we create a section for our induction principle, following the same basic plan as in the last section of this chapter. *)
adamc@35 893
adamc@35 894 Section nat_tree_ind'.
adamc@35 895 Variable P : nat_tree -> Prop.
adamc@35 896
adamc@35 897 Variable NLeaf'_case : P NLeaf'.
adamc@35 898 Variable NNode'_case : forall (n : nat) (ls : list nat_tree),
adamc@35 899 All P ls -> P (NNode' n ls).
adamc@35 900
adamc@35 901 (** A first attempt at writing the induction principle itself follows the intuition that nested inductive type definitions are expanded into mutual inductive definitions.
adamc@35 902
adamc@35 903 [[
adamc@35 904
adamc@35 905 Fixpoint nat_tree_ind' (tr : nat_tree) : P tr :=
adamc@35 906 match tr return (P tr) with
adamc@35 907 | NLeaf' => NLeaf'_case
adamc@35 908 | NNode' n ls => NNode'_case n ls (list_nat_tree_ind ls)
adamc@35 909 end
adamc@35 910
adamc@35 911 with list_nat_tree_ind (ls : list nat_tree) : All P ls :=
adamc@35 912 match ls return (All P ls) with
adamc@35 913 | Nil => I
adamc@35 914 | Cons tr rest => conj (nat_tree_ind' tr) (list_nat_tree_ind rest)
adamc@35 915 end.
adamc@35 916
adamc@35 917 Coq rejects this definition, saying "Recursive call to nat_tree_ind' has principal argument equal to "tr" instead of rest." The term "nested inductive type" hints at the solution to the problem. Just like true mutually-inductive types require mutually-recursive induction principles, nested types require nested recursion. *)
adamc@35 918
adamc@35 919 Fixpoint nat_tree_ind' (tr : nat_tree) : P tr :=
adamc@35 920 match tr return (P tr) with
adamc@35 921 | NLeaf' => NLeaf'_case
adamc@35 922 | NNode' n ls => NNode'_case n ls
adamc@35 923 ((fix list_nat_tree_ind (ls : list nat_tree) : All P ls :=
adamc@35 924 match ls return (All P ls) with
adamc@35 925 | Nil => I
adamc@35 926 | Cons tr rest => conj (nat_tree_ind' tr) (list_nat_tree_ind rest)
adamc@35 927 end) ls)
adamc@35 928 end.
adamc@35 929
adamc@35 930 (** We include an anonymous [fix] version of [list_nat_tree_ind] that is literally %\textit{%#<i>#nested#</i>#%}% inside the definition of the recursive function corresponding to the inductive definition that had the nested use of [list]. *)
adamc@35 931
adamc@35 932 End nat_tree_ind'.
adamc@35 933
adamc@35 934 (** We can try our induction principle out by defining some recursive functions on [nat_tree]s and proving a theorem about them. First, we define some helper functions that operate on lists. *)
adamc@35 935
adamc@35 936 Section map.
adamc@35 937 Variables T T' : Set.
adamc@35 938 Variable f : T -> T'.
adamc@35 939
adamc@35 940 Fixpoint map (ls : list T) : list T' :=
adamc@35 941 match ls with
adamc@35 942 | Nil => Nil
adamc@35 943 | Cons h t => Cons (f h) (map t)
adamc@35 944 end.
adamc@35 945 End map.
adamc@35 946
adamc@35 947 Fixpoint sum (ls : list nat) : nat :=
adamc@35 948 match ls with
adamc@35 949 | Nil => O
adamc@35 950 | Cons h t => plus h (sum t)
adamc@35 951 end.
adamc@35 952
adamc@35 953 (** Now we can define a size function over our trees. *)
adamc@35 954
adamc@35 955 Fixpoint ntsize (tr : nat_tree) : nat :=
adamc@35 956 match tr with
adamc@35 957 | NLeaf' => S O
adamc@35 958 | NNode' _ trs => S (sum (map ntsize trs))
adamc@35 959 end.
adamc@35 960
adamc@35 961 (** Notice that Coq was smart enough to expand the definition of [map] to verify that we are using proper nested recursion, even through a use of a higher-order function. *)
adamc@35 962
adamc@35 963 Fixpoint ntsplice (tr1 tr2 : nat_tree) {struct tr1} : nat_tree :=
adamc@35 964 match tr1 with
adamc@35 965 | NLeaf' => NNode' O (Cons tr2 Nil)
adamc@35 966 | NNode' n Nil => NNode' n (Cons tr2 Nil)
adamc@35 967 | NNode' n (Cons tr trs) => NNode' n (Cons (ntsplice tr tr2) trs)
adamc@35 968 end.
adamc@35 969
adamc@35 970 (** We have defined another arbitrary notion of tree splicing, similar to before, and we can prove an analogous theorem about its relationship with tree size. We start with a useful lemma about addition. *)
adamc@35 971
adamc@35 972 Lemma plus_S : forall n1 n2 : nat,
adamc@35 973 plus n1 (S n2) = S (plus n1 n2).
adamc@35 974 induction n1; crush.
adamc@35 975 Qed.
adamc@35 976
adamc@35 977 (** Now we begin the proof of the theorem, adding the lemma [plus_S] as a hint. *)
adamc@35 978
adamc@35 979 Theorem ntsize_ntsplice : forall tr1 tr2 : nat_tree, ntsize (ntsplice tr1 tr2)
adamc@35 980 = plus (ntsize tr2) (ntsize tr1).
adamc@35 981 Hint Rewrite plus_S : cpdt.
adamc@35 982
adamc@35 983 (** We know that the standard induction principle is insufficient for the task, so we need to provide a [using] clause for the [induction] tactic to specify our alternate principle. *)
adamc@35 984 induction tr1 using nat_tree_ind'; crush.
adamc@35 985
adamc@35 986 (** One subgoal remains: [[
adamc@35 987
adamc@35 988 n : nat
adamc@35 989 ls : list nat_tree
adamc@35 990 H : All
adamc@35 991 (fun tr1 : nat_tree =>
adamc@35 992 forall tr2 : nat_tree,
adamc@35 993 ntsize (ntsplice tr1 tr2) = plus (ntsize tr2) (ntsize tr1)) ls
adamc@35 994 tr2 : nat_tree
adamc@35 995 ============================
adamc@35 996 ntsize
adamc@35 997 match ls with
adamc@35 998 | Nil => NNode' n (Cons tr2 Nil)
adamc@35 999 | Cons tr trs => NNode' n (Cons (ntsplice tr tr2) trs)
adamc@35 1000 end = S (plus (ntsize tr2) (sum (map ntsize ls)))
adamc@35 1001 ]]
adamc@35 1002
adamc@35 1003 After a few moments of squinting at this goal, it becomes apparent that we need to do a case analysis on the structure of [ls]. The rest is routine. *)
adamc@35 1004
adamc@35 1005 destruct ls; crush.
adamc@35 1006
adamc@35 1007 (** We can go further in automating the proof by exploiting the hint mechanism further. *)
adamc@35 1008
adamc@35 1009 Restart.
adamc@35 1010 Hint Extern 1 (ntsize (match ?LS with Nil => _ | Cons _ _ => _ end) = _) =>
adamc@35 1011 destruct LS; crush.
adamc@35 1012 induction tr1 using nat_tree_ind'; crush.
adamc@35 1013 Qed.
adamc@35 1014
adamc@35 1015 (** We will go into great detail on hints in a later chapter, but the only important thing to note here is that we register a pattern that describes a conclusion we expect to encounter during the proof. The pattern may contain unification variables, whose names are prefixed with question marks, and we may refer to those bound variables in a tactic that we ask to have run whenever the pattern matches.
adamc@35 1016
adamc@35 1017 The advantage of using the hint is not very clear here, because the original proof was so short. However, the hint has fundamentally improved the readability of our proof. Before, the proof refered to the local variable [ls], which has an automatically-generated name. To a human reading the proof script without stepping through it interactively, it was not clear where [ls] came from. The hint explains to the reader the process for choosing which variables to case analyze on, and the hint can continue working even if the rest of the proof structure changes significantly. *)