annotate src/InductiveTypes.v @ 205:f05514cc6c0d

'make doc' works with 8.2
author Adam Chlipala <adamc@hcoop.net>
date Fri, 06 Nov 2009 12:15:05 -0500
parents 8caa3b3f8fc0
children b9e9ff52913c
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@74 19 (** %\part{Basic Programming and Proving}
adamc@74 20
adamc@74 21 \chapter{Introducing Inductive Types}% *)
adamc@26 22
adamc@45 23 (** 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 for functional programming in Coq. *)
adamc@26 24
adamc@26 25
adamc@26 26 (** * Enumerations *)
adamc@26 27
adamc@26 28 (** 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 29
adamc@26 30 The singleton type [unit] is an inductive type: *)
adamc@26 31
adamc@26 32 Inductive unit : Set :=
adamc@26 33 | tt.
adamc@26 34
adamc@26 35 (** 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 36
adamc@26 37 Check unit.
adamc@26 38 (** [[
adamc@26 39
adamc@26 40 unit : Set
adamc@26 41 ]] *)
adamc@26 42 Check tt.
adamc@26 43 (** [[
adamc@26 44
adamc@26 45 tt : unit
adamc@26 46 ]] *)
adamc@26 47
adamc@26 48 (** We can prove that [unit] is a genuine singleton type. *)
adamc@26 49
adamc@26 50 Theorem unit_singleton : forall x : unit, x = tt.
adamc@26 51 (** 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@41 52 (* begin thide *)
adamc@26 53 induction x.
adamc@26 54 (** The goal changes to: [[
adamc@26 55
adamc@26 56 tt = tt
adamc@26 57 ]] *)
adamc@26 58 (** ...which we can discharge trivially. *)
adamc@26 59 reflexivity.
adamc@26 60 Qed.
adamc@41 61 (* end thide *)
adamc@26 62
adamc@26 63 (** 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 64
adamc@26 65 destruct x.
adamc@205 66
adamc@205 67 ]]
adamc@205 68
adamc@26 69 ...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 70
adamc@26 71 What exactly %\textit{%#<i>#is#</i>#%}% the induction principle for [unit]? We can ask Coq: *)
adamc@26 72
adamc@26 73 Check unit_ind.
adamc@26 74 (** [[
adamc@26 75
adamc@26 76 unit_ind : forall P : unit -> Prop, P tt -> forall u : unit, P u
adamc@26 77 ]]
adamc@26 78
adamc@26 79 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 80
adamc@26 81 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 82
adamc@26 83 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 84
adamc@26 85 %\medskip%
adamc@26 86
adamc@26 87 We can define an inductive type even simpler than [unit]: *)
adamc@26 88
adamc@26 89 Inductive Empty_set : Set := .
adamc@26 90
adamc@26 91 (** [Empty_set] has no elements. We can prove fun theorems about it: *)
adamc@26 92
adamc@26 93 Theorem the_sky_is_falling : forall x : Empty_set, 2 + 2 = 5.
adamc@41 94 (* begin thide *)
adamc@26 95 destruct 1.
adamc@26 96 Qed.
adamc@41 97 (* end thide *)
adamc@26 98
adamc@32 99 (** 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 100
adamc@26 101 We can see the induction principle that made this proof so easy: *)
adamc@26 102
adamc@26 103 Check Empty_set_ind.
adamc@26 104 (** [[
adamc@26 105
adamc@26 106 Empty_set_ind : forall (P : Empty_set -> Prop) (e : Empty_set), P e
adamc@26 107 ]]
adamc@26 108
adamc@26 109 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 110
adamc@26 111 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 112
adamc@26 113 Definition e2u (e : Empty_set) : unit := match e with end.
adamc@26 114
adamc@26 115 (** 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 116
adamc@26 117 %\medskip%
adamc@26 118
adamc@26 119 Moving up the ladder of complexity, we can define the booleans: *)
adamc@26 120
adamc@26 121 Inductive bool : Set :=
adamc@26 122 | true
adamc@26 123 | false.
adamc@26 124
adamc@26 125 (** We can use less vacuous pattern matching to define boolean negation. *)
adamc@26 126
adamc@26 127 Definition not (b : bool) : bool :=
adamc@26 128 match b with
adamc@26 129 | true => false
adamc@26 130 | false => true
adamc@26 131 end.
adamc@26 132
adamc@27 133 (** An alternative definition desugars to the above: *)
adamc@27 134
adamc@27 135 Definition not' (b : bool) : bool :=
adamc@27 136 if b then false else true.
adamc@27 137
adamc@26 138 (** We might want to prove that [not] is its own inverse operation. *)
adamc@26 139
adamc@26 140 Theorem not_inverse : forall b : bool, not (not b) = b.
adamc@41 141 (* begin thide *)
adamc@26 142 destruct b.
adamc@26 143
adamc@26 144 (** After we case analyze on [b], we are left with one subgoal for each constructor of [bool].
adamc@26 145
adamc@26 146 [[
adamc@26 147
adamc@26 148 2 subgoals
adamc@26 149
adamc@26 150 ============================
adamc@26 151 not (not true) = true
adamc@26 152 ]]
adamc@26 153
adamc@26 154 [[
adamc@26 155 subgoal 2 is:
adamc@26 156 not (not false) = false
adamc@26 157 ]]
adamc@26 158
adamc@26 159 The first subgoal follows by Coq's rules of computation, so we can dispatch it easily: *)
adamc@26 160
adamc@26 161 reflexivity.
adamc@26 162
adamc@26 163 (** Likewise for the second subgoal, so we can restart the proof and give a very compact justification. *)
adamc@26 164
adamc@26 165 Restart.
adamc@26 166 destruct b; reflexivity.
adamc@26 167 Qed.
adamc@41 168 (* end thide *)
adamc@27 169
adamc@27 170 (** Another theorem about booleans illustrates another useful tactic. *)
adamc@27 171
adamc@27 172 Theorem not_ineq : forall b : bool, not b <> b.
adamc@41 173 (* begin thide *)
adamc@27 174 destruct b; discriminate.
adamc@27 175 Qed.
adamc@41 176 (* end thide *)
adamc@27 177
adamc@27 178 (** [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 179
adamc@27 180 At this point, it is probably not hard to guess what the underlying induction principle for [bool] is. *)
adamc@27 181
adamc@27 182 Check bool_ind.
adamc@27 183 (** [[
adamc@27 184
adamc@27 185 bool_ind : forall P : bool -> Prop, P true -> P false -> forall b : bool, P b
adamc@27 186 ]] *)
adamc@28 187
adamc@28 188
adamc@28 189 (** * Simple Recursive Types *)
adamc@28 190
adamc@28 191 (** The natural numbers are the simplest common example of an inductive type that actually deserves the name. *)
adamc@28 192
adamc@28 193 Inductive nat : Set :=
adamc@28 194 | O : nat
adamc@28 195 | S : nat -> nat.
adamc@28 196
adamc@28 197 (** [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 198
adamc@28 199 Pattern matching works as we demonstrated in the last chapter: *)
adamc@28 200
adamc@28 201 Definition isZero (n : nat) : bool :=
adamc@28 202 match n with
adamc@28 203 | O => true
adamc@28 204 | S _ => false
adamc@28 205 end.
adamc@28 206
adamc@28 207 Definition pred (n : nat) : nat :=
adamc@28 208 match n with
adamc@28 209 | O => O
adamc@28 210 | S n' => n'
adamc@28 211 end.
adamc@28 212
adamc@28 213 (** We can prove theorems by case analysis: *)
adamc@28 214
adamc@28 215 Theorem S_isZero : forall n : nat, isZero (pred (S (S n))) = false.
adamc@41 216 (* begin thide *)
adamc@28 217 destruct n; reflexivity.
adamc@28 218 Qed.
adamc@41 219 (* end thide *)
adamc@28 220
adamc@28 221 (** We can also now get into genuine inductive theorems. First, we will need a recursive function, to make things interesting. *)
adamc@28 222
adamc@28 223 Fixpoint plus (n m : nat) {struct n} : nat :=
adamc@28 224 match n with
adamc@28 225 | O => m
adamc@28 226 | S n' => S (plus n' m)
adamc@28 227 end.
adamc@28 228
adamc@28 229 (** 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 230
adamc@28 231 Some theorems about [plus] can be proved without induction. *)
adamc@28 232
adamc@28 233 Theorem O_plus_n : forall n : nat, plus O n = n.
adamc@41 234 (* begin thide *)
adamc@28 235 intro; reflexivity.
adamc@28 236 Qed.
adamc@41 237 (* end thide *)
adamc@28 238
adamc@28 239 (** 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 240
adamc@28 241 Theorem n_plus_O : forall n : nat, plus n O = n.
adamc@41 242 (* begin thide *)
adamc@28 243 induction n.
adamc@28 244
adamc@28 245 (** Our first subgoal is [plus O O = O], which %\textit{%#<i>#is#</i>#%}% trivial by computation. *)
adamc@28 246
adamc@28 247 reflexivity.
adamc@28 248
adamc@28 249 (** Our second subgoal is more work and also demonstrates our first inductive hypothesis.
adamc@28 250
adamc@28 251 [[
adamc@28 252
adamc@28 253 n : nat
adamc@28 254 IHn : plus n O = n
adamc@28 255 ============================
adamc@28 256 plus (S n) O = S n
adamc@28 257 ]]
adamc@28 258
adamc@28 259 We can start out by using computation to simplify the goal as far as we can. *)
adamc@28 260
adamc@28 261 simpl.
adamc@28 262
adamc@28 263 (** Now the conclusion is [S (plus n O) = S n]. Using our inductive hypothesis: *)
adamc@28 264
adamc@28 265 rewrite IHn.
adamc@28 266
adamc@28 267 (** ...we get a trivial conclusion [S n = S n]. *)
adamc@28 268
adamc@28 269 reflexivity.
adamc@28 270
adamc@28 271 (** Not much really went on in this proof, so the [crush] tactic from the [Tactics] module can prove this theorem automatically. *)
adamc@28 272
adamc@28 273 Restart.
adamc@28 274 induction n; crush.
adamc@28 275 Qed.
adamc@41 276 (* end thide *)
adamc@28 277
adamc@28 278 (** We can check out the induction principle at work here: *)
adamc@28 279
adamc@28 280 Check nat_ind.
adamc@28 281 (** [[
adamc@28 282
adamc@28 283 nat_ind : forall P : nat -> Prop,
adamc@28 284 P O -> (forall n : nat, P n -> P (S n)) -> forall n : nat, P n
adamc@28 285 ]]
adamc@28 286
adamc@28 287 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 288
adamc@28 289 Since [nat] has a constructor that takes an argument, we may sometimes need to know that that constructor is injective. *)
adamc@28 290
adamc@28 291 Theorem S_inj : forall n m : nat, S n = S m -> n = m.
adamc@41 292 (* begin thide *)
adamc@28 293 injection 1; trivial.
adamc@28 294 Qed.
adamc@41 295 (* end thide *)
adamc@28 296
adamc@28 297 (** [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 298
adamc@29 299 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 300
adamc@29 301 %\medskip%
adamc@29 302
adamc@29 303 We can define a type of lists of natural numbers. *)
adamc@29 304
adamc@29 305 Inductive nat_list : Set :=
adamc@29 306 | NNil : nat_list
adamc@29 307 | NCons : nat -> nat_list -> nat_list.
adamc@29 308
adamc@29 309 (** Recursive definitions are straightforward extensions of what we have seen before. *)
adamc@29 310
adamc@29 311 Fixpoint nlength (ls : nat_list) : nat :=
adamc@29 312 match ls with
adamc@29 313 | NNil => O
adamc@29 314 | NCons _ ls' => S (nlength ls')
adamc@29 315 end.
adamc@29 316
adamc@29 317 Fixpoint napp (ls1 ls2 : nat_list) {struct ls1} : nat_list :=
adamc@29 318 match ls1 with
adamc@29 319 | NNil => ls2
adamc@29 320 | NCons n ls1' => NCons n (napp ls1' ls2)
adamc@29 321 end.
adamc@29 322
adamc@29 323 (** Inductive theorem proving can again be automated quite effectively. *)
adamc@29 324
adamc@29 325 Theorem nlength_napp : forall ls1 ls2 : nat_list, nlength (napp ls1 ls2)
adamc@29 326 = plus (nlength ls1) (nlength ls2).
adamc@41 327 (* begin thide *)
adamc@29 328 induction ls1; crush.
adamc@29 329 Qed.
adamc@41 330 (* end thide *)
adamc@29 331
adamc@29 332 Check nat_list_ind.
adamc@29 333 (** [[
adamc@29 334
adamc@29 335 nat_list_ind
adamc@29 336 : forall P : nat_list -> Prop,
adamc@29 337 P NNil ->
adamc@29 338 (forall (n : nat) (n0 : nat_list), P n0 -> P (NCons n n0)) ->
adamc@29 339 forall n : nat_list, P n
adamc@29 340 ]]
adamc@29 341
adamc@29 342 %\medskip%
adamc@29 343
adamc@29 344 In general, we can implement any "tree" types as inductive types. For example, here are binary trees of naturals. *)
adamc@29 345
adamc@29 346 Inductive nat_btree : Set :=
adamc@29 347 | NLeaf : nat_btree
adamc@29 348 | NNode : nat_btree -> nat -> nat_btree -> nat_btree.
adamc@29 349
adamc@29 350 Fixpoint nsize (tr : nat_btree) : nat :=
adamc@29 351 match tr with
adamc@35 352 | NLeaf => S O
adamc@29 353 | NNode tr1 _ tr2 => plus (nsize tr1) (nsize tr2)
adamc@29 354 end.
adamc@29 355
adamc@29 356 Fixpoint nsplice (tr1 tr2 : nat_btree) {struct tr1} : nat_btree :=
adamc@29 357 match tr1 with
adamc@35 358 | NLeaf => NNode tr2 O NLeaf
adamc@29 359 | NNode tr1' n tr2' => NNode (nsplice tr1' tr2) n tr2'
adamc@29 360 end.
adamc@29 361
adamc@29 362 Theorem plus_assoc : forall n1 n2 n3 : nat, plus (plus n1 n2) n3 = plus n1 (plus n2 n3).
adamc@41 363 (* begin thide *)
adamc@29 364 induction n1; crush.
adamc@29 365 Qed.
adamc@41 366 (* end thide *)
adamc@29 367
adamc@29 368 Theorem nsize_nsplice : forall tr1 tr2 : nat_btree, nsize (nsplice tr1 tr2)
adamc@29 369 = plus (nsize tr2) (nsize tr1).
adamc@41 370 (* begin thide *)
adamc@29 371 Hint Rewrite n_plus_O plus_assoc : cpdt.
adamc@29 372
adamc@29 373 induction tr1; crush.
adamc@29 374 Qed.
adamc@41 375 (* end thide *)
adamc@29 376
adamc@29 377 Check nat_btree_ind.
adamc@29 378 (** [[
adamc@29 379
adamc@29 380 nat_btree_ind
adamc@29 381 : forall P : nat_btree -> Prop,
adamc@29 382 P NLeaf ->
adamc@29 383 (forall n : nat_btree,
adamc@29 384 P n -> forall (n0 : nat) (n1 : nat_btree), P n1 -> P (NNode n n0 n1)) ->
adamc@29 385 forall n : nat_btree, P n
adamc@29 386 ]] *)
adamc@30 387
adamc@30 388
adamc@30 389 (** * Parameterized Types *)
adamc@30 390
adamc@30 391 (** We can also define polymorphic inductive types, as with algebraic datatypes in Haskell and ML. *)
adamc@30 392
adamc@30 393 Inductive list (T : Set) : Set :=
adamc@30 394 | Nil : list T
adamc@30 395 | Cons : T -> list T -> list T.
adamc@30 396
adamc@30 397 Fixpoint length T (ls : list T) : nat :=
adamc@30 398 match ls with
adamc@30 399 | Nil => O
adamc@30 400 | Cons _ ls' => S (length ls')
adamc@30 401 end.
adamc@30 402
adamc@30 403 Fixpoint app T (ls1 ls2 : list T) {struct ls1} : list T :=
adamc@30 404 match ls1 with
adamc@30 405 | Nil => ls2
adamc@30 406 | Cons x ls1' => Cons x (app ls1' ls2)
adamc@30 407 end.
adamc@30 408
adamc@30 409 Theorem length_app : forall T (ls1 ls2 : list T), length (app ls1 ls2)
adamc@30 410 = plus (length ls1) (length ls2).
adamc@41 411 (* begin thide *)
adamc@30 412 induction ls1; crush.
adamc@30 413 Qed.
adamc@41 414 (* end thide *)
adamc@30 415
adamc@30 416 (** 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 417
adamc@30 418 (* begin hide *)
adamc@30 419 Reset list.
adamc@30 420 (* end hide *)
adamc@30 421
adamc@30 422 Section list.
adamc@30 423 Variable T : Set.
adamc@30 424
adamc@30 425 Inductive list : Set :=
adamc@30 426 | Nil : list
adamc@30 427 | Cons : T -> list -> list.
adamc@30 428
adamc@30 429 Fixpoint length (ls : list) : nat :=
adamc@30 430 match ls with
adamc@30 431 | Nil => O
adamc@30 432 | Cons _ ls' => S (length ls')
adamc@30 433 end.
adamc@30 434
adamc@30 435 Fixpoint app (ls1 ls2 : list) {struct ls1} : list :=
adamc@30 436 match ls1 with
adamc@30 437 | Nil => ls2
adamc@30 438 | Cons x ls1' => Cons x (app ls1' ls2)
adamc@30 439 end.
adamc@30 440
adamc@30 441 Theorem length_app : forall ls1 ls2 : list, length (app ls1 ls2)
adamc@30 442 = plus (length ls1) (length ls2).
adamc@41 443 (* begin thide *)
adamc@30 444 induction ls1; crush.
adamc@30 445 Qed.
adamc@41 446 (* end thide *)
adamc@30 447 End list.
adamc@30 448
adamc@35 449 (* begin hide *)
adamc@35 450 Implicit Arguments Nil [T].
adamc@35 451 (* end hide *)
adamc@35 452
adamc@30 453 (** After we end the section, the [Variable]s we used are added as extra function parameters for each defined identifier, as needed. *)
adamc@30 454
adamc@202 455 Print list.
adamc@30 456 (** [[
adamc@30 457
adamc@30 458
adamc@202 459 Inductive list (T : Set) : Set :=
adamc@202 460 Nil : list T | Cons : T -> list T -> list Tlist
adamc@202 461 ]]
adamc@30 462
adamc@202 463 The final definition is the same as what we wrote manually before. The other elements of the section are altered similarly, turning out exactly as they were before, though we managed to write their definitions more succinctly. *)
adamc@30 464
adamc@30 465 Check length.
adamc@30 466 (** [[
adamc@30 467
adamc@30 468 length
adamc@30 469 : forall T : Set, list T -> nat
adamc@30 470 ]]
adamc@30 471
adamc@202 472 The parameter [T] is treated as a new argument to the induction principle, too. *)
adamc@30 473
adamc@30 474 Check list_ind.
adamc@30 475 (** [[
adamc@30 476
adamc@30 477 list_ind
adamc@30 478 : forall (T : Set) (P : list T -> Prop),
adamc@30 479 P (Nil T) ->
adamc@30 480 (forall (t : T) (l : list T), P l -> P (Cons t l)) ->
adamc@30 481 forall l : list T, P l
adamc@30 482 ]]
adamc@30 483
adamc@30 484 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 485
adamc@31 486
adamc@31 487 (** * Mutually Inductive Types *)
adamc@31 488
adamc@31 489 (** We can define inductive types that refer to each other: *)
adamc@31 490
adamc@31 491 Inductive even_list : Set :=
adamc@31 492 | ENil : even_list
adamc@31 493 | ECons : nat -> odd_list -> even_list
adamc@31 494
adamc@31 495 with odd_list : Set :=
adamc@31 496 | OCons : nat -> even_list -> odd_list.
adamc@31 497
adamc@31 498 Fixpoint elength (el : even_list) : nat :=
adamc@31 499 match el with
adamc@31 500 | ENil => O
adamc@31 501 | ECons _ ol => S (olength ol)
adamc@31 502 end
adamc@31 503
adamc@31 504 with olength (ol : odd_list) : nat :=
adamc@31 505 match ol with
adamc@31 506 | OCons _ el => S (elength el)
adamc@31 507 end.
adamc@31 508
adamc@31 509 Fixpoint eapp (el1 el2 : even_list) {struct el1} : even_list :=
adamc@31 510 match el1 with
adamc@31 511 | ENil => el2
adamc@31 512 | ECons n ol => ECons n (oapp ol el2)
adamc@31 513 end
adamc@31 514
adamc@31 515 with oapp (ol : odd_list) (el : even_list) {struct ol} : odd_list :=
adamc@31 516 match ol with
adamc@31 517 | OCons n el' => OCons n (eapp el' el)
adamc@31 518 end.
adamc@31 519
adamc@31 520 (** 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 521
adamc@31 522 Theorem elength_eapp : forall el1 el2 : even_list,
adamc@31 523 elength (eapp el1 el2) = plus (elength el1) (elength el2).
adamc@41 524 (* begin thide *)
adamc@31 525 induction el1; crush.
adamc@31 526
adamc@31 527 (** One goal remains: [[
adamc@31 528
adamc@31 529 n : nat
adamc@31 530 o : odd_list
adamc@31 531 el2 : even_list
adamc@31 532 ============================
adamc@31 533 S (olength (oapp o el2)) = S (plus (olength o) (elength el2))
adamc@31 534 ]]
adamc@31 535
adamc@31 536 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 537
adamc@31 538 Abort.
adamc@31 539 Check even_list_ind.
adamc@31 540 (** [[
adamc@31 541
adamc@31 542 even_list_ind
adamc@31 543 : forall P : even_list -> Prop,
adamc@31 544 P ENil ->
adamc@31 545 (forall (n : nat) (o : odd_list), P (ECons n o)) ->
adamc@31 546 forall e : even_list, P e
adamc@31 547 ]]
adamc@31 548
adamc@31 549 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 550
adamc@31 551 Scheme even_list_mut := Induction for even_list Sort Prop
adamc@31 552 with odd_list_mut := Induction for odd_list Sort Prop.
adamc@31 553
adamc@31 554 Check even_list_mut.
adamc@31 555 (** [[
adamc@31 556
adamc@31 557 even_list_mut
adamc@31 558 : forall (P : even_list -> Prop) (P0 : odd_list -> Prop),
adamc@31 559 P ENil ->
adamc@31 560 (forall (n : nat) (o : odd_list), P0 o -> P (ECons n o)) ->
adamc@31 561 (forall (n : nat) (e : even_list), P e -> P0 (OCons n e)) ->
adamc@31 562 forall e : even_list, P e
adamc@31 563 ]]
adamc@31 564
adamc@31 565 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 566
adamc@31 567 Theorem n_plus_O' : forall n : nat, plus n O = n.
adamc@31 568 apply (nat_ind (fun n => plus n O = n)); crush.
adamc@31 569 Qed.
adamc@31 570
adamc@31 571 (** 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 572
adamc@31 573 This technique generalizes to our mutual example: *)
adamc@31 574
adamc@31 575 Theorem elength_eapp : forall el1 el2 : even_list,
adamc@31 576 elength (eapp el1 el2) = plus (elength el1) (elength el2).
adamc@41 577
adamc@31 578 apply (even_list_mut
adamc@31 579 (fun el1 : even_list => forall el2 : even_list,
adamc@31 580 elength (eapp el1 el2) = plus (elength el1) (elength el2))
adamc@31 581 (fun ol : odd_list => forall el : even_list,
adamc@31 582 olength (oapp ol el) = plus (olength ol) (elength el))); crush.
adamc@31 583 Qed.
adamc@41 584 (* end thide *)
adamc@31 585
adamc@31 586 (** 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 587
adamc@33 588
adamc@33 589 (** * Reflexive Types *)
adamc@33 590
adamc@33 591 (** 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 592
adamc@33 593 Inductive formula : Set :=
adamc@33 594 | Eq : nat -> nat -> formula
adamc@33 595 | And : formula -> formula -> formula
adamc@33 596 | Forall : (nat -> formula) -> formula.
adamc@33 597
adamc@33 598 (** 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 599
adamc@33 600 Example forall_refl : formula := Forall (fun x => Eq x x).
adamc@33 601
adamc@33 602 (** We can write recursive functions over reflexive types quite naturally. Here is one translating our formulas into native Coq propositions. *)
adamc@33 603
adamc@33 604 Fixpoint formulaDenote (f : formula) : Prop :=
adamc@33 605 match f with
adamc@33 606 | Eq n1 n2 => n1 = n2
adamc@33 607 | And f1 f2 => formulaDenote f1 /\ formulaDenote f2
adamc@33 608 | Forall f' => forall n : nat, formulaDenote (f' n)
adamc@33 609 end.
adamc@33 610
adamc@33 611 (** We can also encode a trivial formula transformation that swaps the order of equality and conjunction operands. *)
adamc@33 612
adamc@33 613 Fixpoint swapper (f : formula) : formula :=
adamc@33 614 match f with
adamc@33 615 | Eq n1 n2 => Eq n2 n1
adamc@33 616 | And f1 f2 => And (swapper f2) (swapper f1)
adamc@33 617 | Forall f' => Forall (fun n => swapper (f' n))
adamc@33 618 end.
adamc@33 619
adamc@33 620 (** It is helpful to prove that this transformation does not make true formulas false. *)
adamc@33 621
adamc@33 622 Theorem swapper_preserves_truth : forall f, formulaDenote f -> formulaDenote (swapper f).
adamc@41 623 (* begin thide *)
adamc@33 624 induction f; crush.
adamc@33 625 Qed.
adamc@41 626 (* end thide *)
adamc@33 627
adamc@33 628 (** We can take a look at the induction principle behind this proof. *)
adamc@33 629
adamc@33 630 Check formula_ind.
adamc@33 631 (** [[
adamc@33 632
adamc@33 633 formula_ind
adamc@33 634 : forall P : formula -> Prop,
adamc@33 635 (forall n n0 : nat, P (Eq n n0)) ->
adamc@33 636 (forall f0 : formula,
adamc@33 637 P f0 -> forall f1 : formula, P f1 -> P (And f0 f1)) ->
adamc@33 638 (forall f1 : nat -> formula,
adamc@33 639 (forall n : nat, P (f1 n)) -> P (Forall f1)) ->
adamc@33 640 forall f2 : formula, P f2
adamc@33 641 ]] *)
adamc@33 642
adamc@33 643 (** 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 644
adamc@33 645 %\medskip%
adamc@33 646
adamc@33 647 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 648
adamc@33 649 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 650
adamc@33 651 (** [[
adamc@33 652
adamc@33 653 Inductive term : Set :=
adamc@33 654 | App : term -> term -> term
adamc@33 655 | Abs : (term -> term) -> term.
adamc@33 656
adamc@33 657 [[
adamc@33 658 Error: Non strictly positive occurrence of "term" in "(term -> term) -> term"
adamc@33 659 ]]
adamc@33 660
adamc@33 661 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 662
adamc@33 663 Why must Coq enforce this restriction? Imagine that our last definition had been accepted, allowing us to write this function:
adamc@33 664
adamc@33 665 [[
adamc@33 666 Definition uhoh (t : term) : term :=
adamc@33 667 match t with
adamc@33 668 | Abs f => f t
adamc@33 669 | _ => t
adamc@33 670 end.
adamc@33 671
adamc@205 672 ]]
adamc@205 673
adamc@33 674 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 675
adamc@33 676 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 677
adamc@33 678 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 679
adamc@34 680
adamc@34 681 (** * An Interlude on Proof Terms *)
adamc@34 682
adamc@34 683 (** 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 684
adamc@34 685 Print unit_ind.
adamc@34 686 (** [[
adamc@34 687
adamc@34 688 unit_ind =
adamc@34 689 fun P : unit -> Prop => unit_rect P
adamc@34 690 : forall P : unit -> Prop, P tt -> forall u : unit, P u
adamc@34 691 ]]
adamc@34 692
adamc@34 693 We see that this induction principle is defined in terms of a more general principle, [unit_rect]. *)
adamc@34 694
adamc@34 695 Check unit_rect.
adamc@34 696 (** [[
adamc@34 697
adamc@34 698 unit_rect
adamc@34 699 : forall P : unit -> Type, P tt -> forall u : unit, P u
adamc@34 700 ]]
adamc@34 701
adamc@34 702 [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 703
adamc@34 704 Print unit_rec.
adamc@34 705 (** [[
adamc@34 706
adamc@34 707 unit_rec =
adamc@34 708 fun P : unit -> Set => unit_rect P
adamc@34 709 : forall P : unit -> Set, P tt -> forall u : unit, P u
adamc@34 710 ]]
adamc@34 711
adamc@34 712 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 713
adamc@34 714 Definition always_O (u : unit) : nat :=
adamc@34 715 match u with
adamc@34 716 | tt => O
adamc@34 717 end.
adamc@34 718
adamc@34 719 Definition always_O' (u : unit) : nat :=
adamc@34 720 unit_rec (fun _ : unit => nat) O u.
adamc@34 721
adamc@34 722 (** 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 723
adamc@34 724 Print unit_rect.
adamc@34 725
adamc@34 726 (** [[
adamc@34 727
adamc@34 728 unit_rect =
adamc@34 729 fun (P : unit -> Type) (f : P tt) (u : unit) =>
adamc@34 730 match u as u0 return (P u0) with
adamc@34 731 | tt => f
adamc@34 732 end
adamc@34 733 : forall P : unit -> Type, P tt -> forall u : unit, P u
adamc@34 734 ]]
adamc@34 735
adamc@34 736 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 737
adamc@34 738 To prove that [unit_rect] is nothing special, we can reimplement it manually. *)
adamc@34 739
adamc@34 740 Definition unit_rect' (P : unit -> Type) (f : P tt) (u : unit) :=
adamc@34 741 match u return (P u) with
adamc@34 742 | tt => f
adamc@34 743 end.
adamc@34 744
adamc@34 745 (** 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 746
adamc@34 747 We can check the implement of [nat_rect] as well: *)
adamc@34 748
adamc@34 749 Print nat_rect.
adamc@34 750 (** [[
adamc@34 751
adamc@34 752 nat_rect =
adamc@34 753 fun (P : nat -> Type) (f : P O) (f0 : forall n : nat, P n -> P (S n)) =>
adamc@34 754 fix F (n : nat) : P n :=
adamc@34 755 match n as n0 return (P n0) with
adamc@34 756 | O => f
adamc@34 757 | S n0 => f0 n0 (F n0)
adamc@34 758 end
adamc@34 759 : forall P : nat -> Type,
adamc@34 760 P O -> (forall n : nat, P n -> P (S n)) -> forall n : nat, P n
adamc@34 761 ]]
adamc@34 762
adamc@34 763 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 764
adamc@34 765 Section nat_ind'.
adamc@34 766 (** First, we have the property of natural numbers that we aim to prove. *)
adamc@34 767 Variable P : nat -> Prop.
adamc@34 768
adamc@34 769 (** Then we require a proof of the [O] case. *)
adamc@38 770 Hypothesis O_case : P O.
adamc@34 771
adamc@34 772 (** Next is a proof of the [S] case, which may assume an inductive hypothesis. *)
adamc@38 773 Hypothesis S_case : forall n : nat, P n -> P (S n).
adamc@34 774
adamc@34 775 (** Finally, we define a recursive function to tie the pieces together. *)
adamc@34 776 Fixpoint nat_ind' (n : nat) : P n :=
adamc@34 777 match n return (P n) with
adamc@34 778 | O => O_case
adamc@34 779 | S n' => S_case (nat_ind' n')
adamc@34 780 end.
adamc@34 781 End nat_ind'.
adamc@34 782
adamc@38 783 (** Closing the section adds the [Variable]s and [Hypothesis]es 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 784
adamc@34 785 %\medskip%
adamc@34 786
adamc@34 787 We can also examine the definition of [even_list_mut], which we generated with [Scheme] for a mutually-recursive type. *)
adamc@34 788
adamc@34 789 Print even_list_mut.
adamc@34 790 (** [[
adamc@34 791
adamc@34 792 even_list_mut =
adamc@34 793 fun (P : even_list -> Prop) (P0 : odd_list -> Prop)
adamc@34 794 (f : P ENil) (f0 : forall (n : nat) (o : odd_list), P0 o -> P (ECons n o))
adamc@34 795 (f1 : forall (n : nat) (e : even_list), P e -> P0 (OCons n e)) =>
adamc@34 796 fix F (e : even_list) : P e :=
adamc@34 797 match e as e0 return (P e0) with
adamc@34 798 | ENil => f
adamc@34 799 | ECons n o => f0 n o (F0 o)
adamc@34 800 end
adamc@34 801 with F0 (o : odd_list) : P0 o :=
adamc@34 802 match o as o0 return (P0 o0) with
adamc@34 803 | OCons n e => f1 n e (F e)
adamc@34 804 end
adamc@34 805 for F
adamc@34 806 : forall (P : even_list -> Prop) (P0 : odd_list -> Prop),
adamc@34 807 P ENil ->
adamc@34 808 (forall (n : nat) (o : odd_list), P0 o -> P (ECons n o)) ->
adamc@34 809 (forall (n : nat) (e : even_list), P e -> P0 (OCons n e)) ->
adamc@34 810 forall e : even_list, P e
adamc@34 811 ]]
adamc@34 812
adamc@34 813 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 814
adamc@34 815 Section even_list_mut'.
adamc@34 816 (** First, we need the properties that we are proving. *)
adamc@34 817 Variable Peven : even_list -> Prop.
adamc@34 818 Variable Podd : odd_list -> Prop.
adamc@34 819
adamc@34 820 (** Next, we need proofs of the three cases. *)
adamc@38 821 Hypothesis ENil_case : Peven ENil.
adamc@38 822 Hypothesis ECons_case : forall (n : nat) (o : odd_list), Podd o -> Peven (ECons n o).
adamc@38 823 Hypothesis OCons_case : forall (n : nat) (e : even_list), Peven e -> Podd (OCons n e).
adamc@34 824
adamc@34 825 (** Finally, we define the recursive functions. *)
adamc@34 826 Fixpoint even_list_mut' (e : even_list) : Peven e :=
adamc@34 827 match e return (Peven e) with
adamc@34 828 | ENil => ENil_case
adamc@34 829 | ECons n o => ECons_case n (odd_list_mut' o)
adamc@34 830 end
adamc@34 831 with odd_list_mut' (o : odd_list) : Podd o :=
adamc@34 832 match o return (Podd o) with
adamc@34 833 | OCons n e => OCons_case n (even_list_mut' e)
adamc@34 834 end.
adamc@34 835 End even_list_mut'.
adamc@34 836
adamc@34 837 (** 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 838
adamc@34 839 Section formula_ind'.
adamc@34 840 Variable P : formula -> Prop.
adamc@38 841 Hypothesis Eq_case : forall n1 n2 : nat, P (Eq n1 n2).
adamc@38 842 Hypothesis And_case : forall f1 f2 : formula,
adamc@34 843 P f1 -> P f2 -> P (And f1 f2).
adamc@38 844 Hypothesis Forall_case : forall f : nat -> formula,
adamc@34 845 (forall n : nat, P (f n)) -> P (Forall f).
adamc@34 846
adamc@34 847 Fixpoint formula_ind' (f : formula) : P f :=
adamc@34 848 match f return (P f) with
adamc@34 849 | Eq n1 n2 => Eq_case n1 n2
adamc@34 850 | And f1 f2 => And_case (formula_ind' f1) (formula_ind' f2)
adamc@34 851 | Forall f' => Forall_case f' (fun n => formula_ind' (f' n))
adamc@34 852 end.
adamc@34 853 End formula_ind'.
adamc@34 854
adamc@35 855
adamc@35 856 (** * Nested Inductive Types *)
adamc@35 857
adamc@35 858 (** 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 859
adamc@35 860 Inductive nat_tree : Set :=
adamc@35 861 | NLeaf' : nat_tree
adamc@35 862 | NNode' : nat -> list nat_tree -> nat_tree.
adamc@35 863
adamc@35 864 (** 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 865
adamc@35 866 Like we encountered for mutual inductive types, we find that the automatically-generated induction principle for [nat_tree] is too weak. *)
adamc@35 867
adamc@35 868 Check nat_tree_ind.
adamc@35 869 (** [[
adamc@35 870
adamc@35 871 nat_tree_ind
adamc@35 872 : forall P : nat_tree -> Prop,
adamc@35 873 P NLeaf' ->
adamc@35 874 (forall (n : nat) (l : list nat_tree), P (NNode' n l)) ->
adamc@35 875 forall n : nat_tree, P n
adamc@35 876 ]]
adamc@35 877
adamc@35 878 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 879
adamc@35 880 First, we will need an auxiliary definition, characterizing what it means for a property to hold of every element of a list. *)
adamc@35 881
adamc@35 882 Section All.
adamc@35 883 Variable T : Set.
adamc@35 884 Variable P : T -> Prop.
adamc@35 885
adamc@35 886 Fixpoint All (ls : list T) : Prop :=
adamc@35 887 match ls with
adamc@35 888 | Nil => True
adamc@35 889 | Cons h t => P h /\ All t
adamc@35 890 end.
adamc@35 891 End All.
adamc@35 892
adamc@35 893 (** 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 894
adamc@35 895 Print True.
adamc@35 896 (** [[
adamc@35 897
adamc@35 898 Inductive True : Prop := I : True
adamc@35 899 ]]
adamc@35 900
adamc@35 901 That is, [True] is a proposition with exactly one proof, [I], which we may always supply trivially.
adamc@35 902
adamc@35 903 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 904
adamc@35 905 Locate "/\".
adamc@35 906 (** [[
adamc@35 907
adamc@35 908 Notation Scope
adamc@35 909 "A /\ B" := and A B : type_scope
adamc@35 910 (default interpretation)
adamc@35 911 ]] *)
adamc@35 912
adamc@35 913 Print and.
adamc@35 914 (** [[
adamc@35 915
adamc@35 916 Inductive and (A : Prop) (B : Prop) : Prop := conj : A -> B -> A /\ B
adamc@35 917 For conj: Arguments A, B are implicit
adamc@35 918 For and: Argument scopes are [type_scope type_scope]
adamc@35 919 For conj: Argument scopes are [type_scope type_scope _ _]
adamc@35 920 ]]
adamc@35 921
adamc@35 922 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 923
adamc@35 924 %\medskip%
adamc@35 925
adamc@35 926 Now we create a section for our induction principle, following the same basic plan as in the last section of this chapter. *)
adamc@35 927
adamc@35 928 Section nat_tree_ind'.
adamc@35 929 Variable P : nat_tree -> Prop.
adamc@35 930
adamc@38 931 Hypothesis NLeaf'_case : P NLeaf'.
adamc@38 932 Hypothesis NNode'_case : forall (n : nat) (ls : list nat_tree),
adamc@35 933 All P ls -> P (NNode' n ls).
adamc@35 934
adamc@35 935 (** 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 936
adamc@35 937 [[
adamc@35 938
adamc@35 939 Fixpoint nat_tree_ind' (tr : nat_tree) : P tr :=
adamc@35 940 match tr return (P tr) with
adamc@35 941 | NLeaf' => NLeaf'_case
adamc@35 942 | NNode' n ls => NNode'_case n ls (list_nat_tree_ind ls)
adamc@35 943 end
adamc@35 944
adamc@35 945 with list_nat_tree_ind (ls : list nat_tree) : All P ls :=
adamc@35 946 match ls return (All P ls) with
adamc@35 947 | Nil => I
adamc@35 948 | Cons tr rest => conj (nat_tree_ind' tr) (list_nat_tree_ind rest)
adamc@35 949 end.
adamc@35 950
adamc@205 951 ]]
adamc@205 952
adamc@35 953 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 954
adamc@35 955 Fixpoint nat_tree_ind' (tr : nat_tree) : P tr :=
adamc@35 956 match tr return (P tr) with
adamc@35 957 | NLeaf' => NLeaf'_case
adamc@35 958 | NNode' n ls => NNode'_case n ls
adamc@35 959 ((fix list_nat_tree_ind (ls : list nat_tree) : All P ls :=
adamc@35 960 match ls return (All P ls) with
adamc@35 961 | Nil => I
adamc@35 962 | Cons tr rest => conj (nat_tree_ind' tr) (list_nat_tree_ind rest)
adamc@35 963 end) ls)
adamc@35 964 end.
adamc@35 965
adamc@35 966 (** 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 967
adamc@35 968 End nat_tree_ind'.
adamc@35 969
adamc@35 970 (** 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 971
adamc@35 972 Section map.
adamc@35 973 Variables T T' : Set.
adamc@35 974 Variable f : T -> T'.
adamc@35 975
adamc@35 976 Fixpoint map (ls : list T) : list T' :=
adamc@35 977 match ls with
adamc@35 978 | Nil => Nil
adamc@35 979 | Cons h t => Cons (f h) (map t)
adamc@35 980 end.
adamc@35 981 End map.
adamc@35 982
adamc@35 983 Fixpoint sum (ls : list nat) : nat :=
adamc@35 984 match ls with
adamc@35 985 | Nil => O
adamc@35 986 | Cons h t => plus h (sum t)
adamc@35 987 end.
adamc@35 988
adamc@35 989 (** Now we can define a size function over our trees. *)
adamc@35 990
adamc@35 991 Fixpoint ntsize (tr : nat_tree) : nat :=
adamc@35 992 match tr with
adamc@35 993 | NLeaf' => S O
adamc@35 994 | NNode' _ trs => S (sum (map ntsize trs))
adamc@35 995 end.
adamc@35 996
adamc@35 997 (** 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 998
adamc@35 999 Fixpoint ntsplice (tr1 tr2 : nat_tree) {struct tr1} : nat_tree :=
adamc@35 1000 match tr1 with
adamc@35 1001 | NLeaf' => NNode' O (Cons tr2 Nil)
adamc@35 1002 | NNode' n Nil => NNode' n (Cons tr2 Nil)
adamc@35 1003 | NNode' n (Cons tr trs) => NNode' n (Cons (ntsplice tr tr2) trs)
adamc@35 1004 end.
adamc@35 1005
adamc@35 1006 (** 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 1007
adamc@41 1008 (* begin thide *)
adamc@35 1009 Lemma plus_S : forall n1 n2 : nat,
adamc@35 1010 plus n1 (S n2) = S (plus n1 n2).
adamc@35 1011 induction n1; crush.
adamc@35 1012 Qed.
adamc@41 1013 (* end thide *)
adamc@35 1014
adamc@35 1015 (** Now we begin the proof of the theorem, adding the lemma [plus_S] as a hint. *)
adamc@35 1016
adamc@35 1017 Theorem ntsize_ntsplice : forall tr1 tr2 : nat_tree, ntsize (ntsplice tr1 tr2)
adamc@35 1018 = plus (ntsize tr2) (ntsize tr1).
adamc@41 1019 (* begin thide *)
adamc@35 1020 Hint Rewrite plus_S : cpdt.
adamc@35 1021
adamc@35 1022 (** 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 1023 induction tr1 using nat_tree_ind'; crush.
adamc@35 1024
adamc@35 1025 (** One subgoal remains: [[
adamc@35 1026
adamc@35 1027 n : nat
adamc@35 1028 ls : list nat_tree
adamc@35 1029 H : All
adamc@35 1030 (fun tr1 : nat_tree =>
adamc@35 1031 forall tr2 : nat_tree,
adamc@35 1032 ntsize (ntsplice tr1 tr2) = plus (ntsize tr2) (ntsize tr1)) ls
adamc@35 1033 tr2 : nat_tree
adamc@35 1034 ============================
adamc@35 1035 ntsize
adamc@35 1036 match ls with
adamc@35 1037 | Nil => NNode' n (Cons tr2 Nil)
adamc@35 1038 | Cons tr trs => NNode' n (Cons (ntsplice tr tr2) trs)
adamc@35 1039 end = S (plus (ntsize tr2) (sum (map ntsize ls)))
adamc@35 1040 ]]
adamc@35 1041
adamc@35 1042 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 1043
adamc@35 1044 destruct ls; crush.
adamc@35 1045
adamc@36 1046 (** We can go further in automating the proof by exploiting the hint mechanism. *)
adamc@35 1047
adamc@35 1048 Restart.
adamc@35 1049 Hint Extern 1 (ntsize (match ?LS with Nil => _ | Cons _ _ => _ end) = _) =>
adamc@35 1050 destruct LS; crush.
adamc@35 1051 induction tr1 using nat_tree_ind'; crush.
adamc@35 1052 Qed.
adamc@41 1053 (* end thide *)
adamc@35 1054
adamc@35 1055 (** 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 1056
adamc@40 1057 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 referred 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. *)
adamc@36 1058
adamc@36 1059
adamc@36 1060 (** * Manual Proofs About Constructors *)
adamc@36 1061
adamc@36 1062 (** It can be useful to understand how tactics like [discriminate] and [injection] work, so it is worth stepping through a manual proof of each kind. We will start with a proof fit for [discriminate]. *)
adamc@36 1063
adamc@36 1064 Theorem true_neq_false : true <> false.
adamc@41 1065 (* begin thide *)
adamc@36 1066 (** We begin with the tactic [red], which is short for "one step of reduction," to unfold the definition of logical negation. *)
adamc@36 1067
adamc@36 1068 red.
adamc@36 1069 (** [[
adamc@36 1070
adamc@36 1071 ============================
adamc@36 1072 true = false -> False
adamc@36 1073 ]]
adamc@36 1074
adamc@36 1075 The negation is replaced with an implication of falsehood. We use the tactic [intro H] to change the assumption of the implication into a hypothesis named [H]. *)
adamc@36 1076
adamc@36 1077 intro H.
adamc@36 1078 (** [[
adamc@36 1079
adamc@36 1080 H : true = false
adamc@36 1081 ============================
adamc@36 1082 False
adamc@36 1083 ]]
adamc@36 1084
adamc@36 1085 This is the point in the proof where we apply some creativity. We define a function whose utility will become clear soon. *)
adamc@36 1086
adamc@36 1087 Definition f (b : bool) := if b then True else False.
adamc@36 1088
adamc@36 1089 (** It is worth recalling the difference between the lowercase and uppercase versions of truth and falsehood: [True] and [False] are logical propositions, while [true] and [false] are boolean values that we can case-analyze. We have defined [f] such that our conclusion of [False] is computationally equivalent to [f false]. Thus, the [change] tactic will let us change the conclusion to [f false]. *)
adamc@36 1090
adamc@36 1091 change (f false).
adamc@36 1092 (** [[
adamc@36 1093
adamc@36 1094 H : true = false
adamc@36 1095 ============================
adamc@36 1096 f false
adamc@36 1097 ]]
adamc@36 1098
adamc@202 1099 Now the righthand side of [H]'s equality appears in the conclusion, so we can rewrite, using the notation [<-] to request to replace the righthand side the equality with the lefthand side. *)
adamc@36 1100
adamc@36 1101 rewrite <- H.
adamc@36 1102 (** [[
adamc@36 1103
adamc@36 1104 H : true = false
adamc@36 1105 ============================
adamc@36 1106 f true
adamc@36 1107 ]]
adamc@36 1108
adamc@36 1109 We are almost done. Just how close we are to done is revealed by computational simplification. *)
adamc@36 1110
adamc@36 1111 simpl.
adamc@36 1112 (** [[
adamc@36 1113
adamc@36 1114 H : true = false
adamc@36 1115 ============================
adamc@36 1116 True
adamc@36 1117 ]] *)
adamc@36 1118
adamc@36 1119 trivial.
adamc@36 1120 Qed.
adamc@41 1121 (* end thide *)
adamc@36 1122
adamc@36 1123 (** I have no trivial automated version of this proof to suggest, beyond using [discriminate] or [congruence] in the first place.
adamc@36 1124
adamc@36 1125 %\medskip%
adamc@36 1126
adamc@36 1127 We can perform a similar manual proof of injectivity of the constructor [S]. I leave a walk-through of the details to curious readers who want to run the proof script interactively. *)
adamc@36 1128
adamc@36 1129 Theorem S_inj' : forall n m : nat, S n = S m -> n = m.
adamc@41 1130 (* begin thide *)
adamc@36 1131 intros n m H.
adamc@36 1132 change (pred (S n) = pred (S m)).
adamc@36 1133 rewrite H.
adamc@36 1134 reflexivity.
adamc@36 1135 Qed.
adamc@41 1136 (* end thide *)
adamc@36 1137
adamc@37 1138
adamc@37 1139 (** * Exercises *)
adamc@37 1140
adamc@37 1141 (** %\begin{enumerate}%#<ol>#
adamc@37 1142
adamc@201 1143 %\item%#<li># Define an inductive type [truth] with three constructors, [Yes], [No], and [Maybe]. [Yes] stands for certain truth, [No] for certain falsehood, and [Maybe] for an unknown situation. Define "not," "and," and "or" for this replacement boolean algebra. Prove that your implementation of "and" is commutative and distributes over your implementation of "or."#</li>#
adamc@37 1144
adamc@39 1145 %\item%#<li># Modify the first example language of Chapter 2 to include variables, where variables are represented with [nat]. Extend the syntax and semantics of expressions to accommodate the change. Your new [expDenote] function should take as a new extra first argument a value of type [var -> nat], where [var] is a synonym for naturals-as-variables, and the function assigns a value to each variable. Define a constant folding function which does a bottom-up pass over an expression, at each stage replacing every binary operation on constants with an equivalent constant. Prove that constant folding preserves the meanings of expressions.#</li>#
adamc@38 1146
adamc@39 1147 %\item%#<li># Reimplement the second example language of Chapter 2 to use mutually-inductive types instead of dependent types. That is, define two separate (non-dependent) inductive types [nat_exp] and [bool_exp] for expressions of the two different types, rather than a single indexed type. To keep things simple, you may consider only the binary operators that take naturals as operands. Add natural number variables to the language, as in the last exercise, and add an "if" expression form taking as arguments one boolean expression and two natural number expressions. Define semantics and constant-folding functions for this new language. Your constant folding should simplify not just binary operations (returning naturals or booleans) with known arguments, but also "if" expressions with known values for their test expressions but possibly undetermined "then" and "else" cases. Prove that constant-folding a natural number expression preserves its meaning.#</li>#
adamc@38 1148
adamc@38 1149 %\item%#<li># Using a reflexive inductive definition, define a type [nat_tree] of infinitary trees, with natural numbers at their leaves and a countable infinity of new trees branching out of each internal node. Define a function [increment] that increments the number in every leaf of a [nat_tree]. Define a function [leapfrog] over a natural [i] and a tree [nt]. [leapfrog] should recurse into the [i]th child of [nt], the [i+1]st child of that node, the [i+2]nd child of the next node, and so on, until reaching a leaf, in which case [leapfrog] should return the number at that leaf. Prove that the result of any call to [leapfrog] is incremented by one by calling [increment] on the tree.#</li>#
adamc@38 1150
adamc@38 1151 %\item%#<li># Define a type of trees of trees of trees of (repeat to infinity). That is, define an inductive type [trexp], whose members are either base cases containing natural numbers or binary trees of [trexp]s. Base your definition on a parameterized binary tree type [btree] that you will also define, so that [trexp] is defined as a nested inductive type. Define a function [total] that sums all of the naturals at the leaves of a [trexp]. Define a function [increment] that increments every leaf of a [trexp] by one. Prove that, for all [tr], [total (increment tr) >= total tr]. On the way to finishing this proof, you will probably want to prove a lemma and add it as a hint using the syntax [Hint Resolve name_of_lemma.].#</li>#
adamc@38 1152
adamc@38 1153 %\item%#<li># Prove discrimination and injectivity theorems for the [nat_btree] type defined earlier in this chapter. In particular, without using the tactics [discriminate], [injection], or [congruence], prove that no leaf equals any node, and prove that two equal nodes carry the same natural number.#</li>#
adamc@37 1154
adamc@37 1155 #</ol>#%\end{enumerate}% *)