annotate src/Predicates.v @ 48:b262252669ce

First-order logic
author Adam Chlipala <adamc@hcoop.net>
date Sat, 27 Sep 2008 16:10:53 -0400
parents be4efaee371f
children 827d7e8a7d9e
rev   line source
adamc@45 1 (* Copyright (c) 2008, Adam Chlipala
adamc@45 2 *
adamc@45 3 * This work is licensed under a
adamc@45 4 * Creative Commons Attribution-Noncommercial-No Derivative Works 3.0
adamc@45 5 * Unported License.
adamc@45 6 * The license text is available at:
adamc@45 7 * http://creativecommons.org/licenses/by-nc-nd/3.0/
adamc@45 8 *)
adamc@45 9
adamc@45 10 (* begin hide *)
adamc@45 11 Require Import List.
adamc@45 12
adamc@45 13 Require Import Tactics.
adamc@45 14
adamc@45 15 Set Implicit Arguments.
adamc@45 16 (* end hide *)
adamc@45 17
adamc@45 18
adamc@45 19 (** %\chapter{Inductive Predicates}% *)
adamc@45 20
adamc@45 21 (** The so-called "Curry-Howard Correspondence" states a formal connection between functional programs and mathematical proofs. In the last chapter, we snuck in a first introduction to this subject in Coq. Witness the close similarity between the types [unit] and [True] from the standard library: *)
adamc@45 22
adamc@45 23 Print unit.
adamc@45 24 (** [[
adamc@45 25
adamc@45 26 Inductive unit : Set := tt : unit
adamc@45 27 ]] *)
adamc@45 28
adamc@45 29 Print True.
adamc@45 30 (** [[
adamc@45 31
adamc@45 32 Inductive True : Prop := I : True
adamc@45 33 ]] *)
adamc@45 34
adamc@45 35 (** Recall that [unit] is the type with only one value, and [True] is the proposition that always holds. Despite this superficial difference between the two concepts, in both cases we can use the same inductive definition mechanism. The connection goes further than this. We see that we arrive at the definition of [True] by replacing [unit] by [True], [tt] by [I], and [Set] by [Prop]. The first two of these differences are superficial changes of names, while the third difference is the crucial one for separating programs from proofs. A term [T] of type [Set] is a type of programs, and a term of type [T] is a program. A term [T] of type [Prop] is a logical proposition, and its proofs are of type [T].
adamc@45 36
adamc@45 37 [unit] has one value, [tt]. [True] has one proof, [I]. Why distinguish between these two types? Many people who have read about Curry-Howard in an abstract context and not put it to use in proof engineering answer that the two types in fact %\textit{%#<i>#should not#</i>#%}% be distinguished. There is a certain aesthetic appeal to this point of view, but I want to argue that it is best to treat Curry-Howard very loosely in practical proving. There are Coq-specific reasons for preferring the distinction, involving efficient compilation and avoidance of paradoxes in the presence of classical math, but I will argue that there is a more general principle that should lead us to avoid conflating programming and proving.
adamc@45 38
adamc@45 39 The essence of the argument is roughly this: to an engineer, not all functions of type [A -> B] are created equal, but all proofs of a proposition [P -> Q] are. This idea is known as %\textit{%#<i>#proof irrelevance#</i>#%}%, and its formalizations in logics prevent us from distinguishing between alternate proofs of the same proposition. Proof irrelevance is compatible with, but not derivable in, Gallina. Apart from this theoretical concern, I will argue that it is most effective to do engineering with Coq by employing different techniques for programs versus proofs. Most of this book is organized around that distinction, describing how to program, by applying standard functional programming techniques in the presence of dependent types; and how to prove, by writing custom Ltac decision procedures.
adamc@45 40
adamc@45 41 With that perspective in mind, this chapter is sort of a mirror image of the last chapter, introducing how to define predicates with inductive definitions. We will point out similarities in places, but much of the effective Coq user's bag of tricks is disjoint for predicates versus "datatypes." This chapter is also a covert introduction to dependent types, which are the foundation on which interesting inductive predicates are built, though we will rely on tactics to build dependently-typed proof terms for us for now. A future chapter introduces more manual application of dependent types. *)
adamc@45 42
adamc@45 43
adamc@48 44 (** * Propositional Logic *)
adamc@45 45
adamc@45 46 (** Let us begin with a brief tour through the definitions of the connectives for propositional logic. We will work within a Coq section that provides us with a set of propositional variables. In Coq parlance, these are just terms of type [Prop.] *)
adamc@45 47
adamc@45 48 Section Propositional.
adamc@46 49 Variables P Q R : Prop.
adamc@45 50
adamc@45 51 (** In Coq, the most basic propositional connective is implication, written [->], which we have already used in almost every proof. Rather than being defined inductively, implication is built into Coq as the function type constructor.
adamc@45 52
adamc@45 53 We have also already seen the definition of [True]. For a demonstration of a lower-level way of establishing proofs of inductive predicates, we turn to this trivial theorem. *)
adamc@45 54
adamc@45 55 Theorem obvious : True.
adamc@45 56 apply I.
adamc@45 57 Qed.
adamc@45 58
adamc@45 59 (** We may always use the [apply] tactic to take a proof step based on applying a particular constructor of the inductive predicate that we are trying to establish. Sometimes there is only one constructor that could possibly apply, in which case a shortcut is available: *)
adamc@45 60
adamc@45 61 Theorem obvious' : True.
adamc@45 62 constructor.
adamc@45 63 Qed.
adamc@45 64
adamc@45 65 (** There is also a predicate [False], which is the Curry-Howard mirror image of [Empty_set] from the last chapter. *)
adamc@45 66
adamc@45 67 Print False.
adamc@45 68 (** [[
adamc@45 69
adamc@45 70 Inductive False : Prop :=
adamc@45 71 ]] *)
adamc@45 72
adamc@45 73 (** We can conclude anything from [False], doing case analysis on a proof of [False] in the same way we might do case analysis on, say, a natural number. Since there are no cases to consider, any such case analysis succeeds immediately in proving the goal. *)
adamc@45 74
adamc@45 75 Theorem False_imp : False -> 2 + 2 = 5.
adamc@45 76 destruct 1.
adamc@45 77 Qed.
adamc@45 78
adamc@45 79 (** In a consistent context, we can never build a proof of [False]. In inconsistent contexts that appear in the courses of proofs, it is usually easiest to proceed by demonstrating that inconsistency with an explicit proof of [False]. *)
adamc@45 80
adamc@45 81 Theorem arith_neq : 2 + 2 = 5 -> 9 + 9 = 835.
adamc@45 82 intro.
adamc@45 83
adamc@45 84 (** At this point, we have an inconsistent hypothesis [2 + 2 = 5], so the specific conclusion is not important. We use the [elimtype] tactic to state a proposition, telling Coq that we wish to construct a proof of the new proposition and then prove the original goal by case analysis on the structure of the new auxiliary proof. Since [False] has no constructors, [elimtype False] simply leaves us with the obligation to prove [False]. *)
adamc@45 85
adamc@45 86 elimtype False.
adamc@45 87 (** [[
adamc@45 88
adamc@45 89 H : 2 + 2 = 5
adamc@45 90 ============================
adamc@45 91 False
adamc@45 92 ]] *)
adamc@45 93
adamc@45 94 (** For now, we will leave the details of this proof about arithmetic to [crush]. *)
adamc@45 95
adamc@45 96 crush.
adamc@45 97 Qed.
adamc@45 98
adamc@45 99 (** A related notion to [False] is logical negation. *)
adamc@45 100
adamc@45 101 Print not.
adamc@45 102 (** [[
adamc@45 103
adamc@45 104 not = fun A : Prop => A -> False
adamc@45 105 : Prop -> Prop
adamc@45 106 ]] *)
adamc@45 107
adamc@45 108 (** We see that [not] is just shorthand for implication of [False]. We can use that fact explicitly in proofs. The syntax [~P] expands to [not P]. *)
adamc@45 109
adamc@45 110 Theorem arith_neq' : ~ (2 + 2 = 5).
adamc@45 111 unfold not.
adamc@45 112
adamc@45 113 (** [[
adamc@45 114
adamc@45 115 ============================
adamc@45 116 2 + 2 = 5 -> False
adamc@45 117 ]] *)
adamc@45 118
adamc@45 119 crush.
adamc@45 120 Qed.
adamc@45 121
adamc@45 122 (** We also have conjunction, which we introduced in the last chapter. *)
adamc@45 123
adamc@45 124 Print and.
adamc@45 125 (** [[
adamc@45 126
adamc@45 127 Inductive and (A : Prop) (B : Prop) : Prop := conj : A -> B -> A /\ B
adamc@45 128 ]] *)
adamc@45 129
adamc@45 130 (** The interested reader can check that [and] has a Curry-Howard doppleganger called [prod], the type of pairs. However, it is generally most convenient to reason about conjunction using tactics. An explicit proof of commutativity of [and] illustrates the usual suspects for such tasks. [/\] is an infix shorthand for [and]. *)
adamc@45 131
adamc@45 132 Theorem and_comm : P /\ Q -> Q /\ P.
adamc@45 133 (** We start by case analysis on the proof of [P /\ Q]. *)
adamc@45 134
adamc@45 135 destruct 1.
adamc@45 136 (** [[
adamc@45 137
adamc@45 138 H : P
adamc@45 139 H0 : Q
adamc@45 140 ============================
adamc@45 141 Q /\ P
adamc@45 142 ]] *)
adamc@45 143
adamc@45 144 (** Every proof of a conjunction provides proofs for both conjuncts, so we get a single subgoal reflecting that. We can proceed by splitting this subgoal into a case for each conjunct of [Q /\ P]. *)
adamc@45 145
adamc@45 146 split.
adamc@45 147 (** [[
adamc@45 148 2 subgoals
adamc@45 149
adamc@45 150 H : P
adamc@45 151 H0 : Q
adamc@45 152 ============================
adamc@45 153 Q
adamc@45 154
adamc@45 155 subgoal 2 is:
adamc@45 156 P
adamc@45 157 ]] *)
adamc@45 158
adamc@45 159 (** In each case, the conclusion is among our hypotheses, so the [assumption] tactic finishes the process. *)
adamc@45 160
adamc@45 161 assumption.
adamc@45 162 assumption.
adamc@45 163 Qed.
adamc@45 164
adamc@45 165 (** Coq disjunction is called [or] and abbreviated with the infix operator [\/]. *)
adamc@45 166
adamc@45 167 Print or.
adamc@45 168 (** [[
adamc@45 169
adamc@45 170 Inductive or (A : Prop) (B : Prop) : Prop :=
adamc@45 171 or_introl : A -> A \/ B | or_intror : B -> A \/ B
adamc@45 172 ]] *)
adamc@45 173
adamc@45 174 (** We see that there are two ways to prove a disjunction: prove the first disjunct or prove the second. The Curry-Howard analogue of this is the Coq [sum] type. We can demonstrate the main tactics here with another proof of commutativity. *)
adamc@45 175
adamc@45 176 Theorem or_comm : P \/ Q -> Q \/ P.
adamc@45 177 (** As in the proof for [and], we begin with case analysis, though this time we are met by two cases instead of one. *)
adamc@45 178 destruct 1.
adamc@45 179 (** [[
adamc@45 180
adamc@45 181 2 subgoals
adamc@45 182
adamc@45 183 H : P
adamc@45 184 ============================
adamc@45 185 Q \/ P
adamc@45 186
adamc@45 187 subgoal 2 is:
adamc@45 188 Q \/ P
adamc@45 189 ]] *)
adamc@45 190
adamc@45 191 (** We can see that, in the first subgoal, we want to prove the disjunction by proving its second disjunct. The [right] tactic telegraphs this intent. *)
adamc@45 192
adamc@45 193 right; assumption.
adamc@45 194
adamc@45 195 (** The second subgoal has a symmetric proof.
adamc@45 196
adamc@45 197 [[
adamc@45 198
adamc@45 199 1 subgoal
adamc@45 200
adamc@45 201 H : Q
adamc@45 202 ============================
adamc@45 203 Q \/ P
adamc@45 204 ]] *)
adamc@45 205
adamc@45 206 left; assumption.
adamc@45 207 Qed.
adamc@45 208
adamc@46 209
adamc@46 210 (* begin hide *)
adamc@46 211 (* In-class exercises *)
adamc@46 212
adamc@46 213 Theorem contra : P -> ~P -> R.
adamc@46 214 Admitted.
adamc@46 215
adamc@46 216 Theorem and_assoc : (P /\ Q) /\ R -> P /\ (Q /\ R).
adamc@46 217 Admitted.
adamc@46 218
adamc@46 219 Theorem or_assoc : (P \/ Q) \/ R -> P \/ (Q \/ R).
adamc@46 220 Admitted.
adamc@46 221
adamc@46 222 (* end hide *)
adamc@46 223
adamc@46 224
adamc@46 225 (** It would be a shame to have to plod manually through all proofs about propositional logic. Luckily, there is no need. One of the most basic Coq automation tactics is [tauto], which is a complete decision procedure for constructive propositional logic. (More on what "constructive" means in the next section.) We can use [tauto] to dispatch all of the purely propositional theorems we have proved so far. *)
adamc@46 226
adamc@46 227 Theorem or_comm' : P \/ Q -> Q \/ P.
adamc@46 228 tauto.
adamc@46 229 Qed.
adamc@46 230
adamc@46 231 (** Sometimes propositional reasoning forms important plumbing for the proof of a theorem, but we still need to apply some other smarts about, say, arithmetic. [intuition] is a generalization of [tauto] that proves everything it can using propositional reasoning. When some goals remain, it uses propositional laws to simplify them as far as possible. Consider this example, which uses the list concatenation operator [++] from the standard library. *)
adamc@46 232
adamc@46 233 Theorem arith_comm : forall ls1 ls2 : list nat,
adamc@46 234 length ls1 = length ls2 \/ length ls1 + length ls2 = 6
adamc@46 235 -> length (ls1 ++ ls2) = 6 \/ length ls1 = length ls2.
adamc@46 236 intuition.
adamc@46 237
adamc@46 238 (** 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 239
adamc@46 240 (** [[
adamc@46 241
adamc@46 242 ls1 : list nat
adamc@46 243 ls2 : list nat
adamc@46 244 H0 : length ls1 + length ls2 = 6
adamc@46 245 ============================
adamc@46 246 length (ls1 ++ ls2) = 6 \/ length ls1 = length ls2
adamc@46 247 ]] *)
adamc@46 248
adamc@46 249 (** 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 250
adamc@46 251 rewrite app_length.
adamc@46 252 (** [[
adamc@46 253
adamc@46 254 ls1 : list nat
adamc@46 255 ls2 : list nat
adamc@46 256 H0 : length ls1 + length ls2 = 6
adamc@46 257 ============================
adamc@46 258 length ls1 + length ls2 = 6 \/ length ls1 = length ls2
adamc@46 259 ]] *)
adamc@46 260
adamc@46 261 (** 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 262
adamc@46 263 tauto.
adamc@46 264 Qed.
adamc@46 265
adamc@46 266 (** [intuition] is one of the main bits of glue in the implementation of [crush], so, with a little help, we can get a short automated proof of the theorem. *)
adamc@46 267
adamc@46 268 Theorem arith_comm' : forall ls1 ls2 : list nat,
adamc@46 269 length ls1 = length ls2 \/ length ls1 + length ls2 = 6
adamc@46 270 -> length (ls1 ++ ls2) = 6 \/ length ls1 = length ls2.
adamc@46 271 Hint Rewrite app_length : cpdt.
adamc@46 272
adamc@46 273 crush.
adamc@46 274 Qed.
adamc@46 275
adamc@45 276 End Propositional.
adamc@45 277
adamc@46 278
adamc@47 279 (** * What Does It Mean to Be Constructive? *)
adamc@46 280
adamc@47 281 (** 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 282
adamc@47 283 The answer comes from the fact that Coq implements %\textit{%#<i>#constructive#</i>#%}% or %\textit{%#<i>#intuitionistic#</i>#%}% logic, in contrast to the %\textit{%#<i>#classical#</i>#%}% logic that you may be more familiar with. In constructive logic, classical tautologies like [~ ~P -> P] and [P \/ ~P] do not always hold. In general, we can only prove these tautologies when [P] is %\textit{%#<i>#decidable#</i>#%}%, in the sense of computability theory. The Curry-Howard encoding that Coq uses for [or] allows us to extract either a proof of [P] or a proof of [~P] from any proof of [P \/ ~P]. Since our proofs are just functional programs which we can run, this would give us a decision procedure for the halting problem, where the instantiations of [P] would be formulas like "this particular Turing machine halts."
adamc@47 284
adamc@47 285 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 286
adamc@47 287 Constructive logic lets us define all of the logical connectives in an aesthetically-appealing way, with orthogonal inductive definitions. That is, each connective is defined independently using a simple, shared mechanism. Constructivity also enables a trick called %\textit{%#<i>#program extraction#</i>#%}%, where we write programs by phrasing them as theorems to be proved. Since our proofs are just functional programs, we can extract executable programs from our final proofs, which we could not do as naturally with classical proofs.
adamc@47 288
adamc@47 289 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 290
adamc@48 291
adamc@48 292 (** * First-Order Logic *)
adamc@48 293
adamc@48 294 (** The [forall] connective of first-order logic, which we have seen in many examples so far, is built into Coq. Getting ahead of ourselves a bit, we can see it as the dependent function type constructor. In fact, implication and universal quantification are just different syntactic shorthands for the same Coq mechanism. A formula [P -> Q] is equivalent to [forall x : P, Q], where [x] does not appear in [Q]. That is, the "real" type of the implication says "for every proof of [P], there exists a proof of [Q]."
adamc@48 295
adamc@48 296 Existential quantification is defined in the standard library. *)
adamc@48 297
adamc@48 298 Print ex.
adamc@48 299 (** [[
adamc@48 300
adamc@48 301 Inductive ex (A : Type) (P : A -> Prop) : Prop :=
adamc@48 302 ex_intro : forall x : A, P x -> ex P
adamc@48 303 ]] *)
adamc@48 304
adamc@48 305 (** [ex] is parameterized by the type [A] that we quantify over, and by a predicate [P] over [A]s. We prove an existential by exhibiting some [x] of type [A], along with a proof of [P x]. As usual, there are tactics that save us from worrying about the low-level details most of the time. *)
adamc@48 306
adamc@48 307 Theorem exist1 : exists x : nat, x + 1 = 2.
adamc@48 308 (** remove printing exists*)
adamc@48 309 (** We can start this proof with a tactic [exists], which should not be confused with the formula constructor shorthand of the same name. (In the PDF version of this document, the reverse 'E' appears instead of the text "exists.") *)
adamc@48 310 exists 1.
adamc@48 311
adamc@48 312 (** The conclusion is replaced with a version using the existential witness that we announced. *)
adamc@48 313
adamc@48 314 (** [[
adamc@48 315
adamc@48 316 ============================
adamc@48 317 1 + 1 = 2
adamc@48 318 ]] *)
adamc@48 319
adamc@48 320 reflexivity.
adamc@48 321 Qed.
adamc@48 322
adamc@48 323 (** printing exists $\exists$ *)
adamc@48 324
adamc@48 325 (** We can also use tactics to reason about existential hypotheses. *)
adamc@48 326
adamc@48 327 Theorem exist2 : forall n m : nat, (exists x : nat, n + x = m) -> n <= m.
adamc@48 328 (** We start by case analysis on the proof of the existential fact. *)
adamc@48 329 destruct 1.
adamc@48 330 (** [[
adamc@48 331
adamc@48 332 n : nat
adamc@48 333 m : nat
adamc@48 334 x : nat
adamc@48 335 H : n + x = m
adamc@48 336 ============================
adamc@48 337 n <= m
adamc@48 338 ]] *)
adamc@48 339
adamc@48 340 (** 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 341
adamc@48 342 crush.
adamc@48 343 Qed.
adamc@48 344
adamc@48 345
adamc@48 346 (* begin hide *)
adamc@48 347 (* In-class exercises *)
adamc@48 348
adamc@48 349 Theorem forall_exists_commute : forall (A B : Type) (P : A -> B -> Prop),
adamc@48 350 (exists x : A, forall y : B, P x y) -> (forall y : B, exists x : A, P x y).
adamc@48 351 Admitted.
adamc@48 352
adamc@48 353 (* end hide *)
adamc@48 354
adamc@48 355
adamc@48 356 (** The tactic [intuition] has a first-order cousin called [firstorder]. [firstorder] proves many formulas when only first-order reasoning is needed, and it tries to perform first-order simplifications in any case. First-order reasoning is much harder than propositional reasoning, so [firstorder] is much more likely than [intuition] to get stuck in a way that makes it run for long enough to be useless. *)