LogicLogic in Rocq
So far, we have seen...
In this chapter we will introduce several more flavors of both
propositions and proofs.
Like everything in Rocq, well-formed propositions have a type:
- propositions: mathematical statements, so far only of 3 kinds:
- equality propositions (e1 = e2)
- implications (P → Q)
- quantified propositions (∀ x, P)
- proofs: ways of presenting evidence for the truth of a proposition
The Prop Type
Check (∀ n m : nat, n + m = m + n) : Prop.
Note that all syntactically well-formed propositions have type Prop in Rocq, regardless of whether they are true or not.
Check 2 = 2 : Prop.
Check 3 = 2 : Prop.
Check ∀ n : nat, n = 2 : Prop.
Check 3 = 2 : Prop.
Check ∀ n : nat, n = 2 : Prop.
So far, we've seen one primary place where 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 : Prop.
Theorem plus_claim_is_true :
plus_claim.
Proof. reflexivity. Qed.
Check 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 Rocq, functions that return propositions are said to define properties of their arguments.
Definition injective {A B} (f : A → B) : Prop :=
∀ x y : A, f x = f y → x = y.
Lemma succ_inj : injective S.
Proof.
intros x y H. injection H as H1. apply H1.
Qed.
∀ x y : A, f x = f y → x = y.
Lemma succ_inj : injective S.
Proof.
intros x y H. injection H as H1. apply H1.
Qed.
The familiar equality operator = is a (binary) function that returns a Prop.
Check @eq : ∀ A : Type, A → A → Prop.
What is the type of the following expression?
pred (S O) = O (A) Prop
(B) nat→Prop
(C) ∀ n:nat, Prop
(D) nat→nat
(E) Not typeable
pred (S O) = O (A) Prop
Check (pred (S O) = O) : Prop.
What is the type of the following expression?
∀ n:nat, pred (S n) = n (A) Prop
(B) nat→Prop
(C) ∀ n:nat, Prop
(D) nat→nat
(E) Not typeable
∀ n:nat, pred (S n) = n (A) Prop
Check (∀ n:nat, pred (S n) = n) : Prop.
What is the type of the following expression?
∀ n:nat, S (pred n) = n (A) Prop
(B) nat→Prop
(C) nat→nat
(D) Not typeable
∀ n:nat, S (pred n) = n (A) Prop
Check (∀ n:nat, S (pred n) = n) : Prop.
What is the type of the following expression?
∀ n:nat, S (pred n) (A) Prop
(B) nat→Prop
(C) nat→nat
(D) Not typeable
∀ n:nat, S (pred n) (A) Prop
Fail Check (∀ n:nat, S (pred n)).
(* The command has indeed failed with message:
In environment
n : nat
The term "(S (pred n))" has type "nat" which should be Set, Prop or Type. *)
(* The command has indeed failed with message:
In environment
n : nat
The term "(S (pred n))" has type "nat" which should be Set, Prop or Type. *)
What is the type of the following expression?
fun n:nat ⇒ S (pred n) (A) Prop
(B) nat→Prop
(C) nat→nat
(D) Not typeable
fun n:nat ⇒ S (pred n) (A) Prop
Check (fun n:nat ⇒ pred (S n)) : nat → nat.
What is the type of the following expression?
fun n:nat ⇒ S (pred n) = n (A) Prop
(B) nat→Prop
(C) nat→nat
(D) Not typeable
fun n:nat ⇒ S (pred n) = n (A) Prop
Check (fun n:nat ⇒ pred (S n) = n) : nat → Prop.
Which of the following is not a proposition?
(A) 3 + 2 = 4
(B) 3 + 2 = 5
(C) 3 + 2 =? 5
(D) ((3+2) =? 4) = false
(E) ∀ n, (((3+2) =? n) = true) → n = 5
(F) All of these are propositions
Check 3 + 2 =? 5 : bool.
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". *)
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". *)
Logical Connectives
Conjunction
Example and_example : 3 + 4 = 7 ∧ 2 × 2 = 4.
To prove a conjunction, start with the split tactic. This 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.
Check @conj : ∀ A B : Prop, A → B → A ∧ B.
Example plus_is_O :
∀ n m : nat, n + m = 0 → n = 0 ∧ m = 0.
Proof.
(* WORK IN CLASS *) Admitted.
∀ 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 can use our good old 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.
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.
The infix notation ∧ is actually just syntactic sugar for and A B. That is, and is a Rocq operator that takes two propositions as arguments and yields a proposition.
Check and : Prop → Prop → Prop.
Disjunction
To use a disjunctive hypothesis in a proof, we proceed by case analysis -- which, as with other data types like nat, can be done explicitly with destruct or implicitly with an intros pattern:
Lemma factor_is_O:
∀ n m : nat, n = 0 ∨ m = 0 → n × m = 0.
Proof.
(* This intro 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 intro 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.
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
To see how negation works, recall the principle of explosion from the Tactics chapter, which asserts that, if we assume a contradiction, then any other proposition can be derived.
Rocq actually makes an equivalent but slightly different choice, defining ¬ P as P → False, where False is a specific un-provable proposition defined in the standard library.
Definition not (P:Prop) := P → False.
Check not : Prop → Prop.
Notation "~ x" := (not x) : type_scope.
Since False is a contradictory proposition, the principle of explosion also applies to it. If we can get False into the context, we can use destruct on it to complete any goal:
Theorem ex_falso_quodlibet : ∀ (P:Prop),
False → P.
Proof.
intros P contra.
destruct contra. Qed.
False → P.
Proof.
intros P contra.
destruct contra. Qed.
Notation "x <> y" := (~(x = y)) : type_scope.
For example:
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 Rocq. Even though you may see perfectly well why a claim involving negation holds, it can be a little tricky at first to see how to make Rocq understand it!
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.
Theorem not_true_is_false : ∀ b : bool,
b ≠ true → b = false.
b ≠ true → b = false.
Proof.
intros b H. destruct b eqn:HE.
- (* b = true *)
unfold not in H.
apply ex_falso_quodlibet.
apply H. reflexivity.
- (* b = false *)
reflexivity.
Qed.
intros b H. destruct b eqn:HE.
- (* b = true *)
unfold not in H.
apply ex_falso_quodlibet.
apply H. reflexivity.
- (* b = false *)
reflexivity.
Qed.
To prove the following proposition, which tactics will we need
besides intros and apply?
∀ X, ∀ a b : X, (a=b) ∧ (a≠b) → False. (A) destruct, unfold, left and right
(B) destruct and unfold
(C) only destruct
(D) left and/or right
(E) only unfold
(F) none of the above
∀ X, ∀ a b : X, (a=b) ∧ (a≠b) → False. (A) destruct, unfold, left and right
Lemma quiz1: ∀ X, ∀ a b : X, (a=b) ∧ (a≠b) → False.
Proof.
intros X a b [Hab Hnab]. apply Hnab. apply Hab.
Qed.
Proof.
intros X a b [Hab Hnab]. apply Hnab. apply Hab.
Qed.
To prove the following proposition, which tactics will we
need besides intros and apply?
∀ P Q : Prop, P ∨ Q → ~~(P ∨ Q). (A) destruct, unfold, left and right
(B) destruct and unfold
(C) only destruct
(D) left and/or right
(E) only unfold
(F) none of the above
∀ P Q : Prop, P ∨ Q → ~~(P ∨ Q). (A) destruct, unfold, left and right
Lemma quiz2 : ∀ P Q : Prop, P ∨ Q → ~~(P ∨ Q).
Proof.
intros P Q HPQ HnPQ. apply HnPQ in HPQ. apply HPQ.
Qed.
Proof.
intros P Q HPQ HnPQ. apply HnPQ in HPQ. apply HPQ.
Qed.
To prove the following proposition, which tactics will we
need besides intros and apply?
∀ P Q: Prop, P → (P ∨ ~~Q). (A) destruct, unfold, left and right
(B) destruct and unfold
(C) only destruct
(D) left and/or right
(E) only unfold
(F) none of the above
∀ P Q: Prop, P → (P ∨ ~~Q). (A) destruct, unfold, left and right
Lemma quiz3 : ∀ P Q: Prop, P → (P ∨ ~~Q).
Proof.
intros P Q HP. left. apply HP.
Qed.
Proof.
intros P Q HP. left. apply HP.
Qed.
To prove the following proposition, which tactics will we need
besides intros and apply?
∀ P Q: Prop, P ∨ Q → ~~P ∨ ~~Q. (A) destruct, unfold, left and right
(B) destruct and unfold
(C) only destruct
(D) left and/or right
(E) only unfold
(F) none of the above
∀ P Q: Prop, P ∨ Q → ~~P ∨ ~~Q. (A) destruct, unfold, left and right
Lemma quiz4 : ∀ P Q: Prop, P ∨ Q → ~~P ∨ ~~Q.
Proof.
intros P Q [HP | HQ].
- (* left *)
left. intros HnP. apply HnP in HP. apply HP.
- (* right *)
right. intros HnQ. apply HnQ in HQ. apply HQ.
Qed.
Proof.
intros P Q [HP | HQ].
- (* left *)
left. intros HnP. apply HnP in HP. apply HP.
- (* right *)
right. intros HnQ. apply HnQ in HQ. apply HQ.
Qed.
To prove the following proposition, which tactics will we need
besides intros and apply?
∀ A : Prop, 1=0 → (A ∨ ¬A). (A) discriminate, unfold, left and right
(B) discriminate and unfold
(C) only discriminate
(D) left and/or right
(E) only unfold
(F) none of the above
∀ A : Prop, 1=0 → (A ∨ ¬A). (A) discriminate, unfold, left and right
Lemma quiz5 : ∀ A : Prop, 1=0 → (A ∨ ¬A).
Proof.
intros P H. discriminate H.
Qed.
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
relatively rarely: it is trivial (and therefore uninteresting) to
prove as a goal, and it provides no useful information when it
appears as a hypothesis.
Logical Equivalence
Print "<->".
The handy "if and only if" connective, which asserts that two
propositions have the same truth value, is simply the conjunction
of two implications.
Print "<->".
Print "<->".
(* ===>
Notation "A <-> B" := (iff A B)
iff = fun A B : Prop => (A -> B) /\ (B -> A)
: Prop -> Prop -> Prop
Argumments iff (A B)*)
Notation "A <-> B" := (iff A B)
iff = fun A B : Prop => (A -> B) /\ (B -> A)
: Prop -> Prop -> Prop
Argumments iff (A B)*)
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.
intros b. split.
- (* -> *) apply not_true_is_false.
- (* <- *)
intros H. rewrite H. intros H'. discriminate H'.
Qed.
(P ↔ Q) → (Q ↔ P).
Proof.
(* WORK IN CLASS *) Admitted.
Lemma not_true_iff_false : ∀ b,
b ≠ true ↔ b = false.
Proof.
intros b. split.
- (* -> *) apply not_true_is_false.
- (* <- *)
intros H. rewrite H. intros H'. discriminate H'.
Qed.
Lemma apply_iff_example1:
∀ P Q R : Prop, (P ↔ Q) → (Q → R) → (P → R).
Proof.
intros P Q R Hiff H HP. apply H. apply Hiff. apply HP.
Qed.
Lemma apply_iff_example2:
∀ P Q R : Prop, (P ↔ Q) → (P → R) → (Q → R).
Proof.
intros P Q R Hiff H HQ. apply H. apply Hiff. apply HQ.
Qed.
∀ P Q R : Prop, (P ↔ Q) → (Q → R) → (P → R).
Proof.
intros P Q R Hiff H HP. apply H. apply Hiff. apply HP.
Qed.
Lemma apply_iff_example2:
∀ P Q R : Prop, (P ↔ Q) → (P → R) → (Q → R).
Proof.
intros P Q R Hiff H HQ. apply H. apply Hiff. apply HQ.
Qed.
Setoids and Logical Equivalence
From Stdlib Require Import Setoids.Setoid.
A "setoid" is a set equipped with an equivalence relation,
such as = or ↔.
Here is a simple example demonstrating how these tactics work with
iff.
First, let's prove a couple of basic iff equivalences. (For these
proofs we are not using setoids yet.)
Lemma mul_eq_0 : ∀ n m, n × m = 0 ↔ n = 0 ∨ m = 0.
Theorem or_assoc :
∀ P Q R : Prop, P ∨ (Q ∨ R) ↔ (P ∨ Q) ∨ R.
Proof.
split.
- apply mult_is_O.
- apply factor_is_O.
Qed.
split.
- apply mult_is_O.
- apply factor_is_O.
Qed.
Theorem 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.
We can now use these facts with rewrite and reflexivity to
prove a ternary version of the mult_eq_0 fact above without
splitting the top-level iff:
Lemma mul_eq_0_ternary :
∀ n m p, n × m × p = 0 ↔ n = 0 ∨ m = 0 ∨ p = 0.
Proof.
intros n m p.
rewrite mul_eq_0. rewrite mul_eq_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 mul_eq_0. rewrite mul_eq_0. rewrite or_assoc.
reflexivity.
Qed.
Existential Quantification
Definition Even x := ∃ n : nat, x = double n.
Check Even : nat → Prop.
Lemma four_is_Even : Even 4.
Proof.
unfold Even. ∃ 2. reflexivity.
Qed.
Check Even : nat → Prop.
Lemma four_is_Even : Even 4.
Proof.
unfold Even. ∃ 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.
Recap -- Logical connectives in Rocq
- and : Prop → Prop → Prop (conjunction):
- introduced with the split tactic
- eliminated with destruct H as [H1 H2]
- or : Prop → Prop → Prop (disjunction):
- introduced with left and right tactics
- eliminated with destruct H as [H1 | H2]
- False : Prop
- eliminated with destruct H as []
- True : Prop
- introduced with apply I, but not as useful
- ex : ∀ A:Type, (A → Prop) → Prop (existential)
- introduced with ∃ w
- eliminated with destruct H as [x H]
- not : Prop → Prop (negation):
- not P defined as P → False
- iff : Prop → Prop → Prop (logical equivalence):
- iff P Q defined as (P → Q) ∧ (Q → P)
- equality (e1 = e2)
- implication (P → Q)
- universal quantification (∀ x, P)
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 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.
Theorem In_map :
∀ (A B : Type) (f : A → B) (l : list A) (x : A),
In x l →
In (f x) (map f l).
∀ (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.
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 : nat → nat → nat.
Check @rev : ∀ X, list X → list X.
Check @rev : ∀ X, list X → list X.
We can also use it to check what theorem a particular identifier
refers to:
Check add_comm : ∀ n m : nat, n + m = m + n.
Check plus_id_example : ∀ n m : nat, n = m → n + n = m + m.
Check plus_id_example : ∀ n m : nat, n = m → n + n = m + m.
Rocq checks the statements of the add_comm and
plus_id_example theorems in the same way that it checks the
type of any term (e.g., plus). If we leave off the colon and
type, Rocq will print these types for us.
Why?
Similarly, the statement of a theorem tells us what we can use that
theorem for.
This is often handy in proof scripts -- e.g., suppose we want too
prove the following:
The reason is that the identifier add_comm actually refers to a proof object -- a logical derivation establishing the truth of the statement ∀ n m : nat, n + m = m + n. The type of this object is the proposition that it is a proof of.
The type of an ordinary function tells us what we can do with it.
- 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 a term of type ∀ n m, n = m → n + n = m + m and we provide it two numbers n and m and a third "argument" of type n = m, we get back a proof object of type n + n = m + m.
Rocq actually allows us to apply a theorem as if it were a function.
Lemma add_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 add_comm twice to make the two sides match. The
problem is that the second rewrite will undo the effect of the
first.
Proof.
intros x y z.
rewrite add_comm.
rewrite add_comm.
(* We are back where we started... *)
Abort.
intros x y z.
rewrite add_comm.
rewrite add_comm.
(* We are back where we started... *)
Abort.
We can fix this by applying add_comm to the arguments we want it be to instantiated with. Then the rewrite is forced to happen exactly where we want it.
Lemma add_comm3_take3 :
∀ x y z, x + (y + z) = (z + y) + x.
Proof.
intros x y z.
rewrite add_comm.
rewrite (add_comm y z).
reflexivity.
Qed.
∀ x y z, x + (y + z) = (z + y) + x.
Proof.
intros x y z.
rewrite add_comm.
rewrite (add_comm y z).
reflexivity.
Qed.
Here's another example of using a theorem about lists like a function. Suppose we have proved the following simple fact about lists...
Theorem in_not_nil :
∀ A (x : A) (l : list A), In x l → l ≠ [].
∀ A (x : A) (l : list A), In x l → l ≠ [].
Proof.
intros A x l H. unfold not. intro Hl.
rewrite Hl in H.
simpl in H.
apply H.
Qed.
intros A x l H. unfold not. intro Hl.
rewrite Hl in H.
simpl in H.
apply H.
Qed.
Note that one quantified variable (x) does not appear in
the conclusion (l ≠ []).
Intuitively, we should be able to use this theorem to prove the special
case where x is 42. However, simply invoking the tactic apply
in_not_nil will fail because it cannot infer the value of x.
Lemma in_not_nil_42 :
∀ l : list nat, In 42 l → l ≠ [].
Proof.
intros l H.
Fail apply in_not_nil.
Abort.
∀ l : list nat, In 42 l → l ≠ [].
Proof.
intros l H.
Fail apply in_not_nil.
Abort.
Lemma in_not_nil_42_take2 :
∀ l : list nat, In 42 l → l ≠ [].
Proof.
intros l H.
apply in_not_nil with (x := 42).
apply H.
Qed.
∀ l : list nat, In 42 l → l ≠ [].
Proof.
intros l H.
apply in_not_nil with (x := 42).
apply H.
Qed.
Lemma in_not_nil_42_take3 :
∀ l : list nat, In 42 l → l ≠ [].
Proof.
intros l H.
apply in_not_nil in H.
apply H.
Qed.
∀ l : list nat, In 42 l → l ≠ [].
Proof.
intros l H.
apply in_not_nil in H.
apply H.
Qed.
Lemma in_not_nil_42_take4 :
∀ l : list nat, In 42 l → l ≠ [].
Proof.
intros l H.
apply (in_not_nil nat 42).
apply H.
Qed.
∀ l : list nat, In 42 l → l ≠ [].
Proof.
intros l H.
apply (in_not_nil nat 42).
apply H.
Qed.
We can also explicitly apply the lemma to a hypothesis, causing the values of the other parameters to be inferred:
Lemma in_not_nil_42_take5 :
∀ l : list nat, In 42 l → l ≠ [].
Proof.
intros l H.
apply (in_not_nil _ _ _ H).
Qed.
∀ l : list nat, In 42 l → l ≠ [].
Proof.
intros l H.
apply (in_not_nil _ _ _ H).
Qed.
Suppose we have
a, b : nat
H1 : a = b
H2 : b = 42
trans_eq : ∀ (X : Type) (n m o : X),
n = m → m = o → n = o What is the type of this "proof object"?
trans_eq nat a b 42 H1 H2
(A) a = b
(B) 42 = a
(C) a = 42
(D) Does not typecheck
a, b : nat
H1 : a = b
H2 : b = 42
trans_eq : ∀ (X : Type) (n m o : X),
n = m → m = o → n = o What is the type of this "proof object"?
trans_eq nat a b 42 H1 H2
Check trans_eq nat a b 42 H1 H2
: a = 42.
: a = 42.
Suppose, again, that we have
a, b : nat
H1 : a = b
H2 : b = 42
trans_eq : ∀ (X : Type) (n m o : X),
n = m → m = o → n = o What is the type of this proof object?
trans_eq nat _ _ _ H1 H2
(A) a = b
(B) 42 = a
(C) a = 42
(D) Does not typecheck
a, b : nat
H1 : a = b
H2 : b = 42
trans_eq : ∀ (X : Type) (n m o : X),
n = m → m = o → n = o What is the type of this proof object?
trans_eq nat _ _ _ H1 H2
Check trans_eq nat _ _ _ H1 H2
: a = 42.
: a = 42.
Suppose, again, that we have
a, b : nat
H1 : a = b
H2 : b = 42
trans_eq : ∀ (X : Type) (n m o : X),
n = m → m = o → n = o What is the type of this proof object?
trans_eq nat b 42 a H2
(A) b = a
(B) b = a → 42 = a
(C) 42 = a → b = a
(D) Does not typecheck
a, b : nat
H1 : a = b
H2 : b = 42
trans_eq : ∀ (X : Type) (n m o : X),
n = m → m = o → n = o What is the type of this proof object?
trans_eq nat b 42 a H2
Check trans_eq nat b 42 a H2
: 42 = a → b = a.
: 42 = a → b = a.
Suppose, again, that we have
a, b : nat
H1 : a = b
H2 : b = 42
trans_eq : ∀ (X : Type) (n m o : X),
n = m → m = o → n = o What is the type of this proof object?
trans_eq _ 42 a b
(A) a = b → b = 42 → a = 42
(B) 42 = a → a = b → 42 = b
(C) a = 42 → 42 = b → a = b
(D) Does not typecheck
a, b : nat
H1 : a = b
H2 : b = 42
trans_eq : ∀ (X : Type) (n m o : X),
n = m → m = o → n = o What is the type of this proof object?
trans_eq _ 42 a b
Check trans_eq _ 42 a b
: 42 = a → a = b → 42 = b.
: 42 = a → a = b → 42 = b.
Suppose, again, that we have
a, b : nat
H1 : a = b
H2 : b = 42
trans_eq : ∀ (X : Type) (n m o : X),
n = m → m = o → n = o What is the type of this proof object?
trans_eq _ _ _ _ H2 H1
(A) b = a
(B) 42 = a
(C) a = 42
(D) Does not typecheck
a, b : nat
H1 : a = b
H2 : b = 42
trans_eq : ∀ (X : Type) (n m o : X),
n = m → m = o → n = o What is the type of this proof object?
trans_eq _ _ _ _ H2 H1
Fail Check trans_eq _ _ _ _ H2 H1.
Working with Decidable Properties
bool Prop
==== ====
decidable? yes no
useable with match? yes no
works with rewrite tactic? no yes
- For example, even : nat → bool is a decision procedure for the property "is even".
It follows that there are some properties of numbers that we cannot express as functions of type nat → bool.
- For example, the property "is the code of a halting Turing machine" is undecidable, so there is no way to write it as a function of type nat → bool.
- For example, "is the code of a halting Turing machine" is a perfectly legitimate mathematical property, and we can absolutely represent it as a Rocq expression of type nat → Prop.
Example even_42_bool : even 42 = true.
Proof. reflexivity. Qed.
... or that there exists some k such that n = double k.
Example even_42_prop : Even 42.
Proof. unfold Even. ∃ 21. reflexivity. Qed.
Of course, it would be deeply strange if these two
characterizations of evenness did not describe the same set of
natural numbers!
Fortunately, they do!
Theorem even_bool_prop : ∀ n,
even n = true ↔ Even n.
even n = true ↔ Even n.
Proof.
intros n. split.
- intros H. destruct (even_double_conv n) as [k Hk].
rewrite Hk. rewrite H. ∃ k. reflexivity.
- intros [k Hk]. rewrite Hk. apply even_double.
Qed.
intros n. split.
- intros H. destruct (even_double_conv n) as [k Hk].
rewrite Hk. rewrite H. ∃ k. reflexivity.
- intros [k Hk]. rewrite Hk. apply even_double.
Qed.
In view of this theorem, we can say that the boolean computation
even 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.
Fail
Definition is_even_prime n :=
if n = 2 then true
else false.
Definition is_even_prime n :=
if n = 2 then true
else false.
Rather, we have to state this definition using a boolean equality
test.
Definition is_even_prime n :=
if n =? 2 then true
else false.
if n =? 2 then true
else false.
More generally, stating facts using booleans can often enable effective proof automation through computation with Rocq terms, a technique known as proof by reflection.
Example even_1000 : Even 1000.
The most direct way to prove this is to give the value of k
explicitly.
Proof. unfold Even. ∃ 500. reflexivity. Qed.
The proof of the corresponding boolean statement is simpler, because we
don't have to invent the witness 500: Rocq's computation mechanism
does it for us!
Example even_1000' : even 1000 = true.
Proof. reflexivity. Qed.
Proof. reflexivity. Qed.
Now, the useful observation 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'' : Even 1000.
Proof. apply even_bool_prop. reflexivity. Qed.
Proof. apply even_bool_prop. reflexivity. Qed.
Although we haven't gained much in terms of proof-script
line count in this case, larger proofs can often be made considerably
simpler by the use of reflection. As an extreme example, a famous
Rocq proof of the even more famous 4-color theorem uses
reflection to reduce the analysis of hundreds of different cases
to a boolean computation.
Another advantage of booleans is that the negation of a claim
about booleans is straightforward to state and (when true) prove:
simply flip the expected boolean result.
Example not_even_1001 : even 1001 = false.
Proof.
reflexivity.
Qed.
Proof.
reflexivity.
Qed.
In contrast, propositional negation can be difficult to work with directly.
Example not_even_1001' : ~(Even 1001).
Proving this directly -- by assuming that there is some n such that
1001 = double n and then somehow reasoning to a contradiction --
would be rather complicated.
But if we convert it to a claim about the boolean even function, we
can let Rocq do the work for us.
Proof.
(* WORK IN CLASS *) Admitted.
(* 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.
Being able to cross back and forth between the boolean and
propositional worlds will often be convenient in later chapters.
The Logic of Rocq
Functional Extensionality
Example function_equality_ex1 :
(fun x ⇒ 3 + x) = (fun x ⇒ (pred 4) + x).
(fun x ⇒ 3 + x) = (fun x ⇒ (pred 4) + x).
Proof. reflexivity. Qed.
This works when Rocq can simplify the functions to the same expression,
but this doesn't always happen.
In common mathematical practice, two functions f and g are
considered equal if they produce the same output on every input:
(∀ x, f x = g x) → f = g This is known as the principle of functional extensionality.
However, functional extensionality is not part of Rocq's built-in logic.
This means that some intuitively obvious propositions are not
provable.
These two functions are equal just by simplification, but in general functions can be equal for more interesting reasons.
(∀ 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.
Fail reflexivity. Fail rewrite add_comm.
(* Stuck *)
Abort.
(fun x ⇒ plus x 1) = (fun x ⇒ plus 1 x).
Proof.
Fail reflexivity. Fail rewrite add_comm.
(* Stuck *)
Abort.
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.
Defining something as an 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 add_comm.
Qed.
(fun x ⇒ plus x 1) = (fun x ⇒ plus 1 x).
Proof.
apply functional_extensionality. intros x.
apply add_comm.
Qed.
Naturally, we need to be quite careful when adding new axioms into Rocq's logic, as this can render it inconsistent -- that is, it may become 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 *)
Classical vs. Constructive Logic
Definition excluded_middle := ∀ P : Prop,
P ∨ ¬P.
P ∨ ¬P.
Logics like Rocq's, which do not assume the excluded middle, are
referred to as constructive logics.
Logical systems such as ZFC, in which the excluded middle does
hold for arbitrary propositions, are referred to as classical.
