comparison src/Predicates.v @ 401:c898e72b84a3

Typesetting pass over Predicates
author Adam Chlipala <adam@chlipala.net>
date Fri, 08 Jun 2012 13:11:12 -0400
parents 5986e9fd40b5
children 10a6b5414551
comparison
equal deleted inserted replaced
400:d27c0be2c3d4 401:c898e72b84a3
32 32
33 (** %\chapter{Inductive Predicates}% *) 33 (** %\chapter{Inductive Predicates}% *)
34 34
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: *) 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: *)
36 36
37 (* begin hide *)
38 Print unit. 37 Print unit.
39 (* end hide *)
40 (** %\noindent%[Print] %\coqdocinductive{%#<tt>#unit#</tt>#%}%[.] *)
41 (** [[ 38 (** [[
42 Inductive unit : Set := tt : unit 39 Inductive unit : Set := tt : unit
43 ]] 40 ]]
44 *) 41 *)
45 42
46 (* begin hide *)
47 Print True. 43 Print True.
48 (* end hide *)
49 (** %\noindent%[Print] %\coqdocinductive{%#<tt>#True#</tt>#%}%[.] *)
50 (** [[ 44 (** [[
51 Inductive True : Prop := I : True 45 Inductive True : Prop := I : True
52 ]] 46 ]]
53 *) 47 *)
54 48
55 (** Recall that [unit] is the type with only one value, and [True] is the proposition that always holds. Despite this superficial difference between the two concepts, in both cases we can use the same inductive definition mechanism. The connection goes further than this. We see that we arrive at the definition of [True] by replacing [unit] by [True], [tt] by [I], and [Set] by [Prop]. The first two of these differences are superficial changes of names, while the third difference is the crucial one for separating programs from proofs. A term [T] of type [Set] is a type of programs, and a term of type [T] is a program. A term [T] of type [Prop] is a logical proposition, and its proofs are of type [T]. Chapter 12 goes into more detail about the theoretical differences between [Prop] and [Set]. For now, we will simply follow common intuitions about what a proof is. 49 (** 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.
56 50
57 The type [unit] has one value, [tt]. The type [True] has one proof, [I]. Why distinguish between these two types? Many people who have read about Curry-Howard in an abstract context and not put it to use in proof engineering answer that the two types in fact _should not_ be distinguished. There is a certain aesthetic appeal to this point of view, but I want to argue that it is best to treat Curry-Howard very loosely in practical proving. There are Coq-specific reasons for preferring the distinction, involving efficient compilation and avoidance of paradoxes in the presence of classical math, but I will argue that there is a more general principle that should lead us to avoid conflating programming and proving. 51 The type [unit] has one value, [tt]. The type [True] has one proof, [I]. Why distinguish between these two types? Many people who have read about Curry-Howard in an abstract context and not put it to use in proof engineering answer that the two types in fact _should not_ be distinguished. There is a certain aesthetic appeal to this point of view, but I want to argue that it is best to treat Curry-Howard very loosely in practical proving. There are Coq-specific reasons for preferring the distinction, involving efficient compilation and avoidance of paradoxes in the presence of classical math, but I will argue that there is a more general principle that should lead us to avoid conflating programming and proving.
58 52
59 The essence of the argument is roughly this: to an engineer, not all functions of type [A -> B] are created equal, but all proofs of a proposition [P -> Q] are. This idea is known as %\index{proof irrelevance}%_proof irrelevance_, and its formalizations in logics prevent us from distinguishing between alternate proofs of the same proposition. Proof irrelevance is compatible with, but not derivable in, Gallina. Apart from this theoretical concern, I will argue that it is most effective to do engineering with Coq by employing different techniques for programs versus proofs. Most of this book is organized around that distinction, describing how to program, by applying standard functional programming techniques in the presence of dependent types; and how to prove, by writing custom Ltac decision procedures. 53 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.
60 54
61 With that perspective in mind, this chapter is sort of a mirror image of the last chapter, introducing how to define predicates with inductive definitions. We will point out similarities in places, but much of the effective Coq user's bag of tricks is disjoint for predicates versus %``%#"#datatypes.#"#%''% This chapter is also a covert introduction to dependent types, which are the foundation on which interesting inductive predicates are built, though we will rely on tactics to build dependently typed proof terms for us for now. A future chapter introduces more manual application of dependent types. *) 55 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. *)
62 56
63 57
64 (** * Propositional Logic *) 58 (** * Propositional Logic *)
76 (* begin thide *) 70 (* begin thide *)
77 apply I. 71 apply I.
78 (* end thide *) 72 (* end thide *)
79 Qed. 73 Qed.
80 74
81 (** We may always use the [apply] tactic to take a proof step based on applying a particular constructor of the inductive predicate that we are trying to establish. Sometimes there is only one constructor that could possibly apply, in which case a shortcut is available: *) 75 (** 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}% *)
82 76
83 (* begin thide *) 77 (* begin thide *)
84 Theorem obvious' : True. 78 Theorem obvious' : True.
85 (* begin hide *)
86 constructor. 79 constructor.
87 (* end hide *)
88 (** %\hspace{.075in}\coqdockw{%#<tt>#constructor#</tt>#%}%.%\vspace{-.1in}% *)
89
90 Qed. 80 Qed.
91 81
92 (* end thide *) 82 (* end thide *)
93 83
94 (** There is also a predicate [False], which is the Curry-Howard mirror image of [Empty_set] from the last chapter. *) 84 (** There is also a predicate [False], which is the Curry-Howard mirror image of [Empty_set] from the last chapter. *)
95 85
96 (* begin hide *)
97 Print False. 86 Print False.
98 (* end hide *)
99 (** %\noindent%[Print] %\coqdocinductive{%#<tt>#False#</tt>#%}%[.] *)
100 (** [[ 87 (** [[
101 Inductive False : Prop := 88 Inductive False : Prop :=
102 89
103 ]] 90 ]]
104 91
156 (* end thide *) 143 (* end thide *)
157 Qed. 144 Qed.
158 145
159 (** We also have conjunction, which we introduced in the last chapter. *) 146 (** We also have conjunction, which we introduced in the last chapter. *)
160 147
161 (* begin hide *)
162 Print and. 148 Print and.
163 (* end hide *) 149 (** [[
164 (** %\noindent%[Print] %\coqdocinductive{%#<tt>#and#</tt>#%}%[.]
165 [[
166 Inductive and (A : Prop) (B : Prop) : Prop := conj : A -> B -> A /\ B 150 Inductive and (A : Prop) (B : Prop) : Prop := conj : A -> B -> A /\ B
167 151
168 ]] 152 ]]
169 153
170 The interested reader can check that [and] has a Curry-Howard equivalent called %\index{Gallina terms!prod}%[prod], the type of pairs. However, it is generally most convenient to reason about conjunction using tactics. An explicit proof of commutativity of [and] illustrates the usual suspects for such tasks. The operator [/\] is an infix shorthand for [and]. *) 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]. *)
207 (* end thide *) 191 (* end thide *)
208 Qed. 192 Qed.
209 193
210 (** Coq disjunction is called %\index{Gallina terms!or}%[or] and abbreviated with the infix operator [\/]. *) 194 (** Coq disjunction is called %\index{Gallina terms!or}%[or] and abbreviated with the infix operator [\/]. *)
211 195
212 (* begin hide *)
213 Print or. 196 Print or.
214 (* end hide *) 197 (** [[
215 (** %\noindent%[Print] %\coqdocinductive{%#<tt>#or#</tt>#%}%[.]
216 [[
217 Inductive or (A : Prop) (B : Prop) : Prop := 198 Inductive or (A : Prop) (B : Prop) : Prop :=
218 or_introl : A -> A \/ B | or_intror : B -> A \/ B 199 or_introl : A -> A \/ B | or_intror : B -> A \/ B
219 200
220 ]] 201 ]]
221 202
238 [[ 219 [[
239 Q \/ P 220 Q \/ P
240 221
241 ]] 222 ]]
242 223
243 We can see that, in the first subgoal, we want to prove the disjunction by proving its second disjunct. The %\index{tactics!right}\coqdockw{%#<tt>#right#</tt>#%}% tactic telegraphs this intent. *) 224 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. *)
244 225
245 (* begin hide *)
246 right; assumption. 226 right; assumption.
247 (* end hide *)
248 (** %\hspace{.075in}\coqdockw{%#<tt>#right#</tt>#%}%[; assumption.] *)
249 227
250 (** The second subgoal has a symmetric proof.%\index{tactics!left}% 228 (** The second subgoal has a symmetric proof.%\index{tactics!left}%
251 229
252 [[ 230 [[
253 1 subgoal 231 1 subgoal
256 ============================ 234 ============================
257 Q \/ P 235 Q \/ P
258 ]] 236 ]]
259 *) 237 *)
260 238
261 (* begin hide *)
262 left; assumption. 239 left; assumption.
263 (* end hide *)
264 (** %\hspace{.075in}\coqdockw{%#<tt>#left#</tt>#%}%[; assumption.] *)
265 240
266 (* end thide *) 241 (* end thide *)
267 Qed. 242 Qed.
268 243
269 244
318 (* begin thide *) 293 (* begin thide *)
319 tauto. 294 tauto.
320 (* end thide *) 295 (* end thide *)
321 Qed. 296 Qed.
322 297
323 (** Sometimes propositional reasoning forms important plumbing for the proof of a theorem, but we still need to apply some other smarts about, say, arithmetic. %\index{tactics!intuition}%[intuition] is a generalization of [tauto] that proves everything it can using propositional reasoning. When some goals remain, it uses propositional laws to simplify them as far as possible. Consider this example, which uses the list concatenation operator [++] from the standard library. *) 298 (** 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 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. *)
324 299
325 Theorem arith_comm : forall ls1 ls2 : list nat, 300 Theorem arith_comm : forall ls1 ls2 : list nat,
326 length ls1 = length ls2 \/ length ls1 + length ls2 = 6 301 length ls1 = length ls2 \/ length ls1 + length ls2 = 6
327 -> length (ls1 ++ ls2) = 6 \/ length ls1 = length ls2. 302 -> length (ls1 ++ ls2) = 6 \/ length ls1 = length ls2.
328 (* begin thide *) 303 (* begin thide *)
361 336
362 (* begin thide *) 337 (* begin thide *)
363 Theorem arith_comm' : forall ls1 ls2 : list nat, 338 Theorem arith_comm' : forall ls1 ls2 : list nat,
364 length ls1 = length ls2 \/ length ls1 + length ls2 = 6 339 length ls1 = length ls2 \/ length ls1 + length ls2 = 6
365 -> length (ls1 ++ ls2) = 6 \/ length ls1 = length ls2. 340 -> length (ls1 ++ ls2) = 6 \/ length ls1 = length ls2.
366 (* begin hide *)
367 Hint Rewrite app_length. 341 Hint Rewrite app_length.
368 (* end hide *)
369 (** %\hspace{.075in}%[Hint] %\coqdockw{%#<tt>#Rewrite#</tt>#%}% [app_length.] *)
370 342
371 crush. 343 crush.
372 Qed. 344 Qed.
373 (* end thide *) 345 (* end thide *)
374 346
377 (** Ending the section here has the same effect as always. Each of our propositional theorems becomes universally quantified over the propositional variables that we used. *) 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. *)
378 350
379 351
380 (** * What Does It Mean to Be Constructive? *) 352 (** * What Does It Mean to Be Constructive? *)
381 353
382 (** One potential point of confusion in the presentation so far is the distinction between [bool] and [Prop]. [bool] is a datatype whose two values are [true] and [false], while [Prop] is a more primitive type that includes among its members [True] and [False]. Why not collapse these two concepts into one, and why must there be more than two states of mathematical truth? 354 (** 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?
383 355
384 The answer comes from the fact that Coq implements %\index{constructive logic}%_constructive_ or %\index{intuitionistic logic|see{constructive logic}}%_intuitionistic_ logic, in contrast to the %\index{classical logic}%_classical_ logic that you may be more familiar with. In constructive logic, classical tautologies like [~ ~ P -> P] and [P \/ ~ P] do not always hold. In general, we can only prove these tautologies when [P] is %\index{decidability}%_decidable_, in the sense of %\index{computability|see{decidability}}%computability theory. The Curry-Howard encoding that Coq uses for [or] allows us to extract either a proof of [P] or a proof of [~ P] from any proof of [P \/ ~ P]. Since our proofs are just functional programs which we can run, a general %\index{law of the excluded middle}%law of the excluded middle would give us a decision procedure for the halting problem, where the instantiations of [P] would be formulas like %``%#"#this particular Turing machine halts.#"#%''% 356 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.#"#%''%
385 357
386 Hence the distinction between [bool] and [Prop]. Programs of type [bool] are computational by construction; we can always run them to determine their results. Many [Prop]s are undecidable, and so we can write more expressive formulas with [Prop]s than with [bool]s, but the inevitable consequence is that we cannot simply %``%#"#run a [Prop] to determine its truth.#"#%''% 358 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.#"#%''%
387 359
388 Constructive logic lets us define all of the logical connectives in an aesthetically appealing way, with orthogonal inductive definitions. That is, each connective is defined independently using a simple, shared mechanism. Constructivity also enables a trick called %\index{program extraction}%_program extraction_, where we write programs by phrasing them as theorems to be proved. Since our proofs are just functional programs, we can extract executable programs from our final proofs, which we could not do as naturally with classical proofs. 360 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.
389 361
390 We will see more about Coq's program extraction facility in a later chapter. However, I think it is worth interjecting another warning at this point, following up on the prior warning about taking the Curry-Howard correspondence too literally. It is possible to write programs by theorem-proving methods in Coq, but hardly anyone does it. It is almost always most useful to maintain the distinction between programs and proofs. If you write a program by proving a theorem, you are likely to run into algorithmic inefficiencies that you introduced in your proof to make it easier to prove. It is a shame to have to worry about such situations while proving tricky theorems, and it is a happy state of affairs that you almost certainly will not need to, with the ideal of extracting programs from proofs being confined mostly to theoretical studies. *) 362 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. *)
391 363
392 364
393 (** * First-Order Logic *) 365 (** * First-Order Logic *)
394 366
395 (** The %\index{Gallina terms!forall}%[forall] connective of first-order logic, which we have seen in many examples so far, is built into Coq. Getting ahead of ourselves a bit, we can see it as the dependent function type constructor. In fact, implication and universal quantification are just different syntactic shorthands for the same Coq mechanism. A formula [P -> Q] is equivalent to [forall x : P, Q], where [x] does not appear in [Q]. That is, the %``%#"#real#"#%''% type of the implication says %``%#"#for every proof of [P], there exists a proof of [Q].#"#%''% 367 (** 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].#"#%''%
396 368
397 %\index{existential quantification}\index{Gallina terms!exists}\index{Gallina terms!ex}%Existential quantification is defined in the standard library. *) 369 %\index{existential quantification}\index{Gallina terms!exists}\index{Gallina terms!ex}%Existential quantification is defined in the standard library. *)
398 370
399 (* begin hide *)
400 Print ex. 371 Print ex.
401 (* end hide *) 372 (** [[
402 (** %\noindent%[Print] %\coqdocinductive{%#<tt>#ex#</tt>#%}%[.]
403 [[
404 Inductive ex (A : Type) (P : A -> Prop) : Prop := 373 Inductive ex (A : Type) (P : A -> Prop) : Prop :=
405 ex_intro : forall x : A, P x -> ex P 374 ex_intro : forall x : A, P x -> ex P
406 375
407 ]] 376 ]]
408 377
409 The family [ex] is parameterized by the type [A] that we quantify over, and by a predicate [P] over [A]s. We prove an existential by exhibiting some [x] of type [A], along with a proof of [P x]. As usual, there are tactics that save us from worrying about the low-level details most of the time. We use the equality operator [=], which, depending on the settings in which they learned logic, different people will say either is or is not part of first-order logic. For our purposes, it is. *) 378 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. *)
410 379
411 Theorem exist1 : exists x : nat, x + 1 = 2. 380 Theorem exist1 : exists x : nat, x + 1 = 2.
412 (* begin thide *) 381 (* begin thide *)
413 (** remove printing exists *) 382 (** remove printing exists *)
414 (** We can start this proof with a tactic %\index{tactics!exists}\coqdockw{%exists%}%, which should not be confused with the formula constructor shorthand of the same name. (In the PDF version of this document, the reverse %`%#'#E#'#%'% appears instead of the text %``%#"#exists#"#%''% in formulas.) *) 383 (** 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 PDF version of this document, the reverse %`%#'#E#'#%'% appears instead of the text %``%#"#exists#"#%''% in formulas.) *)
415 384
416 (* begin hide *)
417 exists 1. 385 exists 1.
418 (* end hide *)
419 (** %\coqdockw{%#<tt>#exists#</tt>#%}% [1.] *)
420 386
421 (** The conclusion is replaced with a version using the existential witness that we announced. 387 (** The conclusion is replaced with a version using the existential witness that we announced.
422 388
423 [[ 389 [[
424 ============================ 390 ============================
482 Inductive isZero : nat -> Prop := 448 Inductive isZero : nat -> Prop :=
483 | IsZero : isZero 0. 449 | IsZero : isZero 0.
484 450
485 Theorem isZero_zero : isZero 0. 451 Theorem isZero_zero : isZero 0.
486 (* begin thide *) 452 (* begin thide *)
487 (* begin hide *)
488 constructor. 453 constructor.
489 (* end hide *) 454 (* end thide *)
490 (** %\coqdockw{%#<tt>#constructor#</tt>#%}%[.]%\vspace{-.075in}% *) 455 Qed.
491 456
492 (* end thide *) 457 (** 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.
493 Qed.
494
495 (** We can call [isZero] a %\index{judgment}%_judgment_, in the sense often used in the semantics of programming languages. Judgments are typically defined in the style of %\index{natural deduction}%_natural deduction_, where we write a number of %\index{inference rules}%_inference rules_ with premises appearing above a solid line and a conclusion appearing below the line. In this example, the sole constructor [IsZero] of [isZero] can be thought of as the single inference rule for deducing [isZero], with nothing above the line and [isZero 0] below it. The proof of [isZero_zero] demonstrates how we can apply an inference rule.
496 458
497 The definition of [isZero] differs in an important way from all of the other inductive definitions that we have seen in this and the previous chapter. Instead of writing just [Set] or [Prop] after the colon, here we write [nat -> Prop]. We saw examples of parameterized types like [list], but there the parameters appeared with names _before_ the colon. Every constructor of a parameterized inductive type must have a range type that uses the same parameter, whereas the form we use here enables us to use different arguments to the type for different constructors. 459 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.
498 460
499 For instance, our definition [isZero] makes the predicate provable only when the argument is [0]. We can see that the concept of equality is somehow implicit in the inductive definition mechanism. The way this is accomplished is similar to the way that logic variables are used in %\index{Prolog}%Prolog, and it is a very powerful mechanism that forms a foundation for formalizing all of mathematics. In fact, though it is natural to think of inductive types as folding in the functionality of equality, in Coq, the true situation is reversed, with equality defined as just another inductive type!%\index{Gallina terms!eq}\index{Gallina terms!refl\_equal}% *) 461 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}% *)
500 462
501 (* begin hide *)
502 Print eq. 463 Print eq.
503 (* end hide *) 464 (** [[
504 (** %\noindent%[Print] %\coqdocinductive{%#<tt>#eq#</tt>#%}%[.]
505 [[
506 Inductive eq (A : Type) (x : A) : A -> Prop := refl_equal : x = x 465 Inductive eq (A : Type) (x : A) : A -> Prop := refl_equal : x = x
507 466
508 ]] 467 ]]
509 468
510 Behind the scenes, uses of infix [=] are expanded to instances of [eq]. We see that [eq] has both a parameter [x] that is fixed and an extra unnamed argument of the same type. The type of [eq] allows us to state any equalities, even those that are provably false. However, examining the type of equality's sole constructor [refl_equal], we see that we can only _prove_ equality when its two arguments are syntactically equal. This definition turns out to capture all of the basic properties of equality, and the equality-manipulating tactics that we have seen so far, like [reflexivity] and [rewrite], are implemented treating [eq] as just another inductive type with a well-chosen definition. Another way of stating that definition is: equality is defined as the least reflexive relation. 469 Behind the scenes, uses of infix [=] are expanded to instances of [eq]. We see that [eq] has both a parameter [x] that is fixed and an extra unnamed argument of the same type. The type of [eq] allows us to state any equalities, even those that are provably false. However, examining the type of equality's sole constructor [refl_equal], we see that we can only _prove_ equality when its two arguments are syntactically equal. This definition turns out to capture all of the basic properties of equality, and the equality-manipulating tactics that we have seen so far, like [reflexivity] and [rewrite], are implemented treating [eq] as just another inductive type with a well-chosen definition. Another way of stating that definition is: equality is defined as the least reflexive relation.
574 isZero_ind 533 isZero_ind
575 : forall P : nat -> Prop, P 0 -> forall n : nat, isZero n -> P n 534 : forall P : nat -> Prop, P 0 -> forall n : nat, isZero n -> P n
576 535
577 ]] 536 ]]
578 537
579 In our last proof script, [destruct] chose to instantiate [P] as [fun n => S n + S n = S (][S (][S (][S n)))]. You can verify for yourself that this specialization of the principle applies to the goal and that the hypothesis [P 0] then matches the subgoal we saw generated. If you are doing a proof and encounter a strange transmutation like this, there is a good chance that you should go back and replace a use of [destruct] with [inversion]. *) 538 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]. *)
580 539
581 540
582 (* begin hide *) 541 (* begin hide *)
583 (* In-class exercises *) 542 (* In-class exercises *)
584 543
611 570
612 Inductive even : nat -> Prop := 571 Inductive even : nat -> Prop :=
613 | EvenO : even O 572 | EvenO : even O
614 | EvenSS : forall n, even n -> even (S (S n)). 573 | EvenSS : forall n, even n -> even (S (S n)).
615 574
616 (** Think of [even] as another judgment defined by natural deduction rules. [EvenO] is a rule with nothing above the line and [even O] below the line, and [EvenSS] is a rule with [even n] above the line and [even (][S (][S n))] below. 575 (** 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.
617 576
618 The proof techniques of the last section are easily adapted. *) 577 The proof techniques of the last section are easily adapted. *)
619 578
620 Theorem even_0 : even 0. 579 Theorem even_0 : even 0.
621 (* begin thide *) 580 (* begin thide *)
622 (* begin hide *)
623 constructor. 581 constructor.
624 (* end hide *)
625 (** %\coqdockw{%#<tt>#constructor#</tt>#%}%[.]%\vspace{-.075in}% *)
626
627 (* end thide *) 582 (* end thide *)
628 Qed. 583 Qed.
629 584
630 Theorem even_4 : even 4. 585 Theorem even_4 : even 4.
631 (* begin thide *) 586 (* begin thide *)
632 (* begin hide *)
633 constructor; constructor; constructor. 587 constructor; constructor; constructor.
634 (* end hide *)
635 (** %\coqdockw{%#<tt>#constructor#</tt>#%}%[; ]%\coqdockw{%#<tt>#constructor#</tt>#%}%[; ]%\coqdockw{%#<tt>#constructor#</tt>#%}%[.]%\vspace{-.075in}% *)
636
637 (* end thide *) 588 (* end thide *)
638 Qed. 589 Qed.
639 590
640 (** It is not hard to see that sequences of constructor applications like the above can get tedious. We can avoid them using Coq's hint facility, with a new [Hint] variant that asks to consider all constructors of an inductive type during proof search. The tactic %\index{tactics!auto}%[auto] performs exhaustive proof search up to a fixed depth, considering only the proof steps we have registered as hints. *) 591 (** 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. *)
641 592
642 (* begin thide *) 593 (* begin thide *)
643 (* begin hide *)
644 Hint Constructors even. 594 Hint Constructors even.
645 (* end hide *)
646 (** %\noindent%[Hint] %\coqdockw{%#<tt>#Constructors#</tt>#%}% [even.] *)
647 595
648 Theorem even_4' : even 4. 596 Theorem even_4' : even 4.
649 auto. 597 auto.
650 Qed. 598 Qed.
651 599
720 ============================ 668 ============================
721 even (S (S (n0 + m))) 669 even (S (S (n0 + m)))
722 ]] 670 ]]
723 *) 671 *)
724 672
725 (* begin hide *)
726 constructor. 673 constructor.
727 (* end hide *) 674
728 (** %\coqdockw{%#<tt>#constructor#</tt>#%}%[.] 675 (** [[
729
730 [[
731 ============================ 676 ============================
732 even (n0 + m) 677 even (n0 + m)
733 678
734 ]] 679 ]]
735 680
737 682
738 [[ 683 [[
739 IHn : forall m : nat, even n -> even m -> even (n + m) 684 IHn : forall m : nat, even n -> even m -> even (n + m)
740 ]] 685 ]]
741 686
742 Unfortunately, the goal mentions [n0] where it would need to mention [n] to match [IHn]. We could keep looking for a way to finish this proof from here, but it turns out that we can make our lives much easier by changing our basic strategy. Instead of inducting on the structure of [n], we should induct _on the structure of one of the [even] proofs_. This technique is commonly called %\index{rule induction}%_rule induction_ in programming language semantics. In the setting of Coq, we have already seen how predicates are defined using the same inductive type mechanism as datatypes, so the fundamental unity of rule induction with %``%#"#normal#"#%''% induction is apparent. 687 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.
743 688
744 Recall that tactics like [induction] and [destruct] may be passed numbers to refer to unnamed lefthand sides of implications in the conclusion, where the argument [n] refers to the [n]th such hypothesis. *) 689 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. *)
745 690
746 (* begin hide *)
747 Restart. 691 Restart.
748 (* end hide *)
749 (** %\noindent\coqdockw{%#<tt>#Restart#</tt>#%}%[.] *)
750 692
751 induction 1. 693 induction 1.
752 (** [[ 694 (** [[
753 m : nat 695 m : nat
754 ============================ 696 ============================
779 721
780 ]] 722 ]]
781 723
782 We simplify and apply a constructor, as in our last proof attempt. *) 724 We simplify and apply a constructor, as in our last proof attempt. *)
783 725
784 (* begin hide *)
785 simpl; constructor. 726 simpl; constructor.
786 (* end hide *) 727
787 (** [simpl; ]%\coqdockw{%#<tt>#constructor#</tt>#%}%[.] 728 (** [[
788
789 [[
790 ============================ 729 ============================
791 even (n + m) 730 even (n + m)
792 731
793 ]] 732 ]]
794 733
796 735
797 apply IHeven; assumption. 736 apply IHeven; assumption.
798 737
799 (** In fact, [crush] can handle all of the details of the proof once we declare the induction strategy. *) 738 (** In fact, [crush] can handle all of the details of the proof once we declare the induction strategy. *)
800 739
801 (* begin hide *)
802 Restart. 740 Restart.
803 (* end hide *)
804 (** %\noindent\coqdockw{%#<tt>#Restart#</tt>#%}%[.] *)
805 741
806 induction 1; crush. 742 induction 1; crush.
807 (* end thide *) 743 (* end thide *)
808 Qed. 744 Qed.
809 745
846 782
847 ]] 783 ]]
848 784
849 At this point it is useful to use a theorem from the standard library, which we also proved with a different name in the last chapter. We can search for a theorem that allows us to rewrite terms of the form [x + S y]. *) 785 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]. *)
850 786
851 (* begin hide *)
852 SearchRewrite (_ + S _). 787 SearchRewrite (_ + S _).
853 (* end hide *) 788
854 (** %\coqdockw{%#<tt>#SearchRewrite#</tt>#%}% [(_ + S _).] 789 (** [[
855
856 [[
857 plus_n_Sm : forall n m : nat, S (n + m) = n + S m 790 plus_n_Sm : forall n m : nat, S (n + m) = n + S m
858 ]] 791 ]]
859 *) 792 *)
860 793
861 rewrite <- plus_n_Sm in H0. 794 rewrite <- plus_n_Sm in H0.
864 797
865 apply IHeven with n0; assumption. 798 apply IHeven with n0; assumption.
866 799
867 (** As usual, we can rewrite the proof to avoid referencing any locally generated names, which makes our proof script more readable and more robust to changes in the theorem statement. We use the notation [<-] to request a hint that does right-to-left rewriting, just like we can with the [rewrite] tactic. *) 800 (** 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. *)
868 801
869 (* begin hide *)
870 Restart. 802 Restart.
871 (* end hide *) 803
872 (** %\hspace{-.075in}\coqdockw{%#<tt>#Restart#</tt>#%}%[.] *)
873
874 (* begin hide *)
875 Hint Rewrite <- plus_n_Sm. 804 Hint Rewrite <- plus_n_Sm.
876 (* end hide *)
877 (** %\hspace{-.075in}%[Hint] %\noindent\coqdockw{%#<tt>#Rewrite#</tt>#%}% [<- plus_n_sm.] *)
878 805
879 induction 1; crush; 806 induction 1; crush;
880 match goal with 807 match goal with
881 | [ H : S ?N = ?N0 + ?N0 |- _ ] => destruct N; destruct N0 808 | [ H : S ?N = ?N0 + ?N0 |- _ ] => destruct N; destruct N0
882 end; crush; eauto. 809 end; crush; eauto.