comparison src/Predicates.v @ 45:cb135b19adb8

Propositional logic
author Adam Chlipala <adamc@hcoop.net>
date Sat, 27 Sep 2008 14:57:08 -0400
parents
children 29058fd5448b
comparison
equal deleted inserted replaced
44:b1e137b4aafe 45:cb135b19adb8
1 (* Copyright (c) 2008, Adam Chlipala
2 *
3 * This work is licensed under a
4 * Creative Commons Attribution-Noncommercial-No Derivative Works 3.0
5 * Unported License.
6 * The license text is available at:
7 * http://creativecommons.org/licenses/by-nc-nd/3.0/
8 *)
9
10 (* begin hide *)
11 Require Import List.
12
13 Require Import Tactics.
14
15 Set Implicit Arguments.
16 (* end hide *)
17
18
19 (** %\chapter{Inductive Predicates}% *)
20
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: *)
22
23 Print unit.
24 (** [[
25
26 Inductive unit : Set := tt : unit
27 ]] *)
28
29 Print True.
30 (** [[
31
32 Inductive True : Prop := I : True
33 ]] *)
34
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].
36
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.
38
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.
40
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. *)
42
43
44 (** * Propositional logic *)
45
46 (** Let us begin with a brief tour through the definitions of the connectives for propositional logic. We will work within a Coq section that provides us with a set of propositional variables. In Coq parlance, these are just terms of type [Prop.] *)
47
48 Section Propositional.
49 Variables P Q : Prop.
50
51 (** In Coq, the most basic propositional connective is implication, written [->], which we have already used in almost every proof. Rather than being defined inductively, implication is built into Coq as the function type constructor.
52
53 We have also already seen the definition of [True]. For a demonstration of a lower-level way of establishing proofs of inductive predicates, we turn to this trivial theorem. *)
54
55 Theorem obvious : True.
56 apply I.
57 Qed.
58
59 (** We may always use the [apply] tactic to take a proof step based on applying a particular constructor of the inductive predicate that we are trying to establish. Sometimes there is only one constructor that could possibly apply, in which case a shortcut is available: *)
60
61 Theorem obvious' : True.
62 constructor.
63 Qed.
64
65 (** There is also a predicate [False], which is the Curry-Howard mirror image of [Empty_set] from the last chapter. *)
66
67 Print False.
68 (** [[
69
70 Inductive False : Prop :=
71 ]] *)
72
73 (** We can conclude anything from [False], doing case analysis on a proof of [False] in the same way we might do case analysis on, say, a natural number. Since there are no cases to consider, any such case analysis succeeds immediately in proving the goal. *)
74
75 Theorem False_imp : False -> 2 + 2 = 5.
76 destruct 1.
77 Qed.
78
79 (** In a consistent context, we can never build a proof of [False]. In inconsistent contexts that appear in the courses of proofs, it is usually easiest to proceed by demonstrating that inconsistency with an explicit proof of [False]. *)
80
81 Theorem arith_neq : 2 + 2 = 5 -> 9 + 9 = 835.
82 intro.
83
84 (** At this point, we have an inconsistent hypothesis [2 + 2 = 5], so the specific conclusion is not important. We use the [elimtype] tactic to state a proposition, telling Coq that we wish to construct a proof of the new proposition and then prove the original goal by case analysis on the structure of the new auxiliary proof. Since [False] has no constructors, [elimtype False] simply leaves us with the obligation to prove [False]. *)
85
86 elimtype False.
87 (** [[
88
89 H : 2 + 2 = 5
90 ============================
91 False
92 ]] *)
93
94 (** For now, we will leave the details of this proof about arithmetic to [crush]. *)
95
96 crush.
97 Qed.
98
99 (** A related notion to [False] is logical negation. *)
100
101 Print not.
102 (** [[
103
104 not = fun A : Prop => A -> False
105 : Prop -> Prop
106 ]] *)
107
108 (** We see that [not] is just shorthand for implication of [False]. We can use that fact explicitly in proofs. The syntax [~P] expands to [not P]. *)
109
110 Theorem arith_neq' : ~ (2 + 2 = 5).
111 unfold not.
112
113 (** [[
114
115 ============================
116 2 + 2 = 5 -> False
117 ]] *)
118
119 crush.
120 Qed.
121
122 (** We also have conjunction, which we introduced in the last chapter. *)
123
124 Print and.
125 (** [[
126
127 Inductive and (A : Prop) (B : Prop) : Prop := conj : A -> B -> A /\ B
128 ]] *)
129
130 (** The interested reader can check that [and] has a Curry-Howard doppleganger called [prod], the type of pairs. However, it is generally most convenient to reason about conjunction using tactics. An explicit proof of commutativity of [and] illustrates the usual suspects for such tasks. [/\] is an infix shorthand for [and]. *)
131
132 Theorem and_comm : P /\ Q -> Q /\ P.
133 (** We start by case analysis on the proof of [P /\ Q]. *)
134
135 destruct 1.
136 (** [[
137
138 H : P
139 H0 : Q
140 ============================
141 Q /\ P
142 ]] *)
143
144 (** Every proof of a conjunction provides proofs for both conjuncts, so we get a single subgoal reflecting that. We can proceed by splitting this subgoal into a case for each conjunct of [Q /\ P]. *)
145
146 split.
147 (** [[
148 2 subgoals
149
150 H : P
151 H0 : Q
152 ============================
153 Q
154
155 subgoal 2 is:
156 P
157 ]] *)
158
159 (** In each case, the conclusion is among our hypotheses, so the [assumption] tactic finishes the process. *)
160
161 assumption.
162 assumption.
163 Qed.
164
165 (** Coq disjunction is called [or] and abbreviated with the infix operator [\/]. *)
166
167 Print or.
168 (** [[
169
170 Inductive or (A : Prop) (B : Prop) : Prop :=
171 or_introl : A -> A \/ B | or_intror : B -> A \/ B
172 ]] *)
173
174 (** We see that there are two ways to prove a disjunction: prove the first disjunct or prove the second. The Curry-Howard analogue of this is the Coq [sum] type. We can demonstrate the main tactics here with another proof of commutativity. *)
175
176 Theorem or_comm : P \/ Q -> Q \/ P.
177 (** As in the proof for [and], we begin with case analysis, though this time we are met by two cases instead of one. *)
178 destruct 1.
179 (** [[
180
181 2 subgoals
182
183 H : P
184 ============================
185 Q \/ P
186
187 subgoal 2 is:
188 Q \/ P
189 ]] *)
190
191 (** We can see that, in the first subgoal, we want to prove the disjunction by proving its second disjunct. The [right] tactic telegraphs this intent. *)
192
193 right; assumption.
194
195 (** The second subgoal has a symmetric proof.
196
197 [[
198
199 1 subgoal
200
201 H : Q
202 ============================
203 Q \/ P
204 ]] *)
205
206 left; assumption.
207 Qed.
208
209 End Propositional.
210