Predicates.v 41.2 KB
Newer Older
1
(* Copyright (c) 2008-2012, Adam Chlipala
Adam Chlipala's avatar
Adam Chlipala committed
2 3 4 5 6 7 8 9 10 11 12
 * 
 * This work is licensed under a
 * Creative Commons Attribution-Noncommercial-No Derivative Works 3.0
 * Unported License.
 * The license text is available at:
 *   http://creativecommons.org/licenses/by-nc-nd/3.0/
 *)

(* begin hide *)
Require Import List.

13
Require Import CpdtTactics.
Adam Chlipala's avatar
Adam Chlipala committed
14 15 16

Set Implicit Arguments.

Adam Chlipala's avatar
Adam Chlipala committed
17 18
(* Extra definitions to get coqdoc to choose the right fonts. *)

19
(* begin thide *)
Adam Chlipala's avatar
Adam Chlipala committed
20 21 22 23 24 25 26 27
Inductive unit := tt.
Inductive Empty_set := .
Inductive bool := true | false.
Inductive sum := .
Inductive prod := .
Inductive and := conj.
Inductive or := or_introl | or_intror.
Inductive ex := ex_intro.
28
Inductive eq := eq_refl.
Adam Chlipala's avatar
Adam Chlipala committed
29
Reset unit.
30
(* end thide *)
Adam Chlipala's avatar
Adam Chlipala committed
31
(* end hide *)
Adam Chlipala's avatar
Adam Chlipala committed
32 33 34

(** %\chapter{Inductive Predicates}% *)

Adam Chlipala's avatar
Adam Chlipala committed
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: *)
Adam Chlipala's avatar
Adam Chlipala committed
36 37

Print unit.
38
(** %\vspace{-.15in}%[[
Adam Chlipala's avatar
Adam Chlipala committed
39
  Inductive unit : Set :=  tt : unit
40 41
]]
*)
Adam Chlipala's avatar
Adam Chlipala committed
42 43

Print True.
44
(** %\vspace{-.15in}%[[
Adam Chlipala's avatar
Adam Chlipala committed
45
  Inductive True : Prop :=  I : True
Adam Chlipala's avatar
Adam Chlipala committed
46
  ]]
Adam Chlipala's avatar
Adam Chlipala committed
47

48
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.
Adam Chlipala's avatar
Adam Chlipala committed
49

50
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 but who have 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.
Adam Chlipala's avatar
Adam Chlipala committed
51

52
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.
Adam Chlipala's avatar
Adam Chlipala committed
53

54
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. *)
Adam Chlipala's avatar
Adam Chlipala committed
55 56


Adam Chlipala's avatar
Adam Chlipala committed
57
(** * Propositional Logic *)
Adam Chlipala's avatar
Adam Chlipala committed
58

59
(** 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 variables of type [Prop]. *)
Adam Chlipala's avatar
Adam Chlipala committed
60 61

Section Propositional.
Adam Chlipala's avatar
Adam Chlipala committed
62
  Variables P Q R : Prop.
Adam Chlipala's avatar
Adam Chlipala committed
63 64 65 66 67 68

  (** 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.

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. *)
  
  Theorem obvious : True.
Adam Chlipala's avatar
Adam Chlipala committed
69
(* begin thide *)
Adam Chlipala's avatar
Adam Chlipala committed
70
    apply I.
Adam Chlipala's avatar
Adam Chlipala committed
71
(* end thide *)
Adam Chlipala's avatar
Adam Chlipala committed
72 73
  Qed.

74
  (** 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:%\index{tactics!constructor}% *)
Adam Chlipala's avatar
Adam Chlipala committed
75

Adam Chlipala's avatar
Adam Chlipala committed
76
(* begin thide *)
Adam Chlipala's avatar
Adam Chlipala committed
77 78 79 80
  Theorem obvious' : True.
    constructor.
  Qed.

Adam Chlipala's avatar
Adam Chlipala committed
81 82
(* end thide *)

Adam Chlipala's avatar
Adam Chlipala committed
83 84 85
  (** There is also a predicate [False], which is the Curry-Howard mirror image of [Empty_set] from the last chapter. *)

  Print False.
86
  (** %\vspace{-.15in}%[[
Adam Chlipala's avatar
Adam Chlipala committed
87 88
  Inductive False : Prop :=
  ]]
89

90
  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. *)
Adam Chlipala's avatar
Adam Chlipala committed
91 92

  Theorem False_imp : False -> 2 + 2 = 5.
Adam Chlipala's avatar
Adam Chlipala committed
93
(* begin thide *)
Adam Chlipala's avatar
Adam Chlipala committed
94
    destruct 1.
Adam Chlipala's avatar
Adam Chlipala committed
95
(* end thide *)
Adam Chlipala's avatar
Adam Chlipala committed
96 97
  Qed.

98
  (** 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 the inconsistency with an explicit proof of [False]. *)
Adam Chlipala's avatar
Adam Chlipala committed
99 100

  Theorem arith_neq : 2 + 2 = 5 -> 9 + 9 = 835.
Adam Chlipala's avatar
Adam Chlipala committed
101
(* begin thide *)
Adam Chlipala's avatar
Adam Chlipala committed
102 103
    intro.

104
    (** 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.  For a full description of it, see the Coq manual.  For our purposes, we only need the variant [elimtype False], which lets us replace any conclusion formula with [False], because any fact follows from an inconsistent context. *)
