annotate src/Predicates.v @ 322:98e0d957df16

Pass through Chapter 4
author Adam Chlipala <adam@chlipala.net>
date Mon, 19 Sep 2011 14:04:09 -0400
parents d5787b70cf48
children 3513d8b0531a
rev   line source
adam@322 1 (* Copyright (c) 2008-2011, 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@322 19 Inductive unit := tt.
adam@322 20 Inductive Empty_set := .
adam@322 21 Inductive bool := true | false.
adam@322 22 Inductive sum := .
adam@322 23 Inductive prod := .
adam@322 24 Inductive and := conj.
adam@322 25 Inductive or := or_introl | or_intror.
adam@322 26 Inductive ex := ex_intro.
adam@322 27 Inductive eq := refl_equal.
adam@322 28 Reset unit.
adamc@45 29 (* end hide *)
adamc@45 30
adamc@45 31 (** %\chapter{Inductive Predicates}% *)
adamc@45 32
adam@322 33 (** 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 34
adam@322 35 (* begin hide *)
adamc@45 36 Print unit.
adam@322 37 (* end hide *)
adam@322 38 (** %\noindent%[Print] %\coqdocinductive{%#<tt>#unit#</tt>#%}%[.] *)
adam@322 39 (** [[
adamc@209 40 Inductive unit : Set := tt : unit
adam@302 41 ]]
adam@302 42 *)
adamc@45 43
adam@322 44 (* begin hide *)
adamc@45 45 Print True.
adam@322 46 (* end hide *)
adam@322 47 (** %\noindent%[Print] %\coqdocinductive{%#<tt>#True#</tt>#%}%[.] *)
adam@322 48 (** [[
adamc@209 49 Inductive True : Prop := I : True
adam@322 50 ]]
adam@302 51 *)
adamc@45 52
adam@322 53 (** 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 11 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 54
adam@322 55 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 %\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 56
adam@322 57 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}\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 58
adam@292 59 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 60
adamc@45 61
adamc@48 62 (** * Propositional Logic *)
adamc@45 63
adamc@45 64 (** 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 65
adamc@45 66 Section Propositional.
adamc@46 67 Variables P Q R : Prop.
adamc@45 68
adamc@45 69 (** 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 70
adamc@45 71 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 72
adamc@45 73 Theorem obvious : True.
adamc@55 74 (* begin thide *)
adamc@45 75 apply I.
adamc@55 76 (* end thide *)
adamc@45 77 Qed.
adamc@45 78
adamc@45 79 (** 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 80
adamc@55 81 (* begin thide *)
adamc@45 82 Theorem obvious' : True.
adam@322 83 (* begin hide *)
adamc@45 84 constructor.
adam@322 85 (* end hide *)
adam@322 86 (** %\hspace{.075in}\coqdockw{%#<tt>#constructor#</tt>#%}%.%\vspace{-.1in}% *)
adam@322 87
adamc@45 88 Qed.
adamc@45 89
adamc@55 90 (* end thide *)
adamc@55 91
adamc@45 92 (** There is also a predicate [False], which is the Curry-Howard mirror image of [Empty_set] from the last chapter. *)
adamc@45 93
adam@322 94 (* begin hide *)
adamc@45 95 Print False.
adam@322 96 (* end hide *)
adam@322 97 (** %\noindent%[Print] %\coqdocinductive{%#<tt>#False#</tt>#%}%[.] *)
adam@322 98 (** [[
adamc@209 99 Inductive False : Prop :=
adamc@209 100
adamc@209 101 ]]
adamc@45 102
adamc@209 103 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 104
adamc@45 105 Theorem False_imp : False -> 2 + 2 = 5.
adamc@55 106 (* begin thide *)
adamc@45 107 destruct 1.
adamc@55 108 (* end thide *)
adamc@45 109 Qed.
adamc@45 110
adamc@45 111 (** 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 112
adamc@45 113 Theorem arith_neq : 2 + 2 = 5 -> 9 + 9 = 835.
adamc@55 114 (* begin thide *)
adamc@45 115 intro.
adamc@45 116
adam@322 117 (** 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 118
adamc@45 119 elimtype False.
adamc@45 120 (** [[
adamc@45 121 H : 2 + 2 = 5
adamc@45 122 ============================
adamc@45 123 False
adamc@209 124
adamc@209 125 ]]
adamc@45 126
adamc@209 127 For now, we will leave the details of this proof about arithmetic to [crush]. *)
adamc@45 128
adamc@45 129 crush.
adamc@55 130 (* end thide *)
adamc@45 131 Qed.
adamc@45 132
adamc@45 133 (** A related notion to [False] is logical negation. *)
adamc@45 134
adamc@45 135 Print not.
adamc@209 136 (** %\vspace{-.15in}% [[
adamc@209 137 not = fun A : Prop => A -> False
adamc@209 138 : Prop -> Prop
adamc@209 139
adamc@209 140 ]]
adamc@45 141
adam@280 142 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 143
adamc@45 144 Theorem arith_neq' : ~ (2 + 2 = 5).
adamc@55 145 (* begin thide *)
adamc@45 146 unfold not.
adamc@45 147 (** [[
adamc@45 148 ============================
adamc@45 149 2 + 2 = 5 -> False
adam@302 150 ]]
adam@302 151 *)
adamc@45 152
adamc@45 153 crush.
adamc@55 154 (* end thide *)
adamc@45 155 Qed.
adamc@45 156
adamc@45 157 (** We also have conjunction, which we introduced in the last chapter. *)
adamc@45 158
adam@322 159 (* begin hide *)
adamc@45 160 Print and.
adam@322 161 (* end hide *)
adam@322 162 (** %\noindent%[Print] %\coqdocinductive{%#<tt>#and#</tt>#%}%[.]
adam@322 163 [[
adam@322 164 Inductive and (A : Prop) (B : Prop) : Prop := conj : A -> B -> A /\ B
adamc@209 165
adamc@209 166 ]]
adamc@209 167
adam@322 168 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 169
adamc@45 170 Theorem and_comm : P /\ Q -> Q /\ P.
adamc@209 171
adamc@55 172 (* begin thide *)
adamc@45 173 (** We start by case analysis on the proof of [P /\ Q]. *)
adamc@45 174
adamc@45 175 destruct 1.
adamc@45 176 (** [[
adamc@45 177 H : P
adamc@45 178 H0 : Q
adamc@45 179 ============================
adamc@45 180 Q /\ P
adamc@209 181
adamc@209 182 ]]
adamc@45 183
adam@322 184 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 185
adamc@45 186 split.
adam@322 187 (** %\vspace{.1in} \noindent 2 \coqdockw{subgoals}\vspace{-.1in}%#<tt>2 subgoals</tt>#
adam@322 188 [[
adamc@45 189
adamc@45 190 H : P
adamc@45 191 H0 : Q
adamc@45 192 ============================
adamc@45 193 Q
adam@322 194 ]]
adam@322 195 %\noindent \coqdockw{subgoal} 2 \coqdockw{is}:%#<tt>subgoal 2 is</tt>#
adam@322 196 [[
adam@322 197 P
adamc@209 198
adamc@209 199 ]]
adamc@45 200
adam@322 201 In each case, the conclusion is among our hypotheses, so the %\index{tactics!assumption}%[assumption] tactic finishes the process. *)
adamc@45 202
adamc@45 203 assumption.
adamc@45 204 assumption.
adamc@55 205 (* end thide *)
adamc@45 206 Qed.
adamc@45 207
adam@322 208 (** Coq disjunction is called %\index{Gallina terms!or}%[or] and abbreviated with the infix operator [\/]. *)
adamc@45 209
adam@322 210 (* begin hide *)
adamc@45 211 Print or.
adam@322 212 (* end hide *)
adam@322 213 (** %\noindent%[Print] %\coqdocinductive{%#<tt>#or#</tt>#%}%[.]
adam@322 214 [[
adamc@209 215 Inductive or (A : Prop) (B : Prop) : Prop :=
adamc@209 216 or_introl : A -> A \/ B | or_intror : B -> A \/ B
adamc@209 217
adamc@209 218 ]]
adamc@45 219
adam@322 220 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 221
adamc@45 222 Theorem or_comm : P \/ Q -> Q \/ P.
adamc@55 223
adamc@55 224 (* begin thide *)
adamc@45 225 (** 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 226
adamc@45 227 destruct 1.
adam@322 228 (** %\vspace{.1in} \noindent 2 \coqdockw{subgoals}\vspace{-.1in}%#<tt>2 subgoals</tt>#
adam@322 229 [[
adamc@45 230
adamc@45 231 H : P
adamc@45 232 ============================
adamc@45 233 Q \/ P
adam@322 234 ]]
adam@322 235 %\noindent \coqdockw{subgoal} 2 \coqdockw{is}:%#<tt>subgoal 2 is</tt>#
adam@322 236 [[
adamc@45 237 Q \/ P
adamc@209 238
adamc@209 239 ]]
adamc@45 240
adam@322 241 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 242
adam@322 243 (* begin hide *)
adamc@45 244 right; assumption.
adam@322 245 (* end hide *)
adam@322 246 (** %\hspace{.075in}\coqdockw{%#<tt>#right#</tt>#%}%[; assumption.] *)
adamc@45 247
adam@322 248 (** The second subgoal has a symmetric proof.%\index{tactics!left}%
adamc@45 249
adamc@45 250 [[
adamc@45 251 1 subgoal
adamc@45 252
adamc@45 253 H : Q
adamc@45 254 ============================
adamc@45 255 Q \/ P
adam@302 256 ]]
adam@302 257 *)
adamc@45 258
adam@322 259 (* begin hide *)
adamc@45 260 left; assumption.
adam@322 261 (* end hide *)
adam@322 262 (** %\hspace{.075in}\coqdockw{%#<tt>#left#</tt>#%}%[; assumption.] *)
adam@322 263
adamc@55 264 (* end thide *)
adamc@45 265 Qed.
adamc@45 266
adamc@46 267
adamc@46 268 (* begin hide *)
adamc@46 269 (* In-class exercises *)
adamc@46 270
adamc@46 271 Theorem contra : P -> ~P -> R.
adamc@52 272 (* begin thide *)
adamc@52 273 unfold not.
adamc@52 274 intros.
adamc@52 275 elimtype False.
adamc@52 276 apply H0.
adamc@52 277 assumption.
adamc@52 278 (* end thide *)
adamc@46 279 Admitted.
adamc@46 280
adamc@46 281 Theorem and_assoc : (P /\ Q) /\ R -> P /\ (Q /\ R).
adamc@52 282 (* begin thide *)
adamc@52 283 intros.
adamc@52 284 destruct H.
adamc@52 285 destruct H.
adamc@52 286 split.
adamc@52 287 assumption.
adamc@52 288 split.
adamc@52 289 assumption.
adamc@52 290 assumption.
adamc@52 291 (* end thide *)
adamc@46 292 Admitted.
adamc@46 293
adamc@46 294 Theorem or_assoc : (P \/ Q) \/ R -> P \/ (Q \/ R).
adamc@52 295 (* begin thide *)
adamc@52 296 intros.
adamc@52 297 destruct H.
adamc@52 298 destruct H.
adamc@52 299 left.
adamc@52 300 assumption.
adamc@52 301 right.
adamc@52 302 left.
adamc@52 303 assumption.
adamc@52 304 right.
adamc@52 305 right.
adamc@52 306 assumption.
adamc@52 307 (* end thide *)
adamc@46 308 Admitted.
adamc@46 309
adamc@46 310 (* end hide *)
adamc@46 311
adamc@46 312
adam@322 313 (** 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 314
adamc@46 315 Theorem or_comm' : P \/ Q -> Q \/ P.
adamc@55 316 (* begin thide *)
adamc@46 317 tauto.
adamc@55 318 (* end thide *)
adamc@46 319 Qed.
adamc@46 320
adam@322 321 (** 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 322
adamc@46 323 Theorem arith_comm : forall ls1 ls2 : list nat,
adamc@46 324 length ls1 = length ls2 \/ length ls1 + length ls2 = 6
adamc@46 325 -> length (ls1 ++ ls2) = 6 \/ length ls1 = length ls2.
adamc@55 326 (* begin thide *)
adamc@46 327 intuition.
adamc@46 328
adamc@46 329 (** 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 330
adamc@46 331 (** [[
adamc@46 332 ls1 : list nat
adamc@46 333 ls2 : list nat
adamc@46 334 H0 : length ls1 + length ls2 = 6
adamc@46 335 ============================
adamc@46 336 length (ls1 ++ ls2) = 6 \/ length ls1 = length ls2
adamc@209 337
adamc@209 338 ]]
adamc@46 339
adamc@209 340 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 341
adamc@46 342 rewrite app_length.
adamc@46 343 (** [[
adamc@46 344 ls1 : list nat
adamc@46 345 ls2 : list nat
adamc@46 346 H0 : length ls1 + length ls2 = 6
adamc@46 347 ============================
adamc@46 348 length ls1 + length ls2 = 6 \/ length ls1 = length ls2
adamc@209 349
adamc@209 350 ]]
adamc@46 351
adamc@209 352 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 353
adamc@46 354 tauto.
adamc@55 355 (* end thide *)
adamc@46 356 Qed.
adamc@46 357
adam@322 358 (** 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 359
adamc@55 360 (* begin thide *)
adamc@46 361 Theorem arith_comm' : forall ls1 ls2 : list nat,
adamc@46 362 length ls1 = length ls2 \/ length ls1 + length ls2 = 6
adamc@46 363 -> length (ls1 ++ ls2) = 6 \/ length ls1 = length ls2.
adam@322 364 (* begin hide *)
adamc@46 365 Hint Rewrite app_length : cpdt.
adam@322 366 (* end hide *)
adam@322 367 (** %\hspace{.075in}%[Hint] %\coqdockw{%#<tt>#Rewrite#</tt>#%}% [app_length : cpdt.] *)
adamc@46 368
adamc@46 369 crush.
adamc@46 370 Qed.
adamc@55 371 (* end thide *)
adamc@46 372
adamc@45 373 End Propositional.
adamc@45 374
adam@322 375 (** 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 376
adamc@46 377
adamc@47 378 (** * What Does It Mean to Be Constructive? *)
adamc@46 379
adamc@47 380 (** 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 381
adam@322 382 The answer comes from the fact that Coq implements %\index{constructive logic}\textit{%#<i>#constructive#</i>#%}% or %\index{intuitionistic logic|see{constructive logic}}\textit{%#<i>#intuitionistic#</i>#%}% logic, in contrast to the %\index{classical logic}\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 %\index{decidability}\textit{%#<i>#decidable#</i>#%}%, 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 383
adam@292 384 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 385
adam@322 386 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}\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 387
adamc@47 388 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 389
adamc@48 390
adamc@48 391 (** * First-Order Logic *)
adamc@48 392
adam@322 393 (** 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 394
adam@322 395 %\index{existential quantification}\index{Gallina terms!exists}\index{Gallina terms!ex}%Existential quantification is defined in the standard library. *)
adamc@48 396
adam@322 397 (* begin hide *)
adam@322 398 Print ex.
adam@322 399 (* end hide *)
adam@322 400 (** %\noindent%[Print] %\coqdocinductive{%#<tt>#ex#</tt>#%}%[.]
adam@322 401 [[
adamc@209 402 Inductive ex (A : Type) (P : A -> Prop) : Prop :=
adamc@209 403 ex_intro : forall x : A, P x -> ex P
adamc@209 404
adamc@209 405 ]]
adamc@48 406
adam@322 407 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 408
adamc@48 409 Theorem exist1 : exists x : nat, x + 1 = 2.
adamc@55 410 (* begin thide *)
adamc@67 411 (** remove printing exists *)
adam@322 412 (** 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 413
adam@322 414 (* begin hide *)
adamc@48 415 exists 1.
adam@322 416 (* end hide *)
adam@322 417 (** %\coqdockw{%#<tt>#exists#</tt>#%}% [1.] *)
adamc@48 418
adamc@209 419 (** The conclusion is replaced with a version using the existential witness that we announced.
adamc@48 420
adamc@209 421 [[
adamc@48 422 ============================
adamc@48 423 1 + 1 = 2
adam@302 424 ]]
adam@302 425 *)
adamc@48 426
adamc@48 427 reflexivity.
adamc@55 428 (* end thide *)
adamc@48 429 Qed.
adamc@48 430
adamc@48 431 (** printing exists $\exists$ *)
adamc@48 432
adamc@48 433 (** We can also use tactics to reason about existential hypotheses. *)
adamc@48 434
adamc@48 435 Theorem exist2 : forall n m : nat, (exists x : nat, n + x = m) -> n <= m.
adamc@55 436 (* begin thide *)
adamc@48 437 (** We start by case analysis on the proof of the existential fact. *)
adamc@209 438
adamc@48 439 destruct 1.
adamc@48 440 (** [[
adamc@48 441 n : nat
adamc@48 442 m : nat
adamc@48 443 x : nat
adamc@48 444 H : n + x = m
adamc@48 445 ============================
adamc@48 446 n <= m
adamc@209 447
adamc@209 448 ]]
adamc@48 449
adamc@209 450 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 451
adamc@48 452 crush.
adamc@55 453 (* end thide *)
adamc@48 454 Qed.
adamc@48 455
adamc@48 456
adamc@48 457 (* begin hide *)
adamc@48 458 (* In-class exercises *)
adamc@48 459
adamc@48 460 Theorem forall_exists_commute : forall (A B : Type) (P : A -> B -> Prop),
adamc@48 461 (exists x : A, forall y : B, P x y) -> (forall y : B, exists x : A, P x y).
adamc@52 462 (* begin thide *)
adamc@52 463 intros.
adamc@52 464 destruct H.
adamc@52 465 exists x.
adamc@52 466 apply H.
adamc@52 467 (* end thide *)
adamc@48 468 Admitted.
adamc@48 469
adamc@48 470 (* end hide *)
adamc@48 471
adamc@48 472
adam@322 473 (** 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 474
adamc@49 475
adamc@49 476 (** * Predicates with Implicit Equality *)
adamc@49 477
adamc@49 478 (** 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 479
adamc@49 480 Inductive isZero : nat -> Prop :=
adamc@49 481 | IsZero : isZero 0.
adamc@49 482
adamc@49 483 Theorem isZero_zero : isZero 0.
adamc@55 484 (* begin thide *)
adam@322 485 (* begin hide *)
adamc@49 486 constructor.
adam@322 487 (* end hide *)
adam@322 488 (** %\coqdockw{%#<tt>#constructor#</tt>#%}%[.]%\vspace{-.075in}% *)
adam@322 489
adamc@55 490 (* end thide *)
adamc@49 491 Qed.
adamc@49 492
adam@322 493 (** We can call [isZero] a %\index{judgment}\textit{%#<i>#judgment#</i>#%}%, in the sense often used in the semantics of programming languages. Judgments are typically defined in the style of %\index{natural deduction}\textit{%#<i>#natural deduction#</i>#%}%, where we write a number of %\index{inference rules}\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 494
adamc@49 495 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 496
adam@322 497 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 498
adam@322 499 (* begin hide *)
adamc@49 500 Print eq.
adam@322 501 (* end hide *)
adam@322 502 (** %\noindent%[Print] %\coqdocinductive{%#<tt>#eq#</tt>#%}%[.]
adam@322 503 [[
adamc@209 504 Inductive eq (A : Type) (x : A) : A -> Prop := refl_equal : x = x
adamc@209 505
adamc@209 506 ]]
adamc@49 507
adam@322 508 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 %\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. Another way of stating that definition is: equality is defined as the least reflexive relation.
adamc@49 509
adam@322 510 Returning to the example of [isZero], we can see how to work with hypotheses that use this predicate. *)
adamc@49 511
adamc@49 512 Theorem isZero_plus : forall n m : nat, isZero m -> n + m = n.
adamc@55 513 (* begin thide *)
adamc@49 514 (** We want to proceed by cases on the proof of the assumption about [isZero]. *)
adamc@209 515
adamc@49 516 destruct 1.
adamc@49 517 (** [[
adamc@49 518 n : nat
adamc@49 519 ============================
adamc@49 520 n + 0 = n
adamc@209 521
adamc@209 522 ]]
adamc@49 523
adamc@209 524 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 525
adamc@49 526 crush.
adamc@55 527 (* end thide *)
adamc@49 528 Qed.
adamc@49 529
adamc@49 530 (** 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 531
adamc@49 532 Theorem isZero_contra : isZero 1 -> False.
adamc@55 533 (* begin thide *)
adamc@49 534 (** Let us try a proof by cases on the assumption, as in the last proof. *)
adamc@209 535
adamc@49 536 destruct 1.
adamc@49 537 (** [[
adamc@49 538 ============================
adamc@49 539 False
adamc@209 540
adamc@209 541 ]]
adamc@49 542
adamc@209 543 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 544
adam@322 545 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 546
adamc@49 547 Undo.
adamc@49 548 inversion 1.
adamc@55 549 (* end thide *)
adamc@49 550 Qed.
adamc@49 551
adamc@49 552 (** 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 553
adamc@49 554 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 555
adamc@49 556 Theorem isZero_contra' : isZero 1 -> 2 + 2 = 5.
adamc@49 557 destruct 1.
adamc@49 558 (** [[
adamc@49 559 ============================
adamc@49 560 1 + 1 = 4
adamc@209 561
adamc@209 562 ]]
adamc@49 563
adam@280 564 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 565
adamc@49 566 Abort.
adamc@49 567
adam@280 568 (** To see more clearly what is happening, we can consider the type of [isZero]'s induction principle. *)
adam@280 569
adam@280 570 Check isZero_ind.
adam@280 571 (** %\vspace{-.15in}% [[
adam@280 572 isZero_ind
adam@280 573 : forall P : nat -> Prop, P 0 -> forall n : nat, isZero n -> P n
adam@280 574
adam@280 575 ]]
adam@280 576
adam@322 577 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 578
adamc@49 579
adamc@49 580 (* begin hide *)
adamc@49 581 (* In-class exercises *)
adamc@49 582
adamc@49 583 (* 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 584
adamc@52 585 (* begin thide *)
adamc@52 586 Section twoEls.
adamc@52 587 Variable A : Type.
adamc@52 588
adamc@52 589 Inductive twoEls : list A -> Prop :=
adamc@52 590 | TwoEls : forall x y, twoEls (x :: y :: nil).
adamc@52 591
adamc@52 592 Theorem twoEls_nil : twoEls nil -> False.
adamc@52 593 inversion 1.
adamc@52 594 Qed.
adamc@52 595
adamc@52 596 Theorem twoEls_two : forall ls, twoEls ls -> length ls = 2.
adamc@52 597 inversion 1.
adamc@52 598 reflexivity.
adamc@52 599 Qed.
adamc@52 600 End twoEls.
adamc@52 601 (* end thide *)
adamc@52 602
adamc@49 603 (* end hide *)
adamc@49 604
adamc@50 605
adamc@50 606 (** * Recursive Predicates *)
adamc@50 607
adamc@50 608 (** We have already seen all of the ingredients we need to build interesting recursive predicates, like this predicate capturing even-ness. *)
adamc@50 609
adamc@50 610 Inductive even : nat -> Prop :=
adamc@50 611 | EvenO : even O
adamc@50 612 | EvenSS : forall n, even n -> even (S (S n)).
adamc@50 613
adam@322 614 (** 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 615
adamc@50 616 The proof techniques of the last section are easily adapted. *)
adamc@50 617
adamc@50 618 Theorem even_0 : even 0.
adamc@55 619 (* begin thide *)
adam@322 620 (* begin hide *)
adamc@50 621 constructor.
adam@322 622 (* end hide *)
adam@322 623 (** %\coqdockw{%#<tt>#constructor#</tt>#%}%[.]%\vspace{-.075in}% *)
adam@322 624
adamc@55 625 (* end thide *)
adamc@50 626 Qed.
adamc@50 627
adamc@50 628 Theorem even_4 : even 4.
adamc@55 629 (* begin thide *)
adam@322 630 (* begin hide *)
adamc@50 631 constructor; constructor; constructor.
adam@322 632 (* end hide *)
adam@322 633 (** %\coqdockw{%#<tt>#constructor#</tt>#%}%[; ]%\coqdockw{%#<tt>#constructor#</tt>#%}%[; ]%\coqdockw{%#<tt>#constructor#</tt>#%}%[.]%\vspace{-.075in}% *)
adam@322 634
adamc@55 635 (* end thide *)
adamc@50 636 Qed.
adamc@50 637
adam@322 638 (** 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. Note that this sort of hint may be placed in a default database, such that a command has no equivalent to the [: cpdt] from our earlier rewrite hints.%\index{Hint Constructirs}% 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 639
adamc@55 640 (* begin thide *)
adam@322 641 (* begin hide *)
adamc@50 642 Hint Constructors even.
adam@322 643 (* end hide *)
adam@322 644 (** %\noindent%[Hint] %\coqdockw{%#<tt>#Constructors#</tt>#%}% [even.] *)
adamc@50 645
adamc@50 646 Theorem even_4' : even 4.
adamc@50 647 auto.
adamc@50 648 Qed.
adamc@50 649
adamc@55 650 (* end thide *)
adamc@55 651
adam@322 652 (** We may also use [inversion] with [even]. *)
adam@322 653
adamc@50 654 Theorem even_1_contra : even 1 -> False.
adamc@55 655 (* begin thide *)
adamc@50 656 inversion 1.
adamc@55 657 (* end thide *)
adamc@50 658 Qed.
adamc@50 659
adamc@50 660 Theorem even_3_contra : even 3 -> False.
adamc@55 661 (* begin thide *)
adamc@50 662 inversion 1.
adamc@50 663 (** [[
adamc@50 664 H : even 3
adamc@50 665 n : nat
adamc@50 666 H1 : even 1
adamc@50 667 H0 : n = 1
adamc@50 668 ============================
adamc@50 669 False
adamc@209 670
adamc@209 671 ]]
adamc@50 672
adam@322 673 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 674
adamc@50 675 inversion H1.
adamc@55 676 (* end thide *)
adamc@50 677 Qed.
adamc@50 678
adamc@50 679 (** We can also do inductive proofs about [even]. *)
adamc@50 680
adamc@50 681 Theorem even_plus : forall n m, even n -> even m -> even (n + m).
adamc@55 682 (* begin thide *)
adamc@50 683 (** It seems a reasonable first choice to proceed by induction on [n]. *)
adamc@209 684
adamc@50 685 induction n; crush.
adamc@50 686 (** [[
adamc@50 687 n : nat
adamc@50 688 IHn : forall m : nat, even n -> even m -> even (n + m)
adamc@50 689 m : nat
adamc@50 690 H : even (S n)
adamc@50 691 H0 : even m
adamc@50 692 ============================
adamc@50 693 even (S (n + m))
adamc@209 694
adamc@209 695 ]]
adamc@50 696
adamc@209 697 We will need to use the hypotheses [H] and [H0] somehow. The most natural choice is to invert [H]. *)
adamc@50 698
adamc@50 699 inversion H.
adamc@50 700 (** [[
adamc@50 701 n : nat
adamc@50 702 IHn : forall m : nat, even n -> even m -> even (n + m)
adamc@50 703 m : nat
adamc@50 704 H : even (S n)
adamc@50 705 H0 : even m
adamc@50 706 n0 : nat
adamc@50 707 H2 : even n0
adamc@50 708 H1 : S n0 = n
adamc@50 709 ============================
adamc@50 710 even (S (S n0 + m))
adamc@209 711
adamc@209 712 ]]
adamc@50 713
adamc@209 714 Simplifying the conclusion brings us to a point where we can apply a constructor. *)
adamc@209 715
adamc@50 716 simpl.
adamc@50 717 (** [[
adamc@50 718 ============================
adamc@50 719 even (S (S (n0 + m)))
adam@302 720 ]]
adam@302 721 *)
adamc@50 722
adam@322 723 (* begin hide *)
adamc@50 724 constructor.
adam@322 725 (* end hide *)
adam@322 726 (** %\coqdockw{%#<tt>#constructor#</tt>#%}%[.]
adam@322 727
adam@322 728 [[
adamc@50 729 ============================
adamc@50 730 even (n0 + m)
adamc@209 731
adamc@209 732 ]]
adamc@50 733
adamc@209 734 At this point, we would like to apply the inductive hypothesis, which is:
adamc@209 735
adamc@209 736 [[
adamc@50 737 IHn : forall m : nat, even n -> even m -> even (n + m)
adamc@209 738 ]]
adamc@50 739
adam@322 740 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 %\index{rule induction}\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 741
adam@322 742 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 743
adam@322 744 (* begin hide *)
adamc@50 745 Restart.
adam@322 746 (* end hide *)
adam@322 747 (** %\noindent\coqdockw{%#<tt>#Restart#</tt>#%}%[.] *)
adamc@50 748
adamc@50 749 induction 1.
adamc@50 750 (** [[
adamc@50 751 m : nat
adamc@50 752 ============================
adamc@50 753 even m -> even (0 + m)
adam@322 754 ]]
adamc@50 755
adam@322 756 %\noindent \coqdockw{subgoal} 2 \coqdockw{is}:%#<tt>subgoal 2 is</tt>#
adam@322 757 [[
adamc@50 758 even m -> even (S (S n) + m)
adamc@209 759
adamc@209 760 ]]
adamc@50 761
adamc@209 762 The first case is easily discharged by [crush], based on the hint we added earlier to try the constructors of [even]. *)
adamc@50 763
adamc@50 764 crush.
adamc@50 765
adamc@50 766 (** Now we focus on the second case: *)
adamc@209 767
adamc@50 768 intro.
adamc@50 769 (** [[
adamc@50 770 m : nat
adamc@50 771 n : nat
adamc@50 772 H : even n
adamc@50 773 IHeven : even m -> even (n + m)
adamc@50 774 H0 : even m
adamc@50 775 ============================
adamc@50 776 even (S (S n) + m)
adamc@209 777
adamc@209 778 ]]
adamc@50 779
adamc@209 780 We simplify and apply a constructor, as in our last proof attempt. *)
adamc@50 781
adam@322 782 (* begin hide *)
adamc@50 783 simpl; constructor.
adam@322 784 (* end hide *)
adam@322 785 (** [simpl; ]%\coqdockw{%#<tt>#constructor#</tt>#%}%[.]
adam@322 786
adam@322 787 [[
adamc@50 788 ============================
adamc@50 789 even (n + m)
adamc@209 790
adamc@209 791 ]]
adamc@50 792
adamc@209 793 Now we have an exact match with our inductive hypothesis, and the remainder of the proof is trivial. *)
adamc@50 794
adamc@50 795 apply IHeven; assumption.
adamc@50 796
adamc@50 797 (** In fact, [crush] can handle all of the details of the proof once we declare the induction strategy. *)
adamc@50 798
adam@322 799 (* begin hide *)
adamc@50 800 Restart.
adam@322 801 (* end hide *)
adam@322 802 (** %\noindent\coqdockw{%#<tt>#Restart#</tt>#%}%[.] *)
adam@322 803
adamc@50 804 induction 1; crush.
adamc@55 805 (* end thide *)
adamc@50 806 Qed.
adamc@50 807
adamc@50 808 (** Induction on recursive predicates has similar pitfalls to those we encountered with inversion in the last section. *)
adamc@50 809
adamc@50 810 Theorem even_contra : forall n, even (S (n + n)) -> False.
adamc@55 811 (* begin thide *)
adamc@50 812 induction 1.
adamc@50 813 (** [[
adamc@50 814 n : nat
adamc@50 815 ============================
adamc@50 816 False
adam@322 817 ]]
adamc@50 818
adam@322 819 %\noindent \coqdockw{subgoal} 2 \coqdockw{is}:%#<tt>subgoal 2 is</tt>#
adam@322 820 [[
adamc@50 821 False
adamc@209 822
adamc@209 823 ]]
adamc@50 824
adam@280 825 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 826
adamc@50 827 Abort.
adamc@50 828
adamc@50 829 Lemma even_contra' : forall n', even n' -> forall n, n' = S (n + n) -> False.
adamc@50 830 induction 1; crush.
adamc@50 831
adamc@54 832 (** 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 833
adamc@50 834 destruct n; destruct n0; crush.
adamc@50 835
adamc@50 836 (** [[
adamc@50 837 n : nat
adamc@50 838 H : even (S n)
adamc@50 839 IHeven : forall n0 : nat, S n = S (n0 + n0) -> False
adamc@50 840 n0 : nat
adamc@50 841 H0 : S n = n0 + S n0
adamc@50 842 ============================
adamc@50 843 False
adamc@209 844
adamc@209 845 ]]
adamc@50 846
adam@280 847 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 848
adam@322 849 (* begin hide *)
adam@280 850 SearchRewrite (_ + S _).
adam@322 851 (* end hide *)
adam@322 852 (** %\coqdockw{%#<tt>#SearchRewrite#</tt>#%}% [(_ + S _).]
adam@322 853
adam@322 854 [[
adam@280 855 plus_n_Sm : forall n m : nat, S (n + m) = n + S m
adam@302 856 ]]
adam@302 857 *)
adamc@50 858
adamc@50 859 rewrite <- plus_n_Sm in H0.
adamc@50 860
adam@322 861 (** 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 862
adamc@50 863 apply IHeven with n0; assumption.
adamc@50 864
adam@322 865 (** 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 866
adam@322 867 (* begin hide *)
adamc@209 868 Restart.
adam@322 869 (* end hide *)
adam@322 870 (** %\hspace{-.075in}\coqdockw{%#<tt>#Restart#</tt>#%}%[.] *)
adam@322 871
adam@322 872 (* begin hide *)
adamc@50 873 Hint Rewrite <- plus_n_Sm : cpdt.
adam@322 874 (* end hide *)
adam@322 875 (** %\hspace{-.075in}%[Hint] %\noindent\coqdockw{%#<tt>#Rewrite#</tt>#%}% [<- plus_n_sm : cpdt.] *)
adamc@50 876
adamc@50 877 induction 1; crush;
adamc@50 878 match goal with
adamc@50 879 | [ H : S ?N = ?N0 + ?N0 |- _ ] => destruct N; destruct N0
adamc@50 880 end; crush; eauto.
adamc@50 881 Qed.
adamc@50 882
adam@322 883 (** 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 884
adam@322 885 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 886
adamc@50 887 The original theorem now follows trivially from our lemma. *)
adamc@50 888
adamc@50 889 Theorem even_contra : forall n, even (S (n + n)) -> False.
adamc@52 890 intros; eapply even_contra'; eauto.
adamc@50 891 Qed.
adamc@52 892
adam@322 893 (** 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.
adamc@52 894
adamc@52 895 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 896
adamc@52 897 Lemma even_contra'' : forall n' n, even n' -> n' = S (n + n) -> False.
adamc@52 898 induction 1; crush;
adamc@52 899 match goal with
adamc@52 900 | [ H : S ?N = ?N0 + ?N0 |- _ ] => destruct N; destruct N0
adamc@52 901 end; crush; eauto.
adamc@52 902
adamc@209 903 (** One subgoal remains:
adamc@52 904
adamc@209 905 [[
adamc@52 906 n : nat
adamc@52 907 H : even (S (n + n))
adamc@52 908 IHeven : S (n + n) = S (S (S (n + n))) -> False
adamc@52 909 ============================
adamc@52 910 False
adamc@209 911
adamc@209 912 ]]
adamc@52 913
adam@322 914 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 915
adam@322 916 Abort.
adam@322 917
adam@322 918 (** 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 919 (* end thide *)
adamc@209 920
adam@322 921
adamc@51 922
adamc@52 923
adamc@52 924 (* begin hide *)
adamc@52 925 (* In-class exercises *)
adamc@52 926
adam@292 927 (* 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 928
adamc@52 929 (* begin thide *)
adamc@52 930 Inductive prop : Set :=
adamc@52 931 | Tru : prop
adamc@52 932 | Fals : prop
adamc@52 933 | And : prop -> prop -> prop
adamc@52 934 | Or : prop -> prop -> prop.
adamc@52 935
adamc@52 936 Inductive holds : prop -> Prop :=
adamc@52 937 | HTru : holds Tru
adamc@52 938 | HAnd : forall p1 p2, holds p1 -> holds p2 -> holds (And p1 p2)
adamc@52 939 | HOr1 : forall p1 p2, holds p1 -> holds (Or p1 p2)
adamc@52 940 | HOr2 : forall p1 p2, holds p2 -> holds (Or p1 p2).
adamc@52 941
adamc@52 942 Inductive falseFree : prop -> Prop :=
adamc@52 943 | FFTru : falseFree Tru
adamc@52 944 | FFAnd : forall p1 p2, falseFree p1 -> falseFree p2 -> falseFree (And p1 p2)
adamc@52 945 | FFNot : forall p1 p2, falseFree p1 -> falseFree p2 -> falseFree (Or p1 p2).
adamc@52 946
adamc@52 947 Hint Constructors holds.
adamc@52 948
adamc@52 949 Theorem falseFree_holds : forall p, falseFree p -> holds p.
adamc@52 950 induction 1; crush.
adamc@52 951 Qed.
adamc@52 952 (* end thide *)
adamc@52 953
adamc@52 954
adamc@52 955 (* 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 956
adamc@52 957 (* begin thide *)
adamc@52 958 Inductive prop' : Set :=
adamc@52 959 | Tru' : prop'
adamc@52 960 | And' : prop' -> prop' -> prop'
adamc@52 961 | Or' : prop' -> prop' -> prop'.
adamc@52 962
adamc@52 963 Inductive holds' : prop' -> Prop :=
adamc@52 964 | HTru' : holds' Tru'
adamc@52 965 | HAnd' : forall p1 p2, holds' p1 -> holds' p2 -> holds' (And' p1 p2)
adamc@52 966 | HOr1' : forall p1 p2, holds' p1 -> holds' (Or' p1 p2)
adamc@52 967 | HOr2' : forall p1 p2, holds' p2 -> holds' (Or' p1 p2).
adamc@52 968
adamc@52 969 Fixpoint propify (p : prop') : prop :=
adamc@52 970 match p with
adamc@52 971 | Tru' => Tru
adamc@52 972 | And' p1 p2 => And (propify p1) (propify p2)
adamc@52 973 | Or' p1 p2 => Or (propify p1) (propify p2)
adamc@52 974 end.
adamc@52 975
adamc@52 976 Hint Constructors holds'.
adamc@52 977
adamc@52 978 Lemma propify_holds' : forall p', holds p' -> forall p, p' = propify p -> holds' p.
adamc@52 979 induction 1; crush; destruct p; crush.
adamc@52 980 Qed.
adamc@52 981
adamc@52 982 Theorem propify_holds : forall p, holds (propify p) -> holds' p.
adamc@52 983 intros; eapply propify_holds'; eauto.
adamc@52 984 Qed.
adamc@52 985 (* end thide *)
adamc@52 986
adamc@52 987 (* end hide *)
adamc@58 988
adamc@58 989
adamc@58 990 (** * Exercises *)
adamc@58 991
adamc@58 992 (** %\begin{enumerate}%#<ol>#
adamc@58 993
adam@322 994 %\item%#<li># Prove these tautologies of propositional logic, using only the tactics [apply], [assumption], %\coqdockw{%#<tt>#constructor#</tt>#%}%, [destruct], [intro], [intros], %\coqdockw{%#<tt>#left#</tt>#%}%, %\coqdockw{%#<tt>#right#</tt>#%}%, [split], and [unfold].
adamc@58 995 %\begin{enumerate}%#<ol>#
adam@322 996 %\item%#<li># [(][True \/ False) /\ (][False \/ True)]#</li>#
adamc@209 997 %\item%#<li># [P -> ~ ~ P]#</li>#
adam@322 998 %\item%#<li># [P /\ (][Q \/ R) -> (][P /\ Q) \/ (][P /\ R)]#</li>#
adamc@61 999 #</ol> </li>#%\end{enumerate}%
adamc@58 1000
adam@322 1001 %\item%#<li># Prove the following tautology of first-order logic, using only the tactics [apply], [assert], [assumption], [destruct], [eapply], %\coqdockw{%#<tt>#eassumption#</tt>#%}%, and %\coqdockw{%#<tt>#exists#</tt>#%}%. You will probably find the [assert] tactic useful for stating and proving an intermediate lemma, enabling a kind of %``%#"#forward reasoning,#"#%''% in contrast to the %``%#"#backward reasoning#"#%''% that is the default for Coq tactics. The tactic %\coqdockw{%#<tt>#eassumption#</tt>#%}% is a version of [assumption] that will do matching of unification variables. Let some variable [T] of type [Set] be the set of individuals. [x] is a constant symbol, [p] is a unary predicate symbol, [q] is a binary predicate symbol, and [f] is a unary function symbol.
adamc@61 1002 %\begin{enumerate}%#<ol>#
adam@322 1003 %\item%#<li># [p x -> (][forall x, p x -> exists y, q x y) -> (][forall x y, q x y -> q y (f y)) -> exists z, q z (f z)]#</li>#
adamc@58 1004 #</ol> </li>#%\end{enumerate}%
adamc@58 1005
adam@292 1006 %\item%#<li># Define an inductive predicate capturing when a natural number is an integer multiple of either 6 or 10. Prove that 13 does not satisfy your predicate, and prove that any number satisfying the predicate is not odd. It is probably easiest to prove the second theorem by indicating %``%#"#odd-ness#"#%''% as equality to [2 * n + 1] for some [n].#</li>#
adamc@59 1007
adam@322 1008 %\item%#<li># $\star$ Define a simple programming language, its semantics, and its typing rules, and then prove that well-typed programs cannot go wrong. Specifically:
adamc@60 1009 %\begin{enumerate}%#<ol>#
adamc@60 1010 %\item%#<li># Define [var] as a synonym for the natural numbers.#</li>#
adamc@60 1011 %\item%#<li># Define an inductive type [exp] of expressions, containing natural number constants, natural number addition, pairing of two other expressions, extraction of the first component of a pair, extraction of the second component of a pair, and variables (based on the [var] type you defined).#</li>#
adamc@60 1012 %\item%#<li># Define an inductive type [cmd] of commands, containing expressions and variable assignments. A variable assignment node should contain the variable being assigned, the expression being assigned to it, and the command to run afterward.#</li>#
adamc@60 1013 %\item%#<li># Define an inductive type [val] of values, containing natural number constants and pairings of values.#</li>#
adamc@60 1014 %\item%#<li># Define a type of variable assignments, which assign a value to each variable.#</li>#
adam@292 1015 %\item%#<li># Define a big-step evaluation relation [eval], capturing what it means for an expression to evaluate to a value under a particular variable assignment. %``%#"#Big step#"#%''% means that the evaluation of every expression should be proved with a single instance of the inductive predicate you will define. For instance, %``%#"#[1 + 1] evaluates to [2] under assignment [va]#"#%''% should be derivable for any assignment [va].#</li>#
adamc@60 1016 %\item%#<li># Define a big-step evaluation relation [run], capturing what it means for a command to run to a value under a particular variable assignment. The value of a command is the result of evaluating its final expression.#</li>#
adamc@60 1017 %\item%#<li># Define a type of variable typings, which are like variable assignments, but map variables to types instead of values. You might use polymorphism to share some code with your variable assignments.#</li>#
adamc@60 1018 %\item%#<li># Define typing judgments for expressions, values, and commands. The expression and command cases will be in terms of a typing assignment.#</li>#
adamc@60 1019 %\item%#<li># Define a predicate [varsType] to express when a variable assignment and a variable typing agree on the types of variables.#</li>#
adamc@60 1020 %\item%#<li># Prove that any expression that has type [t] under variable typing [vt] evaluates under variable assignment [va] to some value that also has type [t] in [vt], as long as [va] and [vt] agree.#</li>#
adamc@60 1021 %\item%#<li># Prove that any command that has type [t] under variable typing [vt] evaluates under variable assignment [va] to some value that also has type [t] in [vt], as long as [va] and [vt] agree.#</li>#
adamc@60 1022 #</ol> </li>#%\end{enumerate}%
adamc@60 1023 A few hints that may be helpful:
adamc@60 1024 %\begin{enumerate}%#<ol>#
adamc@60 1025 %\item%#<li># One easy way of defining variable assignments and typings is to define both as instances of a polymorphic map type. The map type at parameter [T] can be defined to be the type of arbitrary functions from variables to [T]. A helpful function for implementing insertion into such a functional map is [eq_nat_dec], which you can make available with [Require Import Arith.]. [eq_nat_dec] has a dependent type that tells you that it makes accurate decisions on whether two natural numbers are equal, but you can use it as if it returned a boolean, e.g., [if eq_nat_dec n m then E1 else E2].#</li>#
adamc@60 1026 %\item%#<li># If you follow the last hint, you may find yourself writing a proof that involves an expression with [eq_nat_dec] that you would like to simplify. Running [destruct] on the particular call to [eq_nat_dec] should do the trick. You can automate this advice with a piece of Ltac: [[
adamc@60 1027 match goal with
adamc@60 1028 | [ |- context[eq_nat_dec ?X ?Y] ] => destruct (eq_nat_dec X Y)
adamc@60 1029 end
adam@302 1030 ]]
adam@302 1031 #</li>#
adamc@60 1032 %\item%#<li># You probably do not want to use an inductive definition for compatibility of variable assignments and typings.#</li>#
adam@322 1033 %\item%#<li># The [CpdtTactics] module from this book contains a variant [crush'] of [crush]. [crush'] takes two arguments. The first argument is a list of lemmas and other functions to be tried automatically in %``%#"#forward reasoning#"#%''% style, where we add new facts without being sure yet that they link into a proof of the conclusion. The second argument is a list of predicates on which inversion should be attempted automatically. For instance, running [crush' (lemma1, lemma2) pred] will search for chances to apply [lemma1] and [lemma2] to hypotheses that are already available, adding the new concluded fact if suitable hypotheses can be found. Inversion will be attempted on any hypothesis using [pred], but only those inversions that narrow the field of possibilities to one possible rule will be kept. The format of the list arguments to [crush'] is that you can pass an empty list as [tt], a singleton list as the unadorned single element, and a multiple-element list as a tuple of the elements.#</li>#
adamc@60 1034 %\item%#<li># If you want [crush'] to apply polymorphic lemmas, you may have to do a little extra work, if the type parameter is not a free variable of your proof context (so that [crush'] does not know to try it). For instance, if you define a polymorphic map insert function [assign] of some type [forall T : Set, ...], and you want particular applications of [assign] added automatically with type parameter [U], you would need to include [assign] in the lemma list as [assign U] (if you have implicit arguments off) or [assign (T := U)] or [@assign U] (if you have implicit arguments on).#</li>#
adamc@60 1035 #</ol> </li>#%\end{enumerate}%
adamc@60 1036
adamc@60 1037 #</li>#
adamc@60 1038
adamc@58 1039 #</ol>#%\end{enumerate}% *)