annotate src/Predicates.v @ 399:5986e9fd40b5

Start figuring out which coqdoc changes will be needed to produce a pretty final version
author Adam Chlipala <adam@chlipala.net>
date Fri, 08 Jun 2012 11:25:11 -0400
parents 05efde66559d
children c898e72b84a3
rev   line source
adam@394 1 (* Copyright (c) 2008-2012, 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
adam@314 13 Require Import CpdtTactics.
adamc@45 14
adamc@45 15 Set Implicit Arguments.
adam@322 16
adam@322 17 (* Extra definitions to get coqdoc to choose the right fonts. *)
adam@322 18
adam@323 19 (* begin thide *)
adam@322 20 Inductive unit := tt.
adam@322 21 Inductive Empty_set := .
adam@322 22 Inductive bool := true | false.
adam@322 23 Inductive sum := .
adam@322 24 Inductive prod := .
adam@322 25 Inductive and := conj.
adam@322 26 Inductive or := or_introl | or_intror.
adam@322 27 Inductive ex := ex_intro.
adam@322 28 Inductive eq := refl_equal.
adam@322 29 Reset unit.
adam@323 30 (* end thide *)
adamc@45 31 (* end hide *)
adamc@45 32
adamc@45 33 (** %\chapter{Inductive Predicates}% *)
adamc@45 34
adam@322 35 (** The so-called %\index{Curry-Howard correspondence}``%#"#Curry-Howard correspondence#"#%''~\cite{Curry,Howard}% 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 36
adam@322 37 (* begin hide *)
adamc@45 38 Print unit.
adam@322 39 (* end hide *)
adam@322 40 (** %\noindent%[Print] %\coqdocinductive{%#<tt>#unit#</tt>#%}%[.] *)
adam@322 41 (** [[
adamc@209 42 Inductive unit : Set := tt : unit
adam@302 43 ]]
adam@302 44 *)
adamc@45 45
adam@322 46 (* begin hide *)
adamc@45 47 Print True.
adam@322 48 (* end hide *)
adam@322 49 (** %\noindent%[Print] %\coqdocinductive{%#<tt>#True#</tt>#%}%[.] *)
adam@322 50 (** [[
adamc@209 51 Inductive True : Prop := I : True
adam@322 52 ]]
adam@302 53 *)
adamc@45 54
adam@350 55 (** 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]. Chapter 12 goes into more detail about the theoretical differences between [Prop] and [Set]. For now, we will simply follow common intuitions about what a proof is.
adamc@45 56
adam@398 57 The type [unit] has one value, [tt]. The type [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 _should not_ 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 58
adam@398 59 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 %\index{proof irrelevance}%_proof irrelevance_, 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 60
adam@399 61 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 62
adamc@45 63
adamc@48 64 (** * Propositional Logic *)
adamc@45 65
adamc@45 66 (** 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 67
adamc@45 68 Section Propositional.
adamc@46 69 Variables P Q R : Prop.
adamc@45 70
adamc@45 71 (** 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 72
adamc@45 73 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 74
adamc@45 75 Theorem obvious : True.
adamc@55 76 (* begin thide *)
adamc@45 77 apply I.
adamc@55 78 (* end thide *)
adamc@45 79 Qed.
adamc@45 80
adamc@45 81 (** 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 82
adamc@55 83 (* begin thide *)
adamc@45 84 Theorem obvious' : True.
adam@322 85 (* begin hide *)
adamc@45 86 constructor.
adam@322 87 (* end hide *)
adam@322 88 (** %\hspace{.075in}\coqdockw{%#<tt>#constructor#</tt>#%}%.%\vspace{-.1in}% *)
adam@322 89
adamc@45 90 Qed.
adamc@45 91
adamc@55 92 (* end thide *)
adamc@55 93
adamc@45 94 (** There is also a predicate [False], which is the Curry-Howard mirror image of [Empty_set] from the last chapter. *)
adamc@45 95
adam@322 96 (* begin hide *)
adamc@45 97 Print False.
adam@322 98 (* end hide *)
adam@322 99 (** %\noindent%[Print] %\coqdocinductive{%#<tt>#False#</tt>#%}%[.] *)
adam@322 100 (** [[
adamc@209 101 Inductive False : Prop :=
adamc@209 102
adamc@209 103 ]]
adamc@45 104
adamc@209 105 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 106
adamc@45 107 Theorem False_imp : False -> 2 + 2 = 5.
adamc@55 108 (* begin thide *)
adamc@45 109 destruct 1.
adamc@55 110 (* end thide *)
adamc@45 111 Qed.
adamc@45 112
adamc@45 113 (** 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 114
adamc@45 115 Theorem arith_neq : 2 + 2 = 5 -> 9 + 9 = 835.
adamc@55 116 (* begin thide *)
adamc@45 117 intro.
adamc@45 118
adam@322 119 (** At this point, we have an inconsistent hypothesis [2 + 2 = 5], so the specific conclusion is not important. We use the %\index{tactics!elimtype}%[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 120
adamc@45 121 elimtype False.
adamc@45 122 (** [[
adamc@45 123 H : 2 + 2 = 5
adamc@45 124 ============================
adamc@45 125 False
adamc@209 126
adamc@209 127 ]]
adamc@45 128
adamc@209 129 For now, we will leave the details of this proof about arithmetic to [crush]. *)
adamc@45 130
adamc@45 131 crush.
adamc@55 132 (* end thide *)
adamc@45 133 Qed.
adamc@45 134
adamc@45 135 (** A related notion to [False] is logical negation. *)
adamc@45 136
adamc@45 137 Print not.
adamc@209 138 (** %\vspace{-.15in}% [[
adamc@209 139 not = fun A : Prop => A -> False
adamc@209 140 : Prop -> Prop
adamc@209 141
adamc@209 142 ]]
adamc@45 143
adam@280 144 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 145
adamc@45 146 Theorem arith_neq' : ~ (2 + 2 = 5).
adamc@55 147 (* begin thide *)
adamc@45 148 unfold not.
adamc@45 149 (** [[
adamc@45 150 ============================
adamc@45 151 2 + 2 = 5 -> False
adam@302 152 ]]
adam@302 153 *)
adamc@45 154
adamc@45 155 crush.
adamc@55 156 (* end thide *)
adamc@45 157 Qed.
adamc@45 158
adamc@45 159 (** We also have conjunction, which we introduced in the last chapter. *)
adamc@45 160
adam@322 161 (* begin hide *)
adamc@45 162 Print and.
adam@322 163 (* end hide *)
adam@322 164 (** %\noindent%[Print] %\coqdocinductive{%#<tt>#and#</tt>#%}%[.]
adam@322 165 [[
adam@322 166 Inductive and (A : Prop) (B : Prop) : Prop := conj : A -> B -> A /\ B
adamc@209 167
adamc@209 168 ]]
adamc@209 169
adam@322 170 The interested reader can check that [and] has a Curry-Howard equivalent called %\index{Gallina terms!prod}%[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. The operator [/\] is an infix shorthand for [and]. *)
adamc@45 171
adamc@45 172 Theorem and_comm : P /\ Q -> Q /\ P.
adamc@209 173
adamc@55 174 (* begin thide *)
adamc@45 175 (** We start by case analysis on the proof of [P /\ Q]. *)
adamc@45 176
adamc@45 177 destruct 1.
adamc@45 178 (** [[
adamc@45 179 H : P
adamc@45 180 H0 : Q
adamc@45 181 ============================
adamc@45 182 Q /\ P
adamc@209 183
adamc@209 184 ]]
adamc@45 185
adam@322 186 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].%\index{tactics!split}% *)
adamc@45 187
adamc@45 188 split.
adam@322 189 (** %\vspace{.1in} \noindent 2 \coqdockw{subgoals}\vspace{-.1in}%#<tt>2 subgoals</tt>#
adam@322 190 [[
adamc@45 191
adamc@45 192 H : P
adamc@45 193 H0 : Q
adamc@45 194 ============================
adamc@45 195 Q
adam@322 196 ]]
adam@322 197 %\noindent \coqdockw{subgoal} 2 \coqdockw{is}:%#<tt>subgoal 2 is</tt>#
adam@322 198 [[
adam@322 199 P
adamc@209 200
adamc@209 201 ]]
adamc@45 202
adam@322 203 In each case, the conclusion is among our hypotheses, so the %\index{tactics!assumption}%[assumption] tactic finishes the process. *)
adamc@45 204
adamc@45 205 assumption.
adamc@45 206 assumption.
adamc@55 207 (* end thide *)
adamc@45 208 Qed.
adamc@45 209
adam@322 210 (** Coq disjunction is called %\index{Gallina terms!or}%[or] and abbreviated with the infix operator [\/]. *)
adamc@45 211
adam@322 212 (* begin hide *)
adamc@45 213 Print or.
adam@322 214 (* end hide *)
adam@322 215 (** %\noindent%[Print] %\coqdocinductive{%#<tt>#or#</tt>#%}%[.]
adam@322 216 [[
adamc@209 217 Inductive or (A : Prop) (B : Prop) : Prop :=
adamc@209 218 or_introl : A -> A \/ B | or_intror : B -> A \/ B
adamc@209 219
adamc@209 220 ]]
adamc@45 221
adam@322 222 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 %\index{Gallina terms!sum}%[sum] type. We can demonstrate the main tactics here with another proof of commutativity. *)
adamc@45 223
adamc@45 224 Theorem or_comm : P \/ Q -> Q \/ P.
adamc@55 225
adamc@55 226 (* begin thide *)
adamc@45 227 (** As in the proof for [and], we begin with case analysis, though this time we are met by two cases instead of one. *)
adamc@209 228
adamc@45 229 destruct 1.
adam@322 230 (** %\vspace{.1in} \noindent 2 \coqdockw{subgoals}\vspace{-.1in}%#<tt>2 subgoals</tt>#
adam@322 231 [[
adamc@45 232
adamc@45 233 H : P
adamc@45 234 ============================
adamc@45 235 Q \/ P
adam@322 236 ]]
adam@322 237 %\noindent \coqdockw{subgoal} 2 \coqdockw{is}:%#<tt>subgoal 2 is</tt>#
adam@322 238 [[
adamc@45 239 Q \/ P
adamc@209 240
adamc@209 241 ]]
adamc@45 242
adam@322 243 We can see that, in the first subgoal, we want to prove the disjunction by proving its second disjunct. The %\index{tactics!right}\coqdockw{%#<tt>#right#</tt>#%}% tactic telegraphs this intent. *)
adam@322 244
adam@322 245 (* begin hide *)
adamc@45 246 right; assumption.
adam@322 247 (* end hide *)
adam@322 248 (** %\hspace{.075in}\coqdockw{%#<tt>#right#</tt>#%}%[; assumption.] *)
adamc@45 249
adam@322 250 (** The second subgoal has a symmetric proof.%\index{tactics!left}%
adamc@45 251
adamc@45 252 [[
adamc@45 253 1 subgoal
adamc@45 254
adamc@45 255 H : Q
adamc@45 256 ============================
adamc@45 257 Q \/ P
adam@302 258 ]]
adam@302 259 *)
adamc@45 260
adam@322 261 (* begin hide *)
adamc@45 262 left; assumption.
adam@322 263 (* end hide *)
adam@322 264 (** %\hspace{.075in}\coqdockw{%#<tt>#left#</tt>#%}%[; assumption.] *)
adam@322 265
adamc@55 266 (* end thide *)
adamc@45 267 Qed.
adamc@45 268
adamc@46 269
adamc@46 270 (* begin hide *)
adamc@46 271 (* In-class exercises *)
adamc@46 272
adamc@46 273 Theorem contra : P -> ~P -> R.
adamc@52 274 (* begin thide *)
adamc@52 275 unfold not.
adamc@52 276 intros.
adamc@52 277 elimtype False.
adamc@52 278 apply H0.
adamc@52 279 assumption.
adamc@52 280 (* end thide *)
adamc@46 281 Admitted.
adamc@46 282
adamc@46 283 Theorem and_assoc : (P /\ Q) /\ R -> P /\ (Q /\ R).
adamc@52 284 (* begin thide *)
adamc@52 285 intros.
adamc@52 286 destruct H.
adamc@52 287 destruct H.
adamc@52 288 split.
adamc@52 289 assumption.
adamc@52 290 split.
adamc@52 291 assumption.
adamc@52 292 assumption.
adamc@52 293 (* end thide *)
adamc@46 294 Admitted.
adamc@46 295
adamc@46 296 Theorem or_assoc : (P \/ Q) \/ R -> P \/ (Q \/ R).
adamc@52 297 (* begin thide *)
adamc@52 298 intros.
adamc@52 299 destruct H.
adamc@52 300 destruct H.
adamc@52 301 left.
adamc@52 302 assumption.
adamc@52 303 right.
adamc@52 304 left.
adamc@52 305 assumption.
adamc@52 306 right.
adamc@52 307 right.
adamc@52 308 assumption.
adamc@52 309 (* end thide *)
adamc@46 310 Admitted.
adamc@46 311
adamc@46 312 (* end hide *)
adamc@46 313
adamc@46 314
adam@322 315 (** 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 %\index{tactics!tauto}%[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 316
adamc@46 317 Theorem or_comm' : P \/ Q -> Q \/ P.
adamc@55 318 (* begin thide *)
adamc@46 319 tauto.
adamc@55 320 (* end thide *)
adamc@46 321 Qed.
adamc@46 322
adam@322 323 (** Sometimes propositional reasoning forms important plumbing for the proof of a theorem, but we still need to apply some other smarts about, say, arithmetic. %\index{tactics!intuition}%[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 324
adamc@46 325 Theorem arith_comm : forall ls1 ls2 : list nat,
adamc@46 326 length ls1 = length ls2 \/ length ls1 + length ls2 = 6
adamc@46 327 -> length (ls1 ++ ls2) = 6 \/ length ls1 = length ls2.
adamc@55 328 (* begin thide *)
adamc@46 329 intuition.
adamc@46 330
adamc@46 331 (** 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 332
adamc@46 333 (** [[
adamc@46 334 ls1 : list nat
adamc@46 335 ls2 : list nat
adamc@46 336 H0 : length ls1 + length ls2 = 6
adamc@46 337 ============================
adamc@46 338 length (ls1 ++ ls2) = 6 \/ length ls1 = length ls2
adamc@209 339
adamc@209 340 ]]
adamc@46 341
adamc@209 342 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 343
adamc@46 344 rewrite app_length.
adamc@46 345 (** [[
adamc@46 346 ls1 : list nat
adamc@46 347 ls2 : list nat
adamc@46 348 H0 : length ls1 + length ls2 = 6
adamc@46 349 ============================
adamc@46 350 length ls1 + length ls2 = 6 \/ length ls1 = length ls2
adamc@209 351
adamc@209 352 ]]
adamc@46 353
adamc@209 354 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 355
adamc@46 356 tauto.
adamc@55 357 (* end thide *)
adamc@46 358 Qed.
adamc@46 359
adam@322 360 (** The [intuition] tactic 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 361
adamc@55 362 (* begin thide *)
adamc@46 363 Theorem arith_comm' : forall ls1 ls2 : list nat,
adamc@46 364 length ls1 = length ls2 \/ length ls1 + length ls2 = 6
adamc@46 365 -> length (ls1 ++ ls2) = 6 \/ length ls1 = length ls2.
adam@322 366 (* begin hide *)
adam@375 367 Hint Rewrite app_length.
adam@322 368 (* end hide *)
adam@375 369 (** %\hspace{.075in}%[Hint] %\coqdockw{%#<tt>#Rewrite#</tt>#%}% [app_length.] *)
adamc@46 370
adamc@46 371 crush.
adamc@46 372 Qed.
adamc@55 373 (* end thide *)
adamc@46 374
adamc@45 375 End Propositional.
adamc@45 376
adam@322 377 (** Ending the section here has the same effect as always. Each of our propositional theorems becomes universally quantified over the propositional variables that we used. *)
adam@322 378
adamc@46 379
adamc@47 380 (** * What Does It Mean to Be Constructive? *)
adamc@46 381
adamc@47 382 (** 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 383
adam@398 384 The answer comes from the fact that Coq implements %\index{constructive logic}%_constructive_ or %\index{intuitionistic logic|see{constructive logic}}%_intuitionistic_ logic, in contrast to the %\index{classical logic}%_classical_ 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 %\index{decidability}%_decidable_, in the sense of %\index{computability|see{decidability}}%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, a general %\index{law of the excluded middle}%law of the excluded middle 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 385
adam@292 386 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 387
adam@399 388 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 %\index{program extraction}%_program extraction_, 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 389
adamc@47 390 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 391
adamc@48 392
adamc@48 393 (** * First-Order Logic *)
adamc@48 394
adam@322 395 (** The %\index{Gallina terms!forall}%[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 396
adam@322 397 %\index{existential quantification}\index{Gallina terms!exists}\index{Gallina terms!ex}%Existential quantification is defined in the standard library. *)
adamc@48 398
adam@322 399 (* begin hide *)
adam@322 400 Print ex.
adam@322 401 (* end hide *)
adam@322 402 (** %\noindent%[Print] %\coqdocinductive{%#<tt>#ex#</tt>#%}%[.]
adam@322 403 [[
adamc@209 404 Inductive ex (A : Type) (P : A -> Prop) : Prop :=
adamc@209 405 ex_intro : forall x : A, P x -> ex P
adamc@209 406
adamc@209 407 ]]
adamc@48 408
adam@322 409 The family [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. We use the equality operator [=], which, depending on the settings in which they learned logic, different people will say either is or is not part of first-order logic. For our purposes, it is. *)
adamc@48 410
adamc@48 411 Theorem exist1 : exists x : nat, x + 1 = 2.
adamc@55 412 (* begin thide *)
adamc@67 413 (** remove printing exists *)
adam@322 414 (** We can start this proof with a tactic %\index{tactics!exists}\coqdockw{%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#"#%''% in formulas.) *)
adamc@209 415
adam@322 416 (* begin hide *)
adamc@48 417 exists 1.
adam@322 418 (* end hide *)
adam@322 419 (** %\coqdockw{%#<tt>#exists#</tt>#%}% [1.] *)
adamc@48 420
adamc@209 421 (** The conclusion is replaced with a version using the existential witness that we announced.
adamc@48 422
adamc@209 423 [[
adamc@48 424 ============================
adamc@48 425 1 + 1 = 2
adam@302 426 ]]
adam@302 427 *)
adamc@48 428
adamc@48 429 reflexivity.
adamc@55 430 (* end thide *)
adamc@48 431 Qed.
adamc@48 432
adamc@48 433 (** printing exists $\exists$ *)
adamc@48 434
adamc@48 435 (** We can also use tactics to reason about existential hypotheses. *)
adamc@48 436
adamc@48 437 Theorem exist2 : forall n m : nat, (exists x : nat, n + x = m) -> n <= m.
adamc@55 438 (* begin thide *)
adamc@48 439 (** We start by case analysis on the proof of the existential fact. *)
adamc@209 440
adamc@48 441 destruct 1.
adamc@48 442 (** [[
adamc@48 443 n : nat
adamc@48 444 m : nat
adamc@48 445 x : nat
adamc@48 446 H : n + x = m
adamc@48 447 ============================
adamc@48 448 n <= m
adamc@209 449
adamc@209 450 ]]
adamc@48 451
adamc@209 452 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 453
adamc@48 454 crush.
adamc@55 455 (* end thide *)
adamc@48 456 Qed.
adamc@48 457
adamc@48 458
adamc@48 459 (* begin hide *)
adamc@48 460 (* In-class exercises *)
adamc@48 461
adamc@48 462 Theorem forall_exists_commute : forall (A B : Type) (P : A -> B -> Prop),
adamc@48 463 (exists x : A, forall y : B, P x y) -> (forall y : B, exists x : A, P x y).
adamc@52 464 (* begin thide *)
adamc@52 465 intros.
adamc@52 466 destruct H.
adamc@52 467 exists x.
adamc@52 468 apply H.
adamc@52 469 (* end thide *)
adamc@48 470 Admitted.
adamc@48 471
adamc@48 472 (* end hide *)
adamc@48 473
adamc@48 474
adam@322 475 (** The tactic [intuition] has a first-order cousin called %\index{tactics!firstorder}%[firstorder], which 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 476
adamc@49 477
adamc@49 478 (** * Predicates with Implicit Equality *)
adamc@49 479
adamc@49 480 (** 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 481
adamc@49 482 Inductive isZero : nat -> Prop :=
adamc@49 483 | IsZero : isZero 0.
adamc@49 484
adamc@49 485 Theorem isZero_zero : isZero 0.
adamc@55 486 (* begin thide *)
adam@322 487 (* begin hide *)
adamc@49 488 constructor.
adam@322 489 (* end hide *)
adam@322 490 (** %\coqdockw{%#<tt>#constructor#</tt>#%}%[.]%\vspace{-.075in}% *)
adam@322 491
adamc@55 492 (* end thide *)
adamc@49 493 Qed.
adamc@49 494
adam@398 495 (** We can call [isZero] a %\index{judgment}%_judgment_, in the sense often used in the semantics of programming languages. Judgments are typically defined in the style of %\index{natural deduction}%_natural deduction_, where we write a number of %\index{inference rules}%_inference rules_ 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 496
adam@398 497 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 _before_ 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 498
adam@322 499 For instance, our definition [isZero] makes the predicate provable only when the argument is [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 %\index{Prolog}%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!%\index{Gallina terms!eq}\index{Gallina terms!refl\_equal}% *)
adamc@49 500
adam@322 501 (* begin hide *)
adamc@49 502 Print eq.
adam@322 503 (* end hide *)
adam@322 504 (** %\noindent%[Print] %\coqdocinductive{%#<tt>#eq#</tt>#%}%[.]
adam@322 505 [[
adamc@209 506 Inductive eq (A : Type) (x : A) : A -> Prop := refl_equal : x = x
adamc@209 507
adamc@209 508 ]]
adamc@49 509
adam@398 510 Behind the scenes, uses of infix [=] are expanded to instances of [eq]. We see that [eq] has both a parameter [x] that is fixed and an extra unnamed argument of the same type. The type of [eq] allows us to state any equalities, even those that are provably false. However, examining the type of equality's sole constructor [refl_equal], we see that we can only _prove_ 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. Another way of stating that definition is: equality is defined as the least reflexive relation.
adamc@49 511
adam@322 512 Returning to the example of [isZero], we can see how to work with hypotheses that use this predicate. *)
adamc@49 513
adamc@49 514 Theorem isZero_plus : forall n m : nat, isZero m -> n + m = n.
adamc@55 515 (* begin thide *)
adamc@49 516 (** We want to proceed by cases on the proof of the assumption about [isZero]. *)
adamc@209 517
adamc@49 518 destruct 1.
adamc@49 519 (** [[
adamc@49 520 n : nat
adamc@49 521 ============================
adamc@49 522 n + 0 = n
adamc@209 523
adamc@209 524 ]]
adamc@49 525
adamc@209 526 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 527
adamc@49 528 crush.
adamc@55 529 (* end thide *)
adamc@49 530 Qed.
adamc@49 531
adamc@49 532 (** 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 533
adamc@49 534 Theorem isZero_contra : isZero 1 -> False.
adamc@55 535 (* begin thide *)
adamc@49 536 (** Let us try a proof by cases on the assumption, as in the last proof. *)
adamc@209 537
adamc@49 538 destruct 1.
adamc@49 539 (** [[
adamc@49 540 ============================
adamc@49 541 False
adamc@209 542
adamc@209 543 ]]
adamc@49 544
adamc@209 545 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 546
adam@322 547 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 %\index{tactics!inversion}%[inversion], which corresponds to the concept of inversion that is frequently used with natural deduction proof systems. *)
adamc@49 548
adamc@49 549 Undo.
adamc@49 550 inversion 1.
adamc@55 551 (* end thide *)
adamc@49 552 Qed.
adamc@49 553
adamc@49 554 (** 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 555
adamc@49 556 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 557
adamc@49 558 Theorem isZero_contra' : isZero 1 -> 2 + 2 = 5.
adamc@49 559 destruct 1.
adamc@49 560 (** [[
adamc@49 561 ============================
adamc@49 562 1 + 1 = 4
adamc@209 563
adamc@209 564 ]]
adamc@49 565
adam@280 566 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. *)
adamc@209 567
adamc@49 568 Abort.
adamc@49 569
adam@280 570 (** To see more clearly what is happening, we can consider the type of [isZero]'s induction principle. *)
adam@280 571
adam@280 572 Check isZero_ind.
adam@280 573 (** %\vspace{-.15in}% [[
adam@280 574 isZero_ind
adam@280 575 : forall P : nat -> Prop, P 0 -> forall n : nat, isZero n -> P n
adam@280 576
adam@280 577 ]]
adam@280 578
adam@322 579 In our last proof script, [destruct] chose to instantiate [P] as [fun n => S n + S n = S (][S (][S (][S n)))]. You can verify for yourself that this specialization of the principle applies to the goal and that the hypothesis [P 0] then matches the subgoal we saw generated. 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]. *)
adam@280 580
adamc@49 581
adamc@49 582 (* begin hide *)
adamc@49 583 (* In-class exercises *)
adamc@49 584
adamc@49 585 (* 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 586
adamc@52 587 (* begin thide *)
adamc@52 588 Section twoEls.
adamc@52 589 Variable A : Type.
adamc@52 590
adamc@52 591 Inductive twoEls : list A -> Prop :=
adamc@52 592 | TwoEls : forall x y, twoEls (x :: y :: nil).
adamc@52 593
adamc@52 594 Theorem twoEls_nil : twoEls nil -> False.
adamc@52 595 inversion 1.
adamc@52 596 Qed.
adamc@52 597
adamc@52 598 Theorem twoEls_two : forall ls, twoEls ls -> length ls = 2.
adamc@52 599 inversion 1.
adamc@52 600 reflexivity.
adamc@52 601 Qed.
adamc@52 602 End twoEls.
adamc@52 603 (* end thide *)
adamc@52 604
adamc@49 605 (* end hide *)
adamc@49 606
adamc@50 607
adamc@50 608 (** * Recursive Predicates *)
adamc@50 609
adamc@50 610 (** We have already seen all of the ingredients we need to build interesting recursive predicates, like this predicate capturing even-ness. *)
adamc@50 611
adamc@50 612 Inductive even : nat -> Prop :=
adamc@50 613 | EvenO : even O
adamc@50 614 | EvenSS : forall n, even n -> even (S (S n)).
adamc@50 615
adam@322 616 (** 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 617
adamc@50 618 The proof techniques of the last section are easily adapted. *)
adamc@50 619
adamc@50 620 Theorem even_0 : even 0.
adamc@55 621 (* begin thide *)
adam@322 622 (* begin hide *)
adamc@50 623 constructor.
adam@322 624 (* end hide *)
adam@322 625 (** %\coqdockw{%#<tt>#constructor#</tt>#%}%[.]%\vspace{-.075in}% *)
adam@322 626
adamc@55 627 (* end thide *)
adamc@50 628 Qed.
adamc@50 629
adamc@50 630 Theorem even_4 : even 4.
adamc@55 631 (* begin thide *)
adam@322 632 (* begin hide *)
adamc@50 633 constructor; constructor; constructor.
adam@322 634 (* end hide *)
adam@322 635 (** %\coqdockw{%#<tt>#constructor#</tt>#%}%[; ]%\coqdockw{%#<tt>#constructor#</tt>#%}%[; ]%\coqdockw{%#<tt>#constructor#</tt>#%}%[.]%\vspace{-.075in}% *)
adam@322 636
adamc@55 637 (* end thide *)
adamc@50 638 Qed.
adamc@50 639
adam@375 640 (** 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, with a new [Hint] variant that asks to consider all constructors of an inductive type during proof search. The tactic %\index{tactics!auto}%[auto] performs exhaustive proof search up to a fixed depth, considering only the proof steps we have registered as hints. *)
adamc@50 641
adamc@55 642 (* begin thide *)
adam@322 643 (* begin hide *)
adamc@50 644 Hint Constructors even.
adam@322 645 (* end hide *)
adam@322 646 (** %\noindent%[Hint] %\coqdockw{%#<tt>#Constructors#</tt>#%}% [even.] *)
adamc@50 647
adamc@50 648 Theorem even_4' : even 4.
adamc@50 649 auto.
adamc@50 650 Qed.
adamc@50 651
adamc@55 652 (* end thide *)
adamc@55 653
adam@322 654 (** We may also use [inversion] with [even]. *)
adam@322 655
adamc@50 656 Theorem even_1_contra : even 1 -> False.
adamc@55 657 (* begin thide *)
adamc@50 658 inversion 1.
adamc@55 659 (* end thide *)
adamc@50 660 Qed.
adamc@50 661
adamc@50 662 Theorem even_3_contra : even 3 -> False.
adamc@55 663 (* begin thide *)
adamc@50 664 inversion 1.
adamc@50 665 (** [[
adamc@50 666 H : even 3
adamc@50 667 n : nat
adamc@50 668 H1 : even 1
adamc@50 669 H0 : n = 1
adamc@50 670 ============================
adamc@50 671 False
adamc@209 672
adamc@209 673 ]]
adamc@50 674
adam@322 675 The [inversion] tactic 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. More complex inductive definitions and theorems can cause [inversion] to generate equalities where neither side is a variable. *)
adamc@50 676
adamc@50 677 inversion H1.
adamc@55 678 (* end thide *)
adamc@50 679 Qed.
adamc@50 680
adamc@50 681 (** We can also do inductive proofs about [even]. *)
adamc@50 682
adamc@50 683 Theorem even_plus : forall n m, even n -> even m -> even (n + m).
adamc@55 684 (* begin thide *)
adamc@50 685 (** It seems a reasonable first choice to proceed by induction on [n]. *)
adamc@209 686
adamc@50 687 induction n; crush.
adamc@50 688 (** [[
adamc@50 689 n : nat
adamc@50 690 IHn : forall m : nat, even n -> even m -> even (n + m)
adamc@50 691 m : nat
adamc@50 692 H : even (S n)
adamc@50 693 H0 : even m
adamc@50 694 ============================
adamc@50 695 even (S (n + m))
adamc@209 696
adamc@209 697 ]]
adamc@50 698
adamc@209 699 We will need to use the hypotheses [H] and [H0] somehow. The most natural choice is to invert [H]. *)
adamc@50 700
adamc@50 701 inversion H.
adamc@50 702 (** [[
adamc@50 703 n : nat
adamc@50 704 IHn : forall m : nat, even n -> even m -> even (n + m)
adamc@50 705 m : nat
adamc@50 706 H : even (S n)
adamc@50 707 H0 : even m
adamc@50 708 n0 : nat
adamc@50 709 H2 : even n0
adamc@50 710 H1 : S n0 = n
adamc@50 711 ============================
adamc@50 712 even (S (S n0 + m))
adamc@209 713
adamc@209 714 ]]
adamc@50 715
adamc@209 716 Simplifying the conclusion brings us to a point where we can apply a constructor. *)
adamc@209 717
adamc@50 718 simpl.
adamc@50 719 (** [[
adamc@50 720 ============================
adamc@50 721 even (S (S (n0 + m)))
adam@302 722 ]]
adam@302 723 *)
adamc@50 724
adam@322 725 (* begin hide *)
adamc@50 726 constructor.
adam@322 727 (* end hide *)
adam@322 728 (** %\coqdockw{%#<tt>#constructor#</tt>#%}%[.]
adam@322 729
adam@322 730 [[
adamc@50 731 ============================
adamc@50 732 even (n0 + m)
adamc@209 733
adamc@209 734 ]]
adamc@50 735
adamc@209 736 At this point, we would like to apply the inductive hypothesis, which is:
adamc@209 737
adamc@209 738 [[
adamc@50 739 IHn : forall m : nat, even n -> even m -> even (n + m)
adamc@209 740 ]]
adamc@50 741
adam@398 742 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 _on the structure of one of the [even] proofs_. This technique is commonly called %\index{rule induction}%_rule induction_ 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 743
adam@322 744 Recall that tactics like [induction] and [destruct] may be passed numbers to refer to unnamed lefthand sides of implications in the conclusion, where the argument [n] refers to the [n]th such hypothesis. *)
adam@322 745
adam@322 746 (* begin hide *)
adamc@50 747 Restart.
adam@322 748 (* end hide *)
adam@322 749 (** %\noindent\coqdockw{%#<tt>#Restart#</tt>#%}%[.] *)
adamc@50 750
adamc@50 751 induction 1.
adamc@50 752 (** [[
adamc@50 753 m : nat
adamc@50 754 ============================
adamc@50 755 even m -> even (0 + m)
adam@322 756 ]]
adamc@50 757
adam@322 758 %\noindent \coqdockw{subgoal} 2 \coqdockw{is}:%#<tt>subgoal 2 is</tt>#
adam@322 759 [[
adamc@50 760 even m -> even (S (S n) + m)
adamc@209 761
adamc@209 762 ]]
adamc@50 763
adamc@209 764 The first case is easily discharged by [crush], based on the hint we added earlier to try the constructors of [even]. *)
adamc@50 765
adamc@50 766 crush.
adamc@50 767
adamc@50 768 (** Now we focus on the second case: *)
adamc@209 769
adamc@50 770 intro.
adamc@50 771 (** [[
adamc@50 772 m : nat
adamc@50 773 n : nat
adamc@50 774 H : even n
adamc@50 775 IHeven : even m -> even (n + m)
adamc@50 776 H0 : even m
adamc@50 777 ============================
adamc@50 778 even (S (S n) + m)
adamc@209 779
adamc@209 780 ]]
adamc@50 781
adamc@209 782 We simplify and apply a constructor, as in our last proof attempt. *)
adamc@50 783
adam@322 784 (* begin hide *)
adamc@50 785 simpl; constructor.
adam@322 786 (* end hide *)
adam@322 787 (** [simpl; ]%\coqdockw{%#<tt>#constructor#</tt>#%}%[.]
adam@322 788
adam@322 789 [[
adamc@50 790 ============================
adamc@50 791 even (n + m)
adamc@209 792
adamc@209 793 ]]
adamc@50 794
adamc@209 795 Now we have an exact match with our inductive hypothesis, and the remainder of the proof is trivial. *)
adamc@50 796
adamc@50 797 apply IHeven; assumption.
adamc@50 798
adamc@50 799 (** In fact, [crush] can handle all of the details of the proof once we declare the induction strategy. *)
adamc@50 800
adam@322 801 (* begin hide *)
adamc@50 802 Restart.
adam@322 803 (* end hide *)
adam@322 804 (** %\noindent\coqdockw{%#<tt>#Restart#</tt>#%}%[.] *)
adam@322 805
adamc@50 806 induction 1; crush.
adamc@55 807 (* end thide *)
adamc@50 808 Qed.
adamc@50 809
adamc@50 810 (** Induction on recursive predicates has similar pitfalls to those we encountered with inversion in the last section. *)
adamc@50 811
adamc@50 812 Theorem even_contra : forall n, even (S (n + n)) -> False.
adamc@55 813 (* begin thide *)
adamc@50 814 induction 1.
adamc@50 815 (** [[
adamc@50 816 n : nat
adamc@50 817 ============================
adamc@50 818 False
adam@322 819 ]]
adamc@50 820
adam@322 821 %\noindent \coqdockw{subgoal} 2 \coqdockw{is}:%#<tt>subgoal 2 is</tt>#
adam@322 822 [[
adamc@50 823 False
adamc@209 824
adamc@209 825 ]]
adamc@50 826
adam@280 827 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 easier 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@209 828
adamc@50 829 Abort.
adamc@50 830
adamc@50 831 Lemma even_contra' : forall n', even n' -> forall n, n' = S (n + n) -> False.
adamc@50 832 induction 1; crush.
adamc@50 833
adamc@54 834 (** 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 trickiness to it. *)
adamc@209 835
adamc@50 836 destruct n; destruct n0; crush.
adamc@50 837
adamc@50 838 (** [[
adamc@50 839 n : nat
adamc@50 840 H : even (S n)
adamc@50 841 IHeven : forall n0 : nat, S n = S (n0 + n0) -> False
adamc@50 842 n0 : nat
adamc@50 843 H0 : S n = n0 + S n0
adamc@50 844 ============================
adamc@50 845 False
adamc@209 846
adamc@209 847 ]]
adamc@50 848
adam@280 849 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. We can search for a theorem that allows us to rewrite terms of the form [x + S y]. *)
adamc@209 850
adam@322 851 (* begin hide *)
adam@280 852 SearchRewrite (_ + S _).
adam@322 853 (* end hide *)
adam@322 854 (** %\coqdockw{%#<tt>#SearchRewrite#</tt>#%}% [(_ + S _).]
adam@322 855
adam@322 856 [[
adam@280 857 plus_n_Sm : forall n m : nat, S (n + m) = n + S m
adam@302 858 ]]
adam@302 859 *)
adamc@50 860
adamc@50 861 rewrite <- plus_n_Sm in H0.
adamc@50 862
adam@322 863 (** The induction hypothesis lets us complete the proof, if we use a variant of [apply] that has a %\index{tactics!with}%[with] clause to give instantiations of quantified variables. *)
adamc@209 864
adamc@50 865 apply IHeven with n0; assumption.
adamc@50 866
adam@322 867 (** As usual, we can rewrite the proof to avoid referencing any locally generated names, which makes our proof script more readable and more robust to changes in the theorem statement. We use the notation [<-] to request a hint that does right-to-left rewriting, just like we can with the [rewrite] tactic. *)
adamc@209 868
adam@322 869 (* begin hide *)
adamc@209 870 Restart.
adam@322 871 (* end hide *)
adam@322 872 (** %\hspace{-.075in}\coqdockw{%#<tt>#Restart#</tt>#%}%[.] *)
adam@322 873
adam@322 874 (* begin hide *)
adam@375 875 Hint Rewrite <- plus_n_Sm.
adam@322 876 (* end hide *)
adam@375 877 (** %\hspace{-.075in}%[Hint] %\noindent\coqdockw{%#<tt>#Rewrite#</tt>#%}% [<- plus_n_sm.] *)
adamc@50 878
adamc@50 879 induction 1; crush;
adamc@50 880 match goal with
adamc@50 881 | [ H : S ?N = ?N0 + ?N0 |- _ ] => destruct N; destruct N0
adamc@50 882 end; crush; eauto.
adamc@50 883 Qed.
adamc@50 884
adam@322 885 (** We write the proof in a way that avoids the use of local variable or hypothesis names, using the %\index{tactics!match}%[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, %\index{tactics!eauto}%[eauto].
adamc@50 886
adam@322 887 The [crush] tactic 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. The [auto] tactic attempts %\index{Prolog}%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. This restriction is relaxed for [eauto], 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 with [eauto] to handle the tricky parts that can benefit from ad-hoc exhaustive search.
adamc@50 888
adamc@50 889 The original theorem now follows trivially from our lemma. *)
adamc@50 890
adamc@50 891 Theorem even_contra : forall n, even (S (n + n)) -> False.
adamc@52 892 intros; eapply even_contra'; eauto.
adamc@50 893 Qed.
adamc@52 894
adam@398 895 (** We use a variant %\index{tactics!apply}%[eapply] of [apply] which has the same relationship to [apply] as [eauto] has to [auto]. An invocation of [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. In this case, [eauto] is able to determine the right values for those unification variables, using (unsurprisingly) a variant of the classic algorithm for _unification_ %\cite{unification}%.
adamc@52 896
adamc@52 897 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 898
adamc@52 899 Lemma even_contra'' : forall n' n, even n' -> n' = S (n + n) -> False.
adamc@52 900 induction 1; crush;
adamc@52 901 match goal with
adamc@52 902 | [ H : S ?N = ?N0 + ?N0 |- _ ] => destruct N; destruct N0
adamc@52 903 end; crush; eauto.
adamc@52 904
adamc@209 905 (** One subgoal remains:
adamc@52 906
adamc@209 907 [[
adamc@52 908 n : nat
adamc@52 909 H : even (S (n + n))
adamc@52 910 IHeven : S (n + n) = S (S (S (n + n))) -> False
adamc@52 911 ============================
adamc@52 912 False
adamc@209 913
adamc@209 914 ]]
adamc@52 915
adam@398 916 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] _appeared after the thing we are inducting on_ 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 917
adam@322 918 Abort.
adam@322 919
adam@322 920 (** 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@55 921 (* end thide *)
adamc@209 922
adam@322 923
adamc@51 924
adamc@52 925
adamc@52 926 (* begin hide *)
adamc@52 927 (* In-class exercises *)
adamc@52 928
adam@292 929 (* 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 930
adamc@52 931 (* begin thide *)
adamc@52 932 Inductive prop : Set :=
adamc@52 933 | Tru : prop
adamc@52 934 | Fals : prop
adamc@52 935 | And : prop -> prop -> prop
adamc@52 936 | Or : prop -> prop -> prop.
adamc@52 937
adamc@52 938 Inductive holds : prop -> Prop :=
adamc@52 939 | HTru : holds Tru
adamc@52 940 | HAnd : forall p1 p2, holds p1 -> holds p2 -> holds (And p1 p2)
adamc@52 941 | HOr1 : forall p1 p2, holds p1 -> holds (Or p1 p2)
adamc@52 942 | HOr2 : forall p1 p2, holds p2 -> holds (Or p1 p2).
adamc@52 943
adamc@52 944 Inductive falseFree : prop -> Prop :=
adamc@52 945 | FFTru : falseFree Tru
adamc@52 946 | FFAnd : forall p1 p2, falseFree p1 -> falseFree p2 -> falseFree (And p1 p2)
adamc@52 947 | FFNot : forall p1 p2, falseFree p1 -> falseFree p2 -> falseFree (Or p1 p2).
adamc@52 948
adamc@52 949 Hint Constructors holds.
adamc@52 950
adamc@52 951 Theorem falseFree_holds : forall p, falseFree p -> holds p.
adamc@52 952 induction 1; crush.
adamc@52 953 Qed.
adamc@52 954 (* end thide *)
adamc@52 955
adamc@52 956
adamc@52 957 (* 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 958
adamc@52 959 (* begin thide *)
adamc@52 960 Inductive prop' : Set :=
adamc@52 961 | Tru' : prop'
adamc@52 962 | And' : prop' -> prop' -> prop'
adamc@52 963 | Or' : prop' -> prop' -> prop'.
adamc@52 964
adamc@52 965 Inductive holds' : prop' -> Prop :=
adamc@52 966 | HTru' : holds' Tru'
adamc@52 967 | HAnd' : forall p1 p2, holds' p1 -> holds' p2 -> holds' (And' p1 p2)
adamc@52 968 | HOr1' : forall p1 p2, holds' p1 -> holds' (Or' p1 p2)
adamc@52 969 | HOr2' : forall p1 p2, holds' p2 -> holds' (Or' p1 p2).
adamc@52 970
adamc@52 971 Fixpoint propify (p : prop') : prop :=
adamc@52 972 match p with
adamc@52 973 | Tru' => Tru
adamc@52 974 | And' p1 p2 => And (propify p1) (propify p2)
adamc@52 975 | Or' p1 p2 => Or (propify p1) (propify p2)
adamc@52 976 end.
adamc@52 977
adamc@52 978 Hint Constructors holds'.
adamc@52 979
adamc@52 980 Lemma propify_holds' : forall p', holds p' -> forall p, p' = propify p -> holds' p.
adamc@52 981 induction 1; crush; destruct p; crush.
adamc@52 982 Qed.
adamc@52 983
adamc@52 984 Theorem propify_holds : forall p, holds (propify p) -> holds' p.
adamc@52 985 intros; eapply propify_holds'; eauto.
adamc@52 986 Qed.
adamc@52 987 (* end thide *)
adamc@52 988
adamc@52 989 (* end hide *)