Adam Chlipala's avatar
Adam Chlipala committed
105 106 107 108 109 110

    elimtype False.
    (** [[
  H : 2 + 2 = 5
  ============================
   False
Adam Chlipala's avatar
Adam Chlipala committed
111 112
 
   ]]
Adam Chlipala's avatar
Adam Chlipala committed
113

Adam Chlipala's avatar
Adam Chlipala committed
114
   For now, we will leave the details of this proof about arithmetic to [crush]. *)
Adam Chlipala's avatar
Adam Chlipala committed
115 116

    crush.
Adam Chlipala's avatar
Adam Chlipala committed
117
(* end thide *)
Adam Chlipala's avatar
Adam Chlipala committed
118 119 120 121
  Qed.

  (** A related notion to [False] is logical negation. *)

122 123 124 125
  (* begin hide *)
  Definition foo := not.
  (* end hide *)

Adam Chlipala's avatar
Adam Chlipala committed
126
  Print not.
Adam Chlipala's avatar
Adam Chlipala committed
127 128 129 130
  (** %\vspace{-.15in}% [[
    not = fun A : Prop => A -> False
      : Prop -> Prop
     ]]
Adam Chlipala's avatar
Adam Chlipala committed
131

132
     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]. *)
Adam Chlipala's avatar
Adam Chlipala committed
133 134

  Theorem arith_neq' : ~ (2 + 2 = 5).
Adam Chlipala's avatar
Adam Chlipala committed
135
(* begin thide *)
Adam Chlipala's avatar
Adam Chlipala committed
136 137 138 139
    unfold not.
    (** [[
  ============================
   2 + 2 = 5 -> False
140 141
   ]]
   *)
Adam Chlipala's avatar
Adam Chlipala committed
142 143

    crush.
Adam Chlipala's avatar
Adam Chlipala committed
144
(* end thide *)
Adam Chlipala's avatar
Adam Chlipala committed
145 146 147 148 149
  Qed.

  (** We also have conjunction, which we introduced in the last chapter. *)

  Print and.
150
(** %\vspace{-.15in}%[[
Adam Chlipala's avatar
Adam Chlipala committed
151
    Inductive and (A : Prop) (B : Prop) : Prop :=  conj : A -> B -> A /\ B
Adam Chlipala's avatar
Adam Chlipala committed
152 153
  ]]
  
154
  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]. *)
Adam Chlipala's avatar
Adam Chlipala committed
155 156

  Theorem and_comm : P /\ Q -> Q /\ P.
Adam Chlipala's avatar
Adam Chlipala committed
157

Adam Chlipala's avatar
Adam Chlipala committed
158
(* begin thide *)
Adam Chlipala's avatar
Adam Chlipala committed
159 160 161 162 163 164 165 166
    (** We start by case analysis on the proof of [P /\ Q]. *)

    destruct 1.
    (** [[
  H : P
  H0 : Q
  ============================
   Q /\ P
Adam Chlipala's avatar
Adam Chlipala committed
167 168
   
   ]]
Adam Chlipala's avatar
Adam Chlipala committed
169

Adam Chlipala's avatar
Adam Chlipala committed
170
    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}% *)
Adam Chlipala's avatar
Adam Chlipala committed
171 172

    split.
173 174
(** [[
2 subgoals
Adam Chlipala's avatar
Adam Chlipala committed
175 176 177 178 179
  
  H : P
  H0 : Q
  ============================
   Q
180 181 182

subgoal 2 is

Adam Chlipala's avatar
Adam Chlipala committed
183
   P
Adam Chlipala's avatar
Adam Chlipala committed
184 185
 
 ]]
Adam Chlipala's avatar
Adam Chlipala committed
186

Adam Chlipala's avatar
Adam Chlipala committed
187
 In each case, the conclusion is among our hypotheses, so the %\index{tactics!assumption}%[assumption] tactic finishes the process. *)
Adam Chlipala's avatar
Adam Chlipala committed
188 189 190

    assumption.
    assumption.
Adam Chlipala's avatar
Adam Chlipala committed
191
(* end thide *)
Adam Chlipala's avatar
Adam Chlipala committed
192 193
  Qed.

Adam Chlipala's avatar
Adam Chlipala committed
194
  (** Coq disjunction is called %\index{Gallina terms!or}%[or] and abbreviated with the infix operator [\/]. *)
Adam Chlipala's avatar
Adam Chlipala committed
195 196

  Print or.
197
(** %\vspace{-.15in}%[[
Adam Chlipala's avatar
Adam Chlipala committed
198
  Inductive or (A : Prop) (B : Prop) : Prop :=
Adam Chlipala's avatar
Adam Chlipala committed
199
    or_introl : A -> A \/ B | or_intror : B -> A \/ B
Adam Chlipala's avatar
Adam Chlipala committed
200
]]
Adam Chlipala's avatar
Adam Chlipala committed
201

202
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. *)
Adam Chlipala's avatar
Adam Chlipala committed
203 204

  Theorem or_comm : P \/ Q -> Q \/ P.
Adam Chlipala's avatar
Adam Chlipala committed
205 206

(* begin thide *)
Adam Chlipala's avatar
Adam Chlipala committed
207
    (** As in the proof for [and], we begin with case analysis, though this time we are met by two cases instead of one. *)
Adam Chlipala's avatar
Adam Chlipala committed
208

Adam Chlipala's avatar
Adam Chlipala committed
209
    destruct 1.
