comparison src/Predicates.v @ 322:98e0d957df16

Pass through Chapter 4
author Adam Chlipala <adam@chlipala.net>
date Mon, 19 Sep 2011 14:04:09 -0400
parents d5787b70cf48
children 3513d8b0531a
comparison
equal deleted inserted replaced
321:a9edc6ecd904 322:98e0d957df16
1 (* Copyright (c) 2008-2010, Adam Chlipala 1 (* Copyright (c) 2008-2011, Adam Chlipala
2 * 2 *
3 * This work is licensed under a 3 * This work is licensed under a
4 * Creative Commons Attribution-Noncommercial-No Derivative Works 3.0 4 * Creative Commons Attribution-Noncommercial-No Derivative Works 3.0
5 * Unported License. 5 * Unported License.
6 * The license text is available at: 6 * The license text is available at:
11 Require Import List. 11 Require Import List.
12 12
13 Require Import CpdtTactics. 13 Require Import CpdtTactics.
14 14
15 Set Implicit Arguments. 15 Set Implicit Arguments.
16 (* end hide *) 16
17 17 (* Extra definitions to get coqdoc to choose the right fonts. *)
18
19 Inductive unit := tt.
20 Inductive Empty_set := .
21 Inductive bool := true | false.
22 Inductive sum := .
23 Inductive prod := .
24 Inductive and := conj.
25 Inductive or := or_introl | or_intror.
26 Inductive ex := ex_intro.
27 Inductive eq := refl_equal.
28 Reset unit.
29 (* end hide *)
18 30
19 (** %\chapter{Inductive Predicates}% *) 31 (** %\chapter{Inductive Predicates}% *)
20 32
21 (** The so-called %``%#"#Curry-Howard Correspondence#"#%''% states a formal connection between functional programs and mathematical proofs. In the last chapter, we snuck in a first introduction to this subject in Coq. Witness the close similarity between the types [unit] and [True] from the standard library: *) 33 (** The so-called %\index{Curry-Howard correspondence}``%#"#Curry-Howard correspondence#"#%''~\cite{Curry,Howard}% states a formal connection between functional programs and mathematical proofs. In the last chapter, we snuck in a first introduction to this subject in Coq. Witness the close similarity between the types [unit] and [True] from the standard library: *)
22 34
35 (* begin hide *)
23 Print unit. 36 Print unit.
24 (** %\vspace{-.15in}% [[ 37 (* end hide *)
38 (** %\noindent%[Print] %\coqdocinductive{%#<tt>#unit#</tt>#%}%[.] *)
39 (** [[
25 Inductive unit : Set := tt : unit 40 Inductive unit : Set := tt : unit
26 ]] 41 ]]
27 *) 42 *)
28 43
44 (* begin hide *)
29 Print True. 45 Print True.
30 (** %\vspace{-.15in}% [[ 46 (* end hide *)
47 (** %\noindent%[Print] %\coqdocinductive{%#<tt>#True#</tt>#%}%[.] *)
48 (** [[
31 Inductive True : Prop := I : True 49 Inductive True : Prop := I : True
32 ]] 50 ]]
33 *) 51 *)
34 52
35 (** Recall that [unit] is the type with only one value, and [True] is the proposition that always holds. Despite this superficial difference between the two concepts, in both cases we can use the same inductive definition mechanism. The connection goes further than this. We see that we arrive at the definition of [True] by replacing [unit] by [True], [tt] by [I], and [Set] by [Prop]. The first two of these differences are superficial changes of names, while the third difference is the crucial one for separating programs from proofs. A term [T] of type [Set] is a type of programs, and a term of type [T] is a program. A term [T] of type [Prop] is a logical proposition, and its proofs are of type [T]. 53 (** Recall that [unit] is the type with only one value, and [True] is the proposition that always holds. Despite this superficial difference between the two concepts, in both cases we can use the same inductive definition mechanism. The connection goes further than this. We see that we arrive at the definition of [True] by replacing [unit] by [True], [tt] by [I], and [Set] by [Prop]. The first two of these differences are superficial changes of names, while the third difference is the crucial one for separating programs from proofs. A term [T] of type [Set] is a type of programs, and a term of type [T] is a program. A term [T] of type [Prop] is a logical proposition, and its proofs are of type [T]. Chapter 11 goes into more detail about the theoretical differences between [Prop] and [Set]. For now, we will simply follow common intuitions about what a proof is.
36 54
37 [unit] has one value, [tt]. [True] has one proof, [I]. Why distinguish between these two types? Many people who have read about Curry-Howard in an abstract context and not put it to use in proof engineering answer that the two types in fact %\textit{%#<i>#should not#</i>#%}% be distinguished. There is a certain aesthetic appeal to this point of view, but I want to argue that it is best to treat Curry-Howard very loosely in practical proving. There are Coq-specific reasons for preferring the distinction, involving efficient compilation and avoidance of paradoxes in the presence of classical math, but I will argue that there is a more general principle that should lead us to avoid conflating programming and proving. 55 The type [unit] has one value, [tt]. The type [True] has one proof, [I]. Why distinguish between these two types? Many people who have read about Curry-Howard in an abstract context and not put it to use in proof engineering answer that the two types in fact %\textit{%#<i>#should not#</i>#%}% be distinguished. There is a certain aesthetic appeal to this point of view, but I want to argue that it is best to treat Curry-Howard very loosely in practical proving. There are Coq-specific reasons for preferring the distinction, involving efficient compilation and avoidance of paradoxes in the presence of classical math, but I will argue that there is a more general principle that should lead us to avoid conflating programming and proving.
38 56
39 The essence of the argument is roughly this: to an engineer, not all functions of type [A -> B] are created equal, but all proofs of a proposition [P -> Q] are. This idea is known as %\textit{%#<i>#proof irrelevance#</i>#%}%, and its formalizations in logics prevent us from distinguishing between alternate proofs of the same proposition. Proof irrelevance is compatible with, but not derivable in, Gallina. Apart from this theoretical concern, I will argue that it is most effective to do engineering with Coq by employing different techniques for programs versus proofs. Most of this book is organized around that distinction, describing how to program, by applying standard functional programming techniques in the presence of dependent types; and how to prove, by writing custom Ltac decision procedures. 57 The essence of the argument is roughly this: to an engineer, not all functions of type [A -> B] are created equal, but all proofs of a proposition [P -> Q] are. This idea is known as %\index{proof irrelevance}\textit{%#<i>#proof irrelevance#</i>#%}%, and its formalizations in logics prevent us from distinguishing between alternate proofs of the same proposition. Proof irrelevance is compatible with, but not derivable in, Gallina. Apart from this theoretical concern, I will argue that it is most effective to do engineering with Coq by employing different techniques for programs versus proofs. Most of this book is organized around that distinction, describing how to program, by applying standard functional programming techniques in the presence of dependent types; and how to prove, by writing custom Ltac decision procedures.
40 58
41 With that perspective in mind, this chapter is sort of a mirror image of the last chapter, introducing how to define predicates with inductive definitions. We will point out similarities in places, but much of the effective Coq user's bag of tricks is disjoint for predicates versus %``%#"#datatypes.#"#%''% This chapter is also a covert introduction to dependent types, which are the foundation on which interesting inductive predicates are built, though we will rely on tactics to build dependently-typed proof terms for us for now. A future chapter introduces more manual application of dependent types. *) 59 With that perspective in mind, this chapter is sort of a mirror image of the last chapter, introducing how to define predicates with inductive definitions. We will point out similarities in places, but much of the effective Coq user's bag of tricks is disjoint for predicates versus %``%#"#datatypes.#"#%''% This chapter is also a covert introduction to dependent types, which are the foundation on which interesting inductive predicates are built, though we will rely on tactics to build dependently-typed proof terms for us for now. A future chapter introduces more manual application of dependent types. *)
42 60
43 61
44 (** * Propositional Logic *) 62 (** * Propositional Logic *)
60 78
61 (** 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: *) 79 (** We may always use the [apply] tactic to take a proof step based on applying a particular constructor of the inductive predicate that we are trying to establish. Sometimes there is only one constructor that could possibly apply, in which case a shortcut is available: *)
62 80
63 (* begin thide *) 81 (* begin thide *)
64 Theorem obvious' : True. 82 Theorem obvious' : True.
83 (* begin hide *)
65 constructor. 84 constructor.
85 (* end hide *)
86 (** %\hspace{.075in}\coqdockw{%#<tt>#constructor#</tt>#%}%.%\vspace{-.1in}% *)
87
66 Qed. 88 Qed.
67 89
68 (* end thide *) 90 (* end thide *)
69 91
70 (** There is also a predicate [False], which is the Curry-Howard mirror image of [Empty_set] from the last chapter. *) 92 (** There is also a predicate [False], which is the Curry-Howard mirror image of [Empty_set] from the last chapter. *)
71 93
94 (* begin hide *)
72 Print False. 95 Print False.
73 (** %\vspace{-.15in}% [[ 96 (* end hide *)
97 (** %\noindent%[Print] %\coqdocinductive{%#<tt>#False#</tt>#%}%[.] *)
98 (** [[
74 Inductive False : Prop := 99 Inductive False : Prop :=
75 100
76 ]] 101 ]]
77 102
78 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. *) 103 We can conclude anything from [False], doing case analysis on a proof of [False] in the same way we might do case analysis on, say, a natural number. Since there are no cases to consider, any such case analysis succeeds immediately in proving the goal. *)
87 112
88 Theorem arith_neq : 2 + 2 = 5 -> 9 + 9 = 835. 113 Theorem arith_neq : 2 + 2 = 5 -> 9 + 9 = 835.
89 (* begin thide *) 114 (* begin thide *)
90 intro. 115 intro.
91 116
92 (** At this point, we have an inconsistent hypothesis [2 + 2 = 5], so the specific conclusion is not important. We use the [elimtype] tactic to state a proposition, telling Coq that we wish to construct a proof of the new proposition and then prove the original goal by case analysis on the structure of the new auxiliary proof. Since [False] has no constructors, [elimtype False] simply leaves us with the obligation to prove [False]. *) 117 (** At this point, we have an inconsistent hypothesis [2 + 2 = 5], so the specific conclusion is not important. We use the %\index{tactics!elimtype}%[elimtype] tactic to state a proposition, telling Coq that we wish to construct a proof of the new proposition and then prove the original goal by case analysis on the structure of the new auxiliary proof. Since [False] has no constructors, [elimtype False] simply leaves us with the obligation to prove [False]. *)
93 118
94 elimtype False. 119 elimtype False.
95 (** [[ 120 (** [[
96 H : 2 + 2 = 5 121 H : 2 + 2 = 5
97 ============================ 122 ============================
129 (* end thide *) 154 (* end thide *)
130 Qed. 155 Qed.
131 156
132 (** We also have conjunction, which we introduced in the last chapter. *) 157 (** We also have conjunction, which we introduced in the last chapter. *)
133 158
159 (* begin hide *)
134 Print and. 160 Print and.
135 (** %\vspace{-.15in}% [[ 161 (* end hide *)
136 Inductive and (A : Prop) (B : Prop) : Prop := conj : A -> B -> A /\ B 162 (** %\noindent%[Print] %\coqdocinductive{%#<tt>#and#</tt>#%}%[.]
163 [[
164 Inductive and (A : Prop) (B : Prop) : Prop := conj : A -> B -> A /\ B
137 165
138 ]] 166 ]]
139 167
140 The interested reader can check that [and] has a Curry-Howard doppelganger called [prod], the type of pairs. However, it is generally most convenient to reason about conjunction using tactics. An explicit proof of commutativity of [and] illustrates the usual suspects for such tasks. [/\] is an infix shorthand for [and]. *) 168 The interested reader can check that [and] has a Curry-Howard equivalent called %\index{Gallina terms!prod}%[prod], the type of pairs. However, it is generally most convenient to reason about conjunction using tactics. An explicit proof of commutativity of [and] illustrates the usual suspects for such tasks. The operator [/\] is an infix shorthand for [and]. *)
141 169
142 Theorem and_comm : P /\ Q -> Q /\ P. 170 Theorem and_comm : P /\ Q -> Q /\ P.
143 171
144 (* begin thide *) 172 (* begin thide *)
145 (** We start by case analysis on the proof of [P /\ Q]. *) 173 (** We start by case analysis on the proof of [P /\ Q]. *)
151 ============================ 179 ============================
152 Q /\ P 180 Q /\ P
153 181
154 ]] 182 ]]
155 183
156 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]. *) 184 Every proof of a conjunction provides proofs for both conjuncts, so we get a single subgoal reflecting that. We can proceed by splitting this subgoal into a case for each conjunct of [Q /\ P].%\index{tactics!split}% *)
157 185
158 split. 186 split.
159 (** [[ 187 (** %\vspace{.1in} \noindent 2 \coqdockw{subgoals}\vspace{-.1in}%#<tt>2 subgoals</tt>#
160 2 subgoals 188 [[
161 189
162 H : P 190 H : P
163 H0 : Q 191 H0 : Q
164 ============================ 192 ============================
165 Q 193 Q
166 194 ]]
167 subgoal 2 is: 195 %\noindent \coqdockw{subgoal} 2 \coqdockw{is}:%#<tt>subgoal 2 is</tt>#
168 P 196 [[
197 P
169 198
170 ]] 199 ]]
171 200
172 In each case, the conclusion is among our hypotheses, so the [assumption] tactic finishes the process. *) 201 In each case, the conclusion is among our hypotheses, so the %\index{tactics!assumption}%[assumption] tactic finishes the process. *)
173 202
174 assumption. 203 assumption.
175 assumption. 204 assumption.
176 (* end thide *) 205 (* end thide *)
177 Qed. 206 Qed.
178 207
179 (** Coq disjunction is called [or] and abbreviated with the infix operator [\/]. *) 208 (** Coq disjunction is called %\index{Gallina terms!or}%[or] and abbreviated with the infix operator [\/]. *)
180 209
210 (* begin hide *)
181 Print or. 211 Print or.
182 (** %\vspace{-.15in}% [[ 212 (* end hide *)
213 (** %\noindent%[Print] %\coqdocinductive{%#<tt>#or#</tt>#%}%[.]
214 [[
183 Inductive or (A : Prop) (B : Prop) : Prop := 215 Inductive or (A : Prop) (B : Prop) : Prop :=
184 or_introl : A -> A \/ B | or_intror : B -> A \/ B 216 or_introl : A -> A \/ B | or_intror : B -> A \/ B
185 217
186 ]] 218 ]]
187 219
188 We see that there are two ways to prove a disjunction: prove the first disjunct or prove the second. The Curry-Howard analogue of this is the Coq [sum] type. We can demonstrate the main tactics here with another proof of commutativity. *) 220 We see that there are two ways to prove a disjunction: prove the first disjunct or prove the second. The Curry-Howard analogue of this is the Coq %\index{Gallina terms!sum}%[sum] type. We can demonstrate the main tactics here with another proof of commutativity. *)
189 221
190 Theorem or_comm : P \/ Q -> Q \/ P. 222 Theorem or_comm : P \/ Q -> Q \/ P.
191 223
192 (* begin thide *) 224 (* begin thide *)
193 (** As in the proof for [and], we begin with case analysis, though this time we are met by two cases instead of one. *) 225 (** As in the proof for [and], we begin with case analysis, though this time we are met by two cases instead of one. *)
194 226
195 destruct 1. 227 destruct 1.
196 (** [[ 228 (** %\vspace{.1in} \noindent 2 \coqdockw{subgoals}\vspace{-.1in}%#<tt>2 subgoals</tt>#
197 2 subgoals 229 [[
198 230
199 H : P 231 H : P
200 ============================ 232 ============================
201 Q \/ P 233 Q \/ P
202 234 ]]
203 subgoal 2 is: 235 %\noindent \coqdockw{subgoal} 2 \coqdockw{is}:%#<tt>subgoal 2 is</tt>#
236 [[
204 Q \/ P 237 Q \/ P
205 238
206 ]] 239 ]]
207 240
208 We can see that, in the first subgoal, we want to prove the disjunction by proving its second disjunct. The [right] tactic telegraphs this intent. *) 241 We can see that, in the first subgoal, we want to prove the disjunction by proving its second disjunct. The %\index{tactics!right}\coqdockw{%#<tt>#right#</tt>#%}% tactic telegraphs this intent. *)
209 242
243 (* begin hide *)
210 right; assumption. 244 right; assumption.
211 245 (* end hide *)
212 (** The second subgoal has a symmetric proof. 246 (** %\hspace{.075in}\coqdockw{%#<tt>#right#</tt>#%}%[; assumption.] *)
247
248 (** The second subgoal has a symmetric proof.%\index{tactics!left}%
213 249
214 [[ 250 [[
215 1 subgoal 251 1 subgoal
216 252
217 H : Q 253 H : Q
218 ============================ 254 ============================
219 Q \/ P 255 Q \/ P
220 ]] 256 ]]
221 *) 257 *)
222 258
259 (* begin hide *)
223 left; assumption. 260 left; assumption.
261 (* end hide *)
262 (** %\hspace{.075in}\coqdockw{%#<tt>#left#</tt>#%}%[; assumption.] *)
263
224 (* end thide *) 264 (* end thide *)
225 Qed. 265 Qed.
226 266
227 267
228 (* begin hide *) 268 (* begin hide *)
268 Admitted. 308 Admitted.
269 309
270 (* end hide *) 310 (* end hide *)
271 311
272 312
273 (** It would be a shame to have to plod manually through all proofs about propositional logic. Luckily, there is no need. One of the most basic Coq automation tactics is [tauto], which is a complete decision procedure for constructive propositional logic. (More on what %``%#"#constructive#"#%''% means in the next section.) We can use [tauto] to dispatch all of the purely propositional theorems we have proved so far. *) 313 (** It would be a shame to have to plod manually through all proofs about propositional logic. Luckily, there is no need. One of the most basic Coq automation tactics is %\index{tactics!tauto}%[tauto], which is a complete decision procedure for constructive propositional logic. (More on what %``%#"#constructive#"#%''% means in the next section.) We can use [tauto] to dispatch all of the purely propositional theorems we have proved so far. *)
274 314
275 Theorem or_comm' : P \/ Q -> Q \/ P. 315 Theorem or_comm' : P \/ Q -> Q \/ P.
276 (* begin thide *) 316 (* begin thide *)
277 tauto. 317 tauto.
278 (* end thide *) 318 (* end thide *)
279 Qed. 319 Qed.
280 320
281 (** Sometimes propositional reasoning forms important plumbing for the proof of a theorem, but we still need to apply some other smarts about, say, arithmetic. [intuition] is a generalization of [tauto] that proves everything it can using propositional reasoning. When some goals remain, it uses propositional laws to simplify them as far as possible. Consider this example, which uses the list concatenation operator [++] from the standard library. *) 321 (** Sometimes propositional reasoning forms important plumbing for the proof of a theorem, but we still need to apply some other smarts about, say, arithmetic. %\index{tactics!intuition}%[intuition] is a generalization of [tauto] that proves everything it can using propositional reasoning. When some goals remain, it uses propositional laws to simplify them as far as possible. Consider this example, which uses the list concatenation operator [++] from the standard library. *)
282 322
283 Theorem arith_comm : forall ls1 ls2 : list nat, 323 Theorem arith_comm : forall ls1 ls2 : list nat,
284 length ls1 = length ls2 \/ length ls1 + length ls2 = 6 324 length ls1 = length ls2 \/ length ls1 + length ls2 = 6
285 -> length (ls1 ++ ls2) = 6 \/ length ls1 = length ls2. 325 -> length (ls1 ++ ls2) = 6 \/ length ls1 = length ls2.
286 (* begin thide *) 326 (* begin thide *)
313 353
314 tauto. 354 tauto.
315 (* end thide *) 355 (* end thide *)
316 Qed. 356 Qed.
317 357
318 (** [intuition] is one of the main bits of glue in the implementation of [crush], so, with a little help, we can get a short automated proof of the theorem. *) 358 (** The [intuition] tactic is one of the main bits of glue in the implementation of [crush], so, with a little help, we can get a short automated proof of the theorem. *)
319 359
320 (* begin thide *) 360 (* begin thide *)
321 Theorem arith_comm' : forall ls1 ls2 : list nat, 361 Theorem arith_comm' : forall ls1 ls2 : list nat,
322 length ls1 = length ls2 \/ length ls1 + length ls2 = 6 362 length ls1 = length ls2 \/ length ls1 + length ls2 = 6
323 -> length (ls1 ++ ls2) = 6 \/ length ls1 = length ls2. 363 -> length (ls1 ++ ls2) = 6 \/ length ls1 = length ls2.
364 (* begin hide *)
324 Hint Rewrite app_length : cpdt. 365 Hint Rewrite app_length : cpdt.
366 (* end hide *)
367 (** %\hspace{.075in}%[Hint] %\coqdockw{%#<tt>#Rewrite#</tt>#%}% [app_length : cpdt.] *)
325 368
326 crush. 369 crush.
327 Qed. 370 Qed.
328 (* end thide *) 371 (* end thide *)
329 372
330 End Propositional. 373 End Propositional.
331 374
375 (** Ending the section here has the same effect as always. Each of our propositional theorems becomes universally quantified over the propositional variables that we used. *)
376
332 377
333 (** * What Does It Mean to Be Constructive? *) 378 (** * What Does It Mean to Be Constructive? *)
334 379
335 (** 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? 380 (** One potential point of confusion in the presentation so far is the distinction between [bool] and [Prop]. [bool] is a datatype whose two values are [true] and [false], while [Prop] is a more primitive type that includes among its members [True] and [False]. Why not collapse these two concepts into one, and why must there be more than two states of mathematical truth?
336 381
337 The answer comes from the fact that Coq implements %\textit{%#<i>#constructive#</i>#%}% or %\textit{%#<i>#intuitionistic#</i>#%}% logic, in contrast to the %\textit{%#<i>#classical#</i>#%}% logic that you may be more familiar with. In constructive logic, classical tautologies like [~ ~ P -> P] and [P \/ ~ P] do not always hold. In general, we can only prove these tautologies when [P] is %\textit{%#<i>#decidable#</i>#%}%, in the sense of computability theory. The Curry-Howard encoding that Coq uses for [or] allows us to extract either a proof of [P] or a proof of [~ P] from any proof of [P \/ ~ P]. Since our proofs are just functional programs which we can run, this would give us a decision procedure for the halting problem, where the instantiations of [P] would be formulas like %``%#"#this particular Turing machine halts.#"#%''% 382 The answer comes from the fact that Coq implements %\index{constructive logic}\textit{%#<i>#constructive#</i>#%}% or %\index{intuitionistic logic|see{constructive logic}}\textit{%#<i>#intuitionistic#</i>#%}% logic, in contrast to the %\index{classical logic}\textit{%#<i>#classical#</i>#%}% logic that you may be more familiar with. In constructive logic, classical tautologies like [~ ~ P -> P] and [P \/ ~ P] do not always hold. In general, we can only prove these tautologies when [P] is %\index{decidability}\textit{%#<i>#decidable#</i>#%}%, in the sense of %\index{computability|see{decidability}}%computability theory. The Curry-Howard encoding that Coq uses for [or] allows us to extract either a proof of [P] or a proof of [~ P] from any proof of [P \/ ~ P]. Since our proofs are just functional programs which we can run, a general %\index{law of the excluded middle}%law of the excluded middle would give us a decision procedure for the halting problem, where the instantiations of [P] would be formulas like %``%#"#this particular Turing machine halts.#"#%''%
338 383
339 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.#"#%''% 384 Hence the distinction between [bool] and [Prop]. Programs of type [bool] are computational by construction; we can always run them to determine their results. Many [Prop]s are undecidable, and so we can write more expressive formulas with [Prop]s than with [bool]s, but the inevitable consequence is that we cannot simply %``%#"#run a [Prop] to determine its truth.#"#%''%
340 385
341 Constructive logic lets us define all of the logical connectives in an aesthetically-appealing way, with orthogonal inductive definitions. That is, each connective is defined independently using a simple, shared mechanism. Constructivity also enables a trick called %\textit{%#<i>#program extraction#</i>#%}%, where we write programs by phrasing them as theorems to be proved. Since our proofs are just functional programs, we can extract executable programs from our final proofs, which we could not do as naturally with classical proofs. 386 Constructive logic lets us define all of the logical connectives in an aesthetically-appealing way, with orthogonal inductive definitions. That is, each connective is defined independently using a simple, shared mechanism. Constructivity also enables a trick called %\index{program extraction}\textit{%#<i>#program extraction#</i>#%}%, where we write programs by phrasing them as theorems to be proved. Since our proofs are just functional programs, we can extract executable programs from our final proofs, which we could not do as naturally with classical proofs.
342 387
343 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. *) 388 We will see more about Coq's program extraction facility in a later chapter. However, I think it is worth interjecting another warning at this point, following up on the prior warning about taking the Curry-Howard correspondence too literally. It is possible to write programs by theorem-proving methods in Coq, but hardly anyone does it. It is almost always most useful to maintain the distinction between programs and proofs. If you write a program by proving a theorem, you are likely to run into algorithmic inefficiencies that you introduced in your proof to make it easier to prove. It is a shame to have to worry about such situations while proving tricky theorems, and it is a happy state of affairs that you almost certainly will not need to, with the ideal of extracting programs from proofs being confined mostly to theoretical studies. *)
344 389
345 390
346 (** * First-Order Logic *) 391 (** * First-Order Logic *)
347 392
348 (** The [forall] connective of first-order logic, which we have seen in many examples so far, is built into Coq. Getting ahead of ourselves a bit, we can see it as the dependent function type constructor. In fact, implication and universal quantification are just different syntactic shorthands for the same Coq mechanism. A formula [P -> Q] is equivalent to [forall x : P, Q], where [x] does not appear in [Q]. That is, the %``%#"#real#"#%''% type of the implication says %``%#"#for every proof of [P], there exists a proof of [Q].#"#%''% 393 (** The %\index{Gallina terms!forall}%[forall] connective of first-order logic, which we have seen in many examples so far, is built into Coq. Getting ahead of ourselves a bit, we can see it as the dependent function type constructor. In fact, implication and universal quantification are just different syntactic shorthands for the same Coq mechanism. A formula [P -> Q] is equivalent to [forall x : P, Q], where [x] does not appear in [Q]. That is, the %``%#"#real#"#%''% type of the implication says %``%#"#for every proof of [P], there exists a proof of [Q].#"#%''%
349 394
350 Existential quantification is defined in the standard library. *) 395 %\index{existential quantification}\index{Gallina terms!exists}\index{Gallina terms!ex}%Existential quantification is defined in the standard library. *)
351 396
352 Print ex. 397 (* begin hide *)
353 (** %\vspace{-.15in}% [[ 398 Print ex.
399 (* end hide *)
400 (** %\noindent%[Print] %\coqdocinductive{%#<tt>#ex#</tt>#%}%[.]
401 [[
354 Inductive ex (A : Type) (P : A -> Prop) : Prop := 402 Inductive ex (A : Type) (P : A -> Prop) : Prop :=
355 ex_intro : forall x : A, P x -> ex P 403 ex_intro : forall x : A, P x -> ex P
356 404
357 ]] 405 ]]
358 406
359 [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. *) 407 The family [ex] is parameterized by the type [A] that we quantify over, and by a predicate [P] over [A]s. We prove an existential by exhibiting some [x] of type [A], along with a proof of [P x]. As usual, there are tactics that save us from worrying about the low-level details most of the time. We use the equality operator [=], which, depending on the settings in which they learned logic, different people will say either is or is not part of first-order logic. For our purposes, it is. *)
360 408
361 Theorem exist1 : exists x : nat, x + 1 = 2. 409 Theorem exist1 : exists x : nat, x + 1 = 2.
362 (* begin thide *) 410 (* begin thide *)
363 (** remove printing exists *) 411 (** remove printing exists *)
364 (** We can start this proof with a tactic [exists], which should not be confused with the formula constructor shorthand of the same name. (In the PDF version of this document, the reverse 'E' appears instead of the text %``%#"#exists#"#%''% in formulas.) *) 412 (** We can start this proof with a tactic %\index{tactics!exists}\coqdockw{%exists%}%, which should not be confused with the formula constructor shorthand of the same name. (In the PDF version of this document, the reverse %`%#'#E#'#%'% appears instead of the text %``%#"#exists#"#%''% in formulas.) *)
365 413
414 (* begin hide *)
366 exists 1. 415 exists 1.
416 (* end hide *)
417 (** %\coqdockw{%#<tt>#exists#</tt>#%}% [1.] *)
367 418
368 (** The conclusion is replaced with a version using the existential witness that we announced. 419 (** The conclusion is replaced with a version using the existential witness that we announced.
369 420
370 [[ 421 [[
371 ============================ 422 ============================
417 Admitted. 468 Admitted.
418 469
419 (* end hide *) 470 (* end hide *)
420 471
421 472
422 (** The tactic [intuition] has a first-order cousin called [firstorder]. [firstorder] proves many formulas when only first-order reasoning is needed, and it tries to perform first-order simplifications in any case. First-order reasoning is much harder than propositional reasoning, so [firstorder] is much more likely than [intuition] to get stuck in a way that makes it run for long enough to be useless. *) 473 (** The tactic [intuition] has a first-order cousin called %\index{tactics!firstorder}%[firstorder], which proves many formulas when only first-order reasoning is needed, and it tries to perform first-order simplifications in any case. First-order reasoning is much harder than propositional reasoning, so [firstorder] is much more likely than [intuition] to get stuck in a way that makes it run for long enough to be useless. *)
423 474
424 475
425 (** * Predicates with Implicit Equality *) 476 (** * Predicates with Implicit Equality *)
426 477
427 (** 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. *) 478 (** We start our exploration of a more complicated class of predicates with a simple example: an alternative way of characterizing when a natural number is zero. *)
429 Inductive isZero : nat -> Prop := 480 Inductive isZero : nat -> Prop :=
430 | IsZero : isZero 0. 481 | IsZero : isZero 0.
431 482
432 Theorem isZero_zero : isZero 0. 483 Theorem isZero_zero : isZero 0.
433 (* begin thide *) 484 (* begin thide *)
485 (* begin hide *)
434 constructor. 486 constructor.
435 (* end thide *) 487 (* end hide *)
436 Qed. 488 (** %\coqdockw{%#<tt>#constructor#</tt>#%}%[.]%\vspace{-.075in}% *)
437 489
438 (** We can call [isZero] a %\textit{%#<i>#judgment#</i>#%}%, in the sense often used in the semantics of programming languages. Judgments are typically defined in the style of %\textit{%#<i>#natural deduction#</i>#%}%, where we write a number of %\textit{%#<i>#inference rules#</i>#%}% with premises appearing above a solid line and a conclusion appearing below the line. In this example, the sole constructor [IsZero] of [isZero] can be thought of as the single inference rule for deducing [isZero], with nothing above the line and [isZero 0] below it. The proof of [isZero_zero] demonstrates how we can apply an inference rule. 490 (* end thide *)
491 Qed.
492
493 (** We can call [isZero] a %\index{judgment}\textit{%#<i>#judgment#</i>#%}%, in the sense often used in the semantics of programming languages. Judgments are typically defined in the style of %\index{natural deduction}\textit{%#<i>#natural deduction#</i>#%}%, where we write a number of %\index{inference rules}\textit{%#<i>#inference rules#</i>#%}% with premises appearing above a solid line and a conclusion appearing below the line. In this example, the sole constructor [IsZero] of [isZero] can be thought of as the single inference rule for deducing [isZero], with nothing above the line and [isZero 0] below it. The proof of [isZero_zero] demonstrates how we can apply an inference rule.
439 494
440 The definition of [isZero] differs in an important way from all of the other inductive definitions that we have seen in this and the previous chapter. Instead of writing just [Set] or [Prop] after the colon, here we write [nat -> Prop]. We saw examples of parameterized types like [list], but there the parameters appeared with names %\textit{%#<i>#before#</i>#%}% the colon. Every constructor of a parameterized inductive type must have a range type that uses the same parameter, whereas the form we use here enables us to use different arguments to the type for different constructors. 495 The definition of [isZero] differs in an important way from all of the other inductive definitions that we have seen in this and the previous chapter. Instead of writing just [Set] or [Prop] after the colon, here we write [nat -> Prop]. We saw examples of parameterized types like [list], but there the parameters appeared with names %\textit{%#<i>#before#</i>#%}% the colon. Every constructor of a parameterized inductive type must have a range type that uses the same parameter, whereas the form we use here enables us to use different arguments to the type for different constructors.
441 496
442 For instance, [isZero] forces its argument to be [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 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! *) 497 For instance, our definition [isZero] makes the predicate provable only when the argument is [0]. We can see that the concept of equality is somehow implicit in the inductive definition mechanism. The way this is accomplished is similar to the way that logic variables are used in %\index{Prolog}%Prolog, and it is a very powerful mechanism that forms a foundation for formalizing all of mathematics. In fact, though it is natural to think of inductive types as folding in the functionality of equality, in Coq, the true situation is reversed, with equality defined as just another inductive type!%\index{Gallina terms!eq}\index{Gallina terms!refl\_equal}% *)
443 498
499 (* begin hide *)
444 Print eq. 500 Print eq.
445 (** %\vspace{-.15in}% [[ 501 (* end hide *)
502 (** %\noindent%[Print] %\coqdocinductive{%#<tt>#eq#</tt>#%}%[.]
503 [[
446 Inductive eq (A : Type) (x : A) : A -> Prop := refl_equal : x = x 504 Inductive eq (A : Type) (x : A) : A -> Prop := refl_equal : x = x
447 505
448 ]] 506 ]]
449 507
450 [eq] is the type we get behind the scenes when uses of infix [=] are expanded. We see that [eq] has both a parameter [x] that is fixed and an extra unnamed argument of the same type. The type of [eq] allows us to state any equalities, even those that are provably false. However, examining the type of equality's sole constructor [refl_equal], we see that we can only %\textit{%#<i>#prove#</i>#%}% equality when its two arguments are syntactically equal. This definition turns out to capture all of the basic properties of equality, and the equality-manipulating tactics that we have seen so far, like [reflexivity] and [rewrite], are implemented treating [eq] as just another inductive type with a well-chosen definition. 508 Behind the scenes, uses of infix [=] are expanded to instances of [eq]. We see that [eq] has both a parameter [x] that is fixed and an extra unnamed argument of the same type. The type of [eq] allows us to state any equalities, even those that are provably false. However, examining the type of equality's sole constructor [refl_equal], we see that we can only %\textit{%#<i>#prove#</i>#%}% equality when its two arguments are syntactically equal. This definition turns out to capture all of the basic properties of equality, and the equality-manipulating tactics that we have seen so far, like [reflexivity] and [rewrite], are implemented treating [eq] as just another inductive type with a well-chosen definition. Another way of stating that definition is: equality is defined as the least reflexive relation.
451 509
452 Returning to the example of [isZero], we can see how to make use of hypotheses that use this predicate. *) 510 Returning to the example of [isZero], we can see how to work with hypotheses that use this predicate. *)
453 511
454 Theorem isZero_plus : forall n m : nat, isZero m -> n + m = n. 512 Theorem isZero_plus : forall n m : nat, isZero m -> n + m = n.
455 (* begin thide *) 513 (* begin thide *)
456 (** We want to proceed by cases on the proof of the assumption about [isZero]. *) 514 (** We want to proceed by cases on the proof of the assumption about [isZero]. *)
457 515
482 540
483 ]] 541 ]]
484 542
485 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]. 543 It seems that case analysis has not helped us much at all! Our sole hypothesis disappears, leaving us, if anything, worse off than we were before. What went wrong? We have met an important restriction in tactics like [destruct] and [induction] when applied to types with arguments. If the arguments are not already free variables, they will be replaced by new free variables internally before doing the case analysis or induction. Since the argument [1] to [isZero] is replaced by a fresh variable, we lose the crucial fact that it is not equal to [0].
486 544
487 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 [inversion], which corresponds to the concept of inversion that is frequently used with natural deduction proof systems. *) 545 Why does Coq use this restriction? We will discuss the issue in detail in a future chapter, when we see the dependently typed programming techniques that would allow us to write this proof term manually. For now, we just say that the algorithmic problem of %``%#"#logically complete case analysis#"#%''% is undecidable when phrased in Coq's logic. A few tactics and design patterns that we will present in this chapter suffice in almost all cases. For the current example, what we want is a tactic called %\index{tactics!inversion}%[inversion], which corresponds to the concept of inversion that is frequently used with natural deduction proof systems. *)
488 546
489 Undo. 547 Undo.
490 inversion 1. 548 inversion 1.
491 (* end thide *) 549 (* end thide *)
492 Qed. 550 Qed.
514 isZero_ind 572 isZero_ind
515 : forall P : nat -> Prop, P 0 -> forall n : nat, isZero n -> P n 573 : forall P : nat -> Prop, P 0 -> forall n : nat, isZero n -> P n
516 574
517 ]] 575 ]]
518 576
519 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]. *) 577 In our last proof script, [destruct] chose to instantiate [P] as [fun n => S n + S n = S (][S (][S (][S n)))]. You can verify for yourself that this specialization of the principle applies to the goal and that the hypothesis [P 0] then matches the subgoal we saw generated. If you are doing a proof and encounter a strange transmutation like this, there is a good chance that you should go back and replace a use of [destruct] with [inversion]. *)
520 578
521 579
522 (* begin hide *) 580 (* begin hide *)
523 (* In-class exercises *) 581 (* In-class exercises *)
524 582
551 609
552 Inductive even : nat -> Prop := 610 Inductive even : nat -> Prop :=
553 | EvenO : even O 611 | EvenO : even O
554 | EvenSS : forall n, even n -> even (S (S n)). 612 | EvenSS : forall n, even n -> even (S (S n)).
555 613
556 (** 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. 614 (** Think of [even] as another judgment defined by natural deduction rules. [EvenO] is a rule with nothing above the line and [even O] below the line, and [EvenSS] is a rule with [even n] above the line and [even (][S (][S n))] below.
557 615
558 The proof techniques of the last section are easily adapted. *) 616 The proof techniques of the last section are easily adapted. *)
559 617
560 Theorem even_0 : even 0. 618 Theorem even_0 : even 0.
561 (* begin thide *) 619 (* begin thide *)
620 (* begin hide *)
562 constructor. 621 constructor.
622 (* end hide *)
623 (** %\coqdockw{%#<tt>#constructor#</tt>#%}%[.]%\vspace{-.075in}% *)
624
563 (* end thide *) 625 (* end thide *)
564 Qed. 626 Qed.
565 627
566 Theorem even_4 : even 4. 628 Theorem even_4 : even 4.
567 (* begin thide *) 629 (* begin thide *)
630 (* begin hide *)
568 constructor; constructor; constructor. 631 constructor; constructor; constructor.
569 (* end thide *) 632 (* end hide *)
570 Qed. 633 (** %\coqdockw{%#<tt>#constructor#</tt>#%}%[; ]%\coqdockw{%#<tt>#constructor#</tt>#%}%[; ]%\coqdockw{%#<tt>#constructor#</tt>#%}%[.]%\vspace{-.075in}% *)
571 634
572 (** 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. *) 635 (* end thide *)
573 636 Qed.
574 (* begin thide *) 637
638 (** It is not hard to see that sequences of constructor applications like the above can get tedious. We can avoid them using Coq's hint facility, with a new [Hint] variant that asks to consider all constructors of an inductive type during proof search. Note that this sort of hint may be placed in a default database, such that a command has no equivalent to the [: cpdt] from our earlier rewrite hints.%\index{Hint Constructirs}% The tactic %\index{tactics!auto}%[auto] performs exhaustive proof search up to a fixed depth, considering only the proof steps we have registered as hints. *)
639
640 (* begin thide *)
641 (* begin hide *)
575 Hint Constructors even. 642 Hint Constructors even.
643 (* end hide *)
644 (** %\noindent%[Hint] %\coqdockw{%#<tt>#Constructors#</tt>#%}% [even.] *)
576 645
577 Theorem even_4' : even 4. 646 Theorem even_4' : even 4.
578 auto. 647 auto.
579 Qed. 648 Qed.
580 649
581 (* end thide *) 650 (* end thide *)
651
652 (** We may also use [inversion] with [even]. *)
582 653
583 Theorem even_1_contra : even 1 -> False. 654 Theorem even_1_contra : even 1 -> False.
584 (* begin thide *) 655 (* begin thide *)
585 inversion 1. 656 inversion 1.
586 (* end thide *) 657 (* end thide *)
597 ============================ 668 ============================
598 False 669 False
599 670
600 ]] 671 ]]
601 672
602 [inversion] 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. *) 673 The [inversion] tactic can be a little overzealous at times, as we can see here with the introduction of the unused variable [n] and an equality hypothesis about it. For more complicated predicates, though, adding such assumptions is critical to dealing with the undecidability of general inversion. More complex inductive definitions and theorems can cause [inversion] to generate equalities where neither side is a variable. *)
603 674
604 inversion H1. 675 inversion H1.
605 (* end thide *) 676 (* end thide *)
606 Qed. 677 Qed.
607 678
647 ============================ 718 ============================
648 even (S (S (n0 + m))) 719 even (S (S (n0 + m)))
649 ]] 720 ]]
650 *) 721 *)
651 722
723 (* begin hide *)
652 constructor. 724 constructor.
653 (** [[ 725 (* end hide *)
726 (** %\coqdockw{%#<tt>#constructor#</tt>#%}%[.]
727
728 [[
654 ============================ 729 ============================
655 even (n0 + m) 730 even (n0 + m)
656 731
657 ]] 732 ]]
658 733
659 At this point, we would like to apply the inductive hypothesis, which is: 734 At this point, we would like to apply the inductive hypothesis, which is:
660 735
661 [[ 736 [[
662
663 IHn : forall m : nat, even n -> even m -> even (n + m) 737 IHn : forall m : nat, even n -> even m -> even (n + m)
664
665 ]] 738 ]]
666 739
667 Unfortunately, the goal mentions [n0] where it would need to mention [n] to match [IHn]. We could keep looking for a way to finish this proof from here, but it turns out that we can make our lives much easier by changing our basic strategy. Instead of inducting on the structure of [n], we should induct %\textit{%#<i>#on the structure of one of the [even] proofs#</i>#%}%. This technique is commonly called %\textit{%#<i>#rule induction#</i>#%}% in programming language semantics. In the setting of Coq, we have already seen how predicates are defined using the same inductive type mechanism as datatypes, so the fundamental unity of rule induction with %``%#"#normal#"#%''% induction is apparent. *) 740 Unfortunately, the goal mentions [n0] where it would need to mention [n] to match [IHn]. We could keep looking for a way to finish this proof from here, but it turns out that we can make our lives much easier by changing our basic strategy. Instead of inducting on the structure of [n], we should induct %\textit{%#<i>#on the structure of one of the [even] proofs#</i>#%}%. This technique is commonly called %\index{rule induction}\textit{%#<i>#rule induction#</i>#%}% in programming language semantics. In the setting of Coq, we have already seen how predicates are defined using the same inductive type mechanism as datatypes, so the fundamental unity of rule induction with %``%#"#normal#"#%''% induction is apparent.
668 741
742 Recall that tactics like [induction] and [destruct] may be passed numbers to refer to unnamed lefthand sides of implications in the conclusion, where the argument [n] refers to the [n]th such hypothesis. *)
743
744 (* begin hide *)
669 Restart. 745 Restart.
746 (* end hide *)
747 (** %\noindent\coqdockw{%#<tt>#Restart#</tt>#%}%[.] *)
670 748
671 induction 1. 749 induction 1.
672 (** [[ 750 (** [[
673 m : nat 751 m : nat
674 ============================ 752 ============================
675 even m -> even (0 + m) 753 even m -> even (0 + m)
676 754 ]]
677 subgoal 2 is: 755
756 %\noindent \coqdockw{subgoal} 2 \coqdockw{is}:%#<tt>subgoal 2 is</tt>#
757 [[
678 even m -> even (S (S n) + m) 758 even m -> even (S (S n) + m)
679 759
680 ]] 760 ]]
681 761
682 The first case is easily discharged by [crush], based on the hint we added earlier to try the constructors of [even]. *) 762 The first case is easily discharged by [crush], based on the hint we added earlier to try the constructors of [even]. *)
684 crush. 764 crush.
685 765
686 (** Now we focus on the second case: *) 766 (** Now we focus on the second case: *)
687 767
688 intro. 768 intro.
689
690 (** [[ 769 (** [[
691 m : nat 770 m : nat
692 n : nat 771 n : nat
693 H : even n 772 H : even n
694 IHeven : even m -> even (n + m) 773 IHeven : even m -> even (n + m)
698 777
699 ]] 778 ]]
700 779
701 We simplify and apply a constructor, as in our last proof attempt. *) 780 We simplify and apply a constructor, as in our last proof attempt. *)
702 781
782 (* begin hide *)
703 simpl; constructor. 783 simpl; constructor.
704 (** [[ 784 (* end hide *)
785 (** [simpl; ]%\coqdockw{%#<tt>#constructor#</tt>#%}%[.]
786
787 [[
705 ============================ 788 ============================
706 even (n + m) 789 even (n + m)
707 790
708 ]] 791 ]]
709 792
711 794
712 apply IHeven; assumption. 795 apply IHeven; assumption.
713 796
714 (** In fact, [crush] can handle all of the details of the proof once we declare the induction strategy. *) 797 (** In fact, [crush] can handle all of the details of the proof once we declare the induction strategy. *)
715 798
799 (* begin hide *)
716 Restart. 800 Restart.
801 (* end hide *)
802 (** %\noindent\coqdockw{%#<tt>#Restart#</tt>#%}%[.] *)
803
717 induction 1; crush. 804 induction 1; crush.
718 (* end thide *) 805 (* end thide *)
719 Qed. 806 Qed.
720 807
721 (** Induction on recursive predicates has similar pitfalls to those we encountered with inversion in the last section. *) 808 (** Induction on recursive predicates has similar pitfalls to those we encountered with inversion in the last section. *)
725 induction 1. 812 induction 1.
726 (** [[ 813 (** [[
727 n : nat 814 n : nat
728 ============================ 815 ============================
729 False 816 False
730 817 ]]
731 subgoal 2 is: 818
819 %\noindent \coqdockw{subgoal} 2 \coqdockw{is}:%#<tt>subgoal 2 is</tt>#
820 [[
732 False 821 False
733 822
734 ]] 823 ]]
735 824
736 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. *) 825 We are already sunk trying to prove the first subgoal, since the argument to [even] was replaced by a fresh variable internally. This time, we find it easier to prove this theorem by way of a lemma. Instead of trusting [induction] to replace expressions with fresh variables, we do it ourselves, explicitly adding the appropriate equalities as new assumptions. *)
755 844
756 ]] 845 ]]
757 846
758 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]. *) 847 At this point it is useful to use a theorem from the standard library, which we also proved with a different name in the last chapter. We can search for a theorem that allows us to rewrite terms of the form [x + S y]. *)
759 848
849 (* begin hide *)
760 SearchRewrite (_ + S _). 850 SearchRewrite (_ + S _).
761 (** %\vspace{-.15in}% [[ 851 (* end hide *)
852 (** %\coqdockw{%#<tt>#SearchRewrite#</tt>#%}% [(_ + S _).]
853
854 [[
762 plus_n_Sm : forall n m : nat, S (n + m) = n + S m 855 plus_n_Sm : forall n m : nat, S (n + m) = n + S m
763 ]] 856 ]]
764 *) 857 *)
765 858
766 rewrite <- plus_n_Sm in H0. 859 rewrite <- plus_n_Sm in H0.
767 860
768 (** The induction hypothesis lets us complete the proof. *) 861 (** The induction hypothesis lets us complete the proof, if we use a variant of [apply] that has a %\index{tactics!with}%[with] clause to give instantiations of quantified variables. *)
769 862
770 apply IHeven with n0; assumption. 863 apply IHeven with n0; assumption.
771 864
772 (** 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. *) 865 (** As usual, we can rewrite the proof to avoid referencing any locally generated names, which makes our proof script more readable and more robust to changes in the theorem statement. We use the notation [<-] to request a hint that does right-to-left rewriting, just like we can with the [rewrite] tactic. *)
773 866
867 (* begin hide *)
774 Restart. 868 Restart.
869 (* end hide *)
870 (** %\hspace{-.075in}\coqdockw{%#<tt>#Restart#</tt>#%}%[.] *)
871
872 (* begin hide *)
775 Hint Rewrite <- plus_n_Sm : cpdt. 873 Hint Rewrite <- plus_n_Sm : cpdt.
874 (* end hide *)
875 (** %\hspace{-.075in}%[Hint] %\noindent\coqdockw{%#<tt>#Rewrite#</tt>#%}% [<- plus_n_sm : cpdt.] *)
776 876
777 induction 1; crush; 877 induction 1; crush;
778 match goal with 878 match goal with
779 | [ H : S ?N = ?N0 + ?N0 |- _ ] => destruct N; destruct N0 879 | [ H : S ?N = ?N0 + ?N0 |- _ ] => destruct N; destruct N0
780 end; crush; eauto. 880 end; crush; eauto.
781 Qed. 881 Qed.
782 882
783 (** We write the proof in a way that avoids the use of local variable or hypothesis names, using the [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, [eauto]. 883 (** We write the proof in a way that avoids the use of local variable or hypothesis names, using the %\index{tactics!match}%[match] tactic form to do pattern-matching on the goal. We use unification variables prefixed by question marks in the pattern, and we take advantage of the possibility to mention a unification variable twice in one pattern, to enforce equality between occurrences. The hint to rewrite with [plus_n_Sm] in a particular direction saves us from having to figure out the right place to apply that theorem, and we also take critical advantage of a new tactic, %\index{tactics!eauto}%[eauto].
784 884
785 [crush] 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. [auto] attempts Prolog-style logic programming, searching through all proof trees up to a certain depth that are built only out of hints that have been registered with [Hint] commands. Compared to Prolog, [auto] places an important restriction: it never introduces new unification variables during search. That is, every time a rule is applied during proof search, all of its arguments must be deducible by studying the form of the goal. [eauto] relaxes this restriction, at the cost of possibly exponentially greater running time. In this particular case, we know that [eauto] has only a small space of proofs to search, so it makes sense to run it. It is common in effectively-automated Coq proofs to see a bag of standard tactics applied to pick off the %``%#"#easy#"#%''% subgoals, finishing with [eauto] to handle the tricky parts that can benefit from ad-hoc exhaustive search. 885 The [crush] tactic uses the tactic [intuition], which, when it runs out of tricks to try using only propositional logic, by default tries the tactic [auto], which we saw in an earlier example. The [auto] tactic attempts %\index{Prolog}%Prolog-style logic programming, searching through all proof trees up to a certain depth that are built only out of hints that have been registered with [Hint] commands. Compared to Prolog, [auto] places an important restriction: it never introduces new unification variables during search. That is, every time a rule is applied during proof search, all of its arguments must be deducible by studying the form of the goal. This restriction is relaxed for [eauto], at the cost of possibly exponentially greater running time. In this particular case, we know that [eauto] has only a small space of proofs to search, so it makes sense to run it. It is common in effectively automated Coq proofs to see a bag of standard tactics applied to pick off the %``%#"#easy#"#%''% subgoals, finishing with [eauto] to handle the tricky parts that can benefit from ad-hoc exhaustive search.
786 886
787 The original theorem now follows trivially from our lemma. *) 887 The original theorem now follows trivially from our lemma. *)
788 888
789 Theorem even_contra : forall n, even (S (n + n)) -> False. 889 Theorem even_contra : forall n, even (S (n + n)) -> False.
790 intros; eapply even_contra'; eauto. 890 intros; eapply even_contra'; eauto.
791 Qed. 891 Qed.
792 892
793 (** We use a variant [eapply] of [apply] which has the same relationship to [apply] as [eauto] has to [auto]. [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. [eauto] is able to determine the right values for those unification variables. 893 (** We use a variant %\index{tactics!apply}%[eapply] of [apply] which has the same relationship to [apply] as [eauto] has to [auto]. An invocation of [apply] only succeeds if all arguments to the rule being used can be determined from the form of the goal, whereas [eapply] will introduce unification variables for undetermined arguments. In this case, [eauto] is able to determine the right values for those unification variables.
794 894
795 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. *) 895 By considering an alternate attempt at proving the lemma, we can see another common pitfall of inductive proofs in Coq. Imagine that we had tried to prove [even_contra'] with all of the [forall] quantifiers moved to the front of the lemma statement. *)
796 896
797 Lemma even_contra'' : forall n' n, even n' -> n' = S (n + n) -> False. 897 Lemma even_contra'' : forall n' n, even n' -> n' = S (n + n) -> False.
798 induction 1; crush; 898 induction 1; crush;
809 ============================ 909 ============================
810 False 910 False
811 911
812 ]] 912 ]]
813 913
814 We are out of luck here. The inductive hypothesis is trivially true, since its assumption is false. In the version of this proof that succeeded, [IHeven] had an explicit quantification over [n]. This is because the quantification of [n] %\textit{%#<i>#appeared after the thing we are inducting on#</i>#%}% in the theorem statement. In general, quantified variables and hypotheses that appear before the induction object in the theorem statement stay fixed throughout the inductive proof. Variables and hypotheses that are quantified after the induction object may be varied explicitly in uses of inductive hypotheses. 914 We are out of luck here. The inductive hypothesis is trivially true, since its assumption is false. In the version of this proof that succeeded, [IHeven] had an explicit quantification over [n]. This is because the quantification of [n] %\textit{%#<i>#appeared after the thing we are inducting on#</i>#%}% in the theorem statement. In general, quantified variables and hypotheses that appear before the induction object in the theorem statement stay fixed throughout the inductive proof. Variables and hypotheses that are quantified after the induction object may be varied explicitly in uses of inductive hypotheses. *)
815
816 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. *)
817 (* end thide *)
818 915
819 Abort. 916 Abort.
917
918 (** Why should Coq implement [induction] this way? One answer is that it avoids burdening this basic tactic with additional heuristic smarts, but that is not the whole picture. Imagine that [induction] analyzed dependencies among variables and reordered quantifiers to preserve as much freedom as possible in later uses of inductive hypotheses. This could make the inductive hypotheses more complex, which could in turn cause particular automation machinery to fail when it would have succeeded before. In general, we want to avoid quantifiers in our proofs whenever we can, and that goal is furthered by the refactoring that the [induction] tactic forces us to do. *)
919 (* end thide *)
920
921
820 922
821 923
822 (* begin hide *) 924 (* begin hide *)
823 (* In-class exercises *) 925 (* In-class exercises *)
824 926
887 989
888 (** * Exercises *) 990 (** * Exercises *)
889 991
890 (** %\begin{enumerate}%#<ol># 992 (** %\begin{enumerate}%#<ol>#
891 993
892 %\item%#<li># Prove these tautologies of propositional logic, using only the tactics [apply], [assumption], [constructor], [destruct], [intro], [intros], [left], [right], [split], and [unfold]. 994 %\item%#<li># Prove these tautologies of propositional logic, using only the tactics [apply], [assumption], %\coqdockw{%#<tt>#constructor#</tt>#%}%, [destruct], [intro], [intros], %\coqdockw{%#<tt>#left#</tt>#%}%, %\coqdockw{%#<tt>#right#</tt>#%}%, [split], and [unfold].
893 %\begin{enumerate}%#<ol># 995 %\begin{enumerate}%#<ol>#
894 %\item%#<li># [(True \/ False) /\ (False \/ True)]#</li># 996 %\item%#<li># [(][True \/ False) /\ (][False \/ True)]#</li>#
895 %\item%#<li># [P -> ~ ~ P]#</li># 997 %\item%#<li># [P -> ~ ~ P]#</li>#
896 %\item%#<li># [P /\ (Q \/ R) -> (P /\ Q) \/ (P /\ R)]#</li># 998 %\item%#<li># [P /\ (][Q \/ R) -> (][P /\ Q) \/ (][P /\ R)]#</li>#
897 #</ol> </li>#%\end{enumerate}% 999 #</ol> </li>#%\end{enumerate}%
898 1000
899 %\item%#<li># Prove the following tautology of first-order logic, using only the tactics [apply], [assert], [assumption], [destruct], [eapply], [eassumption], and %\textit{%#<tt>#exists#</tt>#%}%. You will probably find [assert] useful for stating and proving an intermediate lemma, enabling a kind of %``%#"#forward reasoning,#"#%''% in contrast to the %``%#"#backward reasoning#"#%''% that is the default for Coq tactics. [eassumption] is a version of [assumption] that will do matching of unification variables. Let some variable [T] of type [Set] be the set of individuals. [x] is a constant symbol, [p] is a unary predicate symbol, [q] is a binary predicate symbol, and [f] is a unary function symbol. 1001 %\item%#<li># Prove the following tautology of first-order logic, using only the tactics [apply], [assert], [assumption], [destruct], [eapply], %\coqdockw{%#<tt>#eassumption#</tt>#%}%, and %\coqdockw{%#<tt>#exists#</tt>#%}%. You will probably find the [assert] tactic useful for stating and proving an intermediate lemma, enabling a kind of %``%#"#forward reasoning,#"#%''% in contrast to the %``%#"#backward reasoning#"#%''% that is the default for Coq tactics. The tactic %\coqdockw{%#<tt>#eassumption#</tt>#%}% is a version of [assumption] that will do matching of unification variables. Let some variable [T] of type [Set] be the set of individuals. [x] is a constant symbol, [p] is a unary predicate symbol, [q] is a binary predicate symbol, and [f] is a unary function symbol.
900 %\begin{enumerate}%#<ol># 1002 %\begin{enumerate}%#<ol>#
901 %\item%#<li># [p x -> (forall x, p x -> exists y, q x y) -> (forall x y, q x y -> q y (f y)) -> exists z, q z (f z)]#</li># 1003 %\item%#<li># [p x -> (][forall x, p x -> exists y, q x y) -> (][forall x y, q x y -> q y (f y)) -> exists z, q z (f z)]#</li>#
902 #</ol> </li>#%\end{enumerate}% 1004 #</ol> </li>#%\end{enumerate}%
903 1005
904 %\item%#<li># Define an inductive predicate capturing when a natural number is an integer multiple of either 6 or 10. Prove that 13 does not satisfy your predicate, and prove that any number satisfying the predicate is not odd. It is probably easiest to prove the second theorem by indicating %``%#"#odd-ness#"#%''% as equality to [2 * n + 1] for some [n].#</li># 1006 %\item%#<li># Define an inductive predicate capturing when a natural number is an integer multiple of either 6 or 10. Prove that 13 does not satisfy your predicate, and prove that any number satisfying the predicate is not odd. It is probably easiest to prove the second theorem by indicating %``%#"#odd-ness#"#%''% as equality to [2 * n + 1] for some [n].#</li>#
905 1007
906 %\item%#<li># Define a simple programming language, its semantics, and its typing rules, and then prove that well-typed programs cannot go wrong. Specifically: 1008 %\item%#<li># $\star$ Define a simple programming language, its semantics, and its typing rules, and then prove that well-typed programs cannot go wrong. Specifically:
907 %\begin{enumerate}%#<ol># 1009 %\begin{enumerate}%#<ol>#
908 %\item%#<li># Define [var] as a synonym for the natural numbers.#</li># 1010 %\item%#<li># Define [var] as a synonym for the natural numbers.#</li>#
909 %\item%#<li># Define an inductive type [exp] of expressions, containing natural number constants, natural number addition, pairing of two other expressions, extraction of the first component of a pair, extraction of the second component of a pair, and variables (based on the [var] type you defined).#</li># 1011 %\item%#<li># Define an inductive type [exp] of expressions, containing natural number constants, natural number addition, pairing of two other expressions, extraction of the first component of a pair, extraction of the second component of a pair, and variables (based on the [var] type you defined).#</li>#
910 %\item%#<li># Define an inductive type [cmd] of commands, containing expressions and variable assignments. A variable assignment node should contain the variable being assigned, the expression being assigned to it, and the command to run afterward.#</li># 1012 %\item%#<li># Define an inductive type [cmd] of commands, containing expressions and variable assignments. A variable assignment node should contain the variable being assigned, the expression being assigned to it, and the command to run afterward.#</li>#
911 %\item%#<li># Define an inductive type [val] of values, containing natural number constants and pairings of values.#</li># 1013 %\item%#<li># Define an inductive type [val] of values, containing natural number constants and pairings of values.#</li>#
920 #</ol> </li>#%\end{enumerate}% 1022 #</ol> </li>#%\end{enumerate}%
921 A few hints that may be helpful: 1023 A few hints that may be helpful:
922 %\begin{enumerate}%#<ol># 1024 %\begin{enumerate}%#<ol>#
923 %\item%#<li># One easy way of defining variable assignments and typings is to define both as instances of a polymorphic map type. The map type at parameter [T] can be defined to be the type of arbitrary functions from variables to [T]. A helpful function for implementing insertion into such a functional map is [eq_nat_dec], which you can make available with [Require Import Arith.]. [eq_nat_dec] has a dependent type that tells you that it makes accurate decisions on whether two natural numbers are equal, but you can use it as if it returned a boolean, e.g., [if eq_nat_dec n m then E1 else E2].#</li># 1025 %\item%#<li># One easy way of defining variable assignments and typings is to define both as instances of a polymorphic map type. The map type at parameter [T] can be defined to be the type of arbitrary functions from variables to [T]. A helpful function for implementing insertion into such a functional map is [eq_nat_dec], which you can make available with [Require Import Arith.]. [eq_nat_dec] has a dependent type that tells you that it makes accurate decisions on whether two natural numbers are equal, but you can use it as if it returned a boolean, e.g., [if eq_nat_dec n m then E1 else E2].#</li>#
924 %\item%#<li># If you follow the last hint, you may find yourself writing a proof that involves an expression with [eq_nat_dec] that you would like to simplify. Running [destruct] on the particular call to [eq_nat_dec] should do the trick. You can automate this advice with a piece of Ltac: [[ 1026 %\item%#<li># If you follow the last hint, you may find yourself writing a proof that involves an expression with [eq_nat_dec] that you would like to simplify. Running [destruct] on the particular call to [eq_nat_dec] should do the trick. You can automate this advice with a piece of Ltac: [[
925
926 match goal with 1027 match goal with
927 | [ |- context[eq_nat_dec ?X ?Y] ] => destruct (eq_nat_dec X Y) 1028 | [ |- context[eq_nat_dec ?X ?Y] ] => destruct (eq_nat_dec X Y)
928 end 1029 end
929 ]] 1030 ]]
930 #</li># 1031 #</li>#
931 %\item%#<li># You probably do not want to use an inductive definition for compatibility of variable assignments and typings.#</li># 1032 %\item%#<li># You probably do not want to use an inductive definition for compatibility of variable assignments and typings.#</li>#
932 %\item%#<li># The [Tactics] module from this book contains a variant [crush'] of [crush]. [crush'] takes two arguments. The first argument is a list of lemmas and other functions to be tried automatically in %``%#"#forward reasoning#"#%''% style, where we add new facts without being sure yet that they link into a proof of the conclusion. The second argument is a list of predicates on which inversion should be attempted automatically. For instance, running [crush' (lemma1, lemma2) pred] will search for chances to apply [lemma1] and [lemma2] to hypotheses that are already available, adding the new concluded fact if suitable hypotheses can be found. Inversion will be attempted on any hypothesis using [pred], but only those inversions that narrow the field of possibilities to one possible rule will be kept. The format of the list arguments to [crush'] is that you can pass an empty list as [tt], a singleton list as the unadorned single element, and a multiple-element list as a tuple of the elements.#</li># 1033 %\item%#<li># The [CpdtTactics] module from this book contains a variant [crush'] of [crush]. [crush'] takes two arguments. The first argument is a list of lemmas and other functions to be tried automatically in %``%#"#forward reasoning#"#%''% style, where we add new facts without being sure yet that they link into a proof of the conclusion. The second argument is a list of predicates on which inversion should be attempted automatically. For instance, running [crush' (lemma1, lemma2) pred] will search for chances to apply [lemma1] and [lemma2] to hypotheses that are already available, adding the new concluded fact if suitable hypotheses can be found. Inversion will be attempted on any hypothesis using [pred], but only those inversions that narrow the field of possibilities to one possible rule will be kept. The format of the list arguments to [crush'] is that you can pass an empty list as [tt], a singleton list as the unadorned single element, and a multiple-element list as a tuple of the elements.#</li>#
933 %\item%#<li># If you want [crush'] to apply polymorphic lemmas, you may have to do a little extra work, if the type parameter is not a free variable of your proof context (so that [crush'] does not know to try it). For instance, if you define a polymorphic map insert function [assign] of some type [forall T : Set, ...], and you want particular applications of [assign] added automatically with type parameter [U], you would need to include [assign] in the lemma list as [assign U] (if you have implicit arguments off) or [assign (T := U)] or [@assign U] (if you have implicit arguments on).#</li># 1034 %\item%#<li># If you want [crush'] to apply polymorphic lemmas, you may have to do a little extra work, if the type parameter is not a free variable of your proof context (so that [crush'] does not know to try it). For instance, if you define a polymorphic map insert function [assign] of some type [forall T : Set, ...], and you want particular applications of [assign] added automatically with type parameter [U], you would need to include [assign] in the lemma list as [assign U] (if you have implicit arguments off) or [assign (T := U)] or [@assign U] (if you have implicit arguments on).#</li>#
934 #</ol> </li>#%\end{enumerate}% 1035 #</ol> </li>#%\end{enumerate}%
935 1036
936 #</li># 1037 #</li>#
937 1038