annotate src/InductiveTypes.v @ 31:1a82839f83b7

Mutual induction
author Adam Chlipala <adamc@hcoop.net>
date Mon, 08 Sep 2008 16:00:02 -0400
parents 4887ddb1ad23
children 77e1a7eda0b8
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@26 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 dependent function types.)
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@29 329 | NLeaf => 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@29 335 | NLeaf => tr2
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@30 418 (** After we end the section, the [Variable]s we used are added as extra function parameters for each defined identifier, as needed. *)
adamc@30 419
adamc@30 420 Check list.
adamc@30 421 (** [[
adamc@30 422
adamc@30 423 list
adamc@30 424 : Set -> Set
adamc@30 425 ]] *)
adamc@30 426
adamc@30 427 Check Cons.
adamc@30 428 (** [[
adamc@30 429
adamc@30 430 Cons
adamc@30 431 : forall T : Set, T -> list T -> list T
adamc@30 432 ]] *)
adamc@30 433
adamc@30 434 Check length.
adamc@30 435 (** [[
adamc@30 436
adamc@30 437 length
adamc@30 438 : forall T : Set, list T -> nat
adamc@30 439 ]]
adamc@30 440
adamc@30 441 The extra parameter [T] is treated as a new argument to the induction principle, too. *)
adamc@30 442
adamc@30 443 Check list_ind.
adamc@30 444 (** [[
adamc@30 445
adamc@30 446 list_ind
adamc@30 447 : forall (T : Set) (P : list T -> Prop),
adamc@30 448 P (Nil T) ->
adamc@30 449 (forall (t : T) (l : list T), P l -> P (Cons t l)) ->
adamc@30 450 forall l : list T, P l
adamc@30 451 ]]
adamc@30 452
adamc@30 453 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 454
adamc@31 455
adamc@31 456 (** * Mutually Inductive Types *)
adamc@31 457
adamc@31 458 (** We can define inductive types that refer to each other: *)
adamc@31 459
adamc@31 460 Inductive even_list : Set :=
adamc@31 461 | ENil : even_list
adamc@31 462 | ECons : nat -> odd_list -> even_list
adamc@31 463
adamc@31 464 with odd_list : Set :=
adamc@31 465 | OCons : nat -> even_list -> odd_list.
adamc@31 466
adamc@31 467 Fixpoint elength (el : even_list) : nat :=
adamc@31 468 match el with
adamc@31 469 | ENil => O
adamc@31 470 | ECons _ ol => S (olength ol)
adamc@31 471 end
adamc@31 472
adamc@31 473 with olength (ol : odd_list) : nat :=
adamc@31 474 match ol with
adamc@31 475 | OCons _ el => S (elength el)
adamc@31 476 end.
adamc@31 477
adamc@31 478 Fixpoint eapp (el1 el2 : even_list) {struct el1} : even_list :=
adamc@31 479 match el1 with
adamc@31 480 | ENil => el2
adamc@31 481 | ECons n ol => ECons n (oapp ol el2)
adamc@31 482 end
adamc@31 483
adamc@31 484 with oapp (ol : odd_list) (el : even_list) {struct ol} : odd_list :=
adamc@31 485 match ol with
adamc@31 486 | OCons n el' => OCons n (eapp el' el)
adamc@31 487 end.
adamc@31 488
adamc@31 489 (** 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 490
adamc@31 491 Theorem elength_eapp : forall el1 el2 : even_list,
adamc@31 492 elength (eapp el1 el2) = plus (elength el1) (elength el2).
adamc@31 493 induction el1; crush.
adamc@31 494
adamc@31 495 (** One goal remains: [[
adamc@31 496
adamc@31 497 n : nat
adamc@31 498 o : odd_list
adamc@31 499 el2 : even_list
adamc@31 500 ============================
adamc@31 501 S (olength (oapp o el2)) = S (plus (olength o) (elength el2))
adamc@31 502 ]]
adamc@31 503
adamc@31 504 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 505
adamc@31 506 Abort.
adamc@31 507 Check even_list_ind.
adamc@31 508 (** [[
adamc@31 509
adamc@31 510 even_list_ind
adamc@31 511 : forall P : even_list -> Prop,
adamc@31 512 P ENil ->
adamc@31 513 (forall (n : nat) (o : odd_list), P (ECons n o)) ->
adamc@31 514 forall e : even_list, P e
adamc@31 515 ]]
adamc@31 516
adamc@31 517 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 518
adamc@31 519 Scheme even_list_mut := Induction for even_list Sort Prop
adamc@31 520 with odd_list_mut := Induction for odd_list Sort Prop.
adamc@31 521
adamc@31 522 Check even_list_mut.
adamc@31 523 (** [[
adamc@31 524
adamc@31 525 even_list_mut
adamc@31 526 : forall (P : even_list -> Prop) (P0 : odd_list -> Prop),
adamc@31 527 P ENil ->
adamc@31 528 (forall (n : nat) (o : odd_list), P0 o -> P (ECons n o)) ->
adamc@31 529 (forall (n : nat) (e : even_list), P e -> P0 (OCons n e)) ->
adamc@31 530 forall e : even_list, P e
adamc@31 531 ]]
adamc@31 532
adamc@31 533 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 534
adamc@31 535 Theorem n_plus_O' : forall n : nat, plus n O = n.
adamc@31 536 apply (nat_ind (fun n => plus n O = n)); crush.
adamc@31 537 Qed.
adamc@31 538
adamc@31 539 (** 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 540
adamc@31 541 This technique generalizes to our mutual example: *)
adamc@31 542
adamc@31 543 Theorem elength_eapp : forall el1 el2 : even_list,
adamc@31 544 elength (eapp el1 el2) = plus (elength el1) (elength el2).
adamc@31 545 apply (even_list_mut
adamc@31 546 (fun el1 : even_list => forall el2 : even_list,
adamc@31 547 elength (eapp el1 el2) = plus (elength el1) (elength el2))
adamc@31 548 (fun ol : odd_list => forall el : even_list,
adamc@31 549 olength (oapp ol el) = plus (olength ol) (elength el))); crush.
adamc@31 550 Qed.
adamc@31 551
adamc@31 552 (** 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. *)