210 211
(** [[
2 subgoals
Adam Chlipala's avatar
Adam Chlipala committed
212 213 214 215
  
  H : P
  ============================
   Q \/ P
216 217 218

subgoal 2 is

Adam Chlipala's avatar
Adam Chlipala committed
219
 Q \/ P
Adam Chlipala's avatar
Adam Chlipala committed
220 221
 
 ]]
Adam Chlipala's avatar
Adam Chlipala committed
222

223
 We can see that, in the first subgoal, we want to prove the disjunction by proving its second disjunct.  The %\index{tactics!right}%[right] tactic telegraphs this intent. *)
Adam Chlipala's avatar
Adam Chlipala committed
224

Adam Chlipala's avatar
Adam Chlipala committed
225 226
    right; assumption.

Adam Chlipala's avatar
Adam Chlipala committed
227
    (** The second subgoal has a symmetric proof.%\index{tactics!left}%
Adam Chlipala's avatar
Adam Chlipala committed
228 229 230 231 232 233 234

       [[
1 subgoal
  
  H : Q
  ============================
   Q \/ P
235 236
   ]]
   *)
Adam Chlipala's avatar
Adam Chlipala committed
237 238

    left; assumption.
Adam Chlipala's avatar
Adam Chlipala committed
239

Adam Chlipala's avatar
Adam Chlipala committed
240
(* end thide *)
Adam Chlipala's avatar
Adam Chlipala committed
241 242
  Qed.

Adam Chlipala's avatar
Adam Chlipala committed
243 244 245 246 247

(* begin hide *)
(* In-class exercises *)

  Theorem contra : P -> ~P -> R.
248 249 250 251 252 253 254
(* begin thide *)
    unfold not.
    intros.
    elimtype False.
    apply H0.
    assumption.
(* end thide *)
Adam Chlipala's avatar
Adam Chlipala committed
255 256 257
  Admitted.

  Theorem and_assoc : (P /\ Q) /\ R -> P /\ (Q /\ R).
258 259 260 261 262 263 264 265 266 267
(* begin thide *)
    intros.
    destruct H.
    destruct H.
    split.
    assumption.
    split.
    assumption.
    assumption.
(* end thide *)
Adam Chlipala's avatar
Adam Chlipala committed
268 269 270
  Admitted.

  Theorem or_assoc : (P \/ Q) \/ R -> P \/ (Q \/ R).
271 272 273 274 275 276 277 278 279 280 281 282 283
(* begin thide *)
    intros.
    destruct H.
    destruct H.
    left.
    assumption.
    right.
    left.
    assumption.
    right.
    right.
    assumption.
(* end thide *)
Adam Chlipala's avatar
Adam Chlipala committed
284 285 286 287 288
  Admitted.

(* end hide *)


289
  (** 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. *)
Adam Chlipala's avatar
Adam Chlipala committed
290 291

  Theorem or_comm' : P \/ Q -> Q \/ P.
Adam Chlipala's avatar
Adam Chlipala committed
292
(* begin thide *)
Adam Chlipala's avatar
Adam Chlipala committed
293
    tauto.
Adam Chlipala's avatar
Adam Chlipala committed
294
(* end thide *)
Adam Chlipala's avatar
Adam Chlipala committed
295 296
  Qed.

