LogicLogic in Coq
We have seen...
In this chapter we will introduce several more flavors of both
propositions and proofs.
Like everything in Coq, propositions are typed:
- propositions: factual claims
- equality propositions (e1 = e2)
- implications (P → Q)
- quantified propositions (∀ x, P)
- proofs: ways of presenting evidence for the truth of a proposition
Check 3 = 3.
(* ===> Prop *)
Check ∀n m : nat, n + m = m + n.
(* ===> Prop *)
(* ===> Prop *)
Check ∀n m : nat, n + m = m + n.
(* ===> Prop *)
Note that all syntactically well-formed propositions have type Prop in Coq, regardless of whether they are true.
Check 2 = 2.
(* ===> Prop *)
Check ∀n : nat, n = 2.
(* ===> Prop *)
Check 3 = 4.
(* ===> Prop *)
(* ===> Prop *)
Check ∀n : nat, n = 2.
(* ===> Prop *)
Check 3 = 4.
(* ===> Prop *)
So far, we've seen one primary place that propositions can appear: in Theorem (and Lemma and Example) declarations.
Theorem plus_2_2_is_4 :
2 + 2 = 4.
Proof. reflexivity. Qed.
2 + 2 = 4.
Proof. reflexivity. Qed.
Definition plus_claim : Prop := 2 + 2 = 4.
Check plus_claim.
(* ===> plus_claim : Prop *)
Theorem plus_claim_is_true :
plus_claim.
Proof. reflexivity. Qed.
Check plus_claim.
(* ===> plus_claim : Prop *)
Theorem plus_claim_is_true :
plus_claim.
Proof. reflexivity. Qed.
We can also write parameterized propositions — that is, functions that take arguments of some type and return a proposition.
Definition is_three (n : nat) : Prop :=
n = 3.
Check is_three.
(* ===> nat -> Prop *)
n = 3.
Check is_three.
(* ===> nat -> Prop *)
In Coq, functions that return propositions are said to define properties of their arguments.
Definition injective {A B} (f : A → B) :=
∀x y : A, f x = f y → x = y.
Lemma succ_inj : injective S.
Proof.
intros n m H. injection H as H1. apply H1.
Qed.
∀x y : A, f x = f y → x = y.
Lemma succ_inj : injective S.
Proof.
intros n m H. injection H as H1. apply H1.
Qed.
The equality operator = is also a function that returns a Prop.
Check @eq.
(* ===> forall A : Type, A -> A -> Prop *)
(* ===> forall A : Type, A -> A -> Prop *)
What is the type of the following expression?
(2) nat→Prop
(3) ∀ n:nat, Prop
(4) nat→nat
(5) Not typeable
pred (S O) = O
(1) Prop
Check (pred (S O) = O).
(* ===> Prop *)
(* ===> Prop *)
What is the type of the following expression?
(2) nat→Prop
(3) ∀ n:nat, Prop
(4) nat→nat
(5) Not typeable
∀n:nat, pred (S n) = n
(1) Prop
Check (∀n:nat, pred (S n) = n).
(* ===> Prop *)
(* ===> Prop *)
What is the type of the following expression?
(2) nat→Prop
(3) nat→nat
(4) Not typeable
∀n:nat, S (pred n) = n
(1) Prop
Check (∀n:nat, S (pred n) = n).
(* ===> Prop *)
(* ===> Prop *)
What is the type of the following expression?
(2) nat→Prop
(3) nat→nat
(4) Not typeable
∀n:nat, S (pred n)
(1) Prop
(* Check (forall n:nat, pred (S n)). *)
(* ===> Error: In environment
n : nat
The term "pred (S n)" has type "nat" which should be Set, Prop or Type. *)
(* ===> Error: In environment
n : nat
The term "pred (S n)" has type "nat" which should be Set, Prop or Type. *)
What is the type of the following expression?
(2) nat→Prop
(3) nat→nat
(4) Not typeable
fun n:nat ⇒ S (pred n)
(1) Prop
Check (fun n:nat ⇒ pred (S n)).
(* ===> nat->nat *)
(* ===> nat->nat *)
What is the type of the following expression?
(2) nat→Prop
(3) nat→nat
(4) Not typeable
fun n:nat ⇒ S (pred n) = n
(1) Prop
Check (fun n:nat ⇒ pred (S n) = n).
(* ===> nat->Prop *)
(* ===> nat->Prop *)
Which of the following is not a proposition?
(1) 3 + 2 = 4
(2) 3 + 2 = 5
(3) 3 + 2 =? 5
(4) (3+2) =? 4 = false
(5) ∀ n, (3+2) =? n = true → n = 5
(6) All of these are propositions
Fail Definition bad : Prop := 3 + 2 =? 5.
(* The command has indeed failed with message: *)
(* The term "3 + 2 =? 5" has type "bool" while it is expected to have type "Prop". *)
(* The command has indeed failed with message: *)
(* The term "3 + 2 =? 5" has type "bool" while it is expected to have type "Prop". *)
Logical Connectives
Conjunction
Example and_example : 3 + 4 = 7 ∧ 2 * 2 = 4.
To prove a conjunction, use the split tactic. It will generate
two subgoals, one for each part of the statement:
Proof.
split.
- (* 3 + 4 = 7 *) reflexivity.
- (* 2 + 2 = 4 *) reflexivity.
Qed.
split.
- (* 3 + 4 = 7 *) reflexivity.
- (* 2 + 2 = 4 *) reflexivity.
Qed.
For any propositions A and B, if we assume that A is true and we assume that B is true, we can conclude that A ∧ B is also true.
Lemma and_intro : ∀A B : Prop, A → B → A ∧ B.
Proof.
intros A B HA HB. split.
- apply HA.
- apply HB.
Qed.
Proof.
intros A B HA HB. split.
- apply HA.
- apply HB.
Qed.
Since applying a theorem with hypotheses to some goal has the effect of generating as many subgoals as there are hypotheses for that theorem, we can apply and_intro to achieve the same effect as split.
Example and_example' : 3 + 4 = 7 ∧ 2 * 2 = 4.
Proof.
apply and_intro.
- (* 3 + 4 = 7 *) reflexivity.
- (* 2 + 2 = 4 *) reflexivity.
Qed.
Example and_exercise :
∀n m : nat, n + m = 0 → n = 0 ∧ m = 0.
Proof.
(* WORK IN CLASS *) Admitted.
Proof.
apply and_intro.
- (* 3 + 4 = 7 *) reflexivity.
- (* 2 + 2 = 4 *) reflexivity.
Qed.
Example and_exercise :
∀n m : nat, n + m = 0 → n = 0 ∧ m = 0.
Proof.
(* WORK IN CLASS *) Admitted.
So much for proving conjunctive statements. To go in the other direction — i.e., to use a conjunctive hypothesis to help prove something else — we employ the destruct tactic.
Lemma and_example2 :
∀n m : nat, n = 0 ∧ m = 0 → n + m = 0.
Proof.
(* WORK IN CLASS *) Admitted.
∀n m : nat, n = 0 ∧ m = 0 → n + m = 0.
Proof.
(* WORK IN CLASS *) Admitted.
As usual, we can also destruct H right when we introduce it, instead of introducing and then destructing it:
Lemma and_example2' :
∀n m : nat, n = 0 ∧ m = 0 → n + m = 0.
Proof.
intros n m [Hn Hm].
rewrite Hn. rewrite Hm.
reflexivity.
Qed.
∀n m : nat, n = 0 ∧ m = 0 → n + m = 0.
Proof.
intros n m [Hn Hm].
rewrite Hn. rewrite Hm.
reflexivity.
Qed.
You may wonder why we bothered packing the two hypotheses n = 0 and m = 0 into a single conjunction, since we could have also stated the theorem with two separate premises:
Lemma and_example2'' :
∀n m : nat, n = 0 → m = 0 → n + m = 0.
Proof.
intros n m Hn Hm.
rewrite Hn. rewrite Hm.
reflexivity.
Qed.
∀n m : nat, n = 0 → m = 0 → n + m = 0.
Proof.
intros n m Hn Hm.
rewrite Hn. rewrite Hm.
reflexivity.
Qed.
For the present example, both ways work. But in other
situations we may wind up with a conjunctive hypothesis in the
middle of a proof...
Lemma and_example3 :
∀n m : nat, n + m = 0 → n * m = 0.
Proof.
(* WORK IN CLASS *) Admitted.
∀n m : nat, n + m = 0 → n * m = 0.
Proof.
(* WORK IN CLASS *) Admitted.
By the way, the infix notation ∧ is actually just syntactic sugar for and A B. That is, and is a Coq operator that takes two propositions as arguments and yields a proposition.
Check and.
(* ===> and : Prop -> Prop -> Prop *)
(* ===> and : Prop -> Prop -> Prop *)
Disjunction
To use a disjunctive hypothesis in a proof, we proceed by case analysis, which, as for nat or other data types, can be done explicitly with destruct or implicitly with an intros pattern:
Lemma eq_mult_0 :
∀n m : nat, n = 0 ∨ m = 0 → n * m = 0.
Proof.
(* This pattern implicitly does case analysis on
n = 0 ∨ m = 0 *)
intros n m [Hn | Hm].
- (* Here, n = 0 *)
rewrite Hn. reflexivity.
- (* Here, m = 0 *)
rewrite Hm. rewrite <- mult_n_O.
reflexivity.
Qed.
∀n m : nat, n = 0 ∨ m = 0 → n * m = 0.
Proof.
(* This pattern implicitly does case analysis on
n = 0 ∨ m = 0 *)
intros n m [Hn | Hm].
- (* Here, n = 0 *)
rewrite Hn. reflexivity.
- (* Here, m = 0 *)
rewrite Hm. rewrite <- mult_n_O.
reflexivity.
Qed.
We can see in this example that, when we perform case
analysis on a disjunction A ∨ B, we must separately satisfy two
proof obligations, each showing that the conclusion holds under a
different assumption — A in the first subgoal and B in the
second. Note that the case analysis pattern [Hn | Hm] allows
us to name the hypotheses that are generated in the subgoals.
Conversely, to show that a disjunction holds, we need to show that one of its sides does. This is done via two tactics, left and right. As their names imply, the first one requires proving the left side of the disjunction, while the second requires proving its right side. Here is a trivial use...
Lemma or_intro_l : ∀A B : Prop, A → A ∨ B.
Proof.
intros A B HA.
left.
apply HA.
Qed.
Proof.
intros A B HA.
left.
apply HA.
Qed.
Lemma zero_or_succ :
∀n : nat, n = 0 ∨ n = S (pred n).
Proof.
(* WORK IN CLASS *) Admitted.
∀n : nat, n = 0 ∨ n = S (pred n).
Proof.
(* WORK IN CLASS *) Admitted.
Falsehood and Negation
So far, we have mostly been concerned with proving that certain things are true — addition is commutative, appending lists is associative, etc. Of course, we may also be interested in negative results, showing that some given proposition is not true. In Coq, such statements are expressed with the negation operator ¬.To see how negation works, recall the principle of explosion from the Tactics chapter; it asserts that, if we assume a contradiction, then any other proposition can be derived.
Module MyNot.
Definition not (P:Prop) := P → False.
Notation "¬x" := (not x) : type_scope.
Check not.
(* ===> Prop -> Prop *)
End MyNot.
Definition not (P:Prop) := P → False.
Notation "¬x" := (not x) : type_scope.
Check not.
(* ===> Prop -> Prop *)
End MyNot.
Since False is a contradictory proposition, the principle of explosion also applies to it. If we get False into the proof context, we can use destruct on it to complete any goal:
Theorem ex_falso_quodlibet : ∀(P:Prop),
False → P.
Proof.
(* WORK IN CLASS *) Admitted.
False → P.
Proof.
(* WORK IN CLASS *) Admitted.
Inequality is a frequent enough example of negated statement that there is a special notation for it, x ≠ y:
Notation "x ≠ y" := (~(x = y)).
Theorem zero_not_one : 0 ≠ 1.
Proof.
unfold not.
intros contra.
discriminate contra.
Qed.
Proof.
unfold not.
intros contra.
discriminate contra.
Qed.
It takes a little practice to get used to working with negation in Coq. Even though you can see perfectly well why a statement involving negation is true, it can be a little tricky at first to get things into the right configuration so that Coq can understand it! Here are proofs of a few familiar facts to get you warmed up.
Theorem not_False :
¬False.
Proof.
unfold not. intros H. destruct H. Qed.
¬False.
Proof.
unfold not. intros H. destruct 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.
(P ∧ ¬P) → Q.
Proof.
(* WORK IN CLASS *) Admitted.
Theorem double_neg : ∀P : Prop,
P → ¬¬P.
Proof.
(* WORK IN CLASS *) Admitted.
Similarly, since inequality involves a negation, it requires a little practice to be able to work with it fluently. Here is one useful trick. If you are trying to prove a goal that is nonsensical (e.g., the goal state is false = true), apply ex_falso_quodlibet to change the goal to False. This makes it easier to use assumptions of the form ¬P that may be available in the context — in particular, assumptions of the form x≠y.
Theorem not_true_is_false : ∀b : bool,
b ≠ true → b = false.
Proof.
intros b H.
destruct b.
- (* b = true *)
unfold not in H.
apply ex_falso_quodlibet.
apply H. reflexivity.
- (* b = false *)
reflexivity.
Qed.
b ≠ true → b = false.
Proof.
intros b H.
destruct b.
- (* b = true *)
unfold not in H.
apply ex_falso_quodlibet.
apply H. reflexivity.
- (* b = false *)
reflexivity.
Qed.
Since reasoning with ex_falso_quodlibet is quite common, Coq provides a built-in tactic, exfalso, for applying it.
Theorem not_true_is_false' : ∀b : bool,
b ≠ true → b = false.
Proof.
intros [] H. (* note implicit destruct b here *)
- (* b = true *)
unfold not in H.
exfalso. (* <=== *)
apply H. reflexivity.
- (* b = false *) reflexivity.
Qed.
b ≠ true → b = false.
Proof.
intros [] H. (* note implicit destruct b here *)
- (* b = true *)
unfold not in H.
exfalso. (* <=== *)
apply H. reflexivity.
- (* b = false *) reflexivity.
Qed.
To prove the following proposition, which tactics will we need
besides intros and apply?
(2) destruct and unfold
(3) only destruct
(4) left and/or right
(5) only unfold
(6) none of the above
∀X, ∀a b : X, (a=b) ∧ (a≠b) → False.
(1) destruct, unfold, left and right
Lemma quiz1: ∀X, ∀a b : X, (a=b) ∧ (a≠b) → False.
Proof.
intros X a b [H0 H1]. apply H1 in H0. apply H0.
Qed.
Proof.
intros X a b [H0 H1]. apply H1 in H0. apply H0.
Qed.
To prove the following proposition, which tactics will we
need besides intros and apply?
(2) destruct and unfold
(3) only destruct
(4) left and/or right
(5) only unfold
(6) none of the above
∀P Q : Prop, P ∨ Q → ~~(P ∨ Q).
(1) destruct, unfold, left and right
Lemma quiz2 :
∀P Q : Prop, P ∨ Q → ~~(P ∨ Q).
Proof.
intros P Q H H1. apply H1 in H. apply H.
Qed.
∀P Q : Prop, P ∨ Q → ~~(P ∨ Q).
Proof.
intros P Q H H1. apply H1 in H. apply H.
Qed.
To prove the following proposition, which tactics will we
need besides intros and apply?
(2) destruct and unfold
(3) only destruct
(4) left and/or right
(5) only unfold
(6) none of the above
∀A B: Prop, A → (A ∨ ¬¬B).
(1) destruct, unfold, left and right
Lemma quiz3 :
∀A B: Prop, A → (A ∨ ¬¬B).
Proof.
intros P Q H. left. apply H.
Qed.
∀A B: Prop, A → (A ∨ ¬¬B).
Proof.
intros P Q H. left. apply H.
Qed.
To prove the following proposition, which tactics will we need
besides intros and apply?
(2) destruct and unfold
(3) only destruct
(4) left and/or right
(5) only unfold
(6) none of the above
∀P Q: Prop, P ∨ Q → ¬¬P ∨ ¬¬Q.
(1) destruct, unfold, left and right
Lemma quiz4 :
∀P Q: Prop, P ∨ Q → ¬¬P ∨ ¬¬Q.
Proof.
intros P Q [H0 | H0].
- (* left *)
left. intros H1. apply H1 in H0. apply H0.
- (* right *)
right. intros H1. apply H1 in H0. apply H0.
Qed.
∀P Q: Prop, P ∨ Q → ¬¬P ∨ ¬¬Q.
Proof.
intros P Q [H0 | H0].
- (* left *)
left. intros H1. apply H1 in H0. apply H0.
- (* right *)
right. intros H1. apply H1 in H0. apply H0.
Qed.
To prove the following proposition, which tactics will we need
besides intros and apply?
(2) discriminate and unfold
(3) only discriminate
(4) left and/or right
(5) only unfold
(6) none of the above
∀A : Prop, 1=0 → (A ∨ ¬A).
(1) discriminate, unfold, left and right
Lemma quiz5 :
∀A : Prop, 1=0 → (A ∨ ¬A).
Proof.
intros P H. discriminate H.
Qed.
∀A : Prop, 1=0 → (A ∨ ¬A).
Proof.
intros P H. discriminate H.
Qed.
Truth
Lemma True_is_true : True.
Proof. apply I. Qed.
Proof. apply I. Qed.
Unlike False, which is used extensively, True is used quite
rarely, since it is trivial (and therefore uninteresting) to prove
as a goal, and it carries no useful information as a hypothesis.
The handy "if and only if" connective, which asserts that two
propositions have the same truth value, is just the conjunction of
two implications.
Logical Equivalence
Module MyIff.
Definition iff (P Q : Prop) := (P → Q) ∧ (Q → P).
Notation "P ↔ Q" := (iff P Q)
(at level 95, no associativity)
: type_scope.
End MyIff.
Definition iff (P Q : Prop) := (P → Q) ∧ (Q → P).
Notation "P ↔ Q" := (iff P Q)
(at level 95, no associativity)
: type_scope.
End MyIff.
Theorem iff_sym : ∀P Q : Prop,
(P ↔ Q) → (Q ↔ P).
Proof.
(* WORK IN CLASS *) Admitted.
Lemma not_true_iff_false : ∀b,
b ≠ true ↔ b = false.
Proof.
(* WORK IN CLASS *) Admitted.
(P ↔ Q) → (Q ↔ P).
Proof.
(* WORK IN CLASS *) Admitted.
Lemma not_true_iff_false : ∀b,
b ≠ true ↔ b = false.
Proof.
(* WORK IN CLASS *) Admitted.
Setoids and Logical Equivalence
From Coq Require Import Setoids.Setoid.
A "setoid" is a set equipped with an equivalence relation,
such as = or ↔.
Example: Using rewrite with ↔.
Lemma mult_0 : ∀n m, n * m = 0 ↔ n = 0 ∨ m = 0.
Lemma or_assoc :
∀P Q R : Prop, P ∨ (Q ∨ R) ↔ (P ∨ Q) ∨ R.
Lemma mult_0_3 :
∀n m p, n * m * p = 0 ↔ n = 0 ∨ m = 0 ∨ p = 0.
Proof.
intros n m p.
rewrite mult_0. rewrite mult_0. rewrite or_assoc.
reflexivity.
Qed.
Proof.
split.
- apply mult_eq_0.
- apply eq_mult_0.
Qed.
split.
- apply mult_eq_0.
- apply eq_mult_0.
Qed.
Lemma or_assoc :
∀P Q R : Prop, P ∨ (Q ∨ R) ↔ (P ∨ Q) ∨ R.
Proof.
intros P Q R. split.
- intros [H | [H | H]].
+ left. left. apply H.
+ left. right. apply H.
+ right. apply H.
- intros [[H | H] | H].
+ left. apply H.
+ right. left. apply H.
+ right. right. apply H.
Qed.
intros P Q R. split.
- intros [H | [H | H]].
+ left. left. apply H.
+ left. right. apply H.
+ right. apply H.
- intros [[H | H] | H].
+ left. apply H.
+ right. left. apply H.
+ right. right. apply H.
Qed.
Lemma mult_0_3 :
∀n m p, n * m * p = 0 ↔ n = 0 ∨ m = 0 ∨ p = 0.
Proof.
intros n m p.
rewrite mult_0. rewrite mult_0. rewrite or_assoc.
reflexivity.
Qed.
Example: using apply with ↔. The apply tactic can also be used with ↔. When given an equivalence as its argument, apply tries to guess which direction of the equivalence will be useful..
Lemma apply_iff_example :
∀n m : nat, n * m = 0 → n = 0 ∨ m = 0.
Proof.
intros n m H. apply mult_0. apply H.
Qed.
∀n m : nat, n * m = 0 → n = 0 ∨ m = 0.
Proof.
intros n m H. apply mult_0. apply H.
Qed.
Existential Quantification
Lemma four_is_even : ∃n : nat, 4 = n + n.
Proof.
∃2. reflexivity.
Qed.
Proof.
∃2. reflexivity.
Qed.
Conversely, if we have an existential hypothesis ∃ x, P in the context, we can destruct it to obtain a witness x and a hypothesis stating that P holds of x.
Theorem exists_example_2 : ∀n,
(∃m, n = 4 + m) →
(∃o, n = 2 + o).
Proof.
(* WORK IN CLASS *) Admitted.
(∃m, n = 4 + m) →
(∃o, n = 2 + o).
Proof.
(* WORK IN CLASS *) Admitted.
Programming with Propositions
- If l is the empty list, then x cannot occur in it, so the property "x appears in l" is simply false.
- Otherwise, l has the form x' :: l'. In this case, x occurs in l if either it is equal to x' or it occurs in l'.
Fixpoint In {A : Type} (x : A) (l : list A) : Prop :=
match l with
| [] ⇒ False
| x' :: l' ⇒ x' = x ∨ In x l'
end.
match l with
| [] ⇒ False
| x' :: l' ⇒ x' = x ∨ In x l'
end.
Example In_example_1 : In 4 [1; 2; 3; 4; 5].
Proof.
(* WORK IN CLASS *) Admitted.
Example In_example_2 :
∀n, In n [2; 4] →
∃n', n = 2 * n'.
Proof.
(* WORK IN CLASS *) Admitted.
Proof.
(* WORK IN CLASS *) Admitted.
Example In_example_2 :
∀n, In n [2; 4] →
∃n', n = 2 * n'.
Proof.
(* WORK IN CLASS *) Admitted.
Lemma In_map :
∀(A B : Type) (f : A → B) (l : list A) (x : A),
In x l →
In (f x) (map f l).
Proof.
intros A B f l x.
induction l as [|x' l' IHl'].
- (* l = nil, contradiction *)
simpl. intros [].
- (* l = x' :: l' *)
simpl. intros [H | H].
+ rewrite H. left. reflexivity.
+ right. apply IHl'. apply H.
Qed.
∀(A B : Type) (f : A → B) (l : list A) (x : A),
In x l →
In (f x) (map f l).
Proof.
intros A B f l x.
induction l as [|x' l' IHl'].
- (* l = nil, contradiction *)
simpl. intros [].
- (* l = x' :: l' *)
simpl. intros [H | H].
+ rewrite H. left. reflexivity.
+ right. apply IHl'. apply H.
Qed.
Applying Theorems to Arguments
Check plus_comm.
(* ===> forall n m : nat, n + m = m + n *)
(* ===> forall n m : nat, n + m = m + n *)
Coq prints the statement of the plus_comm theorem in the same
way that it prints the type of any term that we ask it to
Check. Why?
The type of a computational object tells us what we can do
with that object.
Similarly, the statement of a theorem tells us what we can use
that theorem for.
The reason is that the identifier plus_comm actually refers to a proof object — a data structure that represents a logical derivation establishing of the truth of the statement ∀ n m : nat, n + m = m + n. The type of this object is the statement of the theorem that it is a proof of.
- e.g., if we have a term of type nat → nat → nat, we can give it two nats as arguments and get a nat back.
- if we have an object of type n = m → n + n = m + m and we provide it an "argument" of type n = m, we can derive n + n = m + m.
Coq actually allows us to apply a theorem as if it were a function. This is often very handy in proof scripts — e.g., suppose we want too prove the following:
Lemma plus_comm3 :
∀x y z, x + (y + z) = (z + y) + x.
∀x y z, x + (y + z) = (z + y) + x.
It appears at first sight that we ought to be able to prove this
by rewriting with plus_comm twice to make the two sides match.
The problem, however, is that the second rewrite will undo the
effect of the first.
Proof.
(* WORK IN CLASS *) Admitted.
(* WORK IN CLASS *) Admitted.
Let us show another example of using a theorem or lemma like a function. The following theorem says: any list l containing some element must be nonempty.
Lemma in_not_nil :
∀A (x : A) (l : list A), In x l → l ≠ [].
Proof.
intros A x l H. unfold not. intro Hl. destruct l.
- simpl in H. destruct H.
- discriminate Hl.
Qed.
∀A (x : A) (l : list A), In x l → l ≠ [].
Proof.
intros A x l H. unfold not. intro Hl. destruct l.
- simpl in H. destruct H.
- discriminate Hl.
Qed.
We can use this lemma to prove the special case where x
is 42. Naively, the tactic apply in_not_nil will fail because
it cannot infer the value of x. There are several ways to work
around that...
Lemma in_not_nil_42 :
∀l : list nat, In 42 l → l ≠ [].
Proof.
(* WORK IN CLASS *) Admitted.
∀l : list nat, In 42 l → l ≠ [].
Proof.
(* WORK IN CLASS *) Admitted.
Coq vs. Set Theory
Functional Extensionality
In common mathematical practice, two functions f and g are
considered equal if they produce the same outputs:
Functional extensionality is not part of Coq's built-in logic.
This means that some "reasonable" propositions are not provable.
(∀x, f x = g x) → f = g
This is known as the principle of functional extensionality.
Example function_equality_ex2 :
(fun x ⇒ plus x 1) = (fun x ⇒ plus 1 x).
Proof.
(* Stuck *)
Abort.
(fun x ⇒ plus x 1) = (fun x ⇒ plus 1 x).
Proof.
(* Stuck *)
Abort.
However, we can add functional extensionality to Coq's core using
the Axiom command.
Axiom functional_extensionality : ∀{X Y: Type}
{f g : X → Y},
(∀(x:X), f x = g x) → f = g.
{f g : X → Y},
(∀(x:X), f x = g x) → f = g.
Using Axiom has the same effect as stating a theorem and
skipping its proof using Admitted, but it alerts the reader that
this isn't just something we're going to come back and fill in
later!
We can now invoke functional extensionality in proofs:
Example function_equality_ex2 :
(fun x ⇒ plus x 1) = (fun x ⇒ plus 1 x).
Proof.
apply functional_extensionality. intros x.
apply plus_comm.
Qed.
(fun x ⇒ plus x 1) = (fun x ⇒ plus 1 x).
Proof.
apply functional_extensionality. intros x.
apply plus_comm.
Qed.
Naturally, we must be careful when adding new axioms into Coq's logic, as they may render it inconsistent — that is, they may make it possible to prove every proposition, including False, 2+2=5, etc.!
To check whether a particular proof relies on any additional axioms, use the Print Assumptions command.
Print Assumptions function_equality_ex2.
(* ===>
Axioms:
functional_extensionality :
forall (X Y : Type) (f g : X -> Y),
(forall x : X, f x = g x) -> f = g *)
(* ===>
Axioms:
functional_extensionality :
forall (X Y : Type) (f g : X -> Y),
(forall x : X, f x = g x) -> f = g *)
Is this provable by reflexivity, i.e., without
functional_extensionality?
(fun xs ⇒ 1 :: xs) = (fun xs ⇒ [1] ++ xs)
(1) Yes
(2) No
Example cons_1_eq_ex : (fun xs ⇒ 1 :: xs) = (fun xs ⇒ [1] ++ xs).
Proof. reflexivity. Qed.
Proof. reflexivity. Qed.
Is this provable by reflexivity, i.e., without
functional_extensionality?
(fun x y ⇒ x + S y) = (fun x y ⇒ S (x + y))
(1) Yes
(2) No
Example plus_1_eq_ex : (fun x y ⇒ S (x + y)) = (fun x y ⇒ x + S y).
Proof.
assert_fails reflexivity.
(* Unable to unify "x + S y" with "S (x + y)". *)
Abort.
Example plus_1_eq_ex : (fun x y ⇒ S (x + y)) = (fun x y ⇒ x + S y).
Proof.
apply functional_extensionality. intro x.
apply functional_extensionality. intro y.
(* Search (_ + S _) eq. *)
apply plus_n_Sm.
Qed.
Proof.
assert_fails reflexivity.
(* Unable to unify "x + S y" with "S (x + y)". *)
Abort.
Example plus_1_eq_ex : (fun x y ⇒ S (x + y)) = (fun x y ⇒ x + S y).
Proof.
apply functional_extensionality. intro x.
apply functional_extensionality. intro y.
(* Search (_ + S _) eq. *)
apply plus_n_Sm.
Qed.
Propositions and Booleans
Example even_42_bool : evenb 42 = true.
Proof. reflexivity. Qed.
... or that there exists some k such that n = double k.
Example even_42_prop : ∃k, 42 = double k.
Proof. ∃21. reflexivity. Qed.
Of course, it would be pretty strange if these two
characterizations of evenness did not describe the same set of
natural numbers! Fortunately, we can prove that they do...
Theorem even_bool_prop : ∀n,
evenb n = true ↔ ∃k, n = double k.
evenb n = true ↔ ∃k, n = double k.
Proof.
intros n. split.
- intros H. destruct (evenb_double_conv n) as [k Hk].
rewrite Hk. rewrite H. ∃k. reflexivity.
- intros [k Hk]. rewrite Hk. apply evenb_double.
Qed.
intros n. split.
- intros H. destruct (evenb_double_conv n) as [k Hk].
rewrite Hk. rewrite H. ∃k. reflexivity.
- intros [k Hk]. rewrite Hk. apply evenb_double.
Qed.
In view of this theorem, we say that the boolean computation
evenb n is reflected in the truth of the proposition ∃ k,
n = double k.
Similarly, to state that two numbers n and m are equal, we can say either
- (1) that n =? m returns true, or
- (2) that n = m.
Theorem eqb_eq : ∀n1 n2 : nat,
n1 =? n2 = true ↔ n1 = n2.
n1 =? n2 = true ↔ n1 = n2.
Proof.
intros n1 n2. split.
- apply eqb_true.
- intros H. rewrite H. rewrite <- eqb_refl. reflexivity.
Qed.
intros n1 n2. split.
- apply eqb_true.
- intros H. rewrite H. rewrite <- eqb_refl. reflexivity.
Qed.
An important side benefit of stating facts using booleans is enabling some proof automation through computation with Coq terms, a technique known as proof by reflection. Consider the following statement:
Example even_1000 : ∃k, 1000 = double k.
The most direct proof of this fact is to give the value of k
explicitly.
Proof. ∃500. reflexivity. Qed.
On the other hand, the proof of the corresponding boolean
statement is even simpler:
Example even_1000' : evenb 1000 = true.
Proof. reflexivity. Qed.
Proof. reflexivity. Qed.
What is interesting is that, since the two notions are equivalent,
we can use the boolean formulation to prove the other one without
mentioning the value 500 explicitly:
Example even_1000'' : ∃k, 1000 = double k.
Proof. apply even_bool_prop. reflexivity. Qed.
Proof. apply even_bool_prop. reflexivity. Qed.
Example not_even_1001 : evenb 1001 = false.
Proof.
(* WORK IN CLASS *) Admitted.
Proof.
(* WORK IN CLASS *) Admitted.
In contrast, propositional negation may be more difficult
to work with.
Example not_even_1001' : ~(∃k, 1001 = double k).
Proof.
(* WORK IN CLASS *) Admitted.
Proof.
(* WORK IN CLASS *) Admitted.
Lemma plus_eqb_example : ∀n m p : nat,
n =? m = true → n + p =? m + p = true.
Proof.
(* WORK IN CLASS *) Admitted.
n =? m = true → n + p =? m + p = true.
Proof.
(* WORK IN CLASS *) Admitted.
We won't cover reflection in much detail, but it serves as a good
example showing the complementary strengths of booleans and
general propositions.
The following reasoning principle is not derivable in
Coq (though, again, it can consistently be added):
Classical vs. Constructive Logic
Definition excluded_middle := ∀P : Prop,
P ∨ ¬P.
P ∨ ¬P.
Logics like Coq's, which do not assume the excluded middle, are
referred to as constructive logics.
More conventional logical systems such as ZFC, in which the
excluded middle does hold for arbitrary propositions, are referred
to as classical.