LogicLogic in Coq
Set Warnings "-notation-overridden,-parsing".
From LF Require Export Tactics.
From LF Require Export Tactics.
Back to Logic
Lemma and_symm : ∀ P Q, P ∧ Q → Q ∧ P.
intros P Q H.
destruct H as [Hp Hq].
split.
- apply Hq.
- apply Hp.
Qed.
Lemma or_symm : ∀ P Q, P ∨ Q → Q ∨ P.
intros P Q H.
destruct H as [Hp | Hq].
- right. apply Hp.
- left. apply Hq.
Qed.
intros P Q H.
destruct H as [Hp Hq].
split.
- apply Hq.
- apply Hp.
Qed.
Lemma or_symm : ∀ P Q, P ∨ Q → Q ∨ P.
intros P Q H.
destruct H as [Hp | Hq].
- right. apply Hp.
- left. apply Hq.
Qed.
Expressions of the form P → False or m = n → False are common enough to have specific
notations for them:
Locate "¬ _".
Print not.
Locate "_ ≠ _".
Print not.
Locate "_ ≠ _".
We also have a tactic for dealing with ¬A: contradict.
Lemma zero_not_even : ¬ Even 0 → 1 = 2.
Proof.
intros NE.
contradict NE.
apply E_O.
Qed.
Proof.
intros NE.
contradict NE.
apply E_O.
Qed.
This is equivalent to simply applying ex_falso_quodlibet (from ch. 1)
and then applying ¬A.
Lemma seven_not_four : 7 ≠ 4.
Proof.
intros F.
discriminate F.
Qed.
Proof.
intros F.
discriminate F.
Qed.
Theorem contrapositive : ∀ (P Q : Prop),
(P → Q) → (¬Q → ¬P).
Proof.
(* FILL IN HERE *) Admitted.
☐
(P → Q) → (¬Q → ¬P).
Proof.
(* FILL IN HERE *) Admitted.
Theorem not_both_true_and_false : ∀ P : Prop,
¬ (P ∧ ¬P).
Proof.
(* FILL IN HERE *) Admitted.
☐
¬ (P ∧ ¬P).
Proof.
(* FILL IN HERE *) Admitted.
Exercise: 1 star, optional (informal_not_PNP)
Write an informal proof (in English) of the proposition ∀P : Prop, ~(P ∧ ¬P).
(* FILL IN HERE *)
(* Do not modify the following line: *)
Definition manual_grade_for_informal_not_PNP : option (nat*string) := None.
☐
(* Do not modify the following line: *)
Definition manual_grade_for_informal_not_PNP : option (nat*string) := None.
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.
Theorem iff_sym : ∀ P Q : Prop,
(P ↔ Q) → (Q ↔ P).
Proof.
(* WORKED IN CLASS *)
intros P Q [HAB HBA].
split.
- (* -> *) apply HBA.
- (* <- *) apply HAB. Qed.
Lemma not_true_iff_false : ∀ b,
b ≠ true ↔ b = false.
Proof.
(* WORKED IN CLASS *)
intros b. split.
- (* -> *)
intros H. destruct b.
+ contradict H. reflexivity.
+ reflexivity.
- (* <- *)
intros H1 H2.
rewrite H1 in H2.
discriminate H2.
Qed.
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.
(* WORKED IN CLASS *)
intros P Q [HAB HBA].
split.
- (* -> *) apply HBA.
- (* <- *) apply HAB. Qed.
Lemma not_true_iff_false : ∀ b,
b ≠ true ↔ b = false.
Proof.
(* WORKED IN CLASS *)
intros b. split.
- (* -> *)
intros H. destruct b.
+ contradict H. reflexivity.
+ reflexivity.
- (* <- *)
intros H1 H2.
rewrite H1 in H2.
discriminate H2.
Qed.
Theorem or_distributes_over_and : ∀ P Q R : Prop,
P ∨ (Q ∧ R) ↔ (P ∨ Q) ∧ (P ∨ R).
Proof.
(* FILL IN HERE *) Admitted.
☐
P ∨ (Q ∧ R) ↔ (P ∨ Q) ∧ (P ∨ R).
Proof.
(* FILL IN HERE *) Admitted.
Require Import Coq.Setoids.Setoid.
Here is a simple example demonstrating how these tactics work with
iff. First, let's recall one lemma and admit another:
About or_assoc.
Lemma mult_0 : ∀ n m, n * m = 0 ↔ n = 0 ∨ m = 0.
We can now use these facts with rewrite and reflexivity to
give smooth proofs of statements involving equivalences. Here is
a ternary version of the previous mult_0 result:
Another important logical connective is existential
quantification. To say that there is some x of type T such
that some property P holds of x, we write ∃x : T,
P. As with ∀, the type annotation : T can be omitted if
Coq is able to infer from the context what the type of x should
be.
To prove a statement of the form ∃x, P, we must show that
P holds for some specific choice of value for x, known as the
witness of the existential. This is done in two steps: First,
we explicitly tell Coq which witness t we have in mind by
invoking the tactic ∃t. Then we prove that P holds after
all occurrences of x are replaced by t.
The logical connectives that we have seen provide a rich
vocabulary for defining complex propositions from simpler ones.
To illustrate, let's look at how to express the claim that an
element x occurs in a list l. Notice that this property has a
simple recursive structure:
We can translate this directly into a straightforward recursive
function taking an element and a list and returning a proposition:
Drawing inspiration from In, write a recursive function All
stating that some property P holds of all elements of a list
l. To make sure your definition is correct, prove the All_In
lemma below. (Of course, your definition should not just
restate the left-hand side of All_In.)
Coq's logical core, the Calculus of Inductive Constructions,
differs in some important ways from other formal systems that are
used by mathematicians for writing down precise and rigorous
proofs. For example, in the most popular foundation for
mainstream paper-and-pencil mathematics, Zermelo-Fraenkel Set
Theory (ZFC), a mathematical object can potentially be a member of
many different sets; a term in Coq's logic, on the other hand, is
a member of at most one type. This difference often leads to
slightly different ways of capturing informal mathematical
concepts, but these are, by and large, quite natural and easy to
work with. For example, instead of saying that a natural number
n belongs to the set of even numbers, we would say in Coq that
ev n holds, where ev : nat → Prop is a property describing
even numbers.
However, there are some cases where translating standard
mathematical reasoning into Coq can be either cumbersome or
sometimes even impossible, unless we enrich the core logic with
additional axioms. We conclude this chapter with a brief
discussion of some of the most significant differences between the
two worlds.
In common mathematical practice, two functions f and g are
considered equal if they produce the same outputs:
Informally speaking, an "extensional property" is one that
pertains to an object's observable behavior. Thus, functional
extensionality simply means that a function's identity is
completely determined by what we can observe from it — i.e., in
Coq terms, the results we obtain after applying it.
Functional extensionality is not part of Coq's basic axioms. This
means that some "reasonable" propositions are not provable.
We've seen two different ways of encoding logical facts in Coq:
with booleans (of type bool), and with propositions (of type
Prop).
For instance, to claim that a number n is even, we can say
either
Of course, it would be very strange if these two characterizations
of evenness did not describe the same set of natural numbers!
Fortunately, we can prove that they do...
We first need two helper lemmas.
As we've seen earlier, the following proposition is not provable
inside Coq:
Prove that all five propositions (these four plus
excluded_middle) are equivalent.
Proof.
(* FILL IN HERE *) Admitted.
☐
(* FILL IN HERE *) Admitted.
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.
∀ 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.
The apply tactic can also be used with ↔. When given an
equivalence as its argument, apply tries to guess which side of
the equivalence to use.
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.
(* WORKED IN CLASS *)
intros n [m Hm]. (* note implicit destruct here *)
∃ (2 + m).
apply Hm. Qed.
(∃ m, n = 4 + m) →
(∃ o, n = 2 + o).
Proof.
(* WORKED IN CLASS *)
intros n [m Hm]. (* note implicit destruct here *)
∃ (2 + m).
apply Hm. Qed.
Exercise: 1 star, recommended (dist_not_exists)
Prove that "P holds for all x" implies "there is no x for which P does not hold." (Hint: destruct H as [x E] works on existential assumptions!)
Theorem dist_not_exists : ∀ (X:Type) (P : X → Prop),
(∀ x, P x) → ¬ (∃ x, ¬ P x).
Proof.
(* FILL IN HERE *) Admitted.
☐
(∀ x, P x) → ¬ (∃ x, ¬ P x).
Proof.
(* FILL IN HERE *) Admitted.
Exercise: 2 stars (dist_exists_or)
Prove that existential quantification distributes over disjunction.
Theorem dist_exists_or : ∀ (X:Type) (P Q : X → Prop),
(∃ x, P x ∨ Q x) ↔ (∃ x, P x) ∨ (∃ x, Q x).
Proof.
(* FILL IN HERE *) Admitted.
☐
(∃ x, P x ∨ Q x) ↔ (∃ x, P x) ∨ (∃ x, Q x).
Proof.
(* FILL IN HERE *) Admitted.
Programming with Propositions
- If l is the empty list, then x cannot occur on 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.
When In is applied to a concrete list, it expands into a
concrete sequence of nested disjunctions.
Example In_example_1 : In 4 [1; 2; 3; 4; 5].
Proof.
(* WORKED IN CLASS *)
simpl. right. right. right. left. reflexivity.
Qed.
Example In_example_2 :
∀ n, In n [2; 4] →
∃ n', n = 2 * n'.
Proof.
(* WORKED IN CLASS *)
simpl.
intros n [H | [H | []]].
- ∃ 1. rewrite <- H. reflexivity.
- ∃ 2. rewrite <- H. reflexivity.
Qed.
Proof.
(* WORKED IN CLASS *)
simpl. right. right. right. left. reflexivity.
Qed.
Example In_example_2 :
∀ n, In n [2; 4] →
∃ n', n = 2 * n'.
Proof.
(* WORKED IN CLASS *)
simpl.
intros n [H | [H | []]].
- ∃ 1. rewrite <- H. reflexivity.
- ∃ 2. rewrite <- H. reflexivity.
Qed.
(Notice the use of the empty pattern to discharge the last case
en passant.)
We can also prove more generic, higher-level lemmas about In.
Note, in the next, how In starts out applied to a variable and
only gets expanded when we do case analysis on this variable:
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.
This way of defining propositions recursively, though convenient
in some cases, also has some drawbacks. In particular, it is
subject to Coq's usual restrictions regarding the definition of
recursive functions, e.g., the requirement that they be "obviously
terminating." In the next chapter, we will see how to define
propositions inductively, a different technique with its own set
of strengths and limitations.
Exercise: 2 stars, recommended (In_map_iff)
Lemma In_map_iff :
∀ (A B : Type) (f : A → B) (l : list A) (y : B),
In y (map f l) ↔
∃ x, f x = y ∧ In x l.
Proof.
(* FILL IN HERE *) Admitted.
☐
∀ (A B : Type) (f : A → B) (l : list A) (y : B),
In y (map f l) ↔
∃ x, f x = y ∧ In x l.
Proof.
(* FILL IN HERE *) Admitted.
Lemma In_app_iff : ∀ A l l' (a:A),
In a (l++l') ↔ In a l ∨ In a l'.
Proof.
(* FILL IN HERE *) Admitted.
☐
In a (l++l') ↔ In a l ∨ In a l'.
Proof.
(* FILL IN HERE *) Admitted.
Exercise: 3 stars, optional (All)
Recall that functions returning propositions can be seen as properties of their arguments. For instance, if P has type nat → Prop, then P n states that property P holds of n.
Fixpoint All {T : Type} (P : T → Prop) (l : list T) : Prop
(* REPLACE THIS LINE WITH ":= _your_definition_ ." *). Admitted.
Lemma All_In :
∀ T (P : T → Prop) (l : list T),
(∀ x, In x l → P x) ↔
All P l.
Proof.
(* FILL IN HERE *) Admitted.
☐
(* REPLACE THIS LINE WITH ":= _your_definition_ ." *). Admitted.
Lemma All_In :
∀ T (P : T → Prop) (l : list T),
(∀ x, In x l → P x) ↔
All P l.
Proof.
(* FILL IN HERE *) Admitted.
Exercise: 3 stars, recommended (combine_odd_even)
Complete the definition of the combine_odd_even function below. It takes as arguments two properties of numbers, Podd and Peven, and it should return a property P such that P n is equivalent to Podd n when n is odd and equivalent to Peven n otherwise.
Definition combine_odd_even (Podd Peven : nat → Prop) : nat → Prop
(* REPLACE THIS LINE WITH ":= _your_definition_ ." *). Admitted.
(* REPLACE THIS LINE WITH ":= _your_definition_ ." *). Admitted.
To test your definition, prove the following facts:
Theorem combine_odd_even_intro :
∀ (Podd Peven : nat → Prop) (n : nat),
(oddb n = true → Podd n) →
(oddb n = false → Peven n) →
combine_odd_even Podd Peven n.
Proof.
(* FILL IN HERE *) Admitted.
Theorem combine_odd_even_elim_odd :
∀ (Podd Peven : nat → Prop) (n : nat),
combine_odd_even Podd Peven n →
oddb n = true →
Podd n.
Proof.
(* FILL IN HERE *) Admitted.
Theorem combine_odd_even_elim_even :
∀ (Podd Peven : nat → Prop) (n : nat),
combine_odd_even Podd Peven n →
oddb n = false →
Peven n.
Proof.
(* FILL IN HERE *) Admitted.
☐
∀ (Podd Peven : nat → Prop) (n : nat),
(oddb n = true → Podd n) →
(oddb n = false → Peven n) →
combine_odd_even Podd Peven n.
Proof.
(* FILL IN HERE *) Admitted.
Theorem combine_odd_even_elim_odd :
∀ (Podd Peven : nat → Prop) (n : nat),
combine_odd_even Podd Peven n →
oddb n = true →
Podd n.
Proof.
(* FILL IN HERE *) Admitted.
Theorem combine_odd_even_elim_even :
∀ (Podd Peven : nat → Prop) (n : nat),
combine_odd_even Podd Peven n →
oddb n = false →
Peven n.
Proof.
(* FILL IN HERE *) Admitted.
Coq vs. Set Theory
Functional Extensionality
(∀ 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 logic
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!
Unfortunately, there is no simple way of telling whether an axiom
is safe to add: hard work is generally required to establish the
consistency of any particular combination of axioms.
Fortunately, it is known that adding functional extensionality, in
particular, is consistent.
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 *)
Exercise: 4 stars, optional (tr_rev_correct)
One problem with the definition of the list-reversing function rev that we have is that it performs a call to app on each step; running app takes time asymptotically linear in the size of the list, which means that rev has quadratic running time. We can improve this with the following definition:
Fixpoint rev_append {X} (l1 l2 : list X) : list X :=
match l1 with
| [] ⇒ l2
| x :: l1' ⇒ rev_append l1' (x :: l2)
end.
Definition tr_rev {X} (l : list X) : list X :=
rev_append l [].
match l1 with
| [] ⇒ l2
| x :: l1' ⇒ rev_append l1' (x :: l2)
end.
Definition tr_rev {X} (l : list X) : list X :=
rev_append l [].
This version is said to be tail-recursive, because the recursive
call to the function is the last operation that needs to be
performed (i.e., we don't have to execute ++ after the recursive
call); a decent compiler will generate very efficient code in this
case. Prove that the two definitions are indeed equivalent.
Lemma tr_rev_correct : ∀ X, @tr_rev X = @rev X.
(* FILL IN HERE *) Admitted.
☐
(* FILL IN HERE *) Admitted.
Propositions and Booleans
- (1) that evenb n returns true, or
- (2) that there exists some k such that n = double k. Indeed, these two notions of evenness are equivalent, as can easily be shown with a couple of auxiliary lemmas.
Theorem evenb_double : ∀ k, evenb (double k) = true.
Proof.
intros k. induction k as [|k' IHk'].
- reflexivity.
- simpl. apply IHk'.
Qed.
Proof.
intros k. induction k as [|k' IHk'].
- reflexivity.
- simpl. apply IHk'.
Qed.
Theorem evenb_double_conv : ∀ n,
∃ k, n = if evenb n then double k
else S (double k).
Proof.
(* Hint: Use the evenb_S lemma from Induction.v. *)
(* FILL IN HERE *) Admitted.
☐
∃ k, n = if evenb n then double k
else S (double k).
Proof.
(* Hint: Use the evenb_S lemma from Induction.v. *)
(* FILL IN HERE *) Admitted.
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. Again, these two notions are equivalent.
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.
However, even when the boolean and propositional formulations of a
claim are equivalent from a purely logical perspective, they need
not be equivalent operationally.
Equality provides an extreme example: knowing that n =? m =
true is generally of little direct help in the middle of a proof
involving n and m; however, if we convert the statement to the
equivalent form n = m, we can rewrite with it.
The case of even numbers is also interesting. Recall that,
when proving the backwards direction of even_bool_prop (i.e.,
evenb_double, going from the propositional to the boolean
claim), we used a simple induction on k. On the other hand, the
converse (the evenb_double_conv exercise) required a clever
generalization, since we can't directly prove (∃k, n =
double k) → evenb n = true.
For these examples, the propositional claims are more useful than
their boolean counterparts, but this is not always the case. For
instance, we cannot test whether a general proposition is true or
not in a function definition; as a consequence, the following code
fragment is rejected:
Fail Definition is_even_prime n :=
if n = 2 then true
else false.
if n = 2 then true
else false.
Coq complains that n = 2 has type Prop, while it expects an
elements of bool (or some other inductive type with two
elements). It would be nice if Coq provided a mechanism for
testing the truth value of arbitrary propositions but
unfortunately computably theory and Godel's incompleteness
theorems say that this is impossible. (Complexity theory would add
that if we limited that mechanism to computable propositions, it
would still take forever.)
Although general non-computable properties cannot be phrased as
boolean computations, it is worth noting that even many
computable properties are easier to express using Prop than bool,
since recursive function definitions are subject to significant
restrictions in Coq. For instance, the next chapter shows how to
define the property that a regular expression matches a given
string using Prop. Doing the same with bool would amount to writing
a regular expression matcher, which would be more complicated,
harder to understand, and harder to reason about.
Conversely, 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.
Although we haven't gained much in terms of proof size in this
case, larger proofs can often be made considerably simpler by the
use of reflection. As an extreme example, the Coq proof of the
famous 4-color theorem uses reflection to reduce the analysis of
hundreds of different cases to a boolean computation. We won't
cover reflection in great detail, but it serves as a good example
showing the complementary strengths of booleans and general
propositions.
Exercise: 2 stars (logical_connectives)
The following lemmas relate the propositional connectives to the corresponding boolean operations.
Lemma andb_true_iff : ∀ b1 b2:bool,
b1 && b2 = true ↔ b1 = true ∧ b2 = true.
Proof.
(* FILL IN HERE *) Admitted.
Lemma orb_true_iff : ∀ b1 b2,
b1 || b2 = true ↔ b1 = true ∨ b2 = true.
Proof.
(* FILL IN HERE *) Admitted.
☐
b1 && b2 = true ↔ b1 = true ∧ b2 = true.
Proof.
(* FILL IN HERE *) Admitted.
Lemma orb_true_iff : ∀ b1 b2,
b1 || b2 = true ↔ b1 = true ∨ b2 = true.
Proof.
(* FILL IN HERE *) Admitted.
Exercise: 1 star (eqb_neq)
The following theorem is an alternate "negative" formulation of eqb_eq that is more convenient in certain situations (we'll see examples in later chapters).
Theorem eqb_neq : ∀ x y : nat,
x =? y = false ↔ x ≠ y.
Proof.
(* FILL IN HERE *) Admitted.
☐
x =? y = false ↔ x ≠ y.
Proof.
(* FILL IN HERE *) Admitted.
Exercise: 3 stars, optional (eqb_list)
Given a boolean operator eqb for testing equality of elements of some type A, we can define a function eqb_list for testing equality of lists with elements in A. Complete the definition of the eqb_list function below. To make sure that your definition is correct, prove the lemma eqb_list_true_iff.
Fixpoint eqb_list {A : Type} (eqb : A → A → bool)
(l1 l2 : list A) : bool
(* REPLACE THIS LINE WITH ":= _your_definition_ ." *). Admitted.
Lemma eqb_list_true_iff :
∀ A (eqb : A → A → bool),
(∀ a1 a2, eqb a1 a2 = true ↔ a1 = a2) →
∀ l1 l2, eqb_list eqb l1 l2 = true ↔ l1 = l2.
Proof.
(* FILL IN HERE *) Admitted.
☐
(l1 l2 : list A) : bool
(* REPLACE THIS LINE WITH ":= _your_definition_ ." *). Admitted.
Lemma eqb_list_true_iff :
∀ A (eqb : A → A → bool),
(∀ a1 a2, eqb a1 a2 = true ↔ a1 = a2) →
∀ l1 l2, eqb_list eqb l1 l2 = true ↔ l1 = l2.
Proof.
(* FILL IN HERE *) Admitted.
Exercise: 2 stars, recommended (All_forallb)
Recall the function forallb, from the exercise forall_exists_challenge in chapter Tactics:
Fixpoint forallb {X : Type} (test : X → bool) (l : list X) : bool :=
match l with
| [] ⇒ true
| x :: l' ⇒ andb (test x) (forallb test l')
end.
match l with
| [] ⇒ true
| x :: l' ⇒ andb (test x) (forallb test l')
end.
Prove the theorem below, which relates forallb to the All
property of the above exercise.
Theorem forallb_true_iff : ∀ X test (l : list X),
forallb test l = true ↔ All (fun x ⇒ test x = true) l.
Proof.
(* FILL IN HERE *) Admitted.
forallb test l = true ↔ All (fun x ⇒ test x = true) l.
Proof.
(* FILL IN HERE *) Admitted.
Are there any important properties of the function forallb which
are not captured by this specification?
(* FILL IN HERE *)
☐
Classical vs. Constructive Logic
Definition excluded_middle := ∀ P : Prop,
P ∨ ¬ P.
P ∨ ¬ P.
However, if we happen to know that P is reflected in some
boolean term b, then knowing whether it holds or not is trivial:
we just have to check the value of b.
Theorem restricted_excluded_middle : ∀ P b,
(P ↔ b = true) → P ∨ ¬ P.
Proof.
intros P [] H.
- left. rewrite H. reflexivity.
- right. rewrite H. intros contra. discriminate contra.
Qed.
(P ↔ b = true) → P ∨ ¬ P.
Proof.
intros P [] H.
- left. rewrite H. reflexivity.
- right. rewrite H. intros contra. discriminate contra.
Qed.
In particular, the excluded middle is valid for equations n = m,
between natural numbers n and m.
Theorem restricted_excluded_middle_eq : ∀ (n m : nat),
n = m ∨ n ≠ m.
Proof.
intros n m.
apply (restricted_excluded_middle (n = m) (n =? m)).
symmetry.
apply eqb_eq.
Qed.
n = m ∨ n ≠ m.
Proof.
intros n m.
apply (restricted_excluded_middle (n = m) (n =? m)).
symmetry.
apply eqb_eq.
Qed.
It may seem strange that the general excluded middle is not
available by default in Coq; after all, any given claim must be
either true or false. Nonetheless, there is an advantage in not
assuming the excluded middle: statements in Coq can make stronger
claims than the analogous statements in standard mathematics.
Notably, if there is a Coq proof of ∃x, P x, it is
possible to explicitly exhibit a value of x for which we can
prove P x — in other words, every proof of existence is
necessarily constructive.
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.
The following example illustrates why assuming the excluded middle
may lead to non-constructive proofs:
Claim: There exist irrational numbers a and b such that a ^
b is rational.
Proof: It is not difficult to show that sqrt 2 is irrational.
If sqrt 2 ^ sqrt 2 is rational, it suffices to take a = b =
sqrt 2 and we are done. Otherwise, sqrt 2 ^ sqrt 2 is
irrational. In this case, we can take a = sqrt 2 ^ sqrt 2 and
b = sqrt 2, since a ^ b = sqrt 2 ^ (sqrt 2 * sqrt 2) = sqrt 2 ^
2 = 2. ☐
Do you see what happened here? We used the excluded middle to
consider separately the cases where sqrt 2 ^ sqrt 2 is rational
and where it is not, without knowing which one actually holds!
Because of that, we wind up knowing that such a and b exist
but we cannot determine what their actual values are (at least,
using this line of argument).
As useful as constructive logic is, it does have its limitations:
There are many statements that can easily be proven in classical
logic but that have much more complicated constructive proofs, and
there are some that are known to have no constructive proof at
all! Fortunately, like functional extensionality, the excluded
middle is known to be compatible with Coq's logic, allowing us to
add it safely as an axiom. However, we will not need to do so in
this book: the results that we cover can be developed entirely
within constructive logic at negligible extra cost.
It takes some practice to understand which proof techniques must
be avoided in constructive reasoning, but arguments by
contradiction, in particular, are infamous for leading to
non-constructive proofs. Here's a typical example: suppose that
we want to show that there exists x with some property P,
i.e., such that P x. We start by assuming that our conclusion
is false; that is, ¬ ∃x, P x. From this premise, it is not
hard to derive ∀x, ¬ P x. If we manage to show that this
intermediate fact results in a contradiction, we arrive at an
existence proof without ever exhibiting a value of x for which
P x holds!
The technical flaw here, from a constructive standpoint, is that
we claimed to prove ∃x, P x using a proof of
¬ ¬ (∃x, P x). Allowing ourselves to remove double
negations from arbitrary statements is equivalent to assuming the
excluded middle, as shown in one of the exercises below. Thus,
this line of reasoning cannot be encoded in Coq without assuming
additional axioms.
Exercise: 3 stars (excluded_middle_irrefutable)
Proving the consistency of Coq with the general excluded middle axiom requires complicated reasoning that cannot be carried out within Coq itself. However, the following theorem implies that it is always safe to assume 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) (since P implies ¬ ¬ P, by the exercise below), which would be a contradiction. But since we can't, it is safe to add P ∨ ¬P as an axiom.
Theorem excluded_middle_irrefutable: ∀ (P:Prop),
¬ ¬ (P ∨ ¬ P).
Proof.
(* FILL IN HERE *) Admitted.
☐
¬ ¬ (P ∨ ¬ P).
Proof.
(* FILL IN HERE *) Admitted.
Exercise: 3 stars, optional (not_exists_dist)
It is a theorem of classical logic that the following two assertions are equivalent:
¬ (∃ x, ¬ P x)
∀ x, P x
The dist_not_exists theorem above proves one side of this
equivalence. Interestingly, the other direction cannot be proved
in constructive logic. Your job is to show that it is implied by
the excluded middle.
∀ x, P x
Theorem not_exists_dist :
excluded_middle →
∀ (X:Type) (P : X → Prop),
¬ (∃ x, ¬ P x) → (∀ x, P x).
Proof.
(* FILL IN HERE *) Admitted.
☐
excluded_middle →
∀ (X:Type) (P : X → Prop),
¬ (∃ x, ¬ P x) → (∀ x, P x).
Proof.
(* FILL IN HERE *) Admitted.
Exercise: 5 stars, optional (classical_axioms)
For those who like a challenge, here is an exercise taken from the Coq'Art book by Bertot and Casteran (p. 123). Each of the following four statements, together with excluded_middle, can be considered as characterizing classical logic. We can't prove any of them in Coq, but we can consistently add any one of them as an axiom if we wish to work in classical logic.
Definition peirce := ∀ P Q: Prop,
((P→Q)→P)→P.
Definition double_negation_elimination := ∀ 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).
(* FILL IN HERE *)
☐
((P→Q)→P)→P.
Definition double_negation_elimination := ∀ 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).
(* FILL IN HERE *)