annotate src/Predicates.v @ 52:ce0c9d481ee3

What could go wrong; some exercises
author Adam Chlipala <adamc@hcoop.net>
date Sun, 28 Sep 2008 12:42:59 -0400
parents a21eb02cc7c6
children 422865e240b1
rev   line source
adamc@45 1 (* Copyright (c) 2008, Adam Chlipala
adamc@45 2 *
adamc@45 3 * This work is licensed under a
adamc@45 4 * Creative Commons Attribution-Noncommercial-No Derivative Works 3.0
adamc@45 5 * Unported License.
adamc@45 6 * The license text is available at:
adamc@45 7 * http://creativecommons.org/licenses/by-nc-nd/3.0/
adamc@45 8 *)
adamc@45 9
adamc@45 10 (* begin hide *)
adamc@45 11 Require Import List.
adamc@45 12
adamc@45 13 Require Import Tactics.
adamc@45 14
adamc@45 15 Set Implicit Arguments.
adamc@45 16 (* end hide *)
adamc@45 17
adamc@45 18
adamc@45 19 (** %\chapter{Inductive Predicates}% *)
adamc@45 20
adamc@45 21 (** The so-called "Curry-Howard Correspondence" states a formal connection between functional programs and mathematical proofs. In the last chapter, we snuck in a first introduction to this subject in Coq. Witness the close similarity between the types [unit] and [True] from the standard library: *)
adamc@45 22
adamc@45 23 Print unit.
adamc@45 24 (** [[
adamc@45 25
adamc@45 26 Inductive unit : Set := tt : unit
adamc@45 27 ]] *)
adamc@45 28
adamc@45 29 Print True.
adamc@45 30 (** [[
adamc@45 31
adamc@45 32 Inductive True : Prop := I : True
adamc@45 33 ]] *)
adamc@45 34
adamc@45 35 (** Recall that [unit] is the type with only one value, and [True] is the proposition that always holds. Despite this superficial difference between the two concepts, in both cases we can use the same inductive definition mechanism. The connection goes further than this. We see that we arrive at the definition of [True] by replacing [unit] by [True], [tt] by [I], and [Set] by [Prop]. The first two of these differences are superficial changes of names, while the third difference is the crucial one for separating programs from proofs. A term [T] of type [Set] is a type of programs, and a term of type [T] is a program. A term [T] of type [Prop] is a logical proposition, and its proofs are of type [T].
adamc@45 36
adamc@45 37 [unit] has one value, [tt]. [True] has one proof, [I]. Why distinguish between these two types? Many people who have read about Curry-Howard in an abstract context and not put it to use in proof engineering answer that the two types in fact %\textit{%#<i>#should not#</i>#%}% be distinguished. There is a certain aesthetic appeal to this point of view, but I want to argue that it is best to treat Curry-Howard very loosely in practical proving. There are Coq-specific reasons for preferring the distinction, involving efficient compilation and avoidance of paradoxes in the presence of classical math, but I will argue that there is a more general principle that should lead us to avoid conflating programming and proving.
adamc@45 38
adamc@45 39 The essence of the argument is roughly this: to an engineer, not all functions of type [A -> B] are created equal, but all proofs of a proposition [P -> Q] are. This idea is known as %\textit{%#<i>#proof irrelevance#</i>#%}%, and its formalizations in logics prevent us from distinguishing between alternate proofs of the same proposition. Proof irrelevance is compatible with, but not derivable in, Gallina. Apart from this theoretical concern, I will argue that it is most effective to do engineering with Coq by employing different techniques for programs versus proofs. Most of this book is organized around that distinction, describing how to program, by applying standard functional programming techniques in the presence of dependent types; and how to prove, by writing custom Ltac decision procedures.
adamc@45 40
adamc@45 41 With that perspective in mind, this chapter is sort of a mirror image of the last chapter, introducing how to define predicates with inductive definitions. We will point out similarities in places, but much of the effective Coq user's bag of tricks is disjoint for predicates versus "datatypes." This chapter is also a covert introduction to dependent types, which are the foundation on which interesting inductive predicates are built, though we will rely on tactics to build dependently-typed proof terms for us for now. A future chapter introduces more manual application of dependent types. *)
adamc@45 42
adamc@45 43
adamc@48 44 (** * Propositional Logic *)
adamc@45 45
adamc@45 46 (** Let us begin with a brief tour through the definitions of the connectives for propositional logic. We will work within a Coq section that provides us with a set of propositional variables. In Coq parlance, these are just terms of type [Prop.] *)
adamc@45 47
adamc@45 48 Section Propositional.
adamc@46 49 Variables P Q R : Prop.
adamc@45 50
adamc@45 51 (** In Coq, the most basic propositional connective is implication, written [->], which we have already used in almost every proof. Rather than being defined inductively, implication is built into Coq as the function type constructor.
adamc@45 52
adamc@45 53 We have also already seen the definition of [True]. For a demonstration of a lower-level way of establishing proofs of inductive predicates, we turn to this trivial theorem. *)
adamc@45 54
adamc@45 55 Theorem obvious : True.
adamc@45 56 apply I.
adamc@45 57 Qed.
adamc@45 58
adamc@45 59 (** We may always use the [apply] tactic to take a proof step based on applying a particular constructor of the inductive predicate that we are trying to establish. Sometimes there is only one constructor that could possibly apply, in which case a shortcut is available: *)
adamc@45 60
adamc@45 61 Theorem obvious' : True.
adamc@45 62 constructor.
adamc@45 63 Qed.
adamc@45 64
adamc@45 65 (** There is also a predicate [False], which is the Curry-Howard mirror image of [Empty_set] from the last chapter. *)
adamc@45 66
adamc@45 67 Print False.
adamc@45 68 (** [[
adamc@45 69
adamc@45 70 Inductive False : Prop :=
adamc@45 71 ]] *)
adamc@45 72
adamc@45 73 (** We can conclude anything from [False], doing case analysis on a proof of [False] in the same way we might do case analysis on, say, a natural number. Since there are no cases to consider, any such case analysis succeeds immediately in proving the goal. *)
adamc@45 74
adamc@45 75 Theorem False_imp : False -> 2 + 2 = 5.
adamc@45 76 destruct 1.
adamc@45 77 Qed.
adamc@45 78
adamc@45 79 (** In a consistent context, we can never build a proof of [False]. In inconsistent contexts that appear in the courses of proofs, it is usually easiest to proceed by demonstrating that inconsistency with an explicit proof of [False]. *)
adamc@45 80
adamc@45 81 Theorem arith_neq : 2 + 2 = 5 -> 9 + 9 = 835.
adamc@45 82 intro.
adamc@45 83
adamc@45 84 (** At this point, we have an inconsistent hypothesis [2 + 2 = 5], so the specific conclusion is not important. We use the [elimtype] tactic to state a proposition, telling Coq that we wish to construct a proof of the new proposition and then prove the original goal by case analysis on the structure of the new auxiliary proof. Since [False] has no constructors, [elimtype False] simply leaves us with the obligation to prove [False]. *)
adamc@45 85
adamc@45 86 elimtype False.
adamc@45 87 (** [[
adamc@45 88
adamc@45 89 H : 2 + 2 = 5
adamc@45 90 ============================
adamc@45 91 False
adamc@45 92 ]] *)
adamc@45 93
adamc@45 94 (** For now, we will leave the details of this proof about arithmetic to [crush]. *)
adamc@45 95
adamc@45 96 crush.
adamc@45 97 Qed.
adamc@45 98
adamc@45 99 (** A related notion to [False] is logical negation. *)
adamc@45 100
adamc@45 101 Print not.
adamc@45 102 (** [[
adamc@45 103
adamc@45 104 not = fun A : Prop => A -> False
adamc@45 105 : Prop -> Prop
adamc@45 106 ]] *)
adamc@45 107
adamc@45 108 (** We see that [not] is just shorthand for implication of [False]. We can use that fact explicitly in proofs. The syntax [~P] expands to [not P]. *)
adamc@45 109
adamc@45 110 Theorem arith_neq' : ~ (2 + 2 = 5).
adamc@45 111 unfold not.
adamc@45 112
adamc@45 113 (** [[
adamc@45 114
adamc@45 115 ============================
adamc@45 116 2 + 2 = 5 -> False
adamc@45 117 ]] *)
adamc@45 118
adamc@45 119 crush.
adamc@45 120 Qed.
adamc@45 121
adamc@45 122 (** We also have conjunction, which we introduced in the last chapter. *)
adamc@45 123
adamc@45 124 Print and.
adamc@45 125 (** [[
adamc@45 126
adamc@45 127 Inductive and (A : Prop) (B : Prop) : Prop := conj : A -> B -> A /\ B
adamc@45 128 ]] *)
adamc@45 129
adamc@45 130 (** The interested reader can check that [and] has a Curry-Howard doppleganger called [prod], the type of pairs. However, it is generally most convenient to reason about conjunction using tactics. An explicit proof of commutativity of [and] illustrates the usual suspects for such tasks. [/\] is an infix shorthand for [and]. *)
adamc@45 131
adamc@45 132 Theorem and_comm : P /\ Q -> Q /\ P.
adamc@45 133 (** We start by case analysis on the proof of [P /\ Q]. *)
adamc@45 134
adamc@45 135 destruct 1.
adamc@45 136 (** [[
adamc@45 137
adamc@45 138 H : P
adamc@45 139 H0 : Q
adamc@45 140 ============================
adamc@45 141 Q /\ P
adamc@45 142 ]] *)
adamc@45 143
adamc@45 144 (** Every proof of a conjunction provides proofs for both conjuncts, so we get a single subgoal reflecting that. We can proceed by splitting this subgoal into a case for each conjunct of [Q /\ P]. *)
adamc@45 145
adamc@45 146 split.
adamc@45 147 (** [[
adamc@45 148 2 subgoals
adamc@45 149
adamc@45 150 H : P
adamc@45 151 H0 : Q
adamc@45 152 ============================
adamc@45 153 Q
adamc@45 154
adamc@45 155 subgoal 2 is:
adamc@45 156 P
adamc@45 157 ]] *)
adamc@45 158
adamc@45 159 (** In each case, the conclusion is among our hypotheses, so the [assumption] tactic finishes the process. *)
adamc@45 160
adamc@45 161 assumption.
adamc@45 162 assumption.
adamc@45 163 Qed.
adamc@45 164
adamc@45 165 (** Coq disjunction is called [or] and abbreviated with the infix operator [\/]. *)
adamc@45 166
adamc@45 167 Print or.
adamc@45 168 (** [[
adamc@45 169
adamc@45 170 Inductive or (A : Prop) (B : Prop) : Prop :=
adamc@45 171 or_introl : A -> A \/ B | or_intror : B -> A \/ B
adamc@45 172 ]] *)
adamc@45 173
adamc@45 174 (** We see that there are two ways to prove a disjunction: prove the first disjunct or prove the second. The Curry-Howard analogue of this is the Coq [sum] type. We can demonstrate the main tactics here with another proof of commutativity. *)
adamc@45 175
adamc@45 176 Theorem or_comm : P \/ Q -> Q \/ P.
adamc@45 177 (** As in the proof for [and], we begin with case analysis, though this time we are met by two cases instead of one. *)
adamc@45 178 destruct 1.
adamc@45 179 (** [[
adamc@45 180
adamc@45 181 2 subgoals
adamc@45 182
adamc@45 183 H : P
adamc@45 184 ============================
adamc@45 185 Q \/ P
adamc@45 186
adamc@45 187 subgoal 2 is:
adamc@45 188 Q \/ P
adamc@45 189 ]] *)
adamc@45 190
adamc@45 191 (** We can see that, in the first subgoal, we want to prove the disjunction by proving its second disjunct. The [right] tactic telegraphs this intent. *)
adamc@45 192
adamc@45 193 right; assumption.
adamc@45 194
adamc@45 195 (** The second subgoal has a symmetric proof.
adamc@45 196
adamc@45 197 [[
adamc@45 198
adamc@45 199 1 subgoal
adamc@45 200
adamc@45 201 H : Q
adamc@45 202 ============================
adamc@45 203 Q \/ P
adamc@45 204 ]] *)
adamc@45 205
adamc@45 206 left; assumption.
adamc@45 207 Qed.
adamc@45 208
adamc@46 209
adamc@46 210 (* begin hide *)
adamc@46 211 (* In-class exercises *)
adamc@46 212
adamc@46 213 Theorem contra : P -> ~P -> R.
adamc@52 214 (* begin thide *)
adamc@52 215 unfold not.
adamc@52 216 intros.
adamc@52 217 elimtype False.
adamc@52 218 apply H0.
adamc@52 219 assumption.
adamc@52 220 (* end thide *)
adamc@46 221 Admitted.
adamc@46 222
adamc@46 223 Theorem and_assoc : (P /\ Q) /\ R -> P /\ (Q /\ R).
adamc@52 224 (* begin thide *)
adamc@52 225 intros.
adamc@52 226 destruct H.
adamc@52 227 destruct H.
adamc@52 228 split.
adamc@52 229 assumption.
adamc@52 230 split.
adamc@52 231 assumption.
adamc@52 232 assumption.
adamc@52 233 (* end thide *)
adamc@46 234 Admitted.
adamc@46 235
adamc@46 236 Theorem or_assoc : (P \/ Q) \/ R -> P \/ (Q \/ R).
adamc@52 237 (* begin thide *)
adamc@52 238 intros.
adamc@52 239 destruct H.
adamc@52 240 destruct H.
adamc@52 241 left.
adamc@52 242 assumption.
adamc@52 243 right.
adamc@52 244 left.
adamc@52 245 assumption.
adamc@52 246 right.
adamc@52 247 right.
adamc@52 248 assumption.
adamc@52 249 (* end thide *)
adamc@46 250 Admitted.
adamc@46 251
adamc@46 252 (* end hide *)
adamc@46 253
adamc@46 254
adamc@46 255 (** It would be a shame to have to plod manually through all proofs about propositional logic. Luckily, there is no need. One of the most basic Coq automation tactics is [tauto], which is a complete decision procedure for constructive propositional logic. (More on what "constructive" means in the next section.) We can use [tauto] to dispatch all of the purely propositional theorems we have proved so far. *)
adamc@46 256
adamc@46 257 Theorem or_comm' : P \/ Q -> Q \/ P.
adamc@46 258 tauto.
adamc@46 259 Qed.
adamc@46 260
adamc@46 261 (** Sometimes propositional reasoning forms important plumbing for the proof of a theorem, but we still need to apply some other smarts about, say, arithmetic. [intuition] is a generalization of [tauto] that proves everything it can using propositional reasoning. When some goals remain, it uses propositional laws to simplify them as far as possible. Consider this example, which uses the list concatenation operator [++] from the standard library. *)
adamc@46 262
adamc@46 263 Theorem arith_comm : forall ls1 ls2 : list nat,
adamc@46 264 length ls1 = length ls2 \/ length ls1 + length ls2 = 6
adamc@46 265 -> length (ls1 ++ ls2) = 6 \/ length ls1 = length ls2.
adamc@46 266 intuition.
adamc@46 267
adamc@46 268 (** A lot of the proof structure has been generated for us by [intuition], but the final proof depends on a fact about lists. The remaining subgoal hints at what cleverness we need to inject. *)
adamc@46 269
adamc@46 270 (** [[
adamc@46 271
adamc@46 272 ls1 : list nat
adamc@46 273 ls2 : list nat
adamc@46 274 H0 : length ls1 + length ls2 = 6
adamc@46 275 ============================
adamc@46 276 length (ls1 ++ ls2) = 6 \/ length ls1 = length ls2
adamc@46 277 ]] *)
adamc@46 278
adamc@46 279 (** We can see that we need a theorem about lengths of concatenated lists, which we proved last chapter and is also in the standard library. *)
adamc@46 280
adamc@46 281 rewrite app_length.
adamc@46 282 (** [[
adamc@46 283
adamc@46 284 ls1 : list nat
adamc@46 285 ls2 : list nat
adamc@46 286 H0 : length ls1 + length ls2 = 6
adamc@46 287 ============================
adamc@46 288 length ls1 + length ls2 = 6 \/ length ls1 = length ls2
adamc@46 289 ]] *)
adamc@46 290
adamc@46 291 (** Now the subgoal follows by purely propositional reasoning. That is, we could replace [length ls1 + length ls2 = 6] with [P] and [length ls1 = length ls2] with [Q] and arrive at a tautology of propositional logic. *)
adamc@46 292
adamc@46 293 tauto.
adamc@46 294 Qed.
adamc@46 295
adamc@46 296 (** [intuition] is one of the main bits of glue in the implementation of [crush], so, with a little help, we can get a short automated proof of the theorem. *)
adamc@46 297
adamc@46 298 Theorem arith_comm' : forall ls1 ls2 : list nat,
adamc@46 299 length ls1 = length ls2 \/ length ls1 + length ls2 = 6
adamc@46 300 -> length (ls1 ++ ls2) = 6 \/ length ls1 = length ls2.
adamc@46 301 Hint Rewrite app_length : cpdt.
adamc@46 302
adamc@46 303 crush.
adamc@46 304 Qed.
adamc@46 305
adamc@45 306 End Propositional.
adamc@45 307
adamc@46 308
adamc@47 309 (** * What Does It Mean to Be Constructive? *)
adamc@46 310
adamc@47 311 (** One potential point of confusion in the presentation so far is the distinction between [bool] and [Prop]. [bool] is a datatype whose two values are [true] and [false], while [Prop] is a more primitive type that includes among its members [True] and [False]. Why not collapse these two concepts into one, and why must there be more than two states of mathematical truth?
adamc@46 312
adamc@47 313 The answer comes from the fact that Coq implements %\textit{%#<i>#constructive#</i>#%}% or %\textit{%#<i>#intuitionistic#</i>#%}% logic, in contrast to the %\textit{%#<i>#classical#</i>#%}% logic that you may be more familiar with. In constructive logic, classical tautologies like [~ ~P -> P] and [P \/ ~P] do not always hold. In general, we can only prove these tautologies when [P] is %\textit{%#<i>#decidable#</i>#%}%, in the sense of computability theory. The Curry-Howard encoding that Coq uses for [or] allows us to extract either a proof of [P] or a proof of [~P] from any proof of [P \/ ~P]. Since our proofs are just functional programs which we can run, this would give us a decision procedure for the halting problem, where the instantiations of [P] would be formulas like "this particular Turing machine halts."
adamc@47 314
adamc@47 315 Hence the distinction between [bool] and [Prop]. Programs of type [bool] are computational by construction; we can always run them to determine their results. Many [Prop]s are undecidable, and so we can write more expressive formulas with [Prop]s than with [bool]s, but the inevitable consequence is that we cannot simply "run a [Prop] to determine its truth."
adamc@47 316
adamc@47 317 Constructive logic lets us define all of the logical connectives in an aesthetically-appealing way, with orthogonal inductive definitions. That is, each connective is defined independently using a simple, shared mechanism. Constructivity also enables a trick called %\textit{%#<i>#program extraction#</i>#%}%, where we write programs by phrasing them as theorems to be proved. Since our proofs are just functional programs, we can extract executable programs from our final proofs, which we could not do as naturally with classical proofs.
adamc@47 318
adamc@47 319 We will see more about Coq's program extraction facility in a later chapter. However, I think it is worth interjecting another warning at this point, following up on the prior warning about taking the Curry-Howard correspondence too literally. It is possible to write programs by theorem-proving methods in Coq, but hardly anyone does it. It is almost always most useful to maintain the distinction between programs and proofs. If you write a program by proving a theorem, you are likely to run into algorithmic inefficiencies that you introduced in your proof to make it easier to prove. It is a shame to have to worry about such situations while proving tricky theorems, and it is a happy state of affairs that you almost certainly will not need to, with the ideal of extracting programs from proofs being confined mostly to theoretical studies. *)
adamc@48 320
adamc@48 321
adamc@48 322 (** * First-Order Logic *)
adamc@48 323
adamc@48 324 (** The [forall] connective of first-order logic, which we have seen in many examples so far, is built into Coq. Getting ahead of ourselves a bit, we can see it as the dependent function type constructor. In fact, implication and universal quantification are just different syntactic shorthands for the same Coq mechanism. A formula [P -> Q] is equivalent to [forall x : P, Q], where [x] does not appear in [Q]. That is, the "real" type of the implication says "for every proof of [P], there exists a proof of [Q]."
adamc@48 325
adamc@48 326 Existential quantification is defined in the standard library. *)
adamc@48 327
adamc@48 328 Print ex.
adamc@48 329 (** [[
adamc@48 330
adamc@48 331 Inductive ex (A : Type) (P : A -> Prop) : Prop :=
adamc@48 332 ex_intro : forall x : A, P x -> ex P
adamc@48 333 ]] *)
adamc@48 334
adamc@48 335 (** [ex] is parameterized by the type [A] that we quantify over, and by a predicate [P] over [A]s. We prove an existential by exhibiting some [x] of type [A], along with a proof of [P x]. As usual, there are tactics that save us from worrying about the low-level details most of the time. *)
adamc@48 336
adamc@48 337 Theorem exist1 : exists x : nat, x + 1 = 2.
adamc@48 338 (** remove printing exists*)
adamc@48 339 (** We can start this proof with a tactic [exists], which should not be confused with the formula constructor shorthand of the same name. (In the PDF version of this document, the reverse 'E' appears instead of the text "exists.") *)
adamc@48 340 exists 1.
adamc@48 341
adamc@48 342 (** The conclusion is replaced with a version using the existential witness that we announced. *)
adamc@48 343
adamc@48 344 (** [[
adamc@48 345
adamc@48 346 ============================
adamc@48 347 1 + 1 = 2
adamc@48 348 ]] *)
adamc@48 349
adamc@48 350 reflexivity.
adamc@48 351 Qed.
adamc@48 352
adamc@48 353 (** printing exists $\exists$ *)
adamc@48 354
adamc@48 355 (** We can also use tactics to reason about existential hypotheses. *)
adamc@48 356
adamc@48 357 Theorem exist2 : forall n m : nat, (exists x : nat, n + x = m) -> n <= m.
adamc@48 358 (** We start by case analysis on the proof of the existential fact. *)
adamc@48 359 destruct 1.
adamc@48 360 (** [[
adamc@48 361
adamc@48 362 n : nat
adamc@48 363 m : nat
adamc@48 364 x : nat
adamc@48 365 H : n + x = m
adamc@48 366 ============================
adamc@48 367 n <= m
adamc@48 368 ]] *)
adamc@48 369
adamc@48 370 (** The goal has been replaced by a form where there is a new free variable [x], and where we have a new hypothesis that the body of the existential holds with [x] substituted for the old bound variable. From here, the proof is just about arithmetic and is easy to automate. *)
adamc@48 371
adamc@48 372 crush.
adamc@48 373 Qed.
adamc@48 374
adamc@48 375
adamc@48 376 (* begin hide *)
adamc@48 377 (* In-class exercises *)
adamc@48 378
adamc@48 379 Theorem forall_exists_commute : forall (A B : Type) (P : A -> B -> Prop),
adamc@48 380 (exists x : A, forall y : B, P x y) -> (forall y : B, exists x : A, P x y).
adamc@52 381 (* begin thide *)
adamc@52 382 intros.
adamc@52 383 destruct H.
adamc@52 384 exists x.
adamc@52 385 apply H.
adamc@52 386 (* end thide *)
adamc@48 387 Admitted.
adamc@48 388
adamc@48 389 (* end hide *)
adamc@48 390
adamc@48 391
adamc@48 392 (** The tactic [intuition] has a first-order cousin called [firstorder]. [firstorder] proves many formulas when only first-order reasoning is needed, and it tries to perform first-order simplifications in any case. First-order reasoning is much harder than propositional reasoning, so [firstorder] is much more likely than [intuition] to get stuck in a way that makes it run for long enough to be useless. *)
adamc@49 393
adamc@49 394
adamc@49 395 (** * Predicates with Implicit Equality *)
adamc@49 396
adamc@49 397 (** We start our exploration of a more complicated class of predicates with a simple example: an alternative way of characterizing when a natural number is zero. *)
adamc@49 398
adamc@49 399 Inductive isZero : nat -> Prop :=
adamc@49 400 | IsZero : isZero 0.
adamc@49 401
adamc@49 402 Theorem isZero_zero : isZero 0.
adamc@49 403 constructor.
adamc@49 404 Qed.
adamc@49 405
adamc@49 406 (** We can call [isZero] a %\textit{%#<i>#judgment#</i>#%}%, in the sense often used in the semantics of programming languages. Judgments are typically defined in the style of %\textit{%#<i>#natural deduction#</i>#%}%, where we write a number of %\textit{%#<i>#inference rules#</i>#%}% with premises appearing above a solid line and a conclusion appearing below the line. In this example, the sole constructor [IsZero] of [isZero] can be thought of as the single inference rule for deducing [isZero], with nothing above the line and [isZero 0] below it. The proof of [isZero_zero] demonstrates how we can apply an inference rule.
adamc@49 407
adamc@49 408 The definition of [isZero] differs in an important way from all of the other inductive definitions that we have seen in this and the previous chapter. Instead of writing just [Set] or [Prop] after the colon, here we write [nat -> Prop]. We saw examples of parameterized types like [list], but there the parameters appeared with names %\textit{%#<i>#before#</i>#%}% the colon. Every constructor of a parameterized inductive type must have a range type that uses the same parameter, whereas the form we use here enables us to use different arguments to the type for different constructors.
adamc@49 409
adamc@49 410 For instance, [isZero] forces its argument to be [0]. We can see that the concept of equality is somehow implicit in the inductive definition mechanism. The way this is accomplished is similar to the way that logic variables are used in Prolog, and it is a very powerful mechanism that forms a foundation for formalizing all of mathematics. In fact, though it is natural to think of inductive types as folding in the functionality of equality, in Coq, the true situation is reversed, with equality defined as just another inductive type! *)
adamc@49 411
adamc@49 412 Print eq.
adamc@49 413 (** [[
adamc@49 414
adamc@49 415 Inductive eq (A : Type) (x : A) : A -> Prop := refl_equal : x = x
adamc@49 416 ]] *)
adamc@49 417
adamc@49 418 (** [eq] is the type we get behind the scenes when uses of infix [=] are expanded. We see that [eq] has both a parameter [x] that is fixed and an extra unnamed argument of the same type. It is crucial that the second argument is untied to the parameter in the type of [eq]; otherwise, we would have to prove that two values are equal even to be able to state the possibility of equality, which would more or less defeat the purpose of having an equality proposition. However, examining the type of equality's sole constructor [refl_equal], we see that we can only %\textit{%#<i>#prove#</i>#%}% equality when its two arguments are syntactically equal. This definition turns out to capture all of the basic properties of equality, and the equality-manipulating tactics that we have seen so far, like [reflexivity] and [rewrite], are implemented treating [eq] as just another inductive type with a well-chosen definition.
adamc@49 419
adamc@49 420 Returning to the example of [isZero], we can see how to make use of hypotheses that use this predicate. *)
adamc@49 421
adamc@49 422 Theorem isZero_plus : forall n m : nat, isZero m -> n + m = n.
adamc@49 423 (** We want to proceed by cases on the proof of the assumption about [isZero]. *)
adamc@49 424 destruct 1.
adamc@49 425 (** [[
adamc@49 426
adamc@49 427 n : nat
adamc@49 428 ============================
adamc@49 429 n + 0 = n
adamc@49 430 ]] *)
adamc@49 431
adamc@49 432 (** Since [isZero] has only one constructor, we are presented with only one subgoal. The argument [m] to [isZero] is replaced with that type's argument from the single constructor [IsZero]. From this point, the proof is trivial. *)
adamc@49 433
adamc@49 434 crush.
adamc@49 435 Qed.
adamc@49 436
adamc@49 437 (** Another example seems at first like it should admit an analogous proof, but in fact provides a demonstration of one of the most basic gotchas of Coq proving. *)
adamc@49 438
adamc@49 439 Theorem isZero_contra : isZero 1 -> False.
adamc@49 440 (** Let us try a proof by cases on the assumption, as in the last proof. *)
adamc@49 441 destruct 1.
adamc@49 442 (** [[
adamc@49 443
adamc@49 444 ============================
adamc@49 445 False
adamc@49 446 ]] *)
adamc@49 447
adamc@49 448 (** It seems that case analysis has not helped us much at all! Our sole hypothesis disappears, leaving us, if anything, worse off than we were before. What went wrong? We have met an important restriction in tactics like [destruct] and [induction] when applied to types with arguments. If the arguments are not already free variables, they will be replaced by new free variables internally before doing the case analysis or induction. Since the argument [1] to [isZero] is replaced by a fresh variable, we lose the crucial fact that it is not equal to [0].
adamc@49 449
adamc@49 450 Why does Coq use this restriction? We will discuss the issue in detail in a future chapter, when we see the dependently-typed programming techniques that would allow us to write this proof term manually. For now, we just say that the algorithmic problem of "logically complete case analysis" is undecidable when phrased in Coq's logic. A few tactics and design patterns that we will present in this chapter suffice in almost all cases. For the current example, what we want is a tactic called [inversion], which corresponds to the concept of inversion that is frequently used with natural deduction proof systems. *)
adamc@49 451
adamc@49 452 Undo.
adamc@49 453 inversion 1.
adamc@49 454 Qed.
adamc@49 455
adamc@49 456 (** What does [inversion] do? Think of it as a version of [destruct] that does its best to take advantage of the structure of arguments to inductive types. In this case, [inversion] completed the proof immediately, because it was able to detect that we were using [isZero] with an impossible argument.
adamc@49 457
adamc@49 458 Sometimes using [destruct] when you should have used [inversion] can lead to confusing results. To illustrate, consider an alternate proof attempt for the last theorem. *)
adamc@49 459
adamc@49 460 Theorem isZero_contra' : isZero 1 -> 2 + 2 = 5.
adamc@49 461 destruct 1.
adamc@49 462 (** [[
adamc@49 463
adamc@49 464 ============================
adamc@49 465 1 + 1 = 4
adamc@49 466 ]] *)
adamc@49 467
adamc@49 468 (** What on earth happened here? Internally, [destruct] replaced [1] with a fresh variable, and, trying to be helpful, it also replaced the occurrence of [1] within the unary representation of each number in the goal. This has the net effect of decrementing each of these numbers. If you are doing a proof and encounter a strange transmutation like this, there is a good chance that you should go back and replace a use of [destruct] with [inversion]. *)
adamc@49 469 Abort.
adamc@49 470
adamc@49 471
adamc@49 472 (* begin hide *)
adamc@49 473 (* In-class exercises *)
adamc@49 474
adamc@49 475 (* EX: Define an inductive type capturing when a list has exactly two elements. Prove that your predicate does not hold of the empty list, and prove that, whenever it holds of a list, the length of that list is two. *)
adamc@49 476
adamc@52 477 (* begin thide *)
adamc@52 478 Section twoEls.
adamc@52 479 Variable A : Type.
adamc@52 480
adamc@52 481 Inductive twoEls : list A -> Prop :=
adamc@52 482 | TwoEls : forall x y, twoEls (x :: y :: nil).
adamc@52 483
adamc@52 484 Theorem twoEls_nil : twoEls nil -> False.
adamc@52 485 inversion 1.
adamc@52 486 Qed.
adamc@52 487
adamc@52 488 Theorem twoEls_two : forall ls, twoEls ls -> length ls = 2.
adamc@52 489 inversion 1.
adamc@52 490 reflexivity.
adamc@52 491 Qed.
adamc@52 492 End twoEls.
adamc@52 493 (* end thide *)
adamc@52 494
adamc@49 495 (* end hide *)
adamc@49 496
adamc@50 497
adamc@50 498 (** * Recursive Predicates *)
adamc@50 499
adamc@50 500 (** We have already seen all of the ingredients we need to build interesting recursive predicates, like this predicate capturing even-ness. *)
adamc@50 501
adamc@50 502 Inductive even : nat -> Prop :=
adamc@50 503 | EvenO : even O
adamc@50 504 | EvenSS : forall n, even n -> even (S (S n)).
adamc@50 505
adamc@50 506 (** Think of [even] as another judgment defined by natural deduction rules. [EvenO] is a rule with nothing above the line and [even O] below the line, and [EvenSS] is a rule with [even n] above the line and [even (S (S n))] below.
adamc@50 507
adamc@50 508 The proof techniques of the last section are easily adapted. *)
adamc@50 509
adamc@50 510 Theorem even_0 : even 0.
adamc@50 511 constructor.
adamc@50 512 Qed.
adamc@50 513
adamc@50 514 Theorem even_4 : even 4.
adamc@50 515 constructor; constructor; constructor.
adamc@50 516 Qed.
adamc@50 517
adamc@50 518 (** It is not hard to see that sequences of constructor applications like the above can get tedious. We can avoid them using Coq's hint facility. *)
adamc@50 519
adamc@50 520 Hint Constructors even.
adamc@50 521
adamc@50 522 Theorem even_4' : even 4.
adamc@50 523 auto.
adamc@50 524 Qed.
adamc@50 525
adamc@50 526 Theorem even_1_contra : even 1 -> False.
adamc@50 527 inversion 1.
adamc@50 528 Qed.
adamc@50 529
adamc@50 530 Theorem even_3_contra : even 3 -> False.
adamc@50 531 inversion 1.
adamc@50 532 (** [[
adamc@50 533
adamc@50 534 H : even 3
adamc@50 535 n : nat
adamc@50 536 H1 : even 1
adamc@50 537 H0 : n = 1
adamc@50 538 ============================
adamc@50 539 False
adamc@50 540 ]] *)
adamc@50 541
adamc@50 542 (** [inversion] can be a little overzealous at times, as we can see here with the introduction of the unused variable [n] and an equality hypothesis about it. For more complicated predicates, though, adding such assumptions is critical to dealing with the undecidability of general inversion. *)
adamc@50 543
adamc@50 544 inversion H1.
adamc@50 545 Qed.
adamc@50 546
adamc@50 547 (** We can also do inductive proofs about [even]. *)
adamc@50 548
adamc@50 549 Theorem even_plus : forall n m, even n -> even m -> even (n + m).
adamc@50 550 (** It seems a reasonable first choice to proceed by induction on [n]. *)
adamc@50 551 induction n; crush.
adamc@50 552 (** [[
adamc@50 553
adamc@50 554 n : nat
adamc@50 555 IHn : forall m : nat, even n -> even m -> even (n + m)
adamc@50 556 m : nat
adamc@50 557 H : even (S n)
adamc@50 558 H0 : even m
adamc@50 559 ============================
adamc@50 560 even (S (n + m))
adamc@50 561 ]] *)
adamc@50 562
adamc@50 563 (** We will need to use the hypotheses [H] and [H0] somehow. The most natural choice is to invert [H]. *)
adamc@50 564
adamc@50 565 inversion H.
adamc@50 566 (** [[
adamc@50 567
adamc@50 568 n : nat
adamc@50 569 IHn : forall m : nat, even n -> even m -> even (n + m)
adamc@50 570 m : nat
adamc@50 571 H : even (S n)
adamc@50 572 H0 : even m
adamc@50 573 n0 : nat
adamc@50 574 H2 : even n0
adamc@50 575 H1 : S n0 = n
adamc@50 576 ============================
adamc@50 577 even (S (S n0 + m))
adamc@50 578 ]] *)
adamc@50 579
adamc@50 580 (** Simplifying the conclusion brings us to a point where we can apply a constructor. *)
adamc@50 581 simpl.
adamc@50 582 (** [[
adamc@50 583
adamc@50 584 ============================
adamc@50 585 even (S (S (n0 + m)))
adamc@50 586 ]] *)
adamc@50 587
adamc@50 588 constructor.
adamc@50 589 (** [[
adamc@50 590
adamc@50 591 ============================
adamc@50 592 even (n0 + m)
adamc@50 593 ]] *)
adamc@50 594
adamc@50 595 (** At this point, we would like to apply the inductive hypothesis, which is: *)
adamc@50 596 (** [[
adamc@50 597
adamc@50 598 IHn : forall m : nat, even n -> even m -> even (n + m)
adamc@50 599 ]] *)
adamc@50 600
adamc@50 601 (** Unfortunately, the goal mentions [n0] where it would need to mention [n] to match [IHn]. We could keep looking for a way to finish this proof from here, but it turns out that we can make our lives much easier by changing our basic strategy. Instead of inducting on the structure of [n], we should induct %\textit{%#<i>#on the structure of one of the [even] proofs#</i>#%}%. This technique is commonly called %\textit{%#<i>#rule induction#</i>#%}% in programming language semantics. In the setting of Coq, we have already seen how predicates are defined using the same inductive type mechanism as datatypes, so the fundamental unity of rule induction with "normal" induction is apparent. *)
adamc@50 602
adamc@50 603 Restart.
adamc@50 604
adamc@50 605 induction 1.
adamc@50 606 (** [[
adamc@50 607
adamc@50 608 m : nat
adamc@50 609 ============================
adamc@50 610 even m -> even (0 + m)
adamc@50 611
adamc@50 612 subgoal 2 is:
adamc@50 613 even m -> even (S (S n) + m)
adamc@50 614 ]] *)
adamc@50 615
adamc@50 616 (** The first case is easily discharged by [crush], based on the hint we added earlier to try the constructors of [even]. *)
adamc@50 617
adamc@50 618 crush.
adamc@50 619
adamc@50 620 (** Now we focus on the second case: *)
adamc@50 621 intro.
adamc@50 622
adamc@50 623 (** [[
adamc@50 624
adamc@50 625 m : nat
adamc@50 626 n : nat
adamc@50 627 H : even n
adamc@50 628 IHeven : even m -> even (n + m)
adamc@50 629 H0 : even m
adamc@50 630 ============================
adamc@50 631 even (S (S n) + m)
adamc@50 632 ]] *)
adamc@50 633
adamc@50 634 (** We simplify and apply a constructor, as in our last proof attempt. *)
adamc@50 635
adamc@50 636 simpl; constructor.
adamc@50 637 (** [[
adamc@50 638
adamc@50 639 ============================
adamc@50 640 even (n + m)
adamc@50 641 ]] *)
adamc@50 642
adamc@50 643 (** Now we have an exact match with our inductive hypothesis, and the remainder of the proof is trivial. *)
adamc@50 644
adamc@50 645 apply IHeven; assumption.
adamc@50 646
adamc@50 647 (** In fact, [crush] can handle all of the details of the proof once we declare the induction strategy. *)
adamc@50 648
adamc@50 649 Restart.
adamc@50 650 induction 1; crush.
adamc@50 651 Qed.
adamc@50 652
adamc@50 653 (** Induction on recursive predicates has similar pitfalls to those we encountered with inversion in the last section. *)
adamc@50 654
adamc@50 655 Theorem even_contra : forall n, even (S (n + n)) -> False.
adamc@50 656 induction 1.
adamc@50 657 (** [[
adamc@50 658
adamc@50 659 n : nat
adamc@50 660 ============================
adamc@50 661 False
adamc@50 662
adamc@50 663 subgoal 2 is:
adamc@50 664 False
adamc@50 665 ]] *)
adamc@50 666
adamc@50 667 (** We are already sunk trying to prove the first subgoal, since the argument to [even] was replaced by a fresh variable internally. This time, we find it easiest to prove this theorem by way of a lemma. Instead of trusting [induction] to replace expressions with fresh variables, we do it ourselves, explicitly adding the appropriate equalities as new assumptions. *)
adamc@50 668 Abort.
adamc@50 669
adamc@50 670 Lemma even_contra' : forall n', even n' -> forall n, n' = S (n + n) -> False.
adamc@50 671 induction 1; crush.
adamc@50 672
adamc@50 673 (** At this point, it is useful to consider all cases of [n] and [n0] being zero or nonzero. Only one of these cases has any trickyness to it. *)
adamc@50 674 destruct n; destruct n0; crush.
adamc@50 675
adamc@50 676 (** [[
adamc@50 677
adamc@50 678 n : nat
adamc@50 679 H : even (S n)
adamc@50 680 IHeven : forall n0 : nat, S n = S (n0 + n0) -> False
adamc@50 681 n0 : nat
adamc@50 682 H0 : S n = n0 + S n0
adamc@50 683 ============================
adamc@50 684 False
adamc@50 685 ]] *)
adamc@50 686
adamc@50 687 (** At this point it is useful to use a theorem from the standard library, which we also proved with a different name in the last chapter. *)
adamc@50 688 Check plus_n_Sm.
adamc@50 689 (** [[
adamc@50 690
adamc@50 691 plus_n_Sm
adamc@50 692 : forall n m : nat, S (n + m) = n + S m
adamc@50 693 ]] *)
adamc@50 694
adamc@50 695 rewrite <- plus_n_Sm in H0.
adamc@50 696
adamc@50 697 (** The induction hypothesis lets us complete the proof. *)
adamc@50 698 apply IHeven with n0; assumption.
adamc@50 699
adamc@50 700 (** As usual, we can rewrite the proof to avoid referencing any locally-generated names, which makes our proof script more robust to changes in the theorem statement. *)
adamc@50 701 Restart.
adamc@50 702 Hint Rewrite <- plus_n_Sm : cpdt.
adamc@50 703
adamc@50 704 induction 1; crush;
adamc@50 705 match goal with
adamc@50 706 | [ H : S ?N = ?N0 + ?N0 |- _ ] => destruct N; destruct N0
adamc@50 707 end; crush; eauto.
adamc@50 708 Qed.
adamc@50 709
adamc@50 710 (** We write the proof in a way that avoids the use of local variable or hypothesis names, using the [match] tactic form to do pattern-matching on the goal. We use unification variables prefixed by question marks in the pattern, and we take advantage of the possibility to mention a unification variable twice in one pattern, to enforce equality between occurrences. The hint to rewrite with [plus_n_Sm] in a particular direction saves us from having to figure out the right place to apply that theorem, and we also take critical advantage of a new tactic, [eauto].
adamc@50 711
adamc@50 712 [crush] uses the tactic [intuition], which, when it runs out of tricks to try using only propositional logic, by default tries the tactic [auto], which we saw in an earlier example. [auto] attempts Prolog-style logic programming, searching through all proof trees up to a certain depth that are built only out of hints that have been registered with [Hint] commands. Compared to Prolog, [auto] places an important restriction: it never introduces new unification variables during search. That is, every time a rule is applied during proof search, all of its arguments must be deducible by studying the form of the goal. [eauto] relaxes this restriction, at the cost of possibly exponentially greater running time. In this particular case, we know that [eauto] has only a small space of proofs to search, so it makes sense to run it. It is common in effectively-automated Coq proofs to see a bag of standard tactics applied to pick off the "easy" subgoals, finishing off with [eauto] to handle the tricky parts that can benefit from ad-hoc exhaustive search.
adamc@50 713
adamc@50 714 The original theorem now follows trivially from our lemma. *)
adamc@50 715
adamc@50 716 Theorem even_contra : forall n, even (S (n + n)) -> False.
adamc@52 717 intros; eapply even_contra'; eauto.
adamc@50 718 Qed.
adamc@52 719
adamc@52 720 (** We use a variant [eapply] of [apply] which has the same relationship to [apply] as [eauto] has to [auto]. [apply] only succeeds if all arguments to the rule being used can be determined from the form of the goal, whereas [eapply] will introduce unification variables for undetermined arguments. [eauto] is able to determine the right values for those unification variables.
adamc@52 721
adamc@52 722 By considering an alternate attempt at proving the lemma, we can see another common pitfall of inductive proofs in Coq. Imagine that we had tried to prove [even_contra'] with all of the [forall] quantifiers moved to the front of the lemma statement. *)
adamc@52 723
adamc@52 724 Lemma even_contra'' : forall n' n, even n' -> n' = S (n + n) -> False.
adamc@52 725 induction 1; crush;
adamc@52 726 match goal with
adamc@52 727 | [ H : S ?N = ?N0 + ?N0 |- _ ] => destruct N; destruct N0
adamc@52 728 end; crush; eauto.
adamc@52 729
adamc@52 730 (** One subgoal remains: *)
adamc@52 731
adamc@52 732 (** [[
adamc@52 733
adamc@52 734 n : nat
adamc@52 735 H : even (S (n + n))
adamc@52 736 IHeven : S (n + n) = S (S (S (n + n))) -> False
adamc@52 737 ============================
adamc@52 738 False
adamc@52 739 ]] *)
adamc@52 740
adamc@52 741 (** We are out of luck here. The inductive hypothesis is trivially true, since its assumption is false. In the version of this proof that succeeded, [IHeven] had an explicit quantification over [n]. This is because the quantification of [n] %\textit{%#<i>#appeared after the thing we are inducting on#</i>#%}% in the theorem statement. In general, quantified variables and hypotheses that appear before the induction object in the theorem statement stay fixed throughout the inductive proof. Variables and hypotheses that are quantified after the induction object may be varied explicitly in uses of inductive hypotheses.
adamc@52 742
adamc@52 743 Why should Coq implement [induction] this way? One answer is that it avoids burdening this basic tactic with additional heuristic smarts, but that is not the whole picture. Imagine that [induction] analyzed dependencies among variables and reordered quantifiers to preserve as much freedom as possible in later uses of inductive hypotheses. This could make the inductive hypotheses more complex, which could in turn cause particular automation machinery to fail when it would have succeeded before. In general, we want to avoid quantifiers in our proofs whenever we can, and that goal is furthered by the refactoring that the [induction] tactic forces us to do. *)
adamc@52 744
adamc@52 745
adamc@52 746 (* begin hide *)
adamc@52 747 (* In-class exercises *)
adamc@52 748
adamc@52 749 (* EX: Define a type [prop] of simple boolean formulas made up only of truth, falsehood, binary conjunction, and binary disjunction. Define an inductive predicate [holds] that captures when [prop]s are valid, and define a predicate [falseFree] that captures when a [prop] does not contain the "false" formula. Prove that every false-free [prop] is valid. *)
adamc@52 750
adamc@52 751 (* begin thide *)
adamc@52 752 Inductive prop : Set :=
adamc@52 753 | Tru : prop
adamc@52 754 | Fals : prop
adamc@52 755 | And : prop -> prop -> prop
adamc@52 756 | Or : prop -> prop -> prop.
adamc@52 757
adamc@52 758 Inductive holds : prop -> Prop :=
adamc@52 759 | HTru : holds Tru
adamc@52 760 | HAnd : forall p1 p2, holds p1 -> holds p2 -> holds (And p1 p2)
adamc@52 761 | HOr1 : forall p1 p2, holds p1 -> holds (Or p1 p2)
adamc@52 762 | HOr2 : forall p1 p2, holds p2 -> holds (Or p1 p2).
adamc@52 763
adamc@52 764 Inductive falseFree : prop -> Prop :=
adamc@52 765 | FFTru : falseFree Tru
adamc@52 766 | FFAnd : forall p1 p2, falseFree p1 -> falseFree p2 -> falseFree (And p1 p2)
adamc@52 767 | FFNot : forall p1 p2, falseFree p1 -> falseFree p2 -> falseFree (Or p1 p2).
adamc@52 768
adamc@52 769 Hint Constructors holds.
adamc@52 770
adamc@52 771 Theorem falseFree_holds : forall p, falseFree p -> holds p.
adamc@52 772 induction 1; crush.
adamc@52 773 Qed.
adamc@52 774 (* end thide *)
adamc@52 775
adamc@52 776
adamc@52 777 (* EX: Define an inductive type [prop'] that is the same as [prop] but omits the possibility for falsehood. Define a proposition [holds'] for [prop'] that is analogous to [holds]. Define a function [propify] for translating [prop']s to [prop]s. Prove that, for any [prop'] [p], if [propify p] is valid, then so is [p]. *)
adamc@52 778
adamc@52 779 (* begin thide *)
adamc@52 780 Inductive prop' : Set :=
adamc@52 781 | Tru' : prop'
adamc@52 782 | And' : prop' -> prop' -> prop'
adamc@52 783 | Or' : prop' -> prop' -> prop'.
adamc@52 784
adamc@52 785 Inductive holds' : prop' -> Prop :=
adamc@52 786 | HTru' : holds' Tru'
adamc@52 787 | HAnd' : forall p1 p2, holds' p1 -> holds' p2 -> holds' (And' p1 p2)
adamc@52 788 | HOr1' : forall p1 p2, holds' p1 -> holds' (Or' p1 p2)
adamc@52 789 | HOr2' : forall p1 p2, holds' p2 -> holds' (Or' p1 p2).
adamc@52 790
adamc@52 791 Fixpoint propify (p : prop') : prop :=
adamc@52 792 match p with
adamc@52 793 | Tru' => Tru
adamc@52 794 | And' p1 p2 => And (propify p1) (propify p2)
adamc@52 795 | Or' p1 p2 => Or (propify p1) (propify p2)
adamc@52 796 end.
adamc@52 797
adamc@52 798 Hint Constructors holds'.
adamc@52 799
adamc@52 800 Lemma propify_holds' : forall p', holds p' -> forall p, p' = propify p -> holds' p.
adamc@52 801 induction 1; crush; destruct p; crush.
adamc@52 802 Qed.
adamc@52 803
adamc@52 804 Theorem propify_holds : forall p, holds (propify p) -> holds' p.
adamc@52 805 intros; eapply propify_holds'; eauto.
adamc@52 806 Qed.
adamc@52 807 (* end thide *)
adamc@52 808
adamc@52 809 (* end hide *)