annotate src/InductiveTypes.v @ 392:4b1242b277b2

Typo fixes
author Adam Chlipala <adam@chlipala.net>
date Fri, 20 Apr 2012 12:49:47 -0400
parents d1276004eec9
children 05efde66559d
rev   line source
adam@315 1 (* Copyright (c) 2008-2011, 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@315 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.
adam@315 24
adam@315 25 The last chapter took a deep dive into some of the more advanced Coq features, to highlight the unusual approach that I advocate in this book. However, from this point on, we will rewind and go back to basics, presenting the relevant features of Coq in a more bottom-up manner. A useful first step is a discussion of the differences and relationships between proofs and programs in Coq. *)
adam@315 26
adam@315 27
adam@315 28 (** * Proof Terms *)
adam@315 29
adam@315 30 (** Mainstream presentations of mathematics treat proofs as objects that exist outside of the universe of mathematical objects. However, for a variety of reasoning, it is convenient to encode proofs, traditional mathematical objects, and programs within a single formal language. Validity checks on mathematical objects are useful in any setting, to catch typoes and other uninteresting errors. The benefits of static typing for programs are widely recognized, and Coq brings those benefits to both mathematical objects and programs via a uniform mechanism. In fact, from this point on, we will not bother to distinguish between programs and mathematical objects. Many mathematical formalisms are most easily encoded in terms of programs.
adam@315 31
adam@315 32 Proofs are fundamentally different from programs, because any two proofs of a theorem are considered equivalent, from a formal standpoint if not from an engineering standpoint. However, we can use the same type-checking technology to check proofs as we use to validate our programs. This is the %\index{Curry-Howard correspondence}\emph{%#<i>#Curry-Howard correspondence#</i>#%}~\cite{Curry,Howard}%, an approach for relating proofs and programs. We represent mathematical theorems as types, such that a theorem's proofs are exactly those programs that type-check at the corresponding type.
adam@315 33
adam@315 34 The last chapter's example already snuck in an instance of Curry-Howard. We used the token [->] to stand for both function types and logical implications. One reasonable conclusion upon seeing this might be that some fancy overloading of notations is at work. In fact, functions and implications are precisely identical according to Curry-Howard! That is, they are just two ways of describing the same computational phenomenon.
adam@315 35
adam@315 36 A short demonstration should explain how this can be. The identity function over the natural numbers is certainly not a controversial program. *)
adam@315 37
adam@315 38 Check (fun x : nat => x).
adam@315 39 (** [: nat -> nat] *)
adam@315 40
adam@315 41 (** Consider this alternate program, which is almost identical to the last one. *)
adam@315 42
adam@315 43 Check (fun x : True => x).
adam@315 44 (** [: True -> True] *)
adam@315 45
adam@315 46 (** The identity program is interpreted as a proof that %\index{Gallina terms!True}%[True], the always-true proposition, implies itself! What we see is that Curry-Howard interprets implications as functions, where an input is a proposition being assumed and an output is a proposition being deduced. This intuition is not too far from a common one for informal theorem proving, where we might already think of an implication proof as a process for transforming a hypothesis into a conclusion.
adam@315 47
adam@315 48 There are also more primitive proof forms available. For instance, the term %\index{Gallina terms!I}%[I] is the single proof of [True], applicable in any context. *)
adam@315 49
adam@315 50 Check I.
adam@315 51 (** [: True] *)
adam@315 52
adam@315 53 (** With [I], we can prove another simple propositional theorem. *)
adam@315 54
adam@315 55 Check (fun _ : False => I).
adam@315 56 (** [: False -> True] *)
adam@315 57
adam@315 58 (** No proofs of %\index{Gallina terms!False}%[False] exist in the top-level context, but the implication-as-function analogy gives us an easy way to, for example, show that [False] implies itself. *)
adam@315 59
adam@315 60 Check (fun x : False => x).
adam@315 61 (** [: False -> False] *)
adam@315 62
adam@315 63 (** In fact, [False] implies anything, and we can take advantage of this fact with an odd looking [match] expression that has no branches. Since there are no rules for deducing [False], there are no cases to consider! *)
adam@315 64
adam@315 65 Check (fun x : False => match x with end : True).
adam@315 66 (** [: False -> True] *)
adam@315 67
adam@315 68 (** Every one of these example programs whose type looks like a logical formula is a %\index{proof term}\emph{%#<i>#proof term#</i>#%}%. We use that name for any Gallina term of a logical type, and we will elaborate shortly on what makes a type logical.
adam@315 69
adam@317 70 In the rest of this chapter, we will introduce different ways of defining types. Every example type can be interpreted alternatively as a type of programs or %\index{proposition}%propositions (i.e., formulas or theorem statements).
adam@317 71
adam@317 72 One of the first types we introduce will be [bool], with constructors [true] and [false]. Newcomers to Coq often wonder about the distinction between [True] and [true] and the distinction between [False] and [false]. One glib answer is that [True] and [False] are types, but [true] and [false] are not. A more useful answer is that Coq's metatheory guarantees that any term of type [bool] %\emph{%#<i>#evaluates#</i>#%}% to either [true] or [false]. This means that we have an %\emph{%#<i>#algorithm#</i>#%}% for answering any question phrased as an expression of type [bool]. Conversely, most propositions do not evaluate to [True] or [False]; the language of inductively defined propositions is much richer than that. We ought to be glad that we have no algorithm for deciding mathematical truth, since otherwise it would be clear that we could not formalize undecidable properties, like most any properties of general-purpose programs. *)
adamc@26 73
adamc@26 74
adamc@26 75 (** * Enumerations *)
adamc@26 76
adam@315 77 (** Coq inductive types generalize the %\index{algebraic datatypes}%algebraic datatypes found in %\index{Haskell}%Haskell and %\index{ML}%ML. Confusingly enough, inductive types also generalize %\index{generalized algebraic datatypes}%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 78
adam@315 79 The singleton type [unit] is an inductive type:%\index{Gallina terms!unit}\index{Gallina terms!tt}% *)
adamc@26 80
adamc@26 81 Inductive unit : Set :=
adamc@26 82 | tt.
adamc@26 83
adamc@26 84 (** 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 85
adamc@26 86 Check unit.
adamc@208 87 (** [unit : Set] *)
adamc@26 88
adamc@26 89 Check tt.
adamc@208 90 (** [tt : unit] *)
adamc@26 91
adamc@26 92 (** We can prove that [unit] is a genuine singleton type. *)
adamc@26 93
adamc@26 94 Theorem unit_singleton : forall x : unit, x = tt.
adamc@208 95
adam@315 96 (** 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].%\index{tactics!induction}% *)
adamc@208 97
adamc@41 98 (* begin thide *)
adamc@26 99 induction x.
adamc@26 100
adamc@208 101 (** The goal changes to:
adamc@208 102 [[
adamc@26 103 tt = tt
adam@302 104 ]]
adam@302 105 *)
adamc@208 106
adamc@26 107 (** ...which we can discharge trivially. *)
adamc@208 108
adamc@26 109 reflexivity.
adamc@26 110 Qed.
adamc@41 111 (* end thide *)
adamc@26 112
adam@315 113 (** 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:%\index{tactics!destruct}% [[
adamc@26 114
adamc@26 115 destruct x.
adamc@205 116
adamc@205 117 ]]
adamc@205 118
adam@292 119 %\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 120
adam@315 121 What exactly %\textit{%#<i>#is#</i>#%}% the %\index{induction principles}%induction principle for [unit]? We can ask Coq: *)
adamc@26 122
adamc@26 123 Check unit_ind.
adamc@208 124 (** [unit_ind : forall P : unit -> Prop, P tt -> forall u : unit, P u] *)
adamc@26 125
adam@315 126 (** Every [Inductive] command defining a type [T] also defines an induction principle named [T_ind]. Recall from the last section that 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 %\index{Gallina terms!Prop}%[Prop], which appears in our induction principle; and the type %\index{Gallina terms!Set}%[Set], which we have seen a few times already.
adamc@26 127
adam@315 128 The convention goes like this: [Set] is the type of normal types used in programming, 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 129
adam@315 130 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)].
adam@315 131
adam@315 132 The definition of [unit] places the type in [Set]. By replacing [Set] with [Prop], [unit] with [True], and [tt] with [I], we arrive at precisely the definition of [True] that the Coq standard library employs! The program type [unit] is the Curry-Howard equivalent of the proposition [True]. We might make the tongue-in-cheek claim that, while philosophers have expended much ink on the nature of truth, we have now determined that truth is the [unit] type of functional programming.
adamc@26 133
adamc@26 134 %\medskip%
adamc@26 135
adam@315 136 We can define an inductive type even simpler than [unit]:%\index{Gallina terms!Empty\_set}% *)
adamc@26 137
adamc@26 138 Inductive Empty_set : Set := .
adamc@26 139
adamc@26 140 (** [Empty_set] has no elements. We can prove fun theorems about it: *)
adamc@26 141
adamc@26 142 Theorem the_sky_is_falling : forall x : Empty_set, 2 + 2 = 5.
adamc@41 143 (* begin thide *)
adamc@26 144 destruct 1.
adamc@26 145 Qed.
adamc@41 146 (* end thide *)
adamc@26 147
adam@317 148 (** 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. At least within Coq's logical foundation of %\index{constructive logic}%constructive logic, which we elaborate on more in the next chapter, 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 149
adamc@26 150 We can see the induction principle that made this proof so easy: *)
adamc@26 151
adamc@26 152 Check Empty_set_ind.
adam@315 153 (** [Empty_set_ind : forall (][P : Empty_set -> Prop) (e : Empty_set), P e] *)
adamc@26 154
adam@315 155 (** 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 156
adamc@26 157 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 158
adamc@26 159 Definition e2u (e : Empty_set) : unit := match e with end.
adamc@26 160
adam@315 161 (** 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. This idiom may look familiar; we employed it with proofs of [False] in the last section. In fact, [Empty_set] is the Curry-Howard equivalent of [False]. As for why [Empty_set] starts with a capital letter and not a lowercase letter like [unit] does, we must refer the reader to the authors of the Coq standard library, to which we try to be faithful.
adamc@26 162
adamc@26 163 %\medskip%
adamc@26 164
adam@315 165 Moving up the ladder of complexity, we can define the booleans:%\index{Gallina terms!bool}\index{Gallina terms!true}\index{Gallina terms!false}% *)
adamc@26 166
adamc@26 167 Inductive bool : Set :=
adamc@26 168 | true
adamc@26 169 | false.
adamc@26 170
adam@316 171 (** We can use less vacuous pattern matching to define boolean negation.%\index{Gallina terms!negb}% *)
adamc@26 172
adam@279 173 Definition negb (b : bool) : bool :=
adamc@26 174 match b with
adamc@26 175 | true => false
adamc@26 176 | false => true
adamc@26 177 end.
adamc@26 178
adam@317 179 (** An alternative definition desugars to the above, thanks to an %\index{Gallina terms!if}%[if] notation overloaded to work with any inductive type that has exactly two constructors: *)
adamc@27 180
adam@279 181 Definition negb' (b : bool) : bool :=
adamc@27 182 if b then false else true.
adamc@27 183
adam@279 184 (** We might want to prove that [negb] is its own inverse operation. *)
adamc@26 185
adam@279 186 Theorem negb_inverse : forall b : bool, negb (negb b) = b.
adamc@41 187 (* begin thide *)
adamc@26 188 destruct b.
adamc@26 189
adamc@208 190 (** After we case-analyze on [b], we are left with one subgoal for each constructor of [bool].
adamc@26 191
adam@315 192 %\vspace{.1in} \noindent 2 \coqdockw{subgoals}\vspace{-.1in}%#<tt>2 subgoals</tt>#
adam@315 193
adamc@26 194 [[
adamc@26 195 ============================
adam@279 196 negb (negb true) = true
adamc@26 197 ]]
adam@315 198 %\noindent \coqdockw{subgoal} 2 \coqdockw{is}:%#<tt>subgoal 2 is</tt>#
adamc@26 199 [[
adam@279 200 negb (negb false) = false
adamc@208 201
adamc@26 202 ]]
adamc@26 203
adamc@26 204 The first subgoal follows by Coq's rules of computation, so we can dispatch it easily: *)
adamc@26 205
adamc@26 206 reflexivity.
adamc@26 207
adam@315 208 (** Likewise for the second subgoal, so we can restart the proof and give a very compact justification.%\index{Vernacular commands!Restart}% *)
adamc@26 209
adam@315 210 (* begin hide *)
adamc@26 211 Restart.
adam@315 212 (* end hide *)
adam@315 213 (** %\noindent \coqdockw{Restart}%#<tt>Restart</tt>#. *)
adam@315 214
adamc@26 215 destruct b; reflexivity.
adamc@26 216 Qed.
adamc@41 217 (* end thide *)
adamc@27 218
adam@315 219 (** Another theorem about booleans illustrates another useful tactic.%\index{tactics!discriminate}% *)
adamc@27 220
adam@279 221 Theorem negb_ineq : forall b : bool, negb b <> b.
adamc@41 222 (* begin thide *)
adamc@27 223 destruct b; discriminate.
adamc@27 224 Qed.
adamc@41 225 (* end thide *)
adamc@27 226
adamc@27 227 (** [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 228
adamc@27 229 At this point, it is probably not hard to guess what the underlying induction principle for [bool] is. *)
adamc@27 230
adamc@27 231 Check bool_ind.
adamc@208 232 (** [bool_ind : forall P : bool -> Prop, P true -> P false -> forall b : bool, P b] *)
adamc@28 233
adam@315 234 (** That is, to prove that a property describes all [bool]s, prove that it describes both [true] and [false].
adam@315 235
adam@392 236 There is no interesting Curry-Howard analogue of [bool]. Of course, we can define such a type by replacing [Set] by [Prop] above, but the proposition we arrive at is not very useful. It is logically equivalent to [True], but it provides two indistinguishable primitive proofs, [true] and [false]. In the rest of the chapter, we will skip commenting on Curry-Howard versions of inductive definitions where such versions are not interesting. *)
adam@315 237
adamc@28 238
adamc@28 239 (** * Simple Recursive Types *)
adamc@28 240
adam@315 241 (** The natural numbers are the simplest common example of an inductive type that actually deserves the name.%\index{Gallina terms!nat}\index{Gallina terms!O}\index{Gallina terms!S}% *)
adamc@28 242
adamc@28 243 Inductive nat : Set :=
adamc@28 244 | O : nat
adamc@28 245 | S : nat -> nat.
adamc@28 246
adam@315 247 (** [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 248
adam@316 249 Pattern matching works as we demonstrated in the last chapter:%\index{Gallina terms!pred}% *)
adamc@28 250
adamc@28 251 Definition isZero (n : nat) : bool :=
adamc@28 252 match n with
adamc@28 253 | O => true
adamc@28 254 | S _ => false
adamc@28 255 end.
adamc@28 256
adamc@28 257 Definition pred (n : nat) : nat :=
adamc@28 258 match n with
adamc@28 259 | O => O
adamc@28 260 | S n' => n'
adamc@28 261 end.
adamc@28 262
adamc@28 263 (** We can prove theorems by case analysis: *)
adamc@28 264
adamc@28 265 Theorem S_isZero : forall n : nat, isZero (pred (S (S n))) = false.
adamc@41 266 (* begin thide *)
adamc@28 267 destruct n; reflexivity.
adamc@28 268 Qed.
adamc@41 269 (* end thide *)
adamc@28 270
adam@316 271 (** We can also now get into genuine inductive theorems. First, we will need a recursive function, to make things interesting.%\index{Gallina terms!plus}% *)
adamc@28 272
adamc@208 273 Fixpoint plus (n m : nat) : nat :=
adamc@28 274 match n with
adamc@28 275 | O => m
adamc@28 276 | S n' => S (plus n' m)
adamc@28 277 end.
adamc@28 278
adamc@208 279 (** Recall that [Fixpoint] is Coq's mechanism for recursive function definitions. Some theorems about [plus] can be proved without induction. *)
adamc@28 280
adamc@28 281 Theorem O_plus_n : forall n : nat, plus O n = n.
adamc@41 282 (* begin thide *)
adamc@28 283 intro; reflexivity.
adamc@28 284 Qed.
adamc@41 285 (* end thide *)
adamc@28 286
adamc@208 287 (** 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 288
adamc@28 289 Theorem n_plus_O : forall n : nat, plus n O = n.
adamc@41 290 (* begin thide *)
adamc@28 291 induction n.
adamc@28 292
adamc@28 293 (** Our first subgoal is [plus O O = O], which %\textit{%#<i>#is#</i>#%}% trivial by computation. *)
adamc@28 294
adamc@28 295 reflexivity.
adamc@28 296
adamc@28 297 (** Our second subgoal is more work and also demonstrates our first inductive hypothesis.
adamc@28 298
adamc@28 299 [[
adamc@28 300 n : nat
adamc@28 301 IHn : plus n O = n
adamc@28 302 ============================
adamc@28 303 plus (S n) O = S n
adamc@208 304
adamc@28 305 ]]
adamc@28 306
adam@315 307 We can start out by using computation to simplify the goal as far as we can.%\index{tactics!simpl}% *)
adamc@28 308
adamc@28 309 simpl.
adamc@28 310
adam@315 311 (** Now the conclusion is [S (][plus n O) = S n]. Using our inductive hypothesis: *)
adamc@28 312
adamc@28 313 rewrite IHn.
adamc@28 314
adamc@28 315 (** ...we get a trivial conclusion [S n = S n]. *)
adamc@28 316
adamc@28 317 reflexivity.
adamc@28 318
adam@315 319 (** Not much really went on in this proof, so the [crush] tactic from the [CpdtTactics] module can prove this theorem automatically. *)
adamc@28 320
adam@315 321 (* begin hide *)
adamc@28 322 Restart.
adam@315 323 (* end hide *)
adam@315 324 (** %\noindent \coqdockw{Restart}%#<tt>Restart</tt>#. *)
adam@315 325
adamc@28 326 induction n; crush.
adamc@28 327 Qed.
adamc@41 328 (* end thide *)
adamc@28 329
adamc@28 330 (** We can check out the induction principle at work here: *)
adamc@28 331
adamc@28 332 Check nat_ind.
adamc@208 333 (** %\vspace{-.15in}% [[
adamc@208 334 nat_ind : forall P : nat -> Prop,
adamc@208 335 P O -> (forall n : nat, P n -> P (S n)) -> forall n : nat, P n
adamc@208 336
adamc@208 337 ]]
adamc@28 338
adam@315 339 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 340
adam@315 341 Since [nat] has a constructor that takes an argument, we may sometimes need to know that that constructor is injective.%\index{tactics!injection}\index{tactics!trivial}% *)
adamc@28 342
adamc@28 343 Theorem S_inj : forall n m : nat, S n = S m -> n = m.
adamc@41 344 (* begin thide *)
adamc@28 345 injection 1; trivial.
adamc@28 346 Qed.
adamc@41 347 (* end thide *)
adamc@28 348
adamc@28 349 (** [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 350
adam@315 351 There is also a very useful tactic called %\index{tactics!congruence}%[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 %\index{theory of equality and uninterpreted functions}\textit{%#<i>#complete decision procedure for the theory of equality and uninterpreted functions#</i>#%}%, plus some smarts about inductive types.
adamc@29 352
adamc@29 353 %\medskip%
adamc@29 354
adamc@29 355 We can define a type of lists of natural numbers. *)
adamc@29 356
adamc@29 357 Inductive nat_list : Set :=
adamc@29 358 | NNil : nat_list
adamc@29 359 | NCons : nat -> nat_list -> nat_list.
adamc@29 360
adamc@29 361 (** Recursive definitions are straightforward extensions of what we have seen before. *)
adamc@29 362
adamc@29 363 Fixpoint nlength (ls : nat_list) : nat :=
adamc@29 364 match ls with
adamc@29 365 | NNil => O
adamc@29 366 | NCons _ ls' => S (nlength ls')
adamc@29 367 end.
adamc@29 368
adamc@208 369 Fixpoint napp (ls1 ls2 : nat_list) : nat_list :=
adamc@29 370 match ls1 with
adamc@29 371 | NNil => ls2
adamc@29 372 | NCons n ls1' => NCons n (napp ls1' ls2)
adamc@29 373 end.
adamc@29 374
adamc@29 375 (** Inductive theorem proving can again be automated quite effectively. *)
adamc@29 376
adamc@29 377 Theorem nlength_napp : forall ls1 ls2 : nat_list, nlength (napp ls1 ls2)
adamc@29 378 = plus (nlength ls1) (nlength ls2).
adamc@41 379 (* begin thide *)
adamc@29 380 induction ls1; crush.
adamc@29 381 Qed.
adamc@41 382 (* end thide *)
adamc@29 383
adamc@29 384 Check nat_list_ind.
adamc@208 385 (** %\vspace{-.15in}% [[
adamc@208 386 nat_list_ind
adamc@29 387 : forall P : nat_list -> Prop,
adamc@29 388 P NNil ->
adamc@29 389 (forall (n : nat) (n0 : nat_list), P n0 -> P (NCons n n0)) ->
adamc@29 390 forall n : nat_list, P n
adamc@29 391 ]]
adamc@29 392
adamc@29 393 %\medskip%
adamc@29 394
adam@292 395 In general, we can implement any %``%#"#tree#"#%''% types as inductive types. For example, here are binary trees of naturals. *)
adamc@29 396
adamc@29 397 Inductive nat_btree : Set :=
adamc@29 398 | NLeaf : nat_btree
adamc@29 399 | NNode : nat_btree -> nat -> nat_btree -> nat_btree.
adamc@29 400
adamc@29 401 Fixpoint nsize (tr : nat_btree) : nat :=
adamc@29 402 match tr with
adamc@35 403 | NLeaf => S O
adamc@29 404 | NNode tr1 _ tr2 => plus (nsize tr1) (nsize tr2)
adamc@29 405 end.
adamc@29 406
adamc@208 407 Fixpoint nsplice (tr1 tr2 : nat_btree) : nat_btree :=
adamc@29 408 match tr1 with
adamc@35 409 | NLeaf => NNode tr2 O NLeaf
adamc@29 410 | NNode tr1' n tr2' => NNode (nsplice tr1' tr2) n tr2'
adamc@29 411 end.
adamc@29 412
adamc@29 413 Theorem plus_assoc : forall n1 n2 n3 : nat, plus (plus n1 n2) n3 = plus n1 (plus n2 n3).
adamc@41 414 (* begin thide *)
adamc@29 415 induction n1; crush.
adamc@29 416 Qed.
adamc@41 417 (* end thide *)
adamc@29 418
adamc@29 419 Theorem nsize_nsplice : forall tr1 tr2 : nat_btree, nsize (nsplice tr1 tr2)
adamc@29 420 = plus (nsize tr2) (nsize tr1).
adamc@41 421 (* begin thide *)
adam@315 422 (* begin hide *)
adam@375 423 Hint Rewrite n_plus_O plus_assoc.
adam@315 424 (* end hide *)
adam@375 425 (** [Hint] %\coqdockw{%#<tt>#Rewrite#</tt>#%}% [n_plus_O plus_assoc.] *)
adamc@29 426
adamc@29 427 induction tr1; crush.
adamc@29 428 Qed.
adamc@41 429 (* end thide *)
adamc@29 430
adam@315 431 (** It is convenient that these proofs go through so easily, but it is useful to check that the tree induction principle works as usual. *)
adam@315 432
adamc@29 433 Check nat_btree_ind.
adamc@208 434 (** %\vspace{-.15in}% [[
adamc@208 435 nat_btree_ind
adamc@29 436 : forall P : nat_btree -> Prop,
adamc@29 437 P NLeaf ->
adamc@29 438 (forall n : nat_btree,
adamc@29 439 P n -> forall (n0 : nat) (n1 : nat_btree), P n1 -> P (NNode n n0 n1)) ->
adamc@29 440 forall n : nat_btree, P n
adam@302 441 ]]
adam@315 442
adam@315 443 We have the usual two cases, one for each constructor of [nat_btree]. *)
adamc@30 444
adamc@30 445
adamc@30 446 (** * Parameterized Types *)
adamc@30 447
adam@316 448 (** We can also define %\index{polymorphism}%polymorphic inductive types, as with algebraic datatypes in Haskell and ML.%\index{Gallina terms!list}\index{Gallina terms!Nil}\index{Gallina terms!Cons}\index{Gallina terms!length}\index{Gallina terms!app}% *)
adamc@30 449
adamc@30 450 Inductive list (T : Set) : Set :=
adamc@30 451 | Nil : list T
adamc@30 452 | Cons : T -> list T -> list T.
adamc@30 453
adamc@30 454 Fixpoint length T (ls : list T) : nat :=
adamc@30 455 match ls with
adamc@30 456 | Nil => O
adamc@30 457 | Cons _ ls' => S (length ls')
adamc@30 458 end.
adamc@30 459
adamc@208 460 Fixpoint app T (ls1 ls2 : list T) : list T :=
adamc@30 461 match ls1 with
adamc@30 462 | Nil => ls2
adamc@30 463 | Cons x ls1' => Cons x (app ls1' ls2)
adamc@30 464 end.
adamc@30 465
adamc@30 466 Theorem length_app : forall T (ls1 ls2 : list T), length (app ls1 ls2)
adamc@30 467 = plus (length ls1) (length ls2).
adamc@41 468 (* begin thide *)
adamc@30 469 induction ls1; crush.
adamc@30 470 Qed.
adamc@41 471 (* end thide *)
adamc@30 472
adam@316 473 (** There is a useful shorthand for writing many definitions that share the same parameter, based on Coq's %\index{sections}\index{Vernacular commands!Section}\index{Vernacular commands!Variable}\textit{%#<i>#section#</i>#%}% mechanism. The following block of code is equivalent to the above: *)
adamc@30 474
adamc@30 475 (* begin hide *)
adamc@30 476 Reset list.
adamc@30 477 (* end hide *)
adamc@30 478
adamc@30 479 Section list.
adamc@30 480 Variable T : Set.
adamc@30 481
adamc@30 482 Inductive list : Set :=
adamc@30 483 | Nil : list
adamc@30 484 | Cons : T -> list -> list.
adamc@30 485
adamc@30 486 Fixpoint length (ls : list) : nat :=
adamc@30 487 match ls with
adamc@30 488 | Nil => O
adamc@30 489 | Cons _ ls' => S (length ls')
adamc@30 490 end.
adamc@30 491
adamc@208 492 Fixpoint app (ls1 ls2 : list) : list :=
adamc@30 493 match ls1 with
adamc@30 494 | Nil => ls2
adamc@30 495 | Cons x ls1' => Cons x (app ls1' ls2)
adamc@30 496 end.
adamc@30 497
adamc@30 498 Theorem length_app : forall ls1 ls2 : list, length (app ls1 ls2)
adamc@30 499 = plus (length ls1) (length ls2).
adamc@41 500 (* begin thide *)
adamc@30 501 induction ls1; crush.
adamc@30 502 Qed.
adamc@41 503 (* end thide *)
adamc@30 504 End list.
adamc@30 505
adamc@35 506 (* begin hide *)
adamc@35 507 Implicit Arguments Nil [T].
adamc@35 508 (* end hide *)
adamc@35 509
adamc@210 510 (** 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 511
adamc@202 512 Print list.
adamc@208 513 (** %\vspace{-.15in}% [[
adamc@208 514 Inductive list (T : Set) : Set :=
adam@316 515 Nil : list T | Cons : T -> list T -> list T
adamc@208 516
adamc@202 517 ]]
adamc@30 518
adamc@202 519 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 520
adamc@30 521 Check length.
adamc@208 522 (** %\vspace{-.15in}% [[
adamc@208 523 length
adamc@30 524 : forall T : Set, list T -> nat
adamc@30 525 ]]
adamc@30 526
adamc@202 527 The parameter [T] is treated as a new argument to the induction principle, too. *)
adamc@30 528
adamc@30 529 Check list_ind.
adamc@208 530 (** %\vspace{-.15in}% [[
adamc@208 531 list_ind
adamc@30 532 : forall (T : Set) (P : list T -> Prop),
adamc@30 533 P (Nil T) ->
adamc@30 534 (forall (t : T) (l : list T), P l -> P (Cons t l)) ->
adamc@30 535 forall l : list T, P l
adamc@30 536 ]]
adamc@30 537
adamc@30 538 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 539
adamc@31 540
adamc@31 541 (** * Mutually Inductive Types *)
adamc@31 542
adamc@31 543 (** We can define inductive types that refer to each other: *)
adamc@31 544
adamc@31 545 Inductive even_list : Set :=
adamc@31 546 | ENil : even_list
adamc@31 547 | ECons : nat -> odd_list -> even_list
adamc@31 548
adamc@31 549 with odd_list : Set :=
adamc@31 550 | OCons : nat -> even_list -> odd_list.
adamc@31 551
adamc@31 552 Fixpoint elength (el : even_list) : nat :=
adamc@31 553 match el with
adamc@31 554 | ENil => O
adamc@31 555 | ECons _ ol => S (olength ol)
adamc@31 556 end
adamc@31 557
adamc@31 558 with olength (ol : odd_list) : nat :=
adamc@31 559 match ol with
adamc@31 560 | OCons _ el => S (elength el)
adamc@31 561 end.
adamc@31 562
adamc@208 563 Fixpoint eapp (el1 el2 : even_list) : even_list :=
adamc@31 564 match el1 with
adamc@31 565 | ENil => el2
adamc@31 566 | ECons n ol => ECons n (oapp ol el2)
adamc@31 567 end
adamc@31 568
adamc@208 569 with oapp (ol : odd_list) (el : even_list) : odd_list :=
adamc@31 570 match ol with
adamc@31 571 | OCons n el' => OCons n (eapp el' el)
adamc@31 572 end.
adamc@31 573
adamc@31 574 (** 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 575
adamc@31 576 Theorem elength_eapp : forall el1 el2 : even_list,
adamc@31 577 elength (eapp el1 el2) = plus (elength el1) (elength el2).
adamc@41 578 (* begin thide *)
adamc@31 579 induction el1; crush.
adamc@31 580
adamc@31 581 (** One goal remains: [[
adamc@31 582
adamc@31 583 n : nat
adamc@31 584 o : odd_list
adamc@31 585 el2 : even_list
adamc@31 586 ============================
adamc@31 587 S (olength (oapp o el2)) = S (plus (olength o) (elength el2))
adamc@31 588 ]]
adamc@31 589
adamc@31 590 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 591
adamc@31 592 Abort.
adamc@31 593 Check even_list_ind.
adamc@208 594 (** %\vspace{-.15in}% [[
adamc@208 595 even_list_ind
adamc@31 596 : forall P : even_list -> Prop,
adamc@31 597 P ENil ->
adamc@31 598 (forall (n : nat) (o : odd_list), P (ECons n o)) ->
adamc@31 599 forall e : even_list, P e
adamc@208 600
adamc@31 601 ]]
adamc@31 602
adam@316 603 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 %\index{Vernacular commands!Scheme}%[Scheme] command. *)
adamc@31 604
adamc@31 605 Scheme even_list_mut := Induction for even_list Sort Prop
adamc@31 606 with odd_list_mut := Induction for odd_list Sort Prop.
adamc@31 607
adam@316 608 (** This invocation of [Scheme] asks for the creation of induction principles [even_list_mut] for the type [even_list] and [odd_list_mut] for the type [odd_list]. The [Induction] keyword says we want standard induction schemes, since [Scheme] supports more exotic choices. Finally, [Sort Prop] establishes that we really want induction schemes, not recursion schemes, which are the same according to Curry-Howard, save for the [Prop]/[Set] distinction. *)
adam@316 609
adamc@31 610 Check even_list_mut.
adamc@208 611 (** %\vspace{-.15in}% [[
adamc@208 612 even_list_mut
adamc@31 613 : forall (P : even_list -> Prop) (P0 : odd_list -> Prop),
adamc@31 614 P ENil ->
adamc@31 615 (forall (n : nat) (o : odd_list), P0 o -> P (ECons n o)) ->
adamc@31 616 (forall (n : nat) (e : even_list), P e -> P0 (OCons n e)) ->
adamc@31 617 forall e : even_list, P e
adamc@208 618
adamc@31 619 ]]
adamc@31 620
adam@316 621 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.%\index{tactics!apply}% *)
adamc@31 622
adamc@31 623 Theorem n_plus_O' : forall n : nat, plus n O = n.
adamc@31 624 apply (nat_ind (fun n => plus n O = n)); crush.
adamc@31 625 Qed.
adamc@31 626
adamc@31 627 (** 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 628
adamc@31 629 This technique generalizes to our mutual example: *)
adamc@31 630
adamc@31 631 Theorem elength_eapp : forall el1 el2 : even_list,
adamc@31 632 elength (eapp el1 el2) = plus (elength el1) (elength el2).
adamc@41 633
adamc@31 634 apply (even_list_mut
adamc@31 635 (fun el1 : even_list => forall el2 : even_list,
adamc@31 636 elength (eapp el1 el2) = plus (elength el1) (elength el2))
adamc@31 637 (fun ol : odd_list => forall el : even_list,
adamc@31 638 olength (oapp ol el) = plus (olength ol) (elength el))); crush.
adamc@31 639 Qed.
adamc@41 640 (* end thide *)
adamc@31 641
adamc@31 642 (** 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 643
adamc@33 644
adamc@33 645 (** * Reflexive Types *)
adamc@33 646
adam@316 647 (** 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. Our example will be an encoding of the syntax of first-order logic. Since the idea of syntactic encodings of logic may require a bit of acclimation, let us first consider a simpler formula type for a subset of propositional logic. *)
adam@316 648
adam@316 649 Inductive pformula : Set :=
adam@316 650 | Truth : pformula
adam@316 651 | Falsehood : pformula
adam@316 652 | Conjunction : pformula -> pformula -> pformula.
adam@316 653
adam@317 654 (** A key distinction here is between, for instance, the %\emph{%#<i>#syntax#</i>#%}% [Truth] and its %\emph{%#<i>#semantics#</i>#%}% [True]. We can make the semantics explicit with a recursive function. This function uses the infix operator %\index{Gallina operators!/\textbackslash}%[/\], which desugars to uses of the type family %\index{Gallina terms!and}%[and] from the standard library. The family [and] implements conjunction, the [Prop] Curry-Howard analogue of the usual pair type from functional programming (which is the type family %\index{Gallina terms!prod}%[prod] in Coq's standard library). *)
adam@316 655
adam@316 656 Fixpoint pformulaDenote (f : pformula) : Prop :=
adam@316 657 match f with
adam@316 658 | Truth => True
adam@316 659 | Falsehood => False
adam@316 660 | Conjunction f1 f2 => pformulaDenote f1 /\ pformulaDenote f2
adam@316 661 end.
adam@316 662
adam@392 663 (** This is just a warm-up that does not use reflexive types, the new feature we mean to introduce. When we set our sights on first-order logic instead, it becomes very handy to give constructors recursive arguments that are functions. *)
adamc@33 664
adamc@33 665 Inductive formula : Set :=
adamc@33 666 | Eq : nat -> nat -> formula
adamc@33 667 | And : formula -> formula -> formula
adamc@33 668 | Forall : (nat -> formula) -> formula.
adamc@33 669
adam@316 670 (** 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]:%\index{Vernacular commands!Example}% *)
adamc@33 671
adamc@33 672 Example forall_refl : formula := Forall (fun x => Eq x x).
adamc@33 673
adamc@33 674 (** We can write recursive functions over reflexive types quite naturally. Here is one translating our formulas into native Coq propositions. *)
adamc@33 675
adamc@33 676 Fixpoint formulaDenote (f : formula) : Prop :=
adamc@33 677 match f with
adamc@33 678 | Eq n1 n2 => n1 = n2
adamc@33 679 | And f1 f2 => formulaDenote f1 /\ formulaDenote f2
adamc@33 680 | Forall f' => forall n : nat, formulaDenote (f' n)
adamc@33 681 end.
adamc@33 682
adamc@33 683 (** We can also encode a trivial formula transformation that swaps the order of equality and conjunction operands. *)
adamc@33 684
adamc@33 685 Fixpoint swapper (f : formula) : formula :=
adamc@33 686 match f with
adamc@33 687 | Eq n1 n2 => Eq n2 n1
adamc@33 688 | And f1 f2 => And (swapper f2) (swapper f1)
adamc@33 689 | Forall f' => Forall (fun n => swapper (f' n))
adamc@33 690 end.
adamc@33 691
adamc@33 692 (** It is helpful to prove that this transformation does not make true formulas false. *)
adamc@33 693
adamc@33 694 Theorem swapper_preserves_truth : forall f, formulaDenote f -> formulaDenote (swapper f).
adamc@41 695 (* begin thide *)
adamc@33 696 induction f; crush.
adamc@33 697 Qed.
adamc@41 698 (* end thide *)
adamc@33 699
adamc@33 700 (** We can take a look at the induction principle behind this proof. *)
adamc@33 701
adamc@33 702 Check formula_ind.
adamc@208 703 (** %\vspace{-.15in}% [[
adamc@208 704 formula_ind
adamc@33 705 : forall P : formula -> Prop,
adamc@33 706 (forall n n0 : nat, P (Eq n n0)) ->
adamc@33 707 (forall f0 : formula,
adamc@33 708 P f0 -> forall f1 : formula, P f1 -> P (And f0 f1)) ->
adamc@33 709 (forall f1 : nat -> formula,
adamc@33 710 (forall n : nat, P (f1 n)) -> P (Forall f1)) ->
adamc@33 711 forall f2 : formula, P f2
adamc@208 712
adamc@208 713 ]]
adamc@33 714
adamc@208 715 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 716
adamc@33 717 %\medskip%
adamc@33 718
adam@316 719 Up to this point, we have seen how to encode in Coq more and more of what is possible with algebraic datatypes in %\index{Haskell}%Haskell and %\index{ML}%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 720
adam@316 721 Given our last example of an inductive type, many readers are probably eager to try encoding the syntax of %\index{lambda calculus}%lambda calculus. Indeed, the function-based representation technique that we just used, called %\index{higher-order abstract syntax}\index{HOAS|see{higher-order abstract syntax}}\textit{%#<i>#higher-order abstract syntax (HOAS)#</i>#%}~\cite{HOAS}%, is the representation of choice for lambda calculi in %\index{Twelf}%Twelf and in many applications implemented in Haskell and ML. Let us try to import that choice to Coq: *)
adamc@33 722 (** [[
adamc@33 723 Inductive term : Set :=
adamc@33 724 | App : term -> term -> term
adamc@33 725 | Abs : (term -> term) -> term.
adamc@33 726 ]]
adamc@33 727
adam@316 728 <<
adam@316 729 Error: Non strictly positive occurrence of "term" in "(term -> term) -> term"
adam@316 730 >>
adam@316 731
adam@316 732 We have run afoul of the %\index{strict positivity requirement}\index{positivity requirement}\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. Our candidate definition above violates the positivity requirement because it involves an argument of type [term -> term], where the type [term] that we are defining appears to the left of an arrow. The candidate type of [App] is fine, however, since every occurrence of [term] is either a constructor argument or the final result type.
adamc@33 733
adamc@33 734 Why must Coq enforce this restriction? Imagine that our last definition had been accepted, allowing us to write this function:
adamc@33 735
adamc@33 736 [[
adamc@33 737 Definition uhoh (t : term) : term :=
adamc@33 738 match t with
adamc@33 739 | Abs f => f t
adamc@33 740 | _ => t
adamc@33 741 end.
adamc@33 742
adamc@205 743 ]]
adamc@205 744
adamc@33 745 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 746
adam@316 747 %\index{termination checking}%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 748
adamc@33 749 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 750
adamc@34 751
adam@317 752 (** * An Interlude on Induction Principles *)
adamc@34 753
adam@317 754 (** 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 %\index{induction principles}%induction principles we have used. A close look at the details here will help us construct induction principles manually, which we will see is necessary for some more advanced inductive definitions. *)
adamc@34 755
adam@317 756 (* begin hide *)
adamc@34 757 Print unit_ind.
adam@317 758 (* end hide *)
adam@317 759 (** %\noindent%[Print] [unit_ind.] *)
adam@317 760 (** [[
adamc@208 761 unit_ind =
adamc@208 762 fun P : unit -> Prop => unit_rect P
adamc@34 763 : forall P : unit -> Prop, P tt -> forall u : unit, P u
adamc@208 764
adamc@34 765 ]]
adamc@34 766
adam@317 767 We see that this induction principle is defined in terms of a more general principle, [unit_rect]. The %\texttt{%#<tt>#rec#</tt>#%}% stands for %``%#"#recursion principle,#"#%''% and the %\texttt{%#<tt>#t#</tt>#%}% at the end stands for [Type]. *)
adamc@34 768
adamc@34 769 Check unit_rect.
adamc@208 770 (** %\vspace{-.15in}% [[
adamc@208 771 unit_rect
adamc@34 772 : forall P : unit -> Type, P tt -> forall u : unit, P u
adamc@208 773
adamc@34 774 ]]
adamc@34 775
adamc@34 776 [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 777
adam@317 778 (* begin hide *)
adamc@34 779 Print unit_rec.
adam@317 780 (* end hide *)
adam@317 781 (** %\noindent%[Print] [unit_rec.] *)
adam@317 782 (** [[
adamc@208 783 unit_rec =
adamc@208 784 fun P : unit -> Set => unit_rect P
adamc@34 785 : forall P : unit -> Set, P tt -> forall u : unit, P u
adamc@208 786
adamc@34 787 ]]
adamc@34 788
adam@317 789 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 %\index{recursion principles}%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 790
adamc@34 791 Definition always_O (u : unit) : nat :=
adamc@34 792 match u with
adamc@34 793 | tt => O
adamc@34 794 end.
adamc@34 795
adamc@34 796 Definition always_O' (u : unit) : nat :=
adamc@34 797 unit_rec (fun _ : unit => nat) O u.
adamc@34 798
adamc@34 799 (** 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 800
adam@317 801 (* begin hide *)
adamc@34 802 Print unit_rect.
adam@317 803 (* end hide *)
adam@317 804 (** %\noindent%[Print] [unit_rect.] *)
adam@317 805 (** [[
adamc@208 806 unit_rect =
adamc@208 807 fun (P : unit -> Type) (f : P tt) (u : unit) =>
adamc@208 808 match u as u0 return (P u0) with
adamc@208 809 | tt => f
adamc@208 810 end
adamc@34 811 : forall P : unit -> Type, P tt -> forall u : unit, P u
adamc@208 812
adamc@34 813 ]]
adamc@34 814
adam@317 815 The only new wrinkle here is the annotations on the [match] expression. This is a %\index{dependent pattern matching}\emph{%#<i>#dependently typed#</i>#%}% pattern match, because the %\emph{%#<i>#type#</i>#%}% of the expression depends on the %\emph{%#<i>#value#</i>#%}% being matched on. Of course, for this example, the dependency is degenerate; the value being matched on has type [unit], so it may only take on a single known value, [tt]. We will meet more involved examples later, especially in Part II of the book.
adam@317 816
adam@317 817 %\index{type inference}%Type inference for dependent pattern matching is undecidable, which can be proved by reduction from %\index{higher-order unification}%higher-order unification%~\cite{HOU}%. Thus, we often find ourselves needing to annotate our programs in a way that explains dependencies to the type checker. In the example of [unit_rect], we have an %\index{Gallina terms!as}%[as] clause, which binds a name for the discriminee; and a %\index{Gallina terms!return}%[return] clause, which gives a way to compute the [match] result type as a function of the discriminee.
adamc@34 818
adamc@34 819 To prove that [unit_rect] is nothing special, we can reimplement it manually. *)
adamc@34 820
adamc@34 821 Definition unit_rect' (P : unit -> Type) (f : P tt) (u : unit) :=
adamc@208 822 match u with
adamc@34 823 | tt => f
adamc@34 824 end.
adamc@34 825
adam@317 826 (** We rely on Coq's heuristics for inferring [match] annotations, which are not consulted in the pretty-printing of terms.
adamc@34 827
adam@317 828 We can check the implementation of %\coqdocdefinition{%#<tt>#nat_rect#</tt>#%}% as well: *)
adamc@34 829
adam@317 830 (* begin hide *)
adamc@34 831 Print nat_rect.
adam@317 832 (* end hide *)
adam@317 833 (** %\noindent%[Print] %\coqdocdefinition{%#<tt>#nat_rect#</tt>#%}%[.] *)
adam@317 834
adam@317 835 (** %\hspace{-.05in}\coqdocdefinition{%#<tt>#nat_rect#</tt>#%}% [=] *)
adam@317 836 (** %\vspace{-.05in}% [[
adamc@208 837 fun (P : nat -> Type) (f : P O) (f0 : forall n : nat, P n -> P (S n)) =>
adamc@208 838 fix F (n : nat) : P n :=
adamc@208 839 match n as n0 return (P n0) with
adamc@208 840 | O => f
adamc@208 841 | S n0 => f0 n0 (F n0)
adamc@208 842 end
adamc@208 843 : forall P : nat -> Type,
adamc@208 844 P O -> (forall n : nat, P n -> P (S n)) -> forall n : nat, P n
adamc@208 845 ]]
adamc@34 846
adam@317 847 Now we have an actual recursive definition. %\index{Gallina terms!fix}%[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 %\coqdocdefinition{%#<tt>#nat_rect#</tt>#%}% better by reimplementing [nat_ind] using sections. *)
adamc@34 848
adam@317 849 Section nat_ind'.
adamc@208 850 (** First, we have the property of natural numbers that we aim to prove. *)
adamc@34 851
adam@317 852 Variable P : nat -> Prop.
adamc@34 853
adam@317 854 (** Then we require a proof of the [O] case, which we declare with the command %\index{Vernacular commands!Hypothesis}%[Hypothesis], which is a synonym for [Variable] that, by convention, is used for variables whose types are propositions. *)
adamc@34 855
adam@317 856 Hypothesis O_case : P O.
adamc@34 857
adamc@208 858 (** Next is a proof of the [S] case, which may assume an inductive hypothesis. *)
adamc@34 859
adam@317 860 Hypothesis S_case : forall n : nat, P n -> P (S n).
adamc@34 861
adamc@208 862 (** Finally, we define a recursive function to tie the pieces together. *)
adamc@34 863
adam@317 864 Fixpoint nat_ind' (n : nat) : P n :=
adam@317 865 match n with
adam@317 866 | O => O_case
adam@317 867 | S n' => S_case (nat_ind' n')
adam@317 868 end.
adam@317 869 End nat_ind'.
adamc@34 870
adam@317 871 (** 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 %\coqdocdefinition{%#<tt>#nat_rect#</tt>#%}%.
adamc@34 872
adam@317 873 %\medskip%
adamc@34 874
adam@317 875 We can also examine the definition of [even_list_mut], which we generated with [Scheme] for a mutually recursive type. *)
adamc@34 876
adam@317 877 (* begin hide *)
adam@317 878 Print even_list_mut.
adam@317 879 (* end hide *)
adam@317 880 (** %\noindent%[Print] %\coqdocdefinition{%#<tt>#even_list_mut#</tt>#%}%[.] *)
adam@317 881 (** [[
adam@317 882 even_list_mut =
adam@317 883 fun (P : even_list -> Prop) (P0 : odd_list -> Prop)
adam@317 884 (f : P ENil) (f0 : forall (n : nat) (o : odd_list), P0 o -> P (ECons n o))
adam@317 885 (f1 : forall (n : nat) (e : even_list), P e -> P0 (OCons n e)) =>
adam@317 886 fix F (e : even_list) : P e :=
adam@317 887 match e as e0 return (P e0) with
adam@317 888 | ENil => f
adam@317 889 | ECons n o => f0 n o (F0 o)
adam@317 890 end
adam@317 891 with F0 (o : odd_list) : P0 o :=
adam@317 892 match o as o0 return (P0 o0) with
adam@317 893 | OCons n e => f1 n e (F e)
adam@317 894 end
adam@317 895 for F
adam@317 896 : forall (P : even_list -> Prop) (P0 : odd_list -> Prop),
adam@317 897 P ENil ->
adam@317 898 (forall (n : nat) (o : odd_list), P0 o -> P (ECons n o)) ->
adam@317 899 (forall (n : nat) (e : even_list), P e -> P0 (OCons n e)) ->
adam@317 900 forall e : even_list, P e
adamc@34 901
adam@317 902 ]]
adamc@34 903
adam@317 904 We see a mutually recursive [fix], with the different functions separated by %\index{Gallina terms!with}%[with] in the same way that they would be separated by %\texttt{%#<tt>#and#</tt>#%}% in ML. A final %\index{Gallina terms!for}%[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 905
adam@317 906 Section even_list_mut'.
adam@317 907 (** First, we need the properties that we are proving. *)
adamc@208 908
adam@317 909 Variable Peven : even_list -> Prop.
adam@317 910 Variable Podd : odd_list -> Prop.
adamc@208 911
adam@317 912 (** Next, we need proofs of the three cases. *)
adamc@208 913
adam@317 914 Hypothesis ENil_case : Peven ENil.
adam@317 915 Hypothesis ECons_case : forall (n : nat) (o : odd_list), Podd o -> Peven (ECons n o).
adam@317 916 Hypothesis OCons_case : forall (n : nat) (e : even_list), Peven e -> Podd (OCons n e).
adamc@208 917
adam@317 918 (** Finally, we define the recursive functions. *)
adamc@208 919
adam@317 920 Fixpoint even_list_mut' (e : even_list) : Peven e :=
adam@317 921 match e with
adam@317 922 | ENil => ENil_case
adam@317 923 | ECons n o => ECons_case n (odd_list_mut' o)
adam@317 924 end
adam@317 925 with odd_list_mut' (o : odd_list) : Podd o :=
adam@317 926 match o with
adam@317 927 | OCons n e => OCons_case n (even_list_mut' e)
adam@317 928 end.
adamc@34 929 End even_list_mut'.
adamc@34 930
adamc@34 931 (** 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 932
adamc@34 933 Section formula_ind'.
adamc@34 934 Variable P : formula -> Prop.
adamc@38 935 Hypothesis Eq_case : forall n1 n2 : nat, P (Eq n1 n2).
adamc@38 936 Hypothesis And_case : forall f1 f2 : formula,
adamc@34 937 P f1 -> P f2 -> P (And f1 f2).
adamc@38 938 Hypothesis Forall_case : forall f : nat -> formula,
adamc@34 939 (forall n : nat, P (f n)) -> P (Forall f).
adamc@34 940
adamc@34 941 Fixpoint formula_ind' (f : formula) : P f :=
adamc@208 942 match f with
adamc@34 943 | Eq n1 n2 => Eq_case n1 n2
adamc@34 944 | And f1 f2 => And_case (formula_ind' f1) (formula_ind' f2)
adamc@34 945 | Forall f' => Forall_case f' (fun n => formula_ind' (f' n))
adamc@34 946 end.
adamc@34 947 End formula_ind'.
adamc@34 948
adam@317 949 (** It is apparent that induction principle implementations involve some tedium but not terribly much creativity. *)
adam@317 950
adamc@35 951
adamc@35 952 (** * Nested Inductive Types *)
adamc@35 953
adamc@35 954 (** 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 955
adamc@35 956 Inductive nat_tree : Set :=
adamc@35 957 | NLeaf' : nat_tree
adamc@35 958 | NNode' : nat -> list nat_tree -> nat_tree.
adamc@35 959
adam@317 960 (** This is an example of a %\index{nested inductive type}\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 961
adam@317 962 Like we encountered for mutual inductive types, we find that the automatically generated induction principle for [nat_tree] is too weak. *)
adamc@35 963
adamc@35 964 Check nat_tree_ind.
adamc@208 965 (** %\vspace{-.15in}% [[
adamc@208 966 nat_tree_ind
adamc@35 967 : forall P : nat_tree -> Prop,
adamc@35 968 P NLeaf' ->
adamc@35 969 (forall (n : nat) (l : list nat_tree), P (NNode' n l)) ->
adamc@35 970 forall n : nat_tree, P n
adamc@208 971
adamc@35 972 ]]
adamc@35 973
adam@317 974 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 of different type families. This is roughly the same creativity employed in the traditional task of strengthening an induction hypothesis. 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 975
adamc@35 976 First, we will need an auxiliary definition, characterizing what it means for a property to hold of every element of a list. *)
adamc@35 977
adamc@35 978 Section All.
adamc@35 979 Variable T : Set.
adamc@35 980 Variable P : T -> Prop.
adamc@35 981
adamc@35 982 Fixpoint All (ls : list T) : Prop :=
adamc@35 983 match ls with
adamc@35 984 | Nil => True
adamc@35 985 | Cons h t => P h /\ All t
adamc@35 986 end.
adamc@35 987 End All.
adamc@35 988
adam@317 989 (** It will be useful to review the definitions of [True] and [/\], since we will want to write manual proofs of them below. *)
adamc@35 990
adam@317 991 (* begin hide *)
adamc@35 992 Print True.
adam@317 993 (* end hide *)
adam@317 994 (** %\noindent%[Print] %\coqdocinductive{%#<tt>#True#</tt>#%}%[.] *)
adam@317 995 (** [[
adamc@208 996 Inductive True : Prop := I : True
adamc@208 997 ]]
adamc@35 998
adamc@35 999 That is, [True] is a proposition with exactly one proof, [I], which we may always supply trivially.
adamc@35 1000
adam@317 1001 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 %\index{Vernacular commands!Locate}\coqdockw{%#<tt>#Locate#</tt>#%}% command, whose argument may be a parsing token.%\index{Gallina terms!and}% *)
adamc@35 1002
adam@317 1003 (* begin hide *)
adamc@35 1004 Locate "/\".
adam@317 1005 (* end hide *)
adam@317 1006 (** %\noindent \coqdockw{Locate}%#<tt>Locate</tt># ["/\".] *)
adam@317 1007 (** [[
adam@317 1008 "A /\ B" := and A B : type_scope (default interpretation)
adam@302 1009 ]]
adam@302 1010 *)
adamc@35 1011
adam@317 1012 (* begin hide *)
adamc@35 1013 Print and.
adam@317 1014 (* end hide *)
adam@317 1015 (** %\noindent%[Print] %\coqdocinductive{%#<tt>#and#</tt>#%}%[.] *)
adam@317 1016 (** [[
adamc@208 1017 Inductive and (A : Prop) (B : Prop) : Prop := conj : A -> B -> A /\ B
adam@317 1018 ]]
adam@317 1019 %\vspace{-.1in}%
adam@317 1020 <<
adamc@208 1021 For conj: Arguments A, B are implicit
adam@317 1022 >>
adamc@35 1023
adam@317 1024 In addition to the definition of %\coqdocinductive{%#<tt>#and#</tt>#%}% itself, we get information on %\index{implicit arguments}%implicit arguments (and some other information that we omit here). 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 1025
adamc@35 1026 %\medskip%
adamc@35 1027
adamc@35 1028 Now we create a section for our induction principle, following the same basic plan as in the last section of this chapter. *)
adamc@35 1029
adamc@35 1030 Section nat_tree_ind'.
adamc@35 1031 Variable P : nat_tree -> Prop.
adamc@35 1032
adamc@38 1033 Hypothesis NLeaf'_case : P NLeaf'.
adamc@38 1034 Hypothesis NNode'_case : forall (n : nat) (ls : list nat_tree),
adamc@35 1035 All P ls -> P (NNode' n ls).
adamc@35 1036
adamc@35 1037 (** 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 1038
adamc@35 1039 [[
adamc@35 1040 Fixpoint nat_tree_ind' (tr : nat_tree) : P tr :=
adamc@208 1041 match tr with
adamc@35 1042 | NLeaf' => NLeaf'_case
adamc@35 1043 | NNode' n ls => NNode'_case n ls (list_nat_tree_ind ls)
adamc@35 1044 end
adamc@35 1045
adamc@35 1046 with list_nat_tree_ind (ls : list nat_tree) : All P ls :=
adamc@208 1047 match ls with
adamc@35 1048 | Nil => I
adamc@35 1049 | Cons tr rest => conj (nat_tree_ind' tr) (list_nat_tree_ind rest)
adamc@35 1050 end.
adamc@35 1051
adamc@205 1052 ]]
adamc@205 1053
adam@317 1054 Coq rejects this definition, saying
adam@317 1055 <<
adam@317 1056 Recursive call to nat_tree_ind' has principal argument equal to "tr"
adam@317 1057 instead of rest.
adam@317 1058 >>
adam@317 1059
adam@317 1060 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 1061
adamc@35 1062 Fixpoint nat_tree_ind' (tr : nat_tree) : P tr :=
adamc@208 1063 match tr with
adamc@35 1064 | NLeaf' => NLeaf'_case
adamc@35 1065 | NNode' n ls => NNode'_case n ls
adamc@35 1066 ((fix list_nat_tree_ind (ls : list nat_tree) : All P ls :=
adamc@208 1067 match ls with
adamc@35 1068 | Nil => I
adamc@35 1069 | Cons tr rest => conj (nat_tree_ind' tr) (list_nat_tree_ind rest)
adamc@35 1070 end) ls)
adamc@35 1071 end.
adamc@35 1072
adam@279 1073 (** 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 1074
adamc@35 1075 End nat_tree_ind'.
adamc@35 1076
adamc@35 1077 (** 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 1078
adamc@35 1079 Section map.
adamc@35 1080 Variables T T' : Set.
adam@317 1081 Variable F : T -> T'.
adamc@35 1082
adamc@35 1083 Fixpoint map (ls : list T) : list T' :=
adamc@35 1084 match ls with
adamc@35 1085 | Nil => Nil
adam@317 1086 | Cons h t => Cons (F h) (map t)
adamc@35 1087 end.
adamc@35 1088 End map.
adamc@35 1089
adamc@35 1090 Fixpoint sum (ls : list nat) : nat :=
adamc@35 1091 match ls with
adamc@35 1092 | Nil => O
adamc@35 1093 | Cons h t => plus h (sum t)
adamc@35 1094 end.
adamc@35 1095
adamc@35 1096 (** Now we can define a size function over our trees. *)
adamc@35 1097
adamc@35 1098 Fixpoint ntsize (tr : nat_tree) : nat :=
adamc@35 1099 match tr with
adamc@35 1100 | NLeaf' => S O
adamc@35 1101 | NNode' _ trs => S (sum (map ntsize trs))
adamc@35 1102 end.
adamc@35 1103
adamc@35 1104 (** 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 1105
adamc@208 1106 Fixpoint ntsplice (tr1 tr2 : nat_tree) : nat_tree :=
adamc@35 1107 match tr1 with
adamc@35 1108 | NLeaf' => NNode' O (Cons tr2 Nil)
adamc@35 1109 | NNode' n Nil => NNode' n (Cons tr2 Nil)
adamc@35 1110 | NNode' n (Cons tr trs) => NNode' n (Cons (ntsplice tr tr2) trs)
adamc@35 1111 end.
adamc@35 1112
adamc@35 1113 (** 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 1114
adamc@41 1115 (* begin thide *)
adamc@35 1116 Lemma plus_S : forall n1 n2 : nat,
adamc@35 1117 plus n1 (S n2) = S (plus n1 n2).
adamc@35 1118 induction n1; crush.
adamc@35 1119 Qed.
adamc@41 1120 (* end thide *)
adamc@35 1121
adamc@35 1122 (** Now we begin the proof of the theorem, adding the lemma [plus_S] as a hint. *)
adamc@35 1123
adamc@35 1124 Theorem ntsize_ntsplice : forall tr1 tr2 : nat_tree, ntsize (ntsplice tr1 tr2)
adamc@35 1125 = plus (ntsize tr2) (ntsize tr1).
adamc@41 1126 (* begin thide *)
adam@317 1127 (* begin hide *)
adam@375 1128 Hint Rewrite plus_S.
adam@317 1129 (* end hide *)
adam@375 1130 (** [Hint] %\coqdockw{%#<tt>#Rewrite#</tt>#%}% [plus_S.] *)
adamc@35 1131
adam@317 1132 (** We know that the standard induction principle is insufficient for the task, so we need to provide a %\index{tactics!using}%[using] clause for the [induction] tactic to specify our alternate principle. *)
adamc@208 1133
adamc@35 1134 induction tr1 using nat_tree_ind'; crush.
adamc@35 1135
adamc@35 1136 (** One subgoal remains: [[
adamc@35 1137 n : nat
adamc@35 1138 ls : list nat_tree
adamc@35 1139 H : All
adamc@35 1140 (fun tr1 : nat_tree =>
adamc@35 1141 forall tr2 : nat_tree,
adamc@35 1142 ntsize (ntsplice tr1 tr2) = plus (ntsize tr2) (ntsize tr1)) ls
adamc@35 1143 tr2 : nat_tree
adamc@35 1144 ============================
adamc@35 1145 ntsize
adamc@35 1146 match ls with
adamc@35 1147 | Nil => NNode' n (Cons tr2 Nil)
adamc@35 1148 | Cons tr trs => NNode' n (Cons (ntsplice tr tr2) trs)
adamc@35 1149 end = S (plus (ntsize tr2) (sum (map ntsize ls)))
adamc@208 1150
adamc@35 1151 ]]
adamc@35 1152
adamc@35 1153 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 1154
adamc@35 1155 destruct ls; crush.
adamc@35 1156
adam@317 1157 (** We can go further in automating the proof by exploiting the hint mechanism.%\index{Vernacular commands!Hint Extern}% *)
adamc@35 1158
adam@317 1159 (* begin hide *)
adamc@35 1160 Restart.
adam@317 1161 (* end hide *)
adam@317 1162 (** %\hspace{-.075in}\coqdockw{%#<tt>#Restart#</tt>#%}%[.] *)
adam@317 1163
adamc@35 1164 Hint Extern 1 (ntsize (match ?LS with Nil => _ | Cons _ _ => _ end) = _) =>
adamc@35 1165 destruct LS; crush.
adamc@35 1166 induction tr1 using nat_tree_ind'; crush.
adamc@35 1167 Qed.
adamc@41 1168 (* end thide *)
adamc@35 1169
adamc@35 1170 (** 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 1171
adam@317 1172 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, and the hint can continue working even if the rest of the proof structure changes significantly. *)
adamc@36 1173
adamc@36 1174
adamc@36 1175 (** * Manual Proofs About Constructors *)
adamc@36 1176
adam@317 1177 (** It can be useful to understand how tactics like %\index{tactics!discriminate}%[discriminate] and %\index{tactics!injection}%[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 1178
adamc@36 1179 Theorem true_neq_false : true <> false.
adamc@208 1180
adamc@41 1181 (* begin thide *)
adam@317 1182 (** We begin with the tactic %\index{tactics!red}%[red], which is short for %``%#"#one step of reduction,#"#%''% to unfold the definition of logical negation. *)
adamc@36 1183
adamc@36 1184 red.
adamc@36 1185 (** [[
adamc@36 1186 ============================
adamc@36 1187 true = false -> False
adamc@208 1188
adamc@36 1189 ]]
adamc@36 1190
adam@317 1191 The negation is replaced with an implication of falsehood. We use the tactic %\index{tactics!intro}%[intro H] to change the assumption of the implication into a hypothesis named [H]. *)
adamc@36 1192
adamc@36 1193 intro H.
adamc@36 1194 (** [[
adamc@36 1195 H : true = false
adamc@36 1196 ============================
adamc@36 1197 False
adamc@208 1198
adamc@36 1199 ]]
adamc@36 1200
adamc@36 1201 This is the point in the proof where we apply some creativity. We define a function whose utility will become clear soon. *)
adamc@36 1202
adam@317 1203 Definition toProp (b : bool) := if b then True else False.
adamc@36 1204
adam@317 1205 (** 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 [toProp] such that our conclusion of [False] is computationally equivalent to [toProp false]. Thus, the %\index{tactics!change}\coqdockw{%#<tt>#change#</tt>#%}% tactic will let us change the conclusion to [toProp false]. The general form %\coqdockw{%#<tt>#change#</tt>#%}% [e] replaces the conclusion with [e], whenever Coq's built-in computation rules suffice to establish the equivalence of [e] with the original conclusion. *)
adamc@36 1206
adam@317 1207 (* begin hide *)
adam@317 1208 change (toProp false).
adam@317 1209 (* end hide *)
adam@317 1210 (** %\hspace{-.075in}\coqdockw{%#<tt>#change#</tt>#%}% [(][toProp false).] *)
adamc@36 1211 (** [[
adamc@36 1212 H : true = false
adamc@36 1213 ============================
adam@317 1214 toProp false
adamc@208 1215
adamc@36 1216 ]]
adamc@36 1217
adam@317 1218 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.%\index{tactics!rewrite}% *)
adamc@36 1219
adamc@36 1220 rewrite <- H.
adamc@36 1221 (** [[
adamc@36 1222 H : true = false
adamc@36 1223 ============================
adam@317 1224 toProp true
adamc@208 1225
adamc@36 1226 ]]
adamc@36 1227
adamc@36 1228 We are almost done. Just how close we are to done is revealed by computational simplification. *)
adamc@36 1229
adamc@36 1230 simpl.
adamc@36 1231 (** [[
adamc@36 1232 H : true = false
adamc@36 1233 ============================
adamc@36 1234 True
adamc@208 1235
adam@302 1236 ]]
adam@302 1237 *)
adamc@36 1238
adamc@36 1239 trivial.
adamc@36 1240 Qed.
adamc@41 1241 (* end thide *)
adamc@36 1242
adamc@36 1243 (** I have no trivial automated version of this proof to suggest, beyond using [discriminate] or [congruence] in the first place.
adamc@36 1244
adamc@36 1245 %\medskip%
adamc@36 1246
adamc@36 1247 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 1248
adamc@36 1249 Theorem S_inj' : forall n m : nat, S n = S m -> n = m.
adamc@41 1250 (* begin thide *)
adamc@36 1251 intros n m H.
adam@317 1252 (* begin hide *)
adamc@36 1253 change (pred (S n) = pred (S m)).
adam@317 1254 (* end hide *)
adam@317 1255 (** %\hspace{-.075in}\coqdockw{%#<tt>#change#</tt>#%}% [(][pred (][S n) = pred (][S m)).] *)
adam@317 1256
adamc@36 1257 rewrite H.
adamc@36 1258 reflexivity.
adamc@36 1259 Qed.
adamc@41 1260 (* end thide *)
adamc@36 1261
adam@317 1262 (** The key piece of creativity in this theorem comes in the use of the natural number predecessor function [pred]. Embodied in the implementation of %\coqdockw{%#<tt>#injectivity#</tt>#%}% is a generic recipe for writing such type-specific functions.
adam@317 1263
adam@317 1264 The examples in this section illustrate an important aspect of the design philosophy behind Coq. We could certainly design a Gallina replacement that built in rules for constructor discrimination and injectivity, but a simpler alternative is to include a few carefully chosen rules that enable the desired reasoning patterns and many others. A key benefit of this philosophy is that the complexity of proof checking is minimized, which bolsters our confidence that proved theorems are really true. *)