LogicLogic in Coq
Require Export MoreProp.
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.
Inductive and (P Q : Prop) : Prop :=
conj : P → Q → (and P Q).
Intuition: To construct evidence for and P Q, we must
provide evidence for P and evidence for Q. Conversely,
evidence for and P Q must be of the form conj p q, where p
is evidence for P and q is evidence for Q.
More familiar syntax:
Notation "P ∧ Q" := (and P Q) : type_scope.
Consider the "type" of the constructor conj:
Check conj.
(* ===> forall P Q : Prop, P -> Q -> P /λ Q *)
The tactics we already know can be used prove things involving
conjunctions.
Theorem and_example :
(beautiful 0) ∧ (beautiful 3).
Proof.
apply conj.
Case "left". apply b_0.
Case "right". apply b_3. Qed.
Just for convenience, we can use the tactic split as a shorthand for
apply conj.
Theorem and_example' :
(ev 0) ∧ (ev 4).
Proof.
split.
Case "left". apply ev_0.
Case "right". apply ev_SS. apply ev_SS. apply ev_0. Qed.
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.
(* WORK IN CLASS *) Admitted.
Definition iff (P Q : Prop) := (P → Q) ∧ (Q → P).
Notation "P ↔ Q" := (iff P Q)
(at level 95, no associativity)
: type_scope.
Theorem iff_sym : ∀P Q : Prop,
(P ↔ Q) → (Q ↔ P).
Proof.
(* WORK IN CLASS *) Admitted.
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.
Check or_introl.
(* ===> forall P Q : Prop, P -> P λ/ Q *)
Check or_intror.
(* ===> forall P Q : Prop, Q -> P λ/ Q *)
Intuitively, there are two ways of giving evidence for P ∨ Q:
Since P ∨ Q has two constructors, doing inversion on a
hypothesis of type P ∨ Q yields two subgoals.
- 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.
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.
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.
To prove the following proposition, which tactics will we need besides intros and apply?
(1) split, inversion, left and right
(2) inversion, left and right
(3) inversion, and one of left and right
(4) only left
(5) only right
(6) only inversion
(7) none of the above
∀P Q : Prop, P ∧ Q → P ∨ Q.
What about this one?
(1) split, inversion, left and right
(2) inversion, left and right
(3) inversion, and one of left and right
(4) only left
(5) only right
(6) only inversion
(7) none of the above
∀P Q : Prop, P → (P ∨ Q).
What about this one?
(1) split, inversion, left and right
(2) inversion, left and right
(3) inversion, and one of left and right
(4) only left
(5) only right
(6) only inversion
(7) none of the above
∀P Q : Prop, P ∧ Q → P.
What about this one?
(1) split, inversion, left and right
(2) inversion, left and right
(3) inversion, and one of left and right
(4) only left
(5) only right
(6) only inversion
(7) none of the above
∀P Q : Prop, P ∨ Q ↔ Q ∨ P.
Falsehood
Inductive False : Prop := .
Intuition: False is a proposition for which there is no way
to give evidence.
Theorem False_implies_nonsense :
False → 2 + 2 = 5.
Proof.
intros contra.
inversion contra. Qed.
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:
Theorem nonsense_implies_False :
2 + 2 = 5 → False.
Proof.
intros contra.
inversion contra. Qed.
Theorem ex_falso_quodlibet : ∀(P:Prop),
False → P.
Proof.
(* WORK IN CLASS *) Admitted.
Definition not (P:Prop) := P → False.
Intuition: If we could prove P, then we could prove
False (and hence we could prove anything at all).
Notation "¬ x" := (not x) : type_scope.
Check not.
(* ===> Prop -> Prop *)
Theorem not_False :
¬ False.
Proof.
unfold not. intros H. inversion H. Qed.
Theorem contradiction_implies_anything : ∀P Q : Prop,
(P ∧ ¬P) → Q.
Proof.
(* WORK IN CLASS *) Admitted.
Theorem double_neg : ∀P : Prop,
P → ~~P.
Proof.
(* WORK IN CLASS *) Admitted.
Theorem five_not_even :
¬ ev 5.
Proof.
(* WORK IN CLASS *) Admitted.
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...
Theorem classic_double_neg : ∀P : Prop,
~~P → P.
Proof.
(* WORK IN CLASS *) Admitted.
Notation "x ≠ y" := (¬ (x = y)) : type_scope.
A useful proof idiom: If you are trying to prove a
goal that is nonsensical, 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.
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.
To prove the following proposition, which tactics will we need
besides intros and apply?
(1) inversion, unfold, left and right
(2) inversion and unfold
(3) only inversion
(4) one of left and right
(5) only unfold
(6) none of the above
∀X, ∀a b : X, (a=b) ∧ (a≠b) → False.
To prove the following proposition, which tactics will we
need besides intros and apply?
(1) inversion, unfold, left and right
(2) inversion and unfold
(3) only inversion
(4) one of left and right
(5) only unfold
(6) none of the above
∀P Q : Prop, P ∨ Q → ~~(P ∨ Q).
To prove the following proposition, which tactics will we
need besides intros and apply?
(1) inversion, unfold, left and right
(2) inversion and unfold
(3) only inversion
(4) one of left and right
(5) only unfold
(6) none of the above
∀A B: Prop, A → (A ∨ ~~B).
To prove the following proposition, which tactics will we need
besides intros and apply?
(1) inversion, unfold, left and right
(2) inversion and unfold
(3) only inversion
(4) one of left and right
(5) only unfold
(6) none of the above
∀P Q: Prop, P ∨ Q → ~~P ∨ ~~Q.
To prove the following proposition, which tactics will we need
besides intros and apply?
(1) inversion, unfold, left and right
(2) inversion and unfold
(3) only inversion
(4) one of left and right
(5) only unfold
(6) none of the above
∀A : Prop, 1=0 → (A ∨ ¬A).
Existential Quantification
Inductive ex (X:Type) (P : X→Prop) : Prop :=
ex_intro : ∀(witness:X), P witness → ex X P.
That is, ex is a family of propositions indexed by a type X
and a property P over X. In order to give evidence for the
assertion "there exists an x for which the property P holds"
we must actually name a witness — a specific value x — and
then give evidence for P x, i.e., evidence that x has the
property P.
Standard notations for existentials:
Notation "'exists' x , p" := (ex _ (fun x ⇒ p))
(at level 200, x ident, right associativity) : type_scope.
Notation "'exists' x : X , p" := (ex _ (fun x:X ⇒ p))
(at level 200, x ident, right associativity) : type_scope.
The details of how these notations work are not too
important: the critical point is that they allow us to write
∃ x, P or ∃ x:X, P, just as we do with the ∀
quantifier.
To prove an existential statement, use apply ex_intro.
Example exists_example_1 : ∃n, n + (n × n) = 6.
Proof.
apply ex_intro with (witness:=2).
reflexivity. Qed.
Note that we have to explicitly give the witness.
Or, instead of writing apply ex_intro with (witness:=e) all the
time, we can use the convenient shorthand ∃ e, which means
the same thing.
Example exists_example_1' : ∃n, n + (n × n) = 6.
Proof.
∃2.
reflexivity. Qed.
Conversely, if we have an existential hypothesis in the
context, we can eliminate it with inversion. Note the use
of the as... pattern to name the variable that Coq
introduces to name the witness value and get evidence that
the hypothesis holds for the witness. (If we don't
explicitly choose one, Coq will just call it witness, which
makes proofs confusing.)
Theorem exists_example_2 : ∀n,
(∃m, n = 4 + m) →
(∃o, n = 2 + o).
Proof.
intros n H.
inversion H as [m Hm].
∃(2 + m).
apply Hm. Qed.
Equality
Inductive eq {X:Type} : X → X → Prop :=
refl_equal : ∀x, eq x x.
Standard infix notation:
Notation "x = y" := (eq x y)
(at level 70, no associativity)
: type_scope.
The definition of = is a bit subtle. The way to think about it
is that, given a set X, it defines a family of propositions
"x is equal to y," indexed by pairs of values (x and y)
from X. There is just one way of constructing evidence for
members of this family: applying the constructor refl_equal to a
type X and a value x : X yields evidence that x is equal to
x.
We can use
refl_equal to construct evidence that, for example, 2 = 2.
Can we also use it to construct evidence that 1 + 1 = 2? Yes:
indeed, it is the very same piece of evidence! The reason is that
Coq treats as "the same" any two terms that are convertible
according to a simple set of computation rules. These rules,
which are similar to those used by Eval compute, include
evaluation of function application, inlining of definitions, and
simplification of matches.
Lemma four: 2 + 2 = 1 + 3.
Proof.
apply refl_equal.
Qed.
The reflexivity tactic that we have used to prove equalities up
to now is essentially just short-hand for apply refl_equal.
Which of the following is correct proof object for the proposition
(1) refl_equal 4
(2) ex_intro nat (fun z ⇒ (z + 3 = 4)) 1 refl_equal
(3) ex_intro nat (z + 3 = 4) 1 (refl_equal 4)
(4) ex_intro nat (fun z ⇒ (z + 3 = 4)) 1 (refl_equal 4)
(5) ex_intro nat (fun z ⇒ (z + 3 = 4)) 1 (refl_equal 1)
(6) none of the above
∃x, x + 3 = 4
?
Evidence-carrying booleans.
Inductive sumbool (A B : Prop) : Set :=
| left : A → sumbool A B
| right : B → sumbool A B.
Notation "{ A } + { B }" := (sumbool A B) : type_scope.
Think of sumbool as being like the boolean type, but instead
of its values being just true and false, they carry evidence
of truth or falsity. This means that when we destruct them, we
are left with the relevant evidence as a hypothesis — just as with or.
(In fact, the definition of sumbool is almost the same as for or.
The only difference is that values of sumbool are declared to be in
Set rather than in Prop; this is a technical distinction
that allows us to compute with them.)
Here's how we can define a sumbool for equality on nats
Theorem eq_nat_dec : ∀n m : nat, {n = m} + {n ≠ m}.
Proof.
intros n.
induction n as [|n'].
Case "n = 0".
intros m.
destruct m as [|m'].
SCase "m = 0".
left. reflexivity.
SCase "m = S m'".
right. intros contra. inversion contra.
Case "n = S n'".
intros m.
destruct m as [|m'].
SCase "m = 0".
right. intros contra. inversion contra.
SCase "m = S m'".
destruct IHn' with (m := m') as [eq | neq].
left. apply f_equal. apply eq.
right. intros Heq. inversion Heq as [Heq']. apply neq. apply Heq'.
Defined.
Read as a theorem, this says that equality on nats is decidable:
that is, given two nat values, we can always produce either
evidence that they are equal or evidence that they are not.
Read computationally, eq_nat_dec takes two nat values and returns
a sumbool constructed with left if they are equal and right
if they are not; this result can be tested with a match or, better,
with an if-then-else, just like a regular boolean.
(Notice that we ended this proof with Defined rather than Qed.
The only difference this makes is that the proof becomes transparent,
meaning that its definition is available when Coq tries to do reductions,
which is important for the computational interpretation.)
Here's a simple example illustrating the advantages of the sumbool form.
Definition override' {X: Type} (f: nat→X) (k:nat) (x:X) : nat→X:=
fun (k':nat) ⇒ if eq_nat_dec k k' then x else f k'.
Theorem override_same' : ∀(X:Type) x1 k1 k2 (f : nat→X),
f k1 = x1 →
(override' f k1 x1) k2 = f k2.
Proof.
intros X x1 k1 k2 f. intros Hx1.
unfold override'.
destruct (eq_nat_dec k1 k2). (* observe what appears as a hypothesis *)
Case "k1 = k2".
rewrite ← e.
symmetry. apply Hx1.
Case "k1 ≠ k2".
reflexivity. Qed.
Compare this to the more laborious proof (in MoreCoq.v) for the
version of override defined using beq_nat, where we had to
use the auxiliary lemma beq_nat_true to convert a fact about booleans
to a Prop.
Inversion, Again (Advanced)
- takes a hypothesis H whose type P is inductively defined,
and
- for each constructor C in P's definition,
- generates a new subgoal in which we assume H was
built with C,
- adds the arguments (premises) of C to the context of
the subgoal as extra hypotheses,
- matches the conclusion (result type) of C against the
current goal and calculates a set of equalities that must
hold in order for C to be applicable,
- adds these equalities to the context (and, for convenience,
rewrites them in the goal), and
- if the equalities are not satisfiable (e.g., they involve things like S n = O), immediately solves the subgoal.
- generates a new subgoal in which we assume H was
built with C,