297
  (** Sometimes propositional reasoning forms important plumbing for the proof of a theorem, but we still need to apply some other smarts about, say, arithmetic.  The tactic %\index{tactics!intuition}%[intuition] is a generalization of [tauto] that proves everything it can using propositional reasoning.  When some further facts must be established to finish the proof, [intuition] uses propositional laws to simplify them as far as possible.  Consider this example, which uses the list concatenation operator %\coqdocnotation{%#<tt>#++#</tt>#%}% from the standard library. *)
Adam Chlipala's avatar
Adam Chlipala committed
298 299 300 301

  Theorem arith_comm : forall ls1 ls2 : list nat,
    length ls1 = length ls2 \/ length ls1 + length ls2 = 6
    -> length (ls1 ++ ls2) = 6 \/ length ls1 = length ls2.
Adam Chlipala's avatar
Adam Chlipala committed
302
(* begin thide *)
Adam Chlipala's avatar
Adam Chlipala committed
303 304 305 306 307 308 309 310 311 312
    intuition.

    (** 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. *)

    (** [[
  ls1 : list nat
  ls2 : list nat
  H0 : length ls1 + length ls2 = 6
  ============================
   length (ls1 ++ ls2) = 6 \/ length ls1 = length ls2
Adam Chlipala's avatar
Adam Chlipala committed
313 314
 
   ]]
Adam Chlipala's avatar
Adam Chlipala committed
315

Adam Chlipala's avatar
Adam Chlipala committed
316
   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. *)
Adam Chlipala's avatar
Adam Chlipala committed
317 318 319 320 321 322 323 324

    rewrite app_length.
    (** [[
  ls1 : list nat
  ls2 : list nat
  H0 : length ls1 + length ls2 = 6
  ============================
   length ls1 + length ls2 = 6 \/ length ls1 = length ls2
Adam Chlipala's avatar
Adam Chlipala committed
325 326
 
   ]]
Adam Chlipala's avatar
Adam Chlipala committed
327

Adam Chlipala's avatar
Adam Chlipala committed
328
   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. *)
Adam Chlipala's avatar
Adam Chlipala committed
329 330

    tauto.
Adam Chlipala's avatar
Adam Chlipala committed
331
(* end thide *)
Adam Chlipala's avatar
Adam Chlipala committed
332 333
  Qed.

Adam Chlipala's avatar
Adam Chlipala committed
334
  (** 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. *)
Adam Chlipala's avatar
Adam Chlipala committed
335

Adam Chlipala's avatar
Adam Chlipala committed
336
(* begin thide *)
Adam Chlipala's avatar
Adam Chlipala committed
337 338 339
  Theorem arith_comm' : forall ls1 ls2 : list nat,
    length ls1 = length ls2 \/ length ls1 + length ls2 = 6
    -> length (ls1 ++ ls2) = 6 \/ length ls1 = length ls2.
340
    Hint Rewrite app_length.
Adam Chlipala's avatar
Adam Chlipala committed
341 342 343

    crush.
  Qed.
Adam Chlipala's avatar
Adam Chlipala committed
344
(* end thide *)
Adam Chlipala's avatar
Adam Chlipala committed
345

Adam Chlipala's avatar
Adam Chlipala committed
346 347
End Propositional.

Adam Chlipala's avatar
Adam Chlipala committed
348 349
(** 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 Chlipala's avatar
Adam Chlipala committed
350

Adam Chlipala's avatar
Adam Chlipala committed
351
(** * What Does It Mean to Be Constructive? *)
Adam Chlipala's avatar
Adam Chlipala committed
352

353
(** One potential point of confusion in the presentation so far is the distinction between [bool] and [Prop].  The datatype [bool] is built from two values [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, [True] and [False]?
Adam Chlipala's avatar
Adam Chlipala committed
354

355
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."
Adam Chlipala's avatar
Adam Chlipala committed
356

357 358
A similar paradoxical situation would result if every proposition evaluated to either [True] or [False].  Evaluation in Coq is decidable, so we would be limiting ourselves to decidable propositions only.

359
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."
Adam Chlipala's avatar
Adam Chlipala committed
360

361
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.
Adam Chlipala's avatar
Adam Chlipala committed
362 363

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. *)
Adam Chlipala's avatar
Adam Chlipala committed
364 365 366 367


(** * First-Order Logic *)

368
(** 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]."
Adam Chlipala's avatar
Adam Chlipala committed
369

Adam Chlipala's avatar
Adam Chlipala committed
370
%\index{existential quantification}\index{Gallina terms!exists}\index{Gallina terms!ex}%Existential quantification is defined in the standard library. *)
Adam Chlipala's avatar
Adam Chlipala committed
371

Adam Chlipala's avatar
Adam Chlipala committed
372
  Print ex.
373
(** %\vspace{-.15in}%[[
Adam Chlipala's avatar
Adam Chlipala committed
374
  Inductive ex (A : Type) (P : A -> Prop) : Prop :=
Adam Chlipala's avatar
Adam Chlipala committed
375
    ex_intro : forall x : A, P x -> ex P
Adam Chlipala's avatar
Adam Chlipala committed
376
    ]]
Adam Chlipala's avatar
Adam Chlipala committed
377

378 379
  (Note that here, as always, each [forall] quantifier has the largest possible scope, so that the type of [ex_intro] could also be written [forall x : A, (P x -> ex P)].)

380 381 382
  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.

  Here is an example of a theorem statement with existential quantification.  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. *)
Adam Chlipala's avatar
Adam Chlipala committed
383 384

Theorem exist1 : exists x : nat, x + 1 = 2.
Adam Chlipala's avatar
Adam Chlipala committed
385
(* begin thide *)
386
  (** remove printing exists *)
387
  (** We can start this proof with a tactic %\index{tactics!exists}%[exists], which should not be confused with the formula constructor shorthand of the same name.  %In the version of this document that you are reading, the reverse ``E'' appears instead of the text ``exists'' in formulas.% *)
Adam Chlipala's avatar
Adam Chlipala committed
388

Adam Chlipala's avatar
Adam Chlipala committed
389 390
  exists 1.

Adam Chlipala's avatar
Adam Chlipala committed
391
  (** The conclusion is replaced with a version using the existential witness that we announced.
Adam Chlipala's avatar
Adam Chlipala committed
392

Adam Chlipala's avatar
Adam Chlipala committed
393
     [[
Adam Chlipala's avatar
Adam Chlipala committed
394 395
  ============================
   1 + 1 = 2
396 397
   ]]
   *)
Adam Chlipala's avatar
Adam Chlipala committed
398 399

  reflexivity.
Adam Chlipala's avatar
Adam Chlipala committed
400
(* end thide *)
Adam Chlipala's avatar
Adam Chlipala committed
401 402 403 404 405 406 407
Qed.

(** printing exists $\exists$ *)

(** We can also use tactics to reason about existential hypotheses. *)

Theorem exist2 : forall n m : nat, (exists x : nat, n + x = m) -> n <= m.
Adam Chlipala's avatar
Adam Chlipala committed
408
(* begin thide *)
Adam Chlipala's avatar
Adam Chlipala committed
409
  (** We start by case analysis on the proof of the existential fact. *)
Adam Chlipala's avatar
Adam Chlipala committed
410

Adam Chlipala's avatar
Adam Chlipala committed
411 412 413 414 415 416 417 418
  destruct 1.
  (** [[
  n : nat
  m : nat
  x : nat
  H : n + x = m
  ============================
   n <= m
Adam Chlipala's avatar
Adam Chlipala committed
419 420
 
   ]]
Adam Chlipala's avatar
Adam Chlipala committed
421

Adam Chlipala's avatar
Adam Chlipala committed
422
   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. *)
Adam Chlipala's avatar
Adam Chlipala committed
423 424

  crush.
Adam Chlipala's avatar
Adam Chlipala committed
425
(* end thide *)
Adam Chlipala's avatar
Adam Chlipala committed
426 427 428 429 430 431 432 433
Qed.


(* begin hide *)
(* In-class exercises *)

Theorem forall_exists_commute : forall (A B : Type) (P : A -> B -> Prop),
  (exists x : A, forall y : B, P x y) -> (forall y : B, exists x : A, P x y).
434 435 436 437 438 439
(* begin thide *)
  intros.
  destruct H.
  exists x.
  apply H.
(* end thide *)
Adam Chlipala's avatar
Adam Chlipala committed
440 441 442 443 444
Admitted.

(* end hide *)


Adam Chlipala's avatar
Adam Chlipala committed
445
(** 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. *)
446 447 448 449 450 451 452 453 454 455


(** * Predicates with Implicit Equality *)

(** 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. *)

Inductive isZero : nat -> Prop :=
| IsZero : isZero 0.

Theorem isZero_zero : isZero 0.
Adam Chlipala's avatar
Adam Chlipala committed
456
(* begin thide *)
457
  constructor.
Adam Chlipala's avatar
Adam Chlipala committed
458
(* end thide *)
459 460
Qed.

461
(** 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.  (Readers not familiar with formal semantics should not worry about not following this paragraph!)
462

463
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.
464

465
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 (but worry not if not familiar with 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}% *)
466 467

Print eq.
468
(** %\vspace{-.15in}%[[
469
  Inductive eq (A : Type) (x : A) : A -> Prop :=  eq_refl : x = x
Adam Chlipala's avatar
Adam Chlipala committed
470
  ]]
471

472
  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 [eq_refl], 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.
473

Adam Chlipala's avatar
Adam Chlipala committed
474
Returning to the example of [isZero], we can see how to work with hypotheses that use this predicate. *)
475 476

Theorem isZero_plus : forall n m : nat, isZero m -> n + m = n.
Adam Chlipala's avatar
Adam Chlipala committed
477
(* begin thide *)
478
  (** We want to proceed by cases on the proof of the assumption about [isZero]. *)
Adam Chlipala's avatar
Adam Chlipala committed
479

480 481 482 483 484
  destruct 1.
  (** [[
  n : nat
  ============================
   n + 0 = n
Adam Chlipala's avatar
Adam Chlipala committed
485 486
 
   ]]
487

Adam Chlipala's avatar
Adam Chlipala committed
488
   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. *)
489 490

  crush.
Adam Chlipala's avatar
Adam Chlipala committed
491
(* end thide *)
492 493 494 495 496
Qed.

(** 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. *)

Theorem isZero_contra : isZero 1 -> False.
Adam Chlipala's avatar
Adam Chlipala committed
497
(* begin thide *)
498
  (** Let us try a proof by cases on the assumption, as in the last proof. *)
Adam Chlipala's avatar
Adam Chlipala committed
499

500 501 502 503
  destruct 1.
  (** [[
  ============================
   False
Adam Chlipala's avatar
Adam Chlipala committed
504 505
 
   ]]
506

Adam Chlipala's avatar
Adam Chlipala committed
507
   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].
508

509
     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.  (Again, worry not if the semantics-oriented terminology from this last sentence is unfamiliar.) *)
510 511 512

  Undo.
  inversion 1.
Adam Chlipala's avatar
Adam Chlipala committed
513
(* end thide *)
514 515 516 517 518 519 520 521 522 523 524
Qed.

(** 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.

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. *)

Theorem isZero_contra' : isZero 1 -> 2 + 2 = 5.
  destruct 1.
  (** [[
  ============================
   1 + 1 = 4
Adam Chlipala's avatar
Adam Chlipala committed
525 526 527
 
   ]]

528
   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.  Then, within the [O] case of the proof, we replace the fresh variable with [O].  This has the net effect of decrementing each of these numbers. *)
529 530 531

Abort.

Adam Chlipala's avatar
Adam Chlipala committed
532 533 534 535 536 537 538 539
(** To see more clearly what is happening, we can consider the type of [isZero]'s induction principle. *)

Check isZero_ind.
(** %\vspace{-.15in}% [[
isZero_ind
     : forall P : nat -> Prop, P 0 -> forall n : nat, isZero n -> P n
   ]]

540
   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 Chlipala's avatar
Adam Chlipala committed
541

542 543 544 545 546 547

(* begin hide *)
(* In-class exercises *)

(* 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. *)

548 549 550 551 552 553 554 555 556 557 558 559 560 561 562 563 564 565
(* begin thide *)
Section twoEls.
  Variable A : Type.

  Inductive twoEls : list A -> Prop :=
  | TwoEls : forall x y, twoEls (x :: y :: nil).

  Theorem twoEls_nil : twoEls nil -> False.
    inversion 1.
  Qed.

  Theorem twoEls_two : forall ls, twoEls ls -> length ls = 2.
    inversion 1.
    reflexivity.
  Qed.
End twoEls.
(* end thide *)

566 567
(* end hide *)

Adam Chlipala's avatar
Adam Chlipala committed
568 569 570 571 572 573 574 575 576

(** * Recursive Predicates *)

(** We have already seen all of the ingredients we need to build interesting recursive predicates, like this predicate capturing even-ness. *)

Inductive even : nat -> Prop :=
| EvenO : even O
| EvenSS : forall n, even n -> even (S (S n)).

577
(** Think of [even] as another judgment defined by natural deduction rules.  The rule [EvenO] has 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.
Adam Chlipala's avatar
Adam Chlipala committed
578 579 580 581

The proof techniques of the last section are easily adapted. *)

Theorem even_0 : even 0.
Adam Chlipala's avatar
Adam Chlipala committed
582
(* begin thide *)
Adam Chlipala's avatar
Adam Chlipala committed
583
  constructor.
Adam Chlipala's avatar
Adam Chlipala committed
584
(* end thide *)
Adam Chlipala's avatar
Adam Chlipala committed
585 586 587
Qed.

Theorem even_4 : even 4.
Adam Chlipala's avatar
Adam Chlipala committed
588
(* begin thide *)
Adam Chlipala's avatar
Adam Chlipala committed
589
  constructor; constructor; constructor.
Adam Chlipala's avatar
Adam Chlipala committed
590
(* end thide *)
Adam Chlipala's avatar
Adam Chlipala committed
591 592
Qed.

593
(** 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. *)
Adam Chlipala's avatar
Adam Chlipala committed
594

Adam Chlipala's avatar
Adam Chlipala committed
595
(* begin thide *)
Adam Chlipala's avatar
Adam Chlipala committed
596 597 598 599 600 601
Hint Constructors even.

Theorem even_4' : even 4.
  auto.
Qed.

Adam Chlipala's avatar
Adam Chlipala committed
602 603
(* end thide *)

Adam Chlipala's avatar
Adam Chlipala committed
604 605
(** We may also use [inversion] with [even]. *)

Adam Chlipala's avatar
Adam Chlipala committed
606
Theorem even_1_contra : even 1 -> False.
Adam Chlipala's avatar
Adam Chlipala committed
607
(* begin thide *)
Adam Chlipala's avatar
Adam Chlipala committed
608
  inversion 1.
Adam Chlipala's avatar
Adam Chlipala committed
609
(* end thide *)
Adam Chlipala's avatar
Adam Chlipala committed
610 611 612
Qed.

Theorem even_3_contra : even 3 -> False.
Adam Chlipala's avatar
Adam Chlipala committed
613
(* begin thide *)
Adam Chlipala's avatar
Adam Chlipala committed
614 615 616 617 618 619 620 621
  inversion 1.
  (** [[
  H : even 3
  n : nat
  H1 : even 1
  H0 : n = 1
  ============================
   False
Adam Chlipala's avatar
Adam Chlipala committed
622 623
 
   ]]
Adam Chlipala's avatar
Adam Chlipala committed
624

Adam Chlipala's avatar
Adam Chlipala committed
625
   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. *)
Adam Chlipala's avatar
Adam Chlipala committed
626 627

  inversion H1.
Adam Chlipala's avatar
Adam Chlipala committed
628
(* end thide *)
Adam Chlipala's avatar
Adam Chlipala committed
629 630 631 632 633
Qed.

(** We can also do inductive proofs about [even]. *)

Theorem even_plus : forall n m, even n -> even m -> even (n + m).
Adam Chlipala's avatar
Adam Chlipala committed
634
(* begin thide *)
Adam Chlipala's avatar
Adam Chlipala committed
635
  (** It seems a reasonable first choice to proceed by induction on [n]. *)
Adam Chlipala's avatar
Adam Chlipala committed
636

Adam Chlipala's avatar
Adam Chlipala committed
637 638 639 640 641 642 643 644 645
  induction n; crush.
  (** [[
  n : nat
  IHn : forall m : nat, even n -> even m -> even (n + m)
  m : nat
  H : even (S n)
  H0 : even m
  ============================
   even (S (n + m))
Adam Chlipala's avatar
Adam Chlipala committed
646 647
 
   ]]
Adam Chlipala's avatar
Adam Chlipala committed
648

Adam Chlipala's avatar
Adam Chlipala committed
649
   We will need to use the hypotheses [H] and [H0] somehow.  The most natural choice is to invert [H]. *)
Adam Chlipala's avatar
Adam Chlipala committed
650 651 652 653 654 655 656 657 658 659 660 661 662

  inversion H.
  (** [[
  n : nat
  IHn : forall m : nat, even n -> even m -> even (n + m)
  m : nat
  H : even (S n)
  H0 : even m
  n0 : nat
  H2 : even n0
  H1 : S n0 = n
  ============================
   even (S (S n0 + m))
Adam Chlipala's avatar
Adam Chlipala committed
663 664 665 666
 
   ]]

  Simplifying the conclusion brings us to a point where we can apply a constructor. *)
Adam Chlipala's avatar
Adam Chlipala committed
667 668 669 670 671

  simpl.
  (** [[
  ============================
   even (S (S (n0 + m)))
672 673
   ]]
   *)
Adam Chlipala's avatar
Adam Chlipala committed
674 675

  constructor.
Adam Chlipala's avatar
Adam Chlipala committed
676

677
(** [[
Adam Chlipala's avatar
Adam Chlipala committed
678 679
  ============================
   even (n0 + m)
Adam Chlipala's avatar
Adam Chlipala committed
680 681
 
   ]]
Adam Chlipala's avatar
Adam Chlipala committed
682

Adam Chlipala's avatar
Adam Chlipala committed
683 684 685
   At this point, we would like to apply the inductive hypothesis, which is:

   [[
Adam Chlipala's avatar
Adam Chlipala committed
686
  IHn : forall m : nat, even n -> even m -> even (n + m)
687
 
Adam Chlipala's avatar
Adam Chlipala committed
688
  ]]
Adam Chlipala's avatar
Adam Chlipala committed
689

690
  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.
Adam Chlipala's avatar
Adam Chlipala committed
691 692

  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 Chlipala's avatar
Adam Chlipala committed
693 694 695 696 697 698 699 700

Restart.

  induction 1.
(** [[
  m : nat
  ============================
   even m -> even (0 + m)
Adam Chlipala's avatar
Adam Chlipala committed
701
]]
Adam Chlipala's avatar
Adam Chlipala committed
702

Adam Chlipala's avatar
Adam Chlipala committed
703 704
%\noindent \coqdockw{subgoal} 2 \coqdockw{is}:%#<tt>subgoal 2 is</tt>#
[[
Adam Chlipala's avatar
Adam Chlipala committed
705
 even m -> even (S (S n) + m)
Adam Chlipala's avatar
Adam Chlipala committed
706 707
 
 ]]
Adam Chlipala's avatar
Adam Chlipala committed
708

Adam Chlipala's avatar
Adam Chlipala committed
709
 The first case is easily discharged by [crush], based on the hint we added earlier to try the constructors of [even]. *)
Adam Chlipala's avatar
Adam Chlipala committed
710 711 712 713

  crush.

  (** Now we focus on the second case: *)
Adam Chlipala's avatar
Adam Chlipala committed
714

Adam Chlipala's avatar
Adam Chlipala committed
715 716 717 718 719 720 721 722 723
  intro.
  (** [[
  m : nat
  n : nat
  H : even n
  IHeven : even m -> even (n + m)
  H0 : even m
  ============================
   even (S (S n) + m)
Adam Chlipala's avatar
Adam Chlipala committed
724 725
 
   ]]
Adam Chlipala's avatar
Adam Chlipala committed
726

Adam Chlipala's avatar
Adam Chlipala committed
727
   We simplify and apply a constructor, as in our last proof attempt. *)
Adam Chlipala's avatar
Adam Chlipala committed
728 729

  simpl; constructor.
Adam Chlipala's avatar
Adam Chlipala committed
730

731
(** [[
Adam Chlipala's avatar
Adam Chlipala committed
732 733
  ============================
   even (n + m)
Adam Chlipala's avatar
Adam Chlipala committed
734 735
 
   ]]
Adam Chlipala's avatar
Adam Chlipala committed
736

Adam Chlipala's avatar
Adam Chlipala committed
737
   Now we have an exact match with our inductive hypothesis, and the remainder of the proof is trivial. *)
Adam Chlipala's avatar
Adam Chlipala committed
738 739 740 741 742 743

  apply IHeven; assumption.

  (** In fact, [crush] can handle all of the details of the proof once we declare the induction strategy. *)

Restart.
Adam Chlipala's avatar
Adam Chlipala committed
744

Adam Chlipala's avatar
Adam Chlipala committed
745
  induction 1; crush.
Adam Chlipala's avatar
Adam Chlipala committed
746
(* end thide *)
Adam Chlipala's avatar
Adam Chlipala committed
747 748 749 750 751
Qed.

(** Induction on recursive predicates has similar pitfalls to those we encountered with inversion in the last section. *)

Theorem even_contra : forall n, even (S (n + n)) -> False.
Adam Chlipala's avatar
Adam Chlipala committed
752
(* begin thide *)
Adam Chlipala's avatar
Adam Chlipala committed
753 754 755 756 757
  induction 1.
  (** [[
  n : nat
  ============================
   False
Adam Chlipala's avatar
Adam Chlipala committed
758
]]
Adam Chlipala's avatar
Adam Chlipala committed
759

Adam Chlipala's avatar
Adam Chlipala committed
760 761
%\noindent \coqdockw{subgoal} 2 \coqdockw{is}:%#<tt>subgoal 2 is</tt>#
[[
Adam Chlipala's avatar
Adam Chlipala committed
762
 False
Adam Chlipala's avatar
Adam Chlipala committed
763 764 765
 
 ]]

Adam Chlipala's avatar
Adam Chlipala committed
766
 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. *)
Adam Chlipala's avatar
Adam Chlipala committed
767 768 769 770 771 772

Abort.

Lemma even_contra' : forall n', even n' -> forall n, n' = S (n + n) -> False.
  induction 1; crush.

Adam Chlipala's avatar
Adam Chlipala committed
773
  (** 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. *)
Adam Chlipala's avatar
Adam Chlipala committed
774

Adam Chlipala's avatar
Adam Chlipala committed
775 776 777 778 779 780 781 782 783 784
  destruct n; destruct n0; crush.

  (** [[
  n : nat
  H : even (S n)
  IHeven : forall n0 : nat, S n = S (n0 + n0) -> False
  n0 : nat
  H0 : S n = n0 + S n0
  ============================
   False
Adam Chlipala's avatar
Adam Chlipala committed
785 786
 
   ]]
Adam Chlipala's avatar
Adam Chlipala committed
787

Adam Chlipala's avatar
Adam Chlipala committed
788
  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]. *)
Adam Chlipala's avatar
Adam Chlipala committed
789

Adam Chlipala's avatar
Adam Chlipala committed
790
  SearchRewrite (_ + S _).
791
(** %\vspace{-.15in}%[[
Adam Chlipala's avatar
Adam Chlipala committed
792
  plus_n_Sm : forall n m : nat, S (n + m) = n + S m
793 794
     ]]
     *)
Adam Chlipala's avatar
Adam Chlipala committed
795 796 797

  rewrite <- plus_n_Sm in H0.

Adam Chlipala's avatar
Adam Chlipala committed
798
  (** 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. *)
Adam Chlipala's avatar
Adam Chlipala committed
799

Adam Chlipala's avatar
Adam Chlipala committed
800 801
  apply IHeven with n0; assumption.

Adam Chlipala's avatar
Adam Chlipala committed
802
  (** 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. *)
Adam Chlipala's avatar
Adam Chlipala committed
803 804

  Restart.
Adam Chlipala's avatar
Adam Chlipala committed
805

806
  Hint Rewrite <- plus_n_Sm.
Adam Chlipala's avatar
Adam Chlipala committed
807 808 809 810 811 812 813

  induction 1; crush;
    match goal with
      | [ H : S ?N = ?N0 + ?N0 |- _ ] => destruct N; destruct N0
    end; crush; eauto.
Qed.

Adam Chlipala's avatar
Adam Chlipala committed
814
(** 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].
Adam Chlipala's avatar
Adam Chlipala committed
815

816
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.  For now, think of [eauto] as a potentially more expensive version of [auto] that considers more possible proofs; see Chapter 13 for more detail.  The quick summary is that [eauto] considers applying a lemma even when the form of the current goal doesn not uniquely determine the values of all of the lemma's quantified variables.
Adam Chlipala's avatar
Adam Chlipala committed
817 818 819 820

The original theorem now follows trivially from our lemma. *)

Theorem even_contra : forall n, even (S (n + n)) -> False.
821 822 823
  intros; eapply even_contra'; eauto.
Qed.

824
(** 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}%.
825 826 827 828 829 830 831 832 833

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. *)

Lemma even_contra'' : forall n' n, even n' -> n' = S (n + n) -> False.
  induction 1; crush;
    match goal with
      | [ H : S ?N = ?N0 + ?N0 |- _ ] => destruct N; destruct N0
    end; crush; eauto.

Adam Chlipala's avatar
Adam Chlipala committed
834
  (** One subgoal remains:
835

Adam Chlipala's avatar
Adam Chlipala committed
836
     [[
837 838 839 840 841
  n : nat
  H : even (S (n + n))
  IHeven : S (n + n) = S (S (S (n + n))) -> False
  ============================
   False
Adam Chlipala's avatar
Adam Chlipala committed
842 843
 
   ]]
844

845
   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. *)
846

Adam Chlipala's avatar
Adam Chlipala committed
847 848 849
Abort.

(** 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. *)
Adam Chlipala's avatar
Adam Chlipala committed
850
(* end thide *)
Adam Chlipala's avatar
Adam Chlipala committed
851

Adam Chlipala's avatar
Adam Chlipala committed
852

853 854 855 856 857


(* begin hide *)
(* In-class exercises *)

858
(* 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. *)
859 860 861 862 863 864 865 866 867 868 869 870 871 872 873 874 875 876 877 878 879 880 881

(* begin thide *)
Inductive prop : Set :=
| Tru : prop
| Fals : prop
| And : prop -> prop -> prop
| Or : prop -> prop -> prop.

Inductive holds : prop -> Prop :=
| HTru : holds Tru
| HAnd : forall p1 p2, holds p1 -> holds p2 -> holds (And p1 p2)
| HOr1 : forall p1 p2, holds p1 -> holds (Or p1 p2)
| HOr2 : forall p1 p2, holds p2 -> holds (Or p1 p2).

Inductive falseFree : prop -> Prop :=
| FFTru : falseFree Tru
| FFAnd : forall p1 p2, falseFree p1 -> falseFree p2 -> falseFree (And p1 p2)
| FFNot : forall p1 p2, falseFree p1 -> falseFree p2 -> falseFree (Or p1 p2).

Hint Constructors holds.

Theorem falseFree_holds : forall p, falseFree p -> holds p.
  induction 1; crush.
Adam Chlipala's avatar
Adam Chlipala committed
882
Qed.
883 884 885 886 887 888 889 890 891 892 893 894 895 896 897 898 899 900 901 902 903 904 905 906 907 908 909 910 911 912 913 914 915 916 917 918
(* end thide *)


(* 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]. *)

(* begin thide *)
Inductive prop' : Set :=
| Tru' : prop'
| And' : prop' -> prop' -> prop'
| Or' : prop' -> prop' -> prop'.

Inductive holds' : prop' -> Prop :=
| HTru' : holds' Tru'
| HAnd' : forall p1 p2, holds' p1 -> holds' p2 -> holds' (And' p1 p2)
| HOr1' : forall p1 p2, holds' p1 -> holds' (Or' p1 p2)
| HOr2' : forall p1 p2, holds' p2 -> holds' (Or' p1 p2).

Fixpoint propify (p : prop') : prop :=
  match p with
    | Tru' => Tru
    | And' p1 p2 => And (propify p1) (propify p2)
    | Or' p1 p2 => Or (propify p1) (propify p2)
  end.

Hint Constructors holds'.

Lemma propify_holds' : forall p', holds p' -> forall p, p' = propify p -> holds' p.
  induction 1; crush; destruct p; crush.
Qed.

Theorem propify_holds : forall p, holds (propify p) -> holds' p.
  intros; eapply propify_holds'; eauto.
Qed.
(* end thide *)

(* end hide *)