adamc@26: (* Copyright (c) 2008, Adam Chlipala adamc@26: * adamc@26: * This work is licensed under a adamc@26: * Creative Commons Attribution-Noncommercial-No Derivative Works 3.0 adamc@26: * Unported License. adamc@26: * The license text is available at: adamc@26: * http://creativecommons.org/licenses/by-nc-nd/3.0/ adamc@26: *) adamc@26: adamc@26: (* begin hide *) adamc@26: Require Import List. adamc@26: adamc@26: Require Import Tactics. adamc@26: adamc@26: Set Implicit Arguments. adamc@26: (* end hide *) adamc@26: adamc@26: adamc@26: (** %\chapter{Inductive Types}% *) adamc@26: adamc@26: (** In a sense, CIC is built from just two relatively straightforward features: function types and inductive types. From this modest foundation, we can prove effectively all of the theorems of math and carry out effectively all program verifications, with enough effort expended. This chapter introduces induction and recursion in Coq and shares some "design patterns" for overcoming common pitfalls with them. *) adamc@26: adamc@26: adamc@26: (** * Enumerations *) adamc@26: adamc@26: (** Coq inductive types generalize the algebraic datatypes found in Haskell and ML. Confusingly enough, inductive types also generalize generalized algebraic datatypes (GADTs), by adding the possibility for type dependency. Even so, it is worth backing up from the examples of the last chapter and going over basic, algebraic datatype uses of inductive datatypes, because the chance to prove things about the values of these types adds new wrinkles beyond usual practice in Haskell and ML. adamc@26: adamc@26: The singleton type [unit] is an inductive type: *) adamc@26: adamc@26: Inductive unit : Set := adamc@26: | tt. adamc@26: adamc@26: (** 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: adamc@26: Check unit. adamc@26: (** [[ adamc@26: adamc@26: unit : Set adamc@26: ]] *) adamc@26: Check tt. adamc@26: (** [[ adamc@26: adamc@26: tt : unit adamc@26: ]] *) adamc@26: adamc@26: (** We can prove that [unit] is a genuine singleton type. *) adamc@26: adamc@26: Theorem unit_singleton : forall x : unit, x = tt. adamc@26: (** The important thing about an inductive type is, unsurprisingly, that you can do induction over its values, and induction is the key to proving this theorem. We ask to proceed by induction on the variable [x]. *) adamc@26: induction x. adamc@26: (** The goal changes to: [[ adamc@26: adamc@26: tt = tt adamc@26: ]] *) adamc@26: (** ...which we can discharge trivially. *) adamc@26: reflexivity. adamc@26: Qed. adamc@26: adamc@26: (** 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: adamc@26: destruct x. adamc@26: ...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: adamc@26: What exactly %\textit{%##is##%}% the induction principle for [unit]? We can ask Coq: *) adamc@26: adamc@26: Check unit_ind. adamc@26: (** [[ adamc@26: adamc@26: unit_ind : forall P : unit -> Prop, P tt -> forall u : unit, P u adamc@26: ]] adamc@26: adamc@26: 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: adamc@26: 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: adamc@26: 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: adamc@26: %\medskip% adamc@26: adamc@26: We can define an inductive type even simpler than [unit]: *) adamc@26: adamc@26: Inductive Empty_set : Set := . adamc@26: adamc@26: (** [Empty_set] has no elements. We can prove fun theorems about it: *) adamc@26: adamc@26: Theorem the_sky_is_falling : forall x : Empty_set, 2 + 2 = 5. adamc@26: destruct 1. adamc@26: Qed. adamc@26: adamc@26: (** Because [Empty_set] has no elements, the fact of having an element of this type implies anything. We use [destruct 1] instead of [destruct x] in the proof because unused quantified variables are relegated to being referred to by number. (There is a good reason for this, related to the unity of quantifiers and dependent function types.) adamc@26: adamc@26: We can see the induction principle that made this proof so easy: *) adamc@26: adamc@26: Check Empty_set_ind. adamc@26: (** [[ adamc@26: adamc@26: Empty_set_ind : forall (P : Empty_set -> Prop) (e : Empty_set), P e adamc@26: ]] adamc@26: adamc@26: 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: adamc@26: 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: adamc@26: Definition e2u (e : Empty_set) : unit := match e with end. adamc@26: adamc@26: (** 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: adamc@26: %\medskip% adamc@26: adamc@26: Moving up the ladder of complexity, we can define the booleans: *) adamc@26: adamc@26: Inductive bool : Set := adamc@26: | true adamc@26: | false. adamc@26: adamc@26: (** We can use less vacuous pattern matching to define boolean negation. *) adamc@26: adamc@26: Definition not (b : bool) : bool := adamc@26: match b with adamc@26: | true => false adamc@26: | false => true adamc@26: end. adamc@26: adamc@27: (** An alternative definition desugars to the above: *) adamc@27: adamc@27: Definition not' (b : bool) : bool := adamc@27: if b then false else true. adamc@27: adamc@26: (** We might want to prove that [not] is its own inverse operation. *) adamc@26: adamc@26: Theorem not_inverse : forall b : bool, not (not b) = b. adamc@26: destruct b. adamc@26: adamc@26: (** After we case analyze on [b], we are left with one subgoal for each constructor of [bool]. adamc@26: adamc@26: [[ adamc@26: adamc@26: 2 subgoals adamc@26: adamc@26: ============================ adamc@26: not (not true) = true adamc@26: ]] adamc@26: adamc@26: [[ adamc@26: subgoal 2 is: adamc@26: not (not false) = false adamc@26: ]] adamc@26: adamc@26: The first subgoal follows by Coq's rules of computation, so we can dispatch it easily: *) adamc@26: adamc@26: reflexivity. adamc@26: adamc@26: (** Likewise for the second subgoal, so we can restart the proof and give a very compact justification. *) adamc@26: adamc@26: Restart. adamc@26: destruct b; reflexivity. adamc@26: Qed. adamc@27: adamc@27: (** Another theorem about booleans illustrates another useful tactic. *) adamc@27: adamc@27: Theorem not_ineq : forall b : bool, not b <> b. adamc@27: destruct b; discriminate. adamc@27: Qed. adamc@27: adamc@27: (** [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: adamc@27: At this point, it is probably not hard to guess what the underlying induction principle for [bool] is. *) adamc@27: adamc@27: Check bool_ind. adamc@27: (** [[ adamc@27: adamc@27: bool_ind : forall P : bool -> Prop, P true -> P false -> forall b : bool, P b adamc@27: ]] *) adamc@28: adamc@28: adamc@28: (** * Simple Recursive Types *) adamc@28: adamc@28: (** The natural numbers are the simplest common example of an inductive type that actually deserves the name. *) adamc@28: adamc@28: Inductive nat : Set := adamc@28: | O : nat adamc@28: | S : nat -> nat. adamc@28: adamc@28: (** [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: adamc@28: Pattern matching works as we demonstrated in the last chapter: *) adamc@28: adamc@28: Definition isZero (n : nat) : bool := adamc@28: match n with adamc@28: | O => true adamc@28: | S _ => false adamc@28: end. adamc@28: adamc@28: Definition pred (n : nat) : nat := adamc@28: match n with adamc@28: | O => O adamc@28: | S n' => n' adamc@28: end. adamc@28: adamc@28: (** We can prove theorems by case analysis: *) adamc@28: adamc@28: Theorem S_isZero : forall n : nat, isZero (pred (S (S n))) = false. adamc@28: destruct n; reflexivity. adamc@28: Qed. adamc@28: adamc@28: (** We can also now get into genuine inductive theorems. First, we will need a recursive function, to make things interesting. *) adamc@28: adamc@28: Fixpoint plus (n m : nat) {struct n} : nat := adamc@28: match n with adamc@28: | O => m adamc@28: | S n' => S (plus n' m) adamc@28: end. adamc@28: adamc@28: (** 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: adamc@28: Some theorems about [plus] can be proved without induction. *) adamc@28: adamc@28: Theorem O_plus_n : forall n : nat, plus O n = n. adamc@28: intro; reflexivity. adamc@28: Qed. adamc@28: adamc@28: (** 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: adamc@28: Theorem n_plus_O : forall n : nat, plus n O = n. adamc@28: induction n. adamc@28: adamc@28: (** Our first subgoal is [plus O O = O], which %\textit{%##is##%}% trivial by computation. *) adamc@28: adamc@28: reflexivity. adamc@28: adamc@28: (** Our second subgoal is more work and also demonstrates our first inductive hypothesis. adamc@28: adamc@28: [[ adamc@28: adamc@28: n : nat adamc@28: IHn : plus n O = n adamc@28: ============================ adamc@28: plus (S n) O = S n adamc@28: ]] adamc@28: adamc@28: We can start out by using computation to simplify the goal as far as we can. *) adamc@28: adamc@28: simpl. adamc@28: adamc@28: (** Now the conclusion is [S (plus n O) = S n]. Using our inductive hypothesis: *) adamc@28: adamc@28: rewrite IHn. adamc@28: adamc@28: (** ...we get a trivial conclusion [S n = S n]. *) adamc@28: adamc@28: reflexivity. adamc@28: adamc@28: (** Not much really went on in this proof, so the [crush] tactic from the [Tactics] module can prove this theorem automatically. *) adamc@28: adamc@28: Restart. adamc@28: induction n; crush. adamc@28: Qed. adamc@28: adamc@28: (** We can check out the induction principle at work here: *) adamc@28: adamc@28: Check nat_ind. adamc@28: (** [[ adamc@28: adamc@28: nat_ind : forall P : nat -> Prop, adamc@28: P O -> (forall n : nat, P n -> P (S n)) -> forall n : nat, P n adamc@28: ]] adamc@28: adamc@28: 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: adamc@28: Since [nat] has a constructor that takes an argument, we may sometimes need to know that that constructor is injective. *) adamc@28: adamc@28: Theorem S_inj : forall n m : nat, S n = S m -> n = m. adamc@28: injection 1; trivial. adamc@28: Qed. adamc@28: adamc@28: (** [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: adamc@29: 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{%##complete decision procedure for the theory of equality and uninterpreted functions##%}%, plus some smarts about inductive types. adamc@29: adamc@29: %\medskip% adamc@29: adamc@29: We can define a type of lists of natural numbers. *) adamc@29: adamc@29: Inductive nat_list : Set := adamc@29: | NNil : nat_list adamc@29: | NCons : nat -> nat_list -> nat_list. adamc@29: adamc@29: (** Recursive definitions are straightforward extensions of what we have seen before. *) adamc@29: adamc@29: Fixpoint nlength (ls : nat_list) : nat := adamc@29: match ls with adamc@29: | NNil => O adamc@29: | NCons _ ls' => S (nlength ls') adamc@29: end. adamc@29: adamc@29: Fixpoint napp (ls1 ls2 : nat_list) {struct ls1} : nat_list := adamc@29: match ls1 with adamc@29: | NNil => ls2 adamc@29: | NCons n ls1' => NCons n (napp ls1' ls2) adamc@29: end. adamc@29: adamc@29: (** Inductive theorem proving can again be automated quite effectively. *) adamc@29: adamc@29: Theorem nlength_napp : forall ls1 ls2 : nat_list, nlength (napp ls1 ls2) adamc@29: = plus (nlength ls1) (nlength ls2). adamc@29: induction ls1; crush. adamc@29: Qed. adamc@29: adamc@29: Check nat_list_ind. adamc@29: (** [[ adamc@29: adamc@29: nat_list_ind adamc@29: : forall P : nat_list -> Prop, adamc@29: P NNil -> adamc@29: (forall (n : nat) (n0 : nat_list), P n0 -> P (NCons n n0)) -> adamc@29: forall n : nat_list, P n adamc@29: ]] adamc@29: adamc@29: %\medskip% adamc@29: adamc@29: In general, we can implement any "tree" types as inductive types. For example, here are binary trees of naturals. *) adamc@29: adamc@29: Inductive nat_btree : Set := adamc@29: | NLeaf : nat_btree adamc@29: | NNode : nat_btree -> nat -> nat_btree -> nat_btree. adamc@29: adamc@29: Fixpoint nsize (tr : nat_btree) : nat := adamc@29: match tr with adamc@29: | NLeaf => O adamc@29: | NNode tr1 _ tr2 => plus (nsize tr1) (nsize tr2) adamc@29: end. adamc@29: adamc@29: Fixpoint nsplice (tr1 tr2 : nat_btree) {struct tr1} : nat_btree := adamc@29: match tr1 with adamc@29: | NLeaf => tr2 adamc@29: | NNode tr1' n tr2' => NNode (nsplice tr1' tr2) n tr2' adamc@29: end. adamc@29: adamc@29: Theorem plus_assoc : forall n1 n2 n3 : nat, plus (plus n1 n2) n3 = plus n1 (plus n2 n3). adamc@29: induction n1; crush. adamc@29: Qed. adamc@29: adamc@29: Theorem nsize_nsplice : forall tr1 tr2 : nat_btree, nsize (nsplice tr1 tr2) adamc@29: = plus (nsize tr2) (nsize tr1). adamc@29: Hint Rewrite n_plus_O plus_assoc : cpdt. adamc@29: adamc@29: induction tr1; crush. adamc@29: Qed. adamc@29: adamc@29: Check nat_btree_ind. adamc@29: (** [[ adamc@29: adamc@29: nat_btree_ind adamc@29: : forall P : nat_btree -> Prop, adamc@29: P NLeaf -> adamc@29: (forall n : nat_btree, adamc@29: P n -> forall (n0 : nat) (n1 : nat_btree), P n1 -> P (NNode n n0 n1)) -> adamc@29: forall n : nat_btree, P n adamc@29: ]] *)