Library Logic
Coq's built-in logic is very small: the only primitives are
Inductive definitions, universal quantification (∀), and
implication (→), while all the other familiar logical
connectives -- conjunction, disjunction, negation, existential
quantification, even equality -- can be encoded using just these.
This chapter explains the encodings and shows how the tactics
we've seen can be used to carry out standard forms of logical
reasoning involving these connectives.
Propositions
Here is an example of an unprovable proposition:
Proofs and Evidence
We can see which proof term Coq constructs for a given Lemma by
using the Print directive:
Print silly.
Here, the eq_refl proof term witnesses the equality. (More on equality later!)
Just as we can implement natural number multiplication as a
function:
mult : nat → nat → nat
The proof term for an implication P → Q is a function that takes evidence for P as input and produces evidence for Q as its output.
Implications are functions
We can see that the proof term for the above lemma is indeed a
function:
Print silly_implication.
Defining Propositions
- Typically, rules are defined inductively, just like any other datatype.
- Sometimes a proposition is declared to be true without substantiating evidence. Such propositions are called axioms.
Conjunction (Logical "and")
The intuition behind this definition is simple: to
construct evidence for and P Q, we must provide evidence
for P and evidence for Q. More precisely:
Since we'll be using conjunction a lot, let's introduce a more
familiar-looking infix notation for it.
- conj p q can be taken as evidence for and P Q if p
is evidence for P and q is evidence for Q; and
- this is the only way to give evidence for and P Q -- that is, if someone gives us evidence for and P Q, we know it must have the form conj p q, where p is evidence for P and q is evidence for Q.
(The type_scope annotation tells Coq that this notation
will be appearing in propositions, not values.)
Consider the "type" of the constructor conj:
Notice that it takes 4 inputs -- namely the propositions P
and Q and evidence for P and Q -- and returns as output the
evidence of P ∧ Q.
"Introducing" Conjuctions
Besides the elegance of building everything up from a tiny foundation, what's nice about defining conjunction this way is that we can prove statements involving conjunction using the tactics that we already know. For example, if the goal statement is a conjuction, we can prove it by applying the single constructor conj, which (as can be seen from the type of conj) solves the current goal and leaves the two parts of the conjunction as subgoals to be proved separately.
Theorem and_example :
(0 = 0) ∧ (4 = mult 2 2).
Proof.
apply conj.
Case "left". reflexivity.
Case "right". reflexivity. Qed.
(0 = 0) ∧ (4 = mult 2 2).
Proof.
apply conj.
Case "left". reflexivity.
Case "right". reflexivity. Qed.
Theorem and_example' :
(0 = 0) ∧ (4 = mult 2 2).
Proof.
split.
Case "left". reflexivity.
Case "right". reflexivity. Qed.
"Eliminating" conjunctions
Conversely, the inversion tactic can be used to take a conjunction hypothesis in the context, calculate what evidence must have been used to build it, and add variables representing this evidence to the proof context.Theorem proj1 : ∀ P Q : Prop,
P ∧ Q → P.
Proof.
intros P Q H.
inversion H as [HP HQ].
apply HP. Qed.
Theorem and_commut : ∀ P Q : Prop,
P ∧ Q → Q ∧ P.
Proof.
intros P Q H.
inversion H as [HP HQ].
split.
Case "left". apply HQ.
Case "right". apply HP. Qed.
P ∧ Q → Q ∧ P.
Proof.
intros P Q H.
inversion H as [HP HQ].
split.
Case "left". apply HQ.
Case "right". apply HP. Qed.
Exercise: 2 stars (and_assoc)
In the following proof, notice how the nested pattern in the inversion breaks the hypothesis H : P ∧ (Q ∧ R) down into HP: P, HQ : Q, and HR : R. Finish the proof from there:Theorem and_assoc : ∀ P Q R : Prop,
P ∧ (Q ∧ R) → (P ∧ Q) ∧ R.
Proof.
intros P Q R H.
inversion H as [HP [HQ HR]].
split.
Case "left". inversion H. split.
SCase "left". apply HP.
SCase "right". apply HQ.
Case "right". apply HR.
Qed.
☐
Definition iff (P Q : Prop) := (P → Q) ∧ (Q → P).
Notation "P <-> Q" := (iff P Q)
(at level 95, no associativity)
: type_scope.
Theorem iff_implies : ∀ P Q : Prop,
(P ↔ Q) → P → Q.
Proof.
intros P Q H.
inversion H as [HAB HBA]. apply HAB. Qed.
Notation "P <-> Q" := (iff P Q)
(at level 95, no associativity)
: type_scope.
Theorem iff_implies : ∀ P Q : Prop,
(P ↔ Q) → P → Q.
Proof.
intros P Q H.
inversion H as [HAB HBA]. apply HAB. Qed.
Theorem iff_sym : ∀ P Q : Prop,
(P ↔ Q) → (Q ↔ P).
Proof.
intros P Q H.
inversion H as [HAB HBA].
split.
Case "->". apply HBA.
Case "<-". apply HAB. Qed.
(P ↔ Q) → (Q ↔ P).
Proof.
intros P Q H.
inversion H as [HAB HBA].
split.
Case "->". apply HBA.
Case "<-". apply HAB. Qed.
Exercise: 1 star, optional (iff_properties)
Using the above proof that ↔ is symmetric (iff_sym) as a guide, prove that it is also reflexive and transitive.Theorem iff_refl : ∀ P : Prop,
P ↔ P.
Proof.
intros P. split.
Case "->". intros. apply H.
Case "<-". intros. apply H.
Qed.
Theorem iff_trans : ∀ P Q R : Prop,
(P ↔ Q) → (Q ↔ R) → (P ↔ R).
Proof.
intros P Q R H1 H2. split.
Case "->". intros. apply H1 in H. apply H2 in H. apply H.
Case "<-". intros. apply H2 in H. apply H1 in H. apply H.
Qed.
Hint: If you have an iff hypothesis in the context, you can use
inversion to break it into two separate implications. (Think
about why this works.) ☐
Some of Coq's tactics treat iff statements specially, thus
avoiding the need for some low-level manipulation when reasoning
with them. In particular, rewrite can be used with iff
statements, not just equalities.
Disjunction (Logical "or")
Implementing Disjunction
Inductive or (P Q : Prop) : Prop :=
| or_introl : P → or P Q
| or_intror : Q → or P Q.
Notation "P \/ Q" := (or P Q) : type_scope.
| or_introl : P → or P Q
| or_intror : Q → or P Q.
Notation "P \/ Q" := (or P Q) : type_scope.
Consider the "type" of the constructor or_introl:
It takes 3 inputs, namely the propositions P, Q and
evidence of P, and returns, as output, the evidence of P ∨ Q.
Next, look at the type of or_intror:
It is like or_introl but it requires evidence of Q
instead of evidence of P.
Intuitively, there are two ways of giving evidence for P ∨ Q:
- give evidence for P (and say that it is P you are giving
evidence for -- this is the function of the or_introl
constructor), or
- give evidence for Q, tagged with the or_intror constructor.
Since P ∨ Q has two constructors, doing inversion on a hypothesis of type P ∨ Q yields two subgoals.
Theorem or_commut : ∀ P Q : Prop,
P ∨ Q → Q ∨ P.
Proof.
intros P Q H.
inversion H as [HP | HQ].
Case "left". apply or_intror. apply HP.
Case "right". apply or_introl. apply HQ. Qed.
P ∨ Q → Q ∨ P.
Proof.
intros P Q H.
inversion H as [HP | HQ].
Case "left". apply or_intror. apply HP.
Case "right". apply or_introl. apply HQ. Qed.
From here on, we'll use the shorthand tactics left and right
in place of apply or_introl and apply or_intror.
Theorem or_commut' : ∀ P Q : Prop,
P ∨ Q → Q ∨ P.
Proof.
intros P Q H.
inversion H as [HP | HQ].
Case "left". right. apply HP.
Case "right". left. apply HQ. Qed.
Theorem or_distributes_over_and_1 : ∀ P Q R : Prop,
P ∨ (Q ∧ R) → (P ∨ Q) ∧ (P ∨ R).
Proof.
intros P Q R. intros H. inversion H as [HP | [HQ HR]].
Case "left". split.
SCase "left". left. apply HP.
SCase "right". left. apply HP.
Case "right". split.
SCase "left". right. apply HQ.
SCase "right". right. apply HR. Qed.
Theorem or_distributes_over_and_2 : ∀ P Q R : Prop,
(P ∨ Q) ∧ (P ∨ R) → P ∨ (Q ∧ R).
Proof.
intros P Q R H.
inversion H as [Hl Hr].
inversion Hl as [Hll | Hlr].
Case "P". left. apply Hll.
Case "Q". inversion Hr as [Hrl | Hrr].
SCase "P". left. apply Hrl.
SCase "R". right. split. apply Hlr. apply Hrr.
Qed.
(P ∨ Q) ∧ (P ∨ R) → P ∨ (Q ∧ R).
Proof.
intros P Q R H.
inversion H as [Hl Hr].
inversion Hl as [Hll | Hlr].
Case "P". left. apply Hll.
Case "Q". inversion Hr as [Hrl | Hrr].
SCase "P". left. apply Hrl.
SCase "R". right. split. apply Hlr. apply Hrr.
Qed.
Theorem or_distributes_over_and : ∀ P Q R : Prop,
P ∨ (Q ∧ R) ↔ (P ∨ Q) ∧ (P ∨ R).
Proof.
intros P Q R.
split. apply or_distributes_over_and_1. apply or_distributes_over_and_2.
Qed.
P ∨ (Q ∧ R) ↔ (P ∨ Q) ∧ (P ∨ R).
Proof.
intros P Q R.
split. apply or_distributes_over_and_1. apply or_distributes_over_and_2.
Qed.
☐
Relating ∧ and ∨ with andb and orb (advanced)
Theorem andb_prop : ∀ b c,
andb b c = true → b = true ∧ c = true.
Proof.
intros b c H.
destruct b.
Case "b = true". destruct c.
SCase "c = true". apply conj. reflexivity. reflexivity.
SCase "c = false". inversion H.
Case "b = false". inversion H. Qed.
Theorem andb_true_intro : ∀ b c,
b = true ∧ c = true → andb b c = true.
Proof.
intros b c H.
inversion H.
rewrite H0. rewrite H1. reflexivity. Qed.
Theorem andb_false : ∀ b c,
andb b c = false → b = false ∨ c = false.
Proof.
intros b c H.
destruct b. destruct c.
Case "b = true, c = true". inversion H.
Case "b = true, c = false". right. reflexivity.
Case "b = false". left. reflexivity.
Qed.
Theorem orb_prop : ∀ b c,
orb b c = true → b = true ∨ c = true.
Proof.
intros b c H. destruct b.
Case "b = true". left. reflexivity.
Case "b = false". destruct c.
SCase "c = true". right. reflexivity.
SCase "c = false". inversion H.
Qed.
Theorem orb_false_elim : ∀ b c,
orb b c = false → b = false ∧ c = false.
Proof.
intros b c H. destruct b. destruct c.
Case "b = true, c = true". inversion H.
Case "b = true, c = false". inversion H.
Case "b = false". destruct c. inversion H.
split. reflexivity. reflexivity.
Qed.
andb b c = false → b = false ∨ c = false.
Proof.
intros b c H.
destruct b. destruct c.
Case "b = true, c = true". inversion H.
Case "b = true, c = false". right. reflexivity.
Case "b = false". left. reflexivity.
Qed.
Theorem orb_prop : ∀ b c,
orb b c = true → b = true ∨ c = true.
Proof.
intros b c H. destruct b.
Case "b = true". left. reflexivity.
Case "b = false". destruct c.
SCase "c = true". right. reflexivity.
SCase "c = false". inversion H.
Qed.
Theorem orb_false_elim : ∀ b c,
orb b c = false → b = false ∧ c = false.
Proof.
intros b c H. destruct b. destruct c.
Case "b = true, c = true". inversion H.
Case "b = true, c = false". inversion H.
Case "b = false". destruct c. inversion H.
split. reflexivity. reflexivity.
Qed.
☐
Falsehood
Intuition: False is a proposition for which there is no way
to give evidence.
Since False has no constructors, inverting an assumption
of type False always yields zero subgoals, allowing us to
immediately prove any goal.
How does this work? The inversion tactic breaks contra into
each of its possible cases, and yields a subgoal for each case.
As contra is evidence for False, it has no possible cases,
hence, there are no possible subgoals and the proof is done.
Conversely, the only way to prove False is if there is already something nonsensical or contradictory in the context:
Actually, since the proof of False_implies_nonsense
doesn't actually have anything to do with the specific nonsensical
thing being proved; it can easily be generalized to work for an
arbitrary P:
The Latin ex falso quodlibet means, literally, "from
falsehood follows whatever you please." This theorem is also
known as the principle of explosion.
Truth
Since we have defined falsehood in Coq, one might wonder whether it is possible to define truth in the same way. We can.Exercise: 2 stars, advanced (True)
Define True as another inductively defined proposition. (The intution is that True should be a proposition for which it is trivial to give evidence.)
☐
However, unlike False, which we'll use extensively, True is
used fairly rarely. By itself, it is trivial (and therefore
uninteresting) to prove as a goal, and it carries no useful
information as a hypothesis. But it can be useful when defining
complex Props using conditionals, or as a parameter to
higher-order Props.
It takes a little practice to get used to working with
negation in Coq. Even though you can see perfectly well why
something is true, it can be a little hard at first to get things
into the right configuration so that Coq can see it! Here are
proofs of a few familiar facts about negation to get you warmed
up.
Theorem contradiction_implies_anything : ∀ P Q : Prop,
(P ∧ ¬P) → Q.
Proof.
intros P Q H. inversion H as [HP HNA]. unfold not in HNA.
apply HNA in HP. inversion HP. Qed.
Theorem double_neg : ∀ P : Prop,
P → ~~P.
Proof.
intros P H. unfold not. intros G. apply G. apply H. Qed.
(P ∧ ¬P) → Q.
Proof.
intros P Q H. inversion H as [HP HNA]. unfold not in HNA.
apply HNA in HP. inversion HP. Qed.
Theorem double_neg : ∀ P : Prop,
P → ~~P.
Proof.
intros P H. unfold not. intros G. apply G. apply H. Qed.
Exercise: 2 stars, advanced (double_neg_inf)
Write an informal proof of double_neg:Exercise: 2 stars (contrapositive)
Theorem contrapositive : ∀ P Q : Prop,
(P → Q) → (¬Q → ¬P).
Proof.
intros P Q Himpl Hnot.
unfold not. intros HP.
apply Himpl in HP.
apply contradiction_implies_anything with Q.
split. apply HP. apply Hnot.
Qed.
(P → Q) → (¬Q → ¬P).
Proof.
intros P Q Himpl Hnot.
unfold not. intros HP.
apply Himpl in HP.
apply contradiction_implies_anything with Q.
split. apply HP. apply Hnot.
Qed.
Theorem not_both_true_and_false : ∀ P : Prop,
¬ (P ∧ ¬P).
Proof.
intro P. intros contr. apply contradiction_implies_anything with P. apply contr.
Qed.
¬ (P ∧ ¬P).
Proof.
intro P. intros contr. apply contradiction_implies_anything with P. apply contr.
Qed.
☐
Exercise: 1 star, advanced (informal_not_PNP)
Write an informal proof (in English) of the proposition ∀ P : Prop, ~(P ∧ ¬P).
☐
Constructive logic
Note that some theorems that are true in classical logic are not provable in Coq's (constructive) logic. E.g., let's look at how this proof gets stuck...Exercise: 5 stars, advanced, optional (classical_axioms)
For those who like a challenge, here is an exercise taken from the Coq'Art book (p. 123). The following five statements are often considered as characterizations of classical logic (as opposed to constructive logic, which is what is "built in" to Coq). We can't prove them in Coq, but we can consistently add any one of them as an unproven axiom if we wish to work in classical logic. Prove that these five propositions are equivalent.Definition peirce := ∀ P Q: Prop,
((P→Q)->P)->P.
Definition classic := ∀ P:Prop,
~~P → P.
Definition excluded_middle := ∀ P:Prop,
P ∨ ¬P.
Definition de_morgan_not_and_not := ∀ P Q:Prop,
~(~P ∧ ¬Q) → P∨Q.
Definition implies_to_or := ∀ P Q:Prop,
(P→Q) → (¬P∨Q).
☐
Exercise: 3 stars (excluded_middle_irrefutable)
This theorem implies that it is always safe to add a decidability axiom (i.e. an instance of excluded middle) for any particular Prop P. Why? Because we cannot prove the negation of such an axiom; if we could, we would have both ¬ (P ∨ ¬P) and ¬ ¬ (P ∨ ¬P), a contradiction.
Since inequality involves a negation, it again requires
a little practice to be able to work with it fluently. Here
is one very useful trick. If you are trying to prove a goal
that is nonsensical (e.g., the goal state is false = true),
apply the lemma ex_falso_quodlibet to change the goal to
False. This makes it easier to use assumptions of the form
¬P that are available in the context -- in particular,
assumptions of the form x≠y.
Theorem not_false_then_true : ∀ b : bool,
b ≠ false → b = true.
Proof.
intros b H. destruct b.
Case "b = true". reflexivity.
Case "b = false".
unfold not in H.
apply ex_falso_quodlibet.
apply H. reflexivity. Qed.
b ≠ false → b = true.
Proof.
intros b H. destruct b.
Case "b = true". reflexivity.
Case "b = false".
unfold not in H.
apply ex_falso_quodlibet.
apply H. reflexivity. Qed.
☐