Mercurial > cpdt > repo
comparison src/Predicates.v @ 440:f923024bd284
Vertical spacing pass, through end of Subset
author | Adam Chlipala <adam@chlipala.net> |
---|---|
date | Mon, 30 Jul 2012 16:50:02 -0400 |
parents | 393b8ed99c2f |
children | 89c67796754e |
comparison
equal
deleted
inserted
replaced
439:393b8ed99c2f | 440:f923024bd284 |
---|---|
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 Print unit. | 37 Print unit. |
38 (** [[ | 38 (** %\vspace{-.15in}%[[ |
39 Inductive unit : Set := tt : unit | 39 Inductive unit : Set := tt : unit |
40 ]] | 40 ]] |
41 *) | 41 *) |
42 | 42 |
43 Print True. | 43 Print True. |
44 (** [[ | 44 (** %\vspace{-.15in}%[[ |
45 Inductive True : Prop := I : True | 45 Inductive True : Prop := I : True |
46 ]] | 46 ]] |
47 *) | 47 |
48 | 48 %\smallskip{}%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. | |
50 | 49 |
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. | 50 The type [unit] has one value, [tt]. The type [True] has one proof, [I]. Why distinguish between these two types? Many people who have read about Curry-Howard in an abstract context 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. |
52 | 51 |
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. | 52 The essence of the argument is roughly this: to an engineer, not all functions of type [A -> B] are created equal, but all proofs of a proposition [P -> Q] are. This idea is known as%\index{proof irrelevance}% _proof irrelevance_, and its formalizations in logics prevent us from distinguishing between alternate proofs of the same proposition. Proof irrelevance is compatible with, but not derivable in, Gallina. Apart from this theoretical concern, I will argue that it is most effective to do engineering with Coq by employing different techniques for programs versus proofs. Most of this book is organized around that distinction, describing how to program, by applying standard functional programming techniques in the presence of dependent types; and how to prove, by writing custom Ltac decision procedures. |
54 | 53 |
82 (* end thide *) | 81 (* end thide *) |
83 | 82 |
84 (** There is also a predicate [False], which is the Curry-Howard mirror image of [Empty_set] from the last chapter. *) | 83 (** There is also a predicate [False], which is the Curry-Howard mirror image of [Empty_set] from the last chapter. *) |
85 | 84 |
86 Print False. | 85 Print False. |
87 (** [[ | 86 (** %\vspace{-.15in}%[[ |
88 Inductive False : Prop := | 87 Inductive False : Prop := |
89 | |
90 ]] | 88 ]] |
91 | 89 |
92 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. *) | 90 %\smallskip{}%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. *) |
93 | 91 |
94 Theorem False_imp : False -> 2 + 2 = 5. | 92 Theorem False_imp : False -> 2 + 2 = 5. |
95 (* begin thide *) | 93 (* begin thide *) |
96 destruct 1. | 94 destruct 1. |
97 (* end thide *) | 95 (* end thide *) |
127 | 125 |
128 Print not. | 126 Print not. |
129 (** %\vspace{-.15in}% [[ | 127 (** %\vspace{-.15in}% [[ |
130 not = fun A : Prop => A -> False | 128 not = fun A : Prop => A -> False |
131 : Prop -> Prop | 129 : Prop -> Prop |
132 | |
133 ]] | 130 ]] |
134 | 131 |
135 We see that [not] is just shorthand for implication of [False]. We can use that fact explicitly in proofs. The syntax [~ P] expands to [not P]. *) | 132 %\smallskip{}%We see that [not] is just shorthand for implication of [False]. We can use that fact explicitly in proofs. The syntax [~ P] expands to [not P]. *) |
136 | 133 |
137 Theorem arith_neq' : ~ (2 + 2 = 5). | 134 Theorem arith_neq' : ~ (2 + 2 = 5). |
138 (* begin thide *) | 135 (* begin thide *) |
139 unfold not. | 136 unfold not. |
140 (** [[ | 137 (** [[ |
148 Qed. | 145 Qed. |
149 | 146 |
150 (** We also have conjunction, which we introduced in the last chapter. *) | 147 (** We also have conjunction, which we introduced in the last chapter. *) |
151 | 148 |
152 Print and. | 149 Print and. |
153 (** [[ | 150 (** %\vspace{-.15in}%[[ |
154 Inductive and (A : Prop) (B : Prop) : Prop := conj : A -> B -> A /\ B | 151 Inductive and (A : Prop) (B : Prop) : Prop := conj : A -> B -> A /\ B |
155 | |
156 ]] | 152 ]] |
157 | 153 |
158 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 %\smallskip{}%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]. *) |
159 | 155 |
160 Theorem and_comm : P /\ Q -> Q /\ P. | 156 Theorem and_comm : P /\ Q -> Q /\ P. |
161 | 157 |
162 (* begin thide *) | 158 (* begin thide *) |
163 (** We start by case analysis on the proof of [P /\ Q]. *) | 159 (** We start by case analysis on the proof of [P /\ Q]. *) |
172 ]] | 168 ]] |
173 | 169 |
174 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}% *) | 170 Every proof of a conjunction provides proofs for both conjuncts, so we get a single subgoal reflecting that. We can proceed by splitting this subgoal into a case for each conjunct of [Q /\ P].%\index{tactics!split}% *) |
175 | 171 |
176 split. | 172 split. |
177 (** 2 subgoals | 173 (** [[ |
174 2 subgoals | |
178 | 175 |
179 H : P | 176 H : P |
180 H0 : Q | 177 H0 : Q |
181 ============================ | 178 ============================ |
182 Q | 179 Q |
195 Qed. | 192 Qed. |
196 | 193 |
197 (** 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 [\/]. *) |
198 | 195 |
199 Print or. | 196 Print or. |
200 (** [[ | 197 (** %\vspace{-.15in}%[[ |
201 Inductive or (A : Prop) (B : Prop) : Prop := | 198 Inductive or (A : Prop) (B : Prop) : Prop := |
202 or_introl : A -> A \/ B | or_intror : B -> A \/ B | 199 or_introl : A -> A \/ B | or_intror : B -> A \/ B |
203 | |
204 ]] | 200 ]] |
205 | 201 |
206 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. *) | 202 %\smallskip{}%We see that there are two ways to prov 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. *) |
207 | 203 |
208 Theorem or_comm : P \/ Q -> Q \/ P. | 204 Theorem or_comm : P \/ Q -> Q \/ P. |
209 | 205 |
210 (* begin thide *) | 206 (* begin thide *) |
211 (** As in the proof for [and], we begin with case analysis, though this time we are met by two cases instead of one. *) | 207 (** As in the proof for [and], we begin with case analysis, though this time we are met by two cases instead of one. *) |
370 (** 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]." | 366 (** 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]." |
371 | 367 |
372 %\index{existential quantification}\index{Gallina terms!exists}\index{Gallina terms!ex}%Existential quantification is defined in the standard library. *) | 368 %\index{existential quantification}\index{Gallina terms!exists}\index{Gallina terms!ex}%Existential quantification is defined in the standard library. *) |
373 | 369 |
374 Print ex. | 370 Print ex. |
375 (** [[ | 371 (** %\vspace{-.15in}%[[ |
376 Inductive ex (A : Type) (P : A -> Prop) : Prop := | 372 Inductive ex (A : Type) (P : A -> Prop) : Prop := |
377 ex_intro : forall x : A, P x -> ex P | 373 ex_intro : forall x : A, P x -> ex P |
378 | |
379 ]] | 374 ]] |
380 | 375 |
381 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. *) | 376 %\smallskip{}%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. *) |
382 | 377 |
383 Theorem exist1 : exists x : nat, x + 1 = 2. | 378 Theorem exist1 : exists x : nat, x + 1 = 2. |
384 (* begin thide *) | 379 (* begin thide *) |
385 (** remove printing exists *) | 380 (** remove printing exists *) |
386 (** 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.) *) | 381 (** 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.) *) |
462 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. | 457 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. |
463 | 458 |
464 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}% *) | 459 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}% *) |
465 | 460 |
466 Print eq. | 461 Print eq. |
467 (** [[ | 462 (** %\vspace{-.15in}%[[ |
468 Inductive eq (A : Type) (x : A) : A -> Prop := eq_refl : x = x | 463 Inductive eq (A : Type) (x : A) : A -> Prop := eq_refl : x = x |
469 | |
470 ]] | 464 ]] |
471 | 465 |
472 Behind the scenes, uses of infix [=] are expanded to instances of [eq]. We see that [eq] has both a parameter [x] that is fixed and an extra unnamed argument of the same type. The type of [eq] allows us to state any equalities, even those that are provably false. However, examining the type of equality's sole constructor [eq_refl], we see that we can only _prove_ equality when its two arguments are syntactically equal. This definition turns out to capture all of the basic properties of equality, and the equality-manipulating tactics that we have seen so far, like [reflexivity] and [rewrite], are implemented treating [eq] as just another inductive type with a well-chosen definition. Another way of stating that definition is: equality is defined as the least reflexive relation. | 466 %\smallskip{}%Behind the scenes, uses of infix [=] are expanded to instances of [eq]. We see that [eq] has both a parameter [x] that is fixed and an extra unnamed argument of the same type. The type of [eq] allows us to state any equalities, even those that are provably false. However, examining the type of equality's sole constructor [eq_refl], we see that we can only _prove_ equality when its two arguments are syntactically equal. This definition turns out to capture all of the basic properties of equality, and the equality-manipulating tactics that we have seen so far, like [reflexivity] and [rewrite], are implemented treating [eq] as just another inductive type with a well-chosen definition. Another way of stating that definition is: equality is defined as the least reflexive relation. |
473 | 467 |
474 Returning to the example of [isZero], we can see how to work with hypotheses that use this predicate. *) | 468 Returning to the example of [isZero], we can see how to work with hypotheses that use this predicate. *) |
475 | 469 |
476 Theorem isZero_plus : forall n m : nat, isZero m -> n + m = n. | 470 Theorem isZero_plus : forall n m : nat, isZero m -> n + m = n. |
477 (* begin thide *) | 471 (* begin thide *) |
533 | 527 |
534 Check isZero_ind. | 528 Check isZero_ind. |
535 (** %\vspace{-.15in}% [[ | 529 (** %\vspace{-.15in}% [[ |
536 isZero_ind | 530 isZero_ind |
537 : forall P : nat -> Prop, P 0 -> forall n : nat, isZero n -> P n | 531 : forall P : nat -> Prop, P 0 -> forall n : nat, isZero n -> P n |
538 | 532 ]] |
539 ]] | 533 |
540 | 534 %\smallskip{}%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]. *) |
541 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]. *) | |
542 | 535 |
543 | 536 |
544 (* begin hide *) | 537 (* begin hide *) |
545 (* In-class exercises *) | 538 (* In-class exercises *) |
546 | 539 |
683 | 676 |
684 At this point, we would like to apply the inductive hypothesis, which is: | 677 At this point, we would like to apply the inductive hypothesis, which is: |
685 | 678 |
686 [[ | 679 [[ |
687 IHn : forall m : nat, even n -> even m -> even (n + m) | 680 IHn : forall m : nat, even n -> even m -> even (n + m) |
681 | |
688 ]] | 682 ]] |
689 | 683 |
690 Unfortunately, the goal mentions [n0] where it would need to mention [n] to match [IHn]. We could keep looking for a way to finish this proof from here, but it turns out that we can make our lives much easier by changing our basic strategy. Instead of inducting on the structure of [n], we should induct _on the structure of one of the [even] proofs_. This technique is commonly called%\index{rule induction}% _rule induction_ in programming language semantics. In the setting of Coq, we have already seen how predicates are defined using the same inductive type mechanism as datatypes, so the fundamental unity of rule induction with "normal" induction is apparent. | 684 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. |
691 | 685 |
692 Recall that tactics like [induction] and [destruct] may be passed numbers to refer to unnamed lefthand sides of implications in the conclusion, where the argument [n] refers to the [n]th such hypothesis. *) | 686 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. *) |
786 ]] | 780 ]] |
787 | 781 |
788 At this point it is useful to use a theorem from the standard library, which we also proved with a different name in the last chapter. We can search for a theorem that allows us to rewrite terms of the form [x + S y]. *) | 782 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]. *) |
789 | 783 |
790 SearchRewrite (_ + S _). | 784 SearchRewrite (_ + S _). |
791 | 785 (** %\vspace{-.15in}%[[ |
792 (** [[ | |
793 plus_n_Sm : forall n m : nat, S (n + m) = n + S m | 786 plus_n_Sm : forall n m : nat, S (n + m) = n + S m |
794 ]] | 787 ]] |
795 *) | 788 *) |
796 | 789 |
797 rewrite <- plus_n_Sm in H0. | 790 rewrite <- plus_n_Sm in H0. |