TacticsMore Basic Tactics
- how to use auxiliary lemmas in both "forward-style" and "backward-style" proofs;
- how to reason about data constructors (in particular, how to use the fact that they are injective and disjoint);
- how to strengthen an induction hypothesis (and when such strengthening is required); and
- more details on how to reason by case analysis.
The apply Tactic
Theorem silly1 : ∀ (n m o p : nat),
n = m →
[n;o] = [n;p] →
[n;o] = [m;p].
Proof.
intros n m o p eq1 eq2.
rewrite <- eq1.
n = m →
[n;o] = [n;p] →
[n;o] = [m;p].
Proof.
intros n m o p eq1 eq2.
rewrite <- eq1.
Here, we could finish with "rewrite → eq2. reflexivity." as we
have done several times before. We can achieve the same effect in
a single step by using the apply tactic instead:
apply eq2. Qed.
silly1 is the first theorem we've seen that has multiple
hypotheses. How should we parse the following proposition?
The answer is that implication is right associative, that is,
we assume that parentheses go on the right, as in:
If we were to read this proposition aloud, we could naively
read it as:
But we can read it more naturally as:
That is, we've translated nested implications to a conjuction,
with the word "and". We'll talk more about conjunctions in Logic.
The apply tactic also works with conditional hypotheses
and lemmas: if the statement being applied is an implication, then
the premises of this implication will be added to the list of
subgoals needing to be proved. If there are no premises (i.e., H
isn't an implication), then you're done!
n = m →
[n;o] = [n;p] →
[n;o] = [m;p].
[n;o] = [n;p] →
[n;o] = [m;p].
n = m →
([n;o] = [n;p] →
[n;o] = [m;p]).
([n;o] = [n;p] →
[n;o] = [m;p]).
if it is the case that [n = m], then if it is the case that [[n;o] = [n;p]] then it is the case that [[n;o] = [m;p]].
if it is the case that [n = m] and [[n;o] = [n;p]], then it is the case that [[n;o] = [m;p]].
Theorem silly2 : ∀ (n m o p : nat),
n = m →
(∀ (q r : nat), q = r → [q;o] = [r;p]) →
[n;o] = [m;p].
Proof.
intros n m o p eq1 eq2.
apply eq2. apply eq1. Qed.
n = m →
(∀ (q r : nat), q = r → [q;o] = [r;p]) →
[n;o] = [m;p].
Proof.
intros n m o p eq1 eq2.
apply eq2. apply eq1. Qed.
Typically, when we use apply H, the statement H will
begin with a ∀that binds some universal variables. When
Coq matches the current goal against the conclusion of H, it
will try to find appropriate values for these variables. For
example, when we do apply eq2 in the following proof, the
universal variable q in eq2 gets instantiated with n and r
gets instantiated with m.
Theorem silly2a : ∀ (n m : nat),
(n,n) = (m,m) →
(∀ (q r : nat), (q,q) = (r,r) → [q] = [r]) →
[n] = [m].
Proof.
intros n m eq1 eq2.
apply eq2. apply eq1. Qed.
(n,n) = (m,m) →
(∀ (q r : nat), (q,q) = (r,r) → [q] = [r]) →
[n] = [m].
Proof.
intros n m eq1 eq2.
apply eq2. apply eq1. Qed.
Theorem silly_ex :
(∀ n, evenb n = true → oddb (S n) = true) →
evenb 3 = true →
oddb 4 = true.
Proof.
(* FILL IN HERE *) Admitted.
☐
(∀ n, evenb n = true → oddb (S n) = true) →
evenb 3 = true →
oddb 4 = true.
Proof.
(* FILL IN HERE *) Admitted.
Theorem silly3_firsttry : ∀ (n : nat),
true = beq_nat n 5 →
beq_nat (S (S n)) 7 = true.
Proof.
intros n H.
simpl.
true = beq_nat n 5 →
beq_nat (S (S n)) 7 = true.
Proof.
intros n H.
simpl.
Here we cannot use apply directly, but we can use the symmetry
tactic, which switches the left and right sides of an equality in
the goal.
symmetry.
simpl. (* (This simpl is optional, since apply will perform
simplification first, if needed.) *)
apply H. Qed.
simpl. (* (This simpl is optional, since apply will perform
simplification first, if needed.) *)
apply H. Qed.
Exercise: 3 stars (apply_exercise1)
(Hint: You can use apply with previously defined lemmas, not just hypotheses in the context. Remember that Search is your friend.)
Theorem rev_exercise1 : ∀ (l l' : list nat),
l = rev l' →
l' = rev l.
Proof.
(* FILL IN HERE *) Admitted.
☐
l = rev l' →
l' = rev l.
Proof.
(* FILL IN HERE *) Admitted.
Exercise: 1 star, optional (apply_rewrite)
Briefly explain the difference between the tactics apply and rewrite. What are the situations where both can usefully be applied?☐
The apply ... with ... Tactic
Example trans_eq_example : ∀ (a b c d e f : nat),
[a;b] = [c;d] →
[c;d] = [e;f] →
[a;b] = [e;f].
Proof.
intros a b c d e f eq1 eq2.
rewrite → eq1. rewrite → eq2. reflexivity. Qed.
[a;b] = [c;d] →
[c;d] = [e;f] →
[a;b] = [e;f].
Proof.
intros a b c d e f eq1 eq2.
rewrite → eq1. rewrite → eq2. reflexivity. Qed.
Since this is a common pattern, we might like to pull it out
as a lemma recording, once and for all, the fact that equality is
transitive.
Theorem trans_eq : ∀ (X:Type) (n m o : X),
n = m → m = o → n = o.
Proof.
intros X n m o eq1 eq2. rewrite → eq1. rewrite → eq2.
reflexivity. Qed.
n = m → m = o → n = o.
Proof.
intros X n m o eq1 eq2. rewrite → eq1. rewrite → eq2.
reflexivity. Qed.
Now, we should be able to use trans_eq to prove the above
example. However, to do this we need a slight refinement of the
apply tactic.
Example trans_eq_example' : ∀ (a b c d e f : nat),
[a;b] = [c;d] →
[c;d] = [e;f] →
[a;b] = [e;f].
Proof.
intros a b c d e f eq1 eq2.
[a;b] = [c;d] →
[c;d] = [e;f] →
[a;b] = [e;f].
Proof.
intros a b c d e f eq1 eq2.
If we simply tell Coq apply trans_eq at this point, it can
tell (by matching the goal against the conclusion of the lemma)
that it should instantiate X with [nat], n with [a,b], and
o with [e,f]. However, the matching process doesn't determine
an instantiation for m: we have to supply one explicitly by
adding with (m:=[c,d]) to the invocation of apply.
Actually, we usually don't have to include the name m in
the with clause; Coq is often smart enough to figure out which
instantiation we're giving. We could instead write: apply
trans_eq with [c;d].
Exercise: 3 stars, optional (apply_with_exercise)
Example trans_eq_exercise : ∀ (n m o p : nat),
m = (minustwo o) →
(n + p) = m →
(n + p) = (minustwo o).
Proof.
(* FILL IN HERE *) Admitted.
☐
m = (minustwo o) →
(n + p) = m →
(n + p) = (minustwo o).
Proof.
(* FILL IN HERE *) Admitted.
The inversion Tactic
Inductive nat : Type :=
| O : nat
| S : nat → nat.
It is obvious from this definition that every number has one of
two forms: either it is the constructor O or it is built by
applying the constructor S to another number. But there is more
here than meets the eye: implicit in the definition (and in our
informal understanding of how datatype declarations work in other
programming languages) are two more facts:
| O : nat
| S : nat → nat.
- The constructor S is injective. That is, if S n = S m, it
must be the case that n = m.
- The constructors O and S are disjoint. That is, O is not equal to S n for any n.
c a1 a2 ... an = d b1 b2 ... bm
for some constructors c and d and arguments a1 ... an and
b1 ... bm. Then inversion H does one of two things:
- If c and d are the same constructor, then, inversion H adds
a1 = b1, a2 = b2 to the context and tries to use these new
equalities to rewrite the goal.
- If c and d are different constructors, then the hypothesis H is contradictory—-inversion H marks the current goal as completed and pops it off the goal stack.
By writing inversion H at this point, we are asking Coq to
generate all equations that it can infer from H as additional
hypotheses, replacing variables in the goal as it goes. In the
present example, this amounts to adding a new hypothesis H1 : n =
m and replacing n by m in the goal.
inversion H.
reflexivity.
Qed.
reflexivity.
Qed.
Here's a more interesting example that shows how multiple
equations can be derived at once.
Theorem inversion_ex1 : ∀ (n m o : nat),
[n; m] = [o; o] →
[n] = [m].
Proof.
intros n m o H. inversion H. reflexivity. Qed.
[n; m] = [o; o] →
[n] = [m].
Proof.
intros n m o H. inversion H. reflexivity. Qed.
We can name the equations that inversion generates with an
as ... clause:
Theorem inversion_ex2 : ∀ (n m : nat),
[n] = [m] →
n = m.
Proof.
intros n m H. inversion H as [Hnm]. reflexivity. Qed.
[n] = [m] →
n = m.
Proof.
intros n m H. inversion H as [Hnm]. reflexivity. Qed.
Example inversion_ex3 : ∀ (X : Type) (x y z : X) (l j : list X),
x :: y :: l = z :: j →
y :: l = x :: j →
x = y.
Proof.
(* FILL IN HERE *) Admitted.
☐
x :: y :: l = z :: j →
y :: l = x :: j →
x = y.
Proof.
(* FILL IN HERE *) Admitted.
We can proceed by case analysis on n. The first case is
trivial.
destruct n as [| n'].
- (* n = 0 *)
intros H. reflexivity.
- (* n = 0 *)
intros H. reflexivity.
However, the second one doesn't look so simple: assuming
beq_nat 0 (S n') = true, we must show S n' = 0, but the latter
clearly contradictory! The way forward lies in the assumption.
After simplifying the goal state, we see that beq_nat 0 (S n') =
true has become false = true:
- (* n = S n' *)
simpl.
simpl.
If we use inversion on this hypothesis, Coq notices that
the subgoal we are working on is impossible, and therefore removes
it from further consideration.
intros H. inversion H. Qed.
How might such a proof look informally? There's no direct informal
analogue of inversion—-instead, we either use the language of
injectivity (when we would perform inversion on an equality
between the same constructors) or we use the language of
contradiction (when the constructors differ). For example, we can
prove beq_nat_0_l informally as follows.
This is an instance of a logical principle known as the principle
of explosion, which asserts that a contradictory hypothesis
entails anything, even false things!
- Theorem: If beq_nat 0 n = true then n = 0, for all n.
- If n=0, we are done immediately.
- If n=S n', then beq_nat 0 (S n') = false. But we've assumed that beq_nat 0 (S n') = true. It cannot be that true = false, so we've arrived at a contradiction and it is impossible to have n=S n'. Qed.
- If n=S n', then beq_nat 0 (S n') = false. But we've assumed that beq_nat 0 (S n') = true—-a contradiction. Qed.
Theorem inversion_ex4 : ∀ (n : nat),
S n = O →
2 + 2 = 5.
Proof.
intros n contra. inversion contra. Qed.
Theorem inversion_ex5 : ∀ (n m : nat),
false = true →
[n] = [m].
Proof.
intros n m contra. inversion contra. Qed.
S n = O →
2 + 2 = 5.
Proof.
intros n contra. inversion contra. Qed.
Theorem inversion_ex5 : ∀ (n m : nat),
false = true →
[n] = [m].
Proof.
intros n m contra. inversion contra. Qed.
If you find the principle of explosion confusing, remember
that these proofs are not actually showing that the conclusion of
the statement holds. Rather, they are arguing that, if the
nonsensical situation described by the premise did somehow arise,
then the nonsensical conclusion would follow. We'll explore the
principle of explosion of more detail in the next chapter.
Exercise: 1 star (inversion_ex6)
Example inversion_ex6 : ∀ (X : Type)
(x y z : X) (l j : list X),
x :: y :: l = [] →
y :: l = z :: j →
x = z.
Proof.
(* FILL IN HERE *) Admitted.
☐
(x y z : X) (l j : list X),
x :: y :: l = [] →
y :: l = z :: j →
x = z.
Proof.
(* FILL IN HERE *) Admitted.
c a1 a2 ... an = d b1 b2 ... bm
for some constructors c and d and arguments a1 ... an and
b1 ... bm. Then inversion H has the following effect:
- If c and d are the same constructor, then, by the
injectivity of this constructor, we know that a1 = b1, a2 =
b2, etc. The inversion H adds these facts to the context and
tries to use them to rewrite the goal.
- If c and d are different constructors, then the hypothesis H is contradictory, and the current goal doesn't have to be considered at all. In this case, inversion H marks the current goal as completed and pops it off the goal stack.
Theorem f_equal : ∀ (A B : Type) (f: A → B) (x y: A),
x = y → f x = f y.
Proof. intros A B f x y eq. rewrite eq. reflexivity. Qed.
x = y → f x = f y.
Proof. intros A B f x y eq. rewrite eq. reflexivity. Qed.
We can prove a similar theorem for functions of two arguments.
Theorem f_equal2 : ∀ (A1 A2 B:Type) (f:A1 → A2 → B)
(x1 y1:A1) (x2 y2:A2),
x1 = y1 → x2 = y2 → f x1 x2 = f y1 y2.
Proof. intros A1 A2 B F x1 y1 x2 y2 eq1 eq2. rewrite eq1. rewrite eq2. reflexivity. Qed.
(x1 y1:A1) (x2 y2:A2),
x1 = y1 → x2 = y2 → f x1 x2 = f y1 y2.
Proof. intros A1 A2 B F x1 y1 x2 y2 eq1 eq2. rewrite eq1. rewrite eq2. reflexivity. Qed.
Using Tactics on Hypotheses
Theorem S_inj : ∀ (n m : nat) (b : bool),
beq_nat (S n) (S m) = b →
beq_nat n m = b.
Proof.
intros n m b H. simpl in H. apply H. Qed.
beq_nat (S n) (S m) = b →
beq_nat n m = b.
Proof.
intros n m b H. simpl in H. apply H. Qed.
Similarly, apply L in H matches some conditional statement
L (of the form L1 → L2, say) against a hypothesis H in the
context. However, unlike ordinary apply (which rewrites a goal
matching L2 into a subgoal L1), apply L in H matches H
against L1 and, if successful, replaces it with L2.
In other words, apply L in H gives us a form of "forward
reasoning": from L1 → L2 and a hypothesis matching L1, it
produces a hypothesis matching L2. By contrast, apply L is
"backward reasoning": it says that if we know L1→L2 and we are
trying to prove L2, it suffices to prove L1.
Here is a variant of a proof from above, using forward reasoning
throughout instead of backward reasoning.
Theorem silly3' : ∀ (n : nat),
(beq_nat n 5 = true → beq_nat (S (S n)) 7 = true) →
true = beq_nat n 5 →
true = beq_nat (S (S n)) 7.
Proof.
intros n eq H.
symmetry in H. apply eq in H. symmetry in H.
apply H. Qed.
(beq_nat n 5 = true → beq_nat (S (S n)) 7 = true) →
true = beq_nat n 5 →
true = beq_nat (S (S n)) 7.
Proof.
intros n eq H.
symmetry in H. apply eq in H. symmetry in H.
apply H. Qed.
Forward reasoning starts from what is given (premises,
previously proven theorems) and iteratively draws conclusions from
them until the goal is reached. Backward reasoning starts from
the goal, and iteratively reasons about what would imply the
goal, until premises or previously proven theorems are reached.
If you've seen informal proofs before (for example, in a math
class), they probably used forward reasoning. They might have even
told you that backward reasoning was wrong or not allowed! In
general, idiomatic use of Coq tends to favor backward reasoning,
but in some situations the forward style can be easier to think
about.
Exercise: 3 stars, recommended (plus_n_n_injective)
Practice using "in" variants in this proof. (Hint: use plus_n_Sm.)
Theorem plus_n_n_injective : ∀ n m,
n + n = m + m →
n = m.
Proof.
intros n. induction n as [| n'].
(* FILL IN HERE *) Admitted.
☐
n + n = m + m →
n = m.
Proof.
intros n. induction n as [| n'].
(* FILL IN HERE *) Admitted.
Varying the Induction Hypothesis
Theorem double_injective: ∀ n m,
double n = double m → n = m.
The way we start this proof is a bit delicate: if we begin with
double n = double m → n = m.
intros n. induction n.
all is well. But if we begin it with
intros n m. induction n.
we get stuck in the middle of the inductive case...
Theorem double_injective_FAILED : ∀ n m,
double n = double m →
n = m.
Proof.
intros n m. induction n as [| n'].
- (* n = O *) simpl. intros eq. destruct m as [| m'].
+ (* m = O *) reflexivity.
+ (* m = S m' *) inversion eq.
- (* n = S n' *) intros eq. destruct m as [| m'].
+ (* m = O *) inversion eq.
+ (* m = S m' *) apply f_equal.
double n = double m →
n = m.
Proof.
intros n m. induction n as [| n'].
- (* n = O *) simpl. intros eq. destruct m as [| m'].
+ (* m = O *) reflexivity.
+ (* m = S m' *) inversion eq.
- (* n = S n' *) intros eq. destruct m as [| m'].
+ (* m = O *) inversion eq.
+ (* m = S m' *) apply f_equal.
At this point, the induction hypothesis, IHn', does not give us
n' = m' — there is an extra S in the way — so the goal is
not provable.
Abort.
What went wrong?
The problem is that, at the point we invoke the induction
hypothesis, we have already introduced m into the context —
intuitively, we have told Coq, "Let's consider some particular n
and m..." and we now have to prove that, if double n = double
m for these particular n and m, then n = m.
The next tactic, induction n says to Coq: We are going to show
the goal by induction on n. That is, we are going to prove, for
all n, that the proposition
holds, by showing
If we look closely at the second statement, it is saying something
rather strange: it says that, for a particular m, if we know
then we can prove
To see why this is strange, let's think of a particular m —
say, 5. The statement is then saying that, if we know
then we can prove
But knowing Q doesn't give us any help at all with proving
R! (If we tried to prove R from Q, we would start with
something like "Suppose double (S n) = 10..." but then we'd be
stuck: knowing that double (S n) is 10 tells us nothing about
whether double n is 10, so Q is useless.)
Trying to carry out this proof by induction on n when m is
already in the context doesn't work because we are then trying to
prove a relation involving every n but just a single m.
The successful proof of double_injective leaves m in the goal
statement at the point where the induction tactic is invoked on
n:
- P n = "if double n = double m, then n = m"
- P O
- P n → P (S n)
- "if double n = double m then n = m"
- "if double (S n) = double m then S n = m".
- Q = "if double n = 10 then n = 5"
- R = "if double (S n) = 10 then S n = 5".
Theorem double_injective : ∀ n m,
double n = double m →
n = m.
Proof.
intros n. induction n as [| n'].
- (* n = O *) simpl. intros m eq. destruct m as [| m'].
+ (* m = O *) reflexivity.
+ (* m = S m' *) inversion eq.
- (* n = S n' *) simpl.
double n = double m →
n = m.
Proof.
intros n. induction n as [| n'].
- (* n = O *) simpl. intros m eq. destruct m as [| m'].
+ (* m = O *) reflexivity.
+ (* m = S m' *) inversion eq.
- (* n = S n' *) simpl.
Notice that both the goal and the induction hypothesis are
different this time: the goal asks us to prove something more
general (i.e., to prove the statement for every m), but the IH
is correspondingly more flexible, allowing us to choose any m we
like when we apply the IH.
intros m eq.
Now we've chosen a particular m and introduced the assumption
that double n = double m. Since we are doing a case analysis on
n, we also need a case analysis on m to keep the two "in sync."
destruct m as [| m'].
+ (* m = O *) simpl.
+ (* m = O *) simpl.
The 0 case is trivial:
At this point, since we are in the second branch of the destruct
m, the m' mentioned in the context is the predecessor of the
m we started out talking about. Since we are also in the S
branch of the induction, this is perfect: if we instantiate the
generic m in the IH with the current m' (this instantiation is
performed automatically by the apply in the next step), then
IHn' gives us exactly what we need to finish the proof.
apply IHn'. inversion eq. reflexivity. Qed.
What you should take away from all this is that we need to be
careful about using induction to try to prove something too
specific: To prove a property of n and m by induction on n,
it is sometimes important to leave m generic.
In an informal proof, you should simply be careful about what is
given and what isn't. If you're deliberately delaying
introducing a variable, it's good to be explicit. For example, in
an informal analogue of the above, we might say, "by induction on
n, leaving m general".
The following exercise requires the same pattern.
☐
Exercise: 2 stars (beq_nat_true)
Exercise: 2 stars, advanced (beq_nat_true_informal)
Give a careful informal proof of beq_nat_true, being as explicit as possible about quantifiers. If you're not sure about which language to use, there's an example below.
(* FILL IN HERE *)
☐
Theorem double_injective_take2_FAILED : ∀ n m,
double n = double m →
n = m.
Proof.
intros n m. induction m as [| m'].
- (* m = O *) simpl. intros eq. destruct n as [| n'].
+ (* n = O *) reflexivity.
+ (* n = S n' *) inversion eq.
- (* m = S m' *) intros eq. destruct n as [| n'].
+ (* n = O *) inversion eq.
+ (* n = S n' *) apply f_equal.
(* Stuck again here, just like before. *)
Abort.
double n = double m →
n = m.
Proof.
intros n m. induction m as [| m'].
- (* m = O *) simpl. intros eq. destruct n as [| n'].
+ (* n = O *) reflexivity.
+ (* n = S n' *) inversion eq.
- (* m = S m' *) intros eq. destruct n as [| n'].
+ (* n = O *) inversion eq.
+ (* n = S n' *) apply f_equal.
(* Stuck again here, just like before. *)
Abort.
The problem is that, to do induction on m, we must first
introduce n. (If we simply say induction m without
introducing anything first, Coq will automatically introduce n
for us!)
What can we do about this? One possibility is to rewrite the
statement of the lemma so that m is quantified before n. This
works, but it's not nice: We don't want to have to twist the
statements of lemmas to fit the needs of a particular strategy for
proving them! Rather we want to state them in the clearest and
most natural way.
What we can do instead is to first introduce all the quantified
variables and then re-generalize one or more of them,
selectively taking variables out of the context and putting them
back at the beginning of the goal. The generalize dependent
tactic does this.
Theorem double_injective_take2 : ∀ n m,
double n = double m →
n = m.
Proof.
intros n m.
(* n and m are both in the context *)
generalize dependent n.
(* Now n is back in the goal and we can do induction on
m and get a sufficiently general IH. *)
induction m as [| m'].
- (* m = O *) simpl. intros n eq. destruct n as [| n'].
+ (* n = O *) reflexivity.
+ (* n = S n' *) inversion eq.
- (* m = S m' *) intros n eq. destruct n as [| n'].
+ (* n = O *) inversion eq.
+ (* n = S n' *) apply f_equal.
apply IHm'. inversion eq. reflexivity. Qed.
double n = double m →
n = m.
Proof.
intros n m.
(* n and m are both in the context *)
generalize dependent n.
(* Now n is back in the goal and we can do induction on
m and get a sufficiently general IH. *)
induction m as [| m'].
- (* m = O *) simpl. intros n eq. destruct n as [| n'].
+ (* n = O *) reflexivity.
+ (* n = S n' *) inversion eq.
- (* m = S m' *) intros n eq. destruct n as [| n'].
+ (* n = O *) inversion eq.
+ (* n = S n' *) apply f_equal.
apply IHm'. inversion eq. reflexivity. Qed.
Let's look at an informal proof of this theorem. Note that
the proposition we prove by induction leaves n quantified,
corresponding to the use of generalize dependent in our formal
proof.
Theorem: For any nats n and m, if double n = double m, then
n = m.
Proof: Let m be a nat. We prove by induction on m that, for
any n, if double n = double m then n = m.
- First, suppose m = 0, and suppose n is a number such
that double n = double m. We must show that n = 0.
- Second, suppose m = S m' and that n is again a number such
that double n = double m. We must show that n = S m', with
the induction hypothesis that for every number s, if double s =
double m' then s = m'.
Exercise: 3 stars, recommended (gen_dep_practice)
Prove this by induction on l.
Theorem nth_error_after_last: ∀ (n : nat) (X : Type) (l : list X),
length l = n →
nth_error l n = None.
Proof.
(* FILL IN HERE *) Admitted.
☐
length l = n →
nth_error l n = None.
Proof.
(* FILL IN HERE *) Admitted.
Unfolding Definitions
... and try to prove a simple fact about square...
... we get stuck: simpl doesn't simplify anything at this point,
and since we haven't proved any other facts about square, there
is nothing we can apply or rewrite with.
To make progress, we can manually unfold the definition of
square:
Now we have plenty to work with: both sides of the equality are
expressions involving multiplication, and we have lots of facts
about multiplication at our disposal. In particular, we know that
it is commutative and associative, and from these facts it is not
hard to finish the proof.
rewrite mult_assoc.
assert (H : n * m * n = n * n * m).
{ rewrite mult_comm. apply mult_assoc. }
rewrite H. rewrite mult_assoc. reflexivity.
Qed.
assert (H : n * m * n = n * n * m).
{ rewrite mult_comm. apply mult_assoc. }
rewrite H. rewrite mult_assoc. reflexivity.
Qed.
At this point, a deeper discussion of unfolding and simplification
is in order.
You may already have observed that tactics like simpl,
reflexivity, and apply will often unfold the definitions of
functions automatically when this allows them to make progress. For
example, if we define foo m to be the constant 5...
then the simpl in the following proof (or the reflexivity, if
we omit the simpl) will unfold foo m to (fun x ⇒ 5) m and
then further simplify this expression to just 5.
However, this automatic unfolding is rather conservative. For
example, if we define a slightly more complicated function
involving a pattern match...
...then the analogous proof will get stuck:
Fact silly_fact_2_FAILED : ∀ m, bar m + 1 = bar (m + 1) + 1.
Proof.
intros m.
simpl. (* Does nothing! *)
Abort.
Proof.
intros m.
simpl. (* Does nothing! *)
Abort.
The reason that simpl doesn't make progress here is that it
notices that, after tentatively unfolding bar m, it is left with
a match whose scrutinee, m, is a variable, so the match cannot
be simplified further. (It is not smart enough to notice that the
two branches of the match are identical.) So it gives up on
unfolding bar m and leaves it alone. Similarly, tentatively
unfolding bar (m+1) leaves a match whose scrutinee is a
function application (that, itself, cannot be simplified, even
after unfolding the definition of +), so simpl leaves it
alone.
At this point, there are two ways to make progress. One is to use
destruct m to break the proof into two cases, each focusing on a
more concrete choice of m (O vs S _). In each case, the
match inside of bar can now make progress, and the proof is
easy to complete.
Fact silly_fact_2 : ∀ m, bar m + 1 = bar (m + 1) + 1.
Proof.
intros m.
destruct m.
- simpl. reflexivity.
- simpl. reflexivity.
Qed.
Proof.
intros m.
destruct m.
- simpl. reflexivity.
- simpl. reflexivity.
Qed.
This approach works, but it depends on our recognizing that the
match hidden inside bar is what was preventing us from making
progress.
A more straightforward way to make progress is to explicitly tell
Coq to unfold bar.
Now it is apparent that we are stuck on the match expressions on
both sides of the =, and we can use destruct to finish the
proof without thinking too hard.
destruct m.
- reflexivity.
- reflexivity.
Qed.
- reflexivity.
- reflexivity.
Qed.
Using destruct on Compound Expressions
Definition sillyfun (n : nat) : bool :=
if beq_nat n 3 then false
else if beq_nat n 5 then false
else false.
Theorem sillyfun_false : ∀ (n : nat),
sillyfun n = false.
Proof.
intros n. unfold sillyfun.
destruct (beq_nat n 3).
- (* beq_nat n 3 = true *) reflexivity.
- (* beq_nat n 3 = false *) destruct (beq_nat n 5).
+ (* beq_nat n 5 = true *) reflexivity.
+ (* beq_nat n 5 = false *) reflexivity. Qed.
if beq_nat n 3 then false
else if beq_nat n 5 then false
else false.
Theorem sillyfun_false : ∀ (n : nat),
sillyfun n = false.
Proof.
intros n. unfold sillyfun.
destruct (beq_nat n 3).
- (* beq_nat n 3 = true *) reflexivity.
- (* beq_nat n 3 = false *) destruct (beq_nat n 5).
+ (* beq_nat n 5 = true *) reflexivity.
+ (* beq_nat n 5 = false *) reflexivity. Qed.
After unfolding sillyfun in the above proof, we find that
we are stuck on if (beq_nat n 3) then ... else .... But either
n is equal to 3 or it isn't, so we can use destruct (beq_nat
n 3) to let us reason about the two cases.
In general, the destruct tactic can be used to perform case
analysis of the results of arbitrary computations. If e is an
expression whose type is some inductively defined type T, then,
for each constructor c of T, destruct e generates a subgoal
in which all occurrences of e (in the goal and in the context)
are replaced by c.
Exercise: 3 stars, optional (combine_split)
Here is an implementation of the split function mentioned in chapter Poly:
Fixpoint split {X Y : Type} (l : list (X*Y))
: (list X) * (list Y) :=
match l with
| [] ⇒ ([], [])
| (x, y) :: t ⇒
match split t with
| (lx, ly) ⇒ (x :: lx, y :: ly)
end
end.
: (list X) * (list Y) :=
match l with
| [] ⇒ ([], [])
| (x, y) :: t ⇒
match split t with
| (lx, ly) ⇒ (x :: lx, y :: ly)
end
end.
Prove that split and combine are inverses in the following
sense:
Theorem combine_split : ∀ X Y (l : list (X * Y)) l1 l2,
split l = (l1, l2) →
combine l1 l2 = l.
Proof.
(* FILL IN HERE *) Admitted.
☐
split l = (l1, l2) →
combine l1 l2 = l.
Proof.
(* FILL IN HERE *) Admitted.
Definition sillyfun1 (n : nat) : bool :=
if beq_nat n 3 then true
else if beq_nat n 5 then true
else false.
if beq_nat n 3 then true
else if beq_nat n 5 then true
else false.
Now suppose that we want to convince Coq of the (rather
obvious) fact that sillyfun1 n yields true only when n is
odd. By analogy with the proofs we did with sillyfun above, it
is natural to start the proof like this:
Theorem sillyfun1_odd_FAILED : ∀ (n : nat),
sillyfun1 n = true →
oddb n = true.
Proof.
intros n eq. unfold sillyfun1 in eq.
destruct (beq_nat n 3).
(* stuck... *)
Abort.
sillyfun1 n = true →
oddb n = true.
Proof.
intros n eq. unfold sillyfun1 in eq.
destruct (beq_nat n 3).
(* stuck... *)
Abort.
We get stuck at this point because the context does not
contain enough information to prove the goal! The problem is that
the substitution performed by destruct is too brutal — it threw
away every occurrence of beq_nat n 3, but we need to keep some
memory of this expression and how it was destructed, because we
need to be able to reason that, since beq_nat n 3 = true in this
branch of the case analysis, it must be that n = 3, from which
it follows that n is odd.
What we would really like is to substitute away all existing
occurences of beq_nat n 3, but at the same time add an equation
to the context that records which case we are in. The eqn:
qualifier allows us to introduce such an equation, giving it a
name that we choose.
Theorem sillyfun1_odd : ∀ (n : nat),
sillyfun1 n = true →
oddb n = true.
Proof.
intros n eq. unfold sillyfun1 in eq.
destruct (beq_nat n 3) eqn:Heqe3.
(* Now we have the same state as at the point where we got
stuck above, except that the context contains an extra
equality assumption, which is exactly what we need to
make progress. *)
- (* e3 = true *) apply beq_nat_true in Heqe3.
rewrite → Heqe3. reflexivity.
- (* e3 = false *)
(* When we come to the second equality test in the body
of the function we are reasoning about, we can use
eqn: again in the same way, allow us to finish the
proof. *)
destruct (beq_nat n 5) eqn:Heqe5.
+ (* e5 = true *)
apply beq_nat_true in Heqe5.
rewrite → Heqe5. reflexivity.
+ (* e5 = false *) inversion eq. Qed.
sillyfun1 n = true →
oddb n = true.
Proof.
intros n eq. unfold sillyfun1 in eq.
destruct (beq_nat n 3) eqn:Heqe3.
(* Now we have the same state as at the point where we got
stuck above, except that the context contains an extra
equality assumption, which is exactly what we need to
make progress. *)
- (* e3 = true *) apply beq_nat_true in Heqe3.
rewrite → Heqe3. reflexivity.
- (* e3 = false *)
(* When we come to the second equality test in the body
of the function we are reasoning about, we can use
eqn: again in the same way, allow us to finish the
proof. *)
destruct (beq_nat n 5) eqn:Heqe5.
+ (* e5 = true *)
apply beq_nat_true in Heqe5.
rewrite → Heqe5. reflexivity.
+ (* e5 = false *) inversion eq. Qed.
Theorem bool_fn_applied_thrice :
∀ (f : bool → bool) (b : bool),
f (f (f b)) = f b.
Proof.
(* FILL IN HERE *) Admitted.
☐
∀ (f : bool → bool) (b : bool),
f (f (f b)) = f b.
Proof.
(* FILL IN HERE *) Admitted.
Review
- intros: move hypotheses/variables from goal to context
- reflexivity: finish the proof (when the goal looks like e =
e)
- apply: prove goal using a hypothesis, lemma, or constructor
- apply... in H: apply a hypothesis, lemma, or constructor to
a hypothesis in the context (forward reasoning)
- apply... with...: explicitly specify values for variables
that cannot be determined by pattern matching
- simpl: simplify computations in the goal
- simpl in H: ... or a hypothesis
- rewrite: use an equality hypothesis (or lemma) to rewrite
the goal
- rewrite ... in H: ... or a hypothesis
- symmetry: changes a goal of the form t=u into u=t
- symmetry in H: changes a hypothesis of the form t=u into
u=t
- unfold: replace a defined constant by its right-hand side in
the goal
- unfold... in H: ... or a hypothesis
- destruct... as...: case analysis on values of inductively
defined types
- destruct... eqn:...: specify the name of an equation to be
added to the context, recording the result of the case
analysis
- induction... as...: induction on values of inductively
defined types
- inversion: reason by injectivity and distinctness of
constructors
- assert (H: e) (or assert (e) as H): introduce a "local
lemma" e and call it H
- generalize dependent x: move the variable x (and anything else that depends on it) from the context back to an explicit hypothesis in the goal formula
Exercise: 3 stars, advanced (beq_nat_sym_informal)
Give an informal proof of this lemma that corresponds to your formal proof above:☐
Exercise: 3 stars, optional (beq_nat_trans)
Theorem beq_nat_trans : ∀ n m p,
beq_nat n m = true →
beq_nat m p = true →
beq_nat n p = true.
Proof.
(* FILL IN HERE *) Admitted.
☐
beq_nat n m = true →
beq_nat m p = true →
beq_nat n p = true.
Proof.
(* FILL IN HERE *) Admitted.
Exercise: 3 stars, advanced (split_combine)
We proved, in an exercise above, that for all lists of pairs, combine is the inverse of split. How would you formalize the statement that split is the inverse of combine? When is this property true?
Definition split_combine_statement : Prop
(* (": Prop" means that we are giving a name to a
logical proposition here.) *)
(* REPLACE THIS LINE WITH ":= _your_definition_ ." *). Admitted.
Theorem split_combine : split_combine_statement.
Proof.
(* FILL IN HERE *) Admitted.
☐
(* (": Prop" means that we are giving a name to a
logical proposition here.) *)
(* REPLACE THIS LINE WITH ":= _your_definition_ ." *). Admitted.
Theorem split_combine : split_combine_statement.
Proof.
(* FILL IN HERE *) Admitted.
Exercise: 3 stars, advanced (filter_exercise)
This one is a bit challenging. Pay attention to the form of your induction hypothesis.
Theorem filter_exercise : ∀ (X : Type) (test : X → bool)
(x : X) (l lf : list X),
filter test l = x :: lf →
test x = true.
Proof.
(* FILL IN HERE *) Admitted.
☐
(x : X) (l lf : list X),
filter test l = x :: lf →
test x = true.
Proof.
(* FILL IN HERE *) Admitted.
Exercise: 4 stars, advanced, optional (forall_exists_challenge)
Define two recursive Fixpoints, forallb and existsb. The first checks whether every element in a list satisfies a given predicate:
forallb oddb [1;3;5;7;9] = true
forallb negb [false;false] = true
forallb evenb [0;2;4;5] = false
forallb (beq_nat 5) [] = true
The second checks whether there exists an element in the list that
satisfies a given predicate:
forallb negb [false;false] = true
forallb evenb [0;2;4;5] = false
forallb (beq_nat 5) [] = true
existsb (beq_nat 5) [0;2;3;6] = false
existsb (andb true) [true;true;false] = true
existsb oddb [1;0;0;0;0;3] = true
existsb evenb [] = false
Next, define a nonrecursive version of existsb — call it
existsb' — using forallb and negb.
existsb (andb true) [true;true;false] = true
existsb oddb [1;0;0;0;0;3] = true
existsb evenb [] = false
(* FILL IN HERE *)
☐