annotate src/InductiveTypes.v @ 314:d5787b70cf48

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