Built with Alectryon, running Coq+SerAPI v8.10.0+0.7.0. Coq sources are in this panel; goals and messages will appear in the other. Bubbles () indicate interactive fragments: hover for details, tap to reveal contents. Use Ctrl+↑ Ctrl+↓ to navigate, Ctrl+🖱️ to focus.
This chapter introduces several additional proof strategies
and tactics that allow us to begin proving more interesting
properties of functional programs. We will see:
- 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.
Set Warnings "-notation-overridden,-parsing". From LF Require Export Poly. (* ################################################################# *)
We often encounter situations where the goal to be proved is
exactly the same as some hypothesis in the context or some
previously proved lemma.
forall n m o p : nat, n = m -> [n; o] = [n; p] -> [n; o] = [m; p]forall n m o p : nat, n = m -> [n; o] = [n; p] -> [n; o] = [m; p]n, m, o, p:nateq1:n = meq2:[n; o] = [n; p][n; o] = [m; p]n, m, o, p:nateq1:n = meq2:[n; o] = [n; p][n; o] = [n; p]
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.
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.
forall n m o p : nat, n = m -> (forall q r : nat, q = r -> [q; o] = [r; p]) -> [n; o] = [m; p]forall n m o p : nat, n = m -> (forall q r : nat, q = r -> [q; o] = [r; p]) -> [n; o] = [m; p]n, m, o, p:nateq1:n = meq2:forall q r : nat, q = r -> [q; o] = [r; p][n; o] = [m; p]apply eq1. Qed.n, m, o, p:nateq1:n = meq2:forall q r : nat, q = r -> [q; o] = [r; p]n = m
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.
forall n m : nat, (n, n) = (m, m) -> (forall q r : nat, (q, q) = (r, r) -> [q] = [r]) -> [n] = [m]forall n m : nat, (n, n) = (m, m) -> (forall q r : nat, (q, q) = (r, r) -> [q] = [r]) -> [n] = [m]n, m:nateq1:(n, n) = (m, m)eq2:forall q r : nat, (q, q) = (r, r) -> [q] = [r][n] = [m]apply eq1. Qed.n, m:nateq1:(n, n) = (m, m)eq2:forall q r : nat, (q, q) = (r, r) -> [q] = [r](n, n) = (m, m)
(forall n : nat, evenb n = true -> oddb (S n) = true) -> oddb 3 = true -> evenb 4 = true(* FILL IN HERE *) Admitted.(forall n : nat, evenb n = true -> oddb (S n) = true) -> oddb 3 = true -> evenb 4 = true
☐
To use the apply tactic, the (conclusion of the) fact
being applied must match the goal exactly -- for example, apply
will not work if the left and right sides of the equality are
swapped.
forall n : nat, true = (n =? 5) -> (S (S n) =? 7) = trueforall n : nat, true = (n =? 5) -> (S (S n) =? 7) = truen:natH:true = (n =? 5)(S (S n) =? 7) = true
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.
n:natH:true = (n =? 5)true = (S (S n) =? 7)n:natH:true = (n =? 5)true = (n =? 5)
(This simpl is optional, since apply will perform
simplification first, if needed.)
apply H. Qed.
Exercise: 3 stars, standard (apply_exercise1)
forall l l' : list nat, l = rev l' -> l' = rev l(* FILL IN HERE *) Admitted.forall l l' : list nat, l = rev l' -> l' = rev l
☐
Exercise: 1 star, standard, optional (apply_rewrite)
(* FILL IN HERE
[] *)
(* ################################################################# *)
The following silly example uses two rewrites in a row to
get from [a;b] to [e;f].
forall a b c d e f : nat, [a; b] = [c; d] -> [c; d] = [e; f] -> [a; b] = [e; f]forall a b c d e f : nat, [a; b] = [c; d] -> [c; d] = [e; f] -> [a; b] = [e; f]a, b, c, d, e, f:nateq1:[a; b] = [c; d]eq2:[c; d] = [e; f][a; b] = [e; f]a, b, c, d, e, f:nateq1:[a; b] = [c; d]eq2:[c; d] = [e; f][c; d] = [e; f]reflexivity. Qed.a, b, c, d, e, f:nateq1:[a; b] = [c; d]eq2:[c; d] = [e; f][e; f] = [e; f]
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.
forall (X : Type) (n m o : X), n = m -> m = o -> n = oforall (X : Type) (n m o : X), n = m -> m = o -> n = oX:Typen, m, o:Xeq1:n = meq2:m = on = oX:Typen, m, o:Xeq1:n = meq2:m = om = oreflexivity. Qed.X:Typen, m, o:Xeq1:n = meq2:m = oo = o
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.
forall a b c d e f : nat, [a; b] = [c; d] -> [c; d] = [e; f] -> [a; b] = [e; f]forall a b c d e f : nat, [a; b] = [c; d] -> [c; d] = [e; f] -> [a; b] = [e; f]a, b, c, d, e, f:nateq1:[a; b] = [c; d]eq2:[c; d] = [e; f][a; b] = [e; f]
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.
a, b, c, d, e, f:nateq1:[a; b] = [c; d]eq2:[c; d] = [e; f][a; b] = [c; d]a, b, c, d, e, f:nateq1:[a; b] = [c; d]eq2:[c; d] = [e; f][c; d] = [e; f]apply eq2. Qed.a, b, c, d, e, f:nateq1:[a; b] = [c; d]eq2:[c; d] = [e; f][c; d] = [e; f]
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].
forall n m o p : nat, m = minustwo o -> n + p = m -> n + p = minustwo o(* FILL IN HERE *) Admitted.forall n m o p : nat, m = minustwo o -> n + p = m -> n + p = minustwo o
☐
(* ################################################################# *)
Recall the definition of natural numbers:
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:
Similar principles apply to all inductively defined types: all
constructors are injective, and the values built from distinct
constructors are never equal. For lists, the cons constructor
is injective and nil is different from every non-empty list.
For booleans, true and false are different. (Since neither
true nor false take any arguments, their injectivity is not
interesting.) And so on.
- 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.
For example, we can prove the injectivity of S by using the
pred function defined in Basics.v.
forall n m : nat, S n = S m -> n = mforall n m : nat, S n = S m -> n = mn, m:natH1:S n = S mn = mn, m:natH1:S n = S mn = Nat.pred (S n)n, m:natH1:S n = S mH2:n = Nat.pred (S n)n = mreflexivity.n, m:natH1:S n = S mn = Nat.pred (S n)n, m:natH1:S n = S mH2:n = Nat.pred (S n)n = mn, m:natH1:S n = S mH2:n = Nat.pred (S n)Nat.pred (S n) = mreflexivity. Qed.n, m:natH1:S n = S mH2:n = Nat.pred (S n)Nat.pred (S m) = m
This technique can be generalized to any constructor by
writing the equivalent of pred for that constructor -- i.e.,
writing a function that "undoes" one application of the
constructor. As a more convenient alternative, Coq provides a
tactic called injection that allows us to exploit the
injectivity of any constructor. Here is an alternate proof of the
above theorem using injection:
forall n m : nat, S n = S m -> n = mforall n m : nat, S n = S m -> n = mn, m:natH:S n = S mn = m
By writing injection H at this point, we are asking Coq to
generate all equations that it can infer from H using the
injectivity of constructors. Each such equation is added as a
premise to the goal. In the present example, adds the premise
n = m.
n, m:natH:S n = S mn = m -> n = mapply Hnm. Qed.n, m:natH:S n = S mHnm:n = mn = m
Here's a more interesting example that shows how injection can
derive multiple equations at once.
forall n m o : nat, [n; m] = [o; o] -> [n] = [m]forall n m o : nat, [n; m] = [o; o] -> [n] = [m]n, m, o:natH:[n; m] = [o; o][n] = [m]n, m, o:natH:[n; m] = [o; o]m = o -> n = o -> [n] = [m]n, m, o:natH:[n; m] = [o; o]H1:m = oH2:n = o[n] = [m]n, m, o:natH:[n; m] = [o; o]H1:m = oH2:n = o[n] = [o]reflexivity. Qed.n, m, o:natH:[n; m] = [o; o]H1:m = oH2:n = o[o] = [o]
The "as" variant of injection permits us to choose names for
the introduced equations rather than letting Coq do it.
forall n m : nat, [n] = [m] -> n = mforall n m : nat, [n] = [m] -> n = mn, m:natH:[n] = [m]n = mn, m:natHnm:n = mn = mreflexivity. Qed.n, m:natHnm:n = mm = m
forall (X : Type) (x y z : X) (l j : list X), x :: y :: l = z :: j -> y :: l = x :: j -> x = y(* FILL IN HERE *) Admitted.forall (X : Type) (x y z : X) (l j : list X), x :: y :: l = z :: j -> y :: l = x :: j -> x = y
☐
So much for injectivity of constructors. What about disjointness?
The principle of disjointness says that two terms beginning with
different constructors (like O and S, or true and false)
can never be equal. This means that, any time we find ourselves
working in a context where we've assumed that two such terms are
equal, we are justified in concluding anything we want to (because
the assumption is nonsensical).
The discriminate tactic embodies this principle: It is used on a
hypothesis involving an equality between different
constructors (e.g., S n = O), and it solves the current goal
immediately. For example:
forall n : nat, (0 =? n) = true -> n = 0forall n : nat, (0 =? n) = true -> n = 0n:nat(0 =? n) = true -> n = 0
We can proceed by case analysis on n. The first case is
trivial.
n:natE:n = 0(0 =? 0) = true -> 0 = 0n, n':natE:n = S n'(0 =? S n') = true -> S n' = 0n:natE:n = 0(0 =? 0) = true -> 0 = 0reflexivity.n:natE:n = 0H:(0 =? 0) = true0 = 0
However, the second one doesn't look so simple: assuming 0
=? (S n') = true, we must show S n' = 0! The way forward is to
observe that the assumption itself is nonsensical:
n, n':natE:n = S n'(0 =? S n') = true -> S n' = 0n, n':natE:n = S n'false = true -> S n' = 0
If we use discriminate on this hypothesis, Coq confirms
that the subgoal we are working on is impossible and removes it
from further consideration.
discriminate H. Qed.n, n':natE:n = S n'H:false = trueS n' = 0
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!
forall n : nat, S n = 0 -> 2 + 2 = 5forall n : nat, S n = 0 -> 2 + 2 = 5discriminate contra. Qed.n:natcontra:S n = 02 + 2 = 5forall n m : nat, false = true -> [n] = [m]forall n m : nat, false = true -> [n] = [m]discriminate contra. Qed.n, m:natcontra:false = true[n] = [m]
If you find the principle of explosion confusing, remember
that these proofs are not showing that the conclusion of the
statement holds. Rather, they are showing 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.
forall (X : Type) (x y z : X) (l : list X), list X -> x :: y :: l = [ ] -> x = z(* FILL IN HERE *) Admitted.forall (X : Type) (x y z : X) (l : list X), list X -> x :: y :: l = [ ] -> x = z
☐
The injectivity of constructors allows us to reason that
∀ (n m : nat), S n = S m → n = m. The converse of this
implication is an instance of a more general fact about both
constructors and functions, which we will find convenient in a few
places below:
forall (A B : Type) (f : A -> B) (x y : A), x = y -> f x = f yforall (A B : Type) (f : A -> B) (x y : A), x = y -> f x = f yA, B:Typef:A -> Bx, y:Aeq:x = yf x = f yreflexivity. Qed. (* ################################################################# *)A, B:Typef:A -> Bx, y:Aeq:x = yf y = f y
By default, most tactics work on the goal formula and leave
the context unchanged. However, most tactics also have a variant
that performs a similar operation on a statement in the context.
For example, the tactic simpl in H performs simplification in
the hypothesis named H in the context.
forall (n m : nat) (b : bool), (S n =? S m) = b -> (n =? m) = bforall (n m : nat) (b : bool), (S n =? S m) = b -> (n =? m) = bn, m:natb:boolH:(S n =? S m) = b(n =? m) = bapply H. Qed.n, m:natb:boolH:(n =? m) = b(n =? m) = b
Similarly, apply L in H matches some conditional statement
L (of the form X → Y, say) against a hypothesis H in the
context. However, unlike ordinary apply (which rewrites a goal
matching Y into a subgoal X), apply L in H matches H
against X and, if successful, replaces it with Y.
In other words, apply L in H gives us a form of "forward
reasoning": from X → Y and a hypothesis matching X, it
produces a hypothesis matching X. By contrast, apply L is
"backward reasoning": it says that if we know X → Y and we
are trying to prove Y, it suffices to prove X.
Here is a variant of a proof from above, using forward reasoning
throughout instead of backward reasoning.
forall n : nat, ((n =? 5) = true -> (S (S n) =? 7) = true) -> true = (n =? 5) -> true = (S (S n) =? 7)forall n : nat, ((n =? 5) = true -> (S (S n) =? 7) = true) -> true = (n =? 5) -> true = (S (S n) =? 7)n:nateq:(n =? 5) = true -> (S (S n) =? 7) = trueH:true = (n =? 5)true = (S (S n) =? 7)n:nateq:(n =? 5) = true -> (S (S n) =? 7) = trueH:(n =? 5) = truetrue = (S (S n) =? 7)n:nateq:(n =? 5) = true -> (S (S n) =? 7) = trueH:(S (S n) =? 7) = truetrue = (S (S n) =? 7)apply H. Qed.n:nateq:(n =? 5) = true -> (S (S n) =? 7) = trueH:true = (S (S n) =? 7)true = (S (S n) =? 7)
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 or
computer science class), they probably used forward reasoning. 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, standard, recommended (plus_n_n_injective)
forall n m : nat, n + n = m + m -> n = mforall n m : nat, n + n = m + m -> n = mn:natforall m : nat, n + n = m + m -> n = m(* FILL IN HERE *) Admitted.forall m : nat, 0 + 0 = m + m -> 0 = mn':natIHn':forall m : nat, n' + n' = m + m -> n' = mforall m : nat, S n' + S n' = m + m -> S n' = m
☐
(* ################################################################# *)
Sometimes it is important to control the exact form of the
induction hypothesis when carrying out inductive proofs in Coq.
In particular, we need to be careful about which of the
assumptions we move (using intros) from the goal to the context
before invoking the induction tactic. For example, suppose
we want to show that double is injective -- i.e., that it maps
different arguments to different results:
Theorem double_injective: forall n m,
double n = double m -> n = m.
The way we start this proof is a bit delicate: if we begin with
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...
forall n m : nat, double n = double m -> n = mforall n m : nat, double n = double m -> n = mn, m:natdouble n = double m -> n = mm:natdouble 0 = double m -> 0 = mn', m:natIHn':double n' = double m -> n' = mdouble (S n') = double m -> S n' = mm:natdouble 0 = double m -> 0 = mm:nat0 = double m -> 0 = mm:nateq:0 = double m0 = mm:natE:m = 0eq:0 = double 00 = 0m, m':natE:m = S m'eq:0 = double (S m')0 = S m'reflexivity.m:natE:m = 0eq:0 = double 00 = 0discriminate eq.m, m':natE:m = S m'eq:0 = double (S m')0 = S m'n', m:natIHn':double n' = double m -> n' = mdouble (S n') = double m -> S n' = mn', m:natIHn':double n' = double m -> n' = meq:double (S n') = double mS n' = mn', m:natE:m = 0IHn':double n' = double 0 -> n' = 0eq:double (S n') = double 0S n' = 0n', m, m':natE:m = S m'IHn':double n' = double (S m') -> n' = S m'eq:double (S n') = double (S m')S n' = S m'discriminate eq.n', m:natE:m = 0IHn':double n' = double 0 -> n' = 0eq:double (S n') = double 0S n' = 0n', m, m':natE:m = S m'IHn':double n' = double (S m') -> n' = S m'eq:double (S n') = double (S m')S n' = S m'n', m, m':natE:m = S m'IHn':double n' = double (S m') -> n' = S m'eq:double (S n') = double (S m')n' = m'
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.)
- 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".
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 statement 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:
forall n m : nat, double n = double m -> n = mforall n m : nat, double n = double m -> n = mn:natforall m : nat, double n = double m -> n = mforall m : nat, double 0 = double m -> 0 = mn':natIHn':forall m : nat, double n' = double m -> n' = mforall m : nat, double (S n') = double m -> S n' = mforall m : nat, double 0 = double m -> 0 = mforall m : nat, 0 = double m -> 0 = mm:nateq:0 = double m0 = mm:natE:m = 0eq:0 = double 00 = 0m, m':natE:m = S m'eq:0 = double (S m')0 = S m'reflexivity.m:natE:m = 0eq:0 = double 00 = 0discriminate eq.m, m':natE:m = S m'eq:0 = double (S m')0 = S m'n':natIHn':forall m : nat, double n' = double m -> n' = mforall m : nat, double (S n') = double m -> S n' = mn':natIHn':forall m : nat, double n' = double m -> n' = mforall m : nat, S (S (double n')) = double m -> S n' = m
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.
n':natIHn':forall m0 : nat, double n' = double m0 -> n' = m0m:nateq:S (S (double n')) = double mS n' = m
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."
n':natIHn':forall m0 : nat, double n' = double m0 -> n' = m0m:natE:m = 0eq:S (S (double n')) = double 0S n' = 0n':natIHn':forall m0 : nat, double n' = double m0 -> n' = m0m, m':natE:m = S m'eq:S (S (double n')) = double (S m')S n' = S m'n':natIHn':forall m0 : nat, double n' = double m0 -> n' = m0m:natE:m = 0eq:S (S (double n')) = double 0S n' = 0n':natIHn':forall m0 : nat, double n' = double m0 -> n' = m0m:natE:m = 0eq:S (S (double n')) = double 0S n' = 0
The 0 case is trivial:
discriminate eq.n':natIHn':forall m0 : nat, double n' = double m0 -> n' = m0m, m':natE:m = S m'eq:S (S (double n')) = double (S m')S n' = S m'n':natIHn':forall m0 : nat, double n' = double m0 -> n' = m0m, m':natE:m = S m'eq:S (S (double n')) = double (S m')n' = m'
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.
n':natIHn':forall m0 : nat, double n' = double m0 -> n' = m0m, m':natE:m = S m'eq:S (S (double n')) = double (S m')double n' = double m'apply goal. Qed.n':natIHn':forall m0 : nat, double n' = double m0 -> n' = m0m, m':natE:m = S m'goal:double n' = double m'double n' = double m'
What you should take away from all this is that we need to be
careful, when using induction, that we are not trying 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.
The following exercise requires the same pattern.
forall n m : nat, (n =? m) = true -> n = m(* FILL IN HERE *) Admitted.forall n m : nat, (n =? m) = true -> n = m
☐
Exercise: 2 stars, advanced (eqb_true_informal)
(* FILL IN HERE *) (* Do not modify the following line: *) Definition manual_grade_for_informal_proof : option (nat*string) := None.
☐
The strategy of doing fewer intros before an induction to
obtain a more general IH doesn't always work by itself; sometimes
some rearrangement of quantified variables is needed. Suppose,
for example, that we wanted to prove double_injective by
induction on m instead of n.
forall n m : nat, double n = double m -> n = mforall n m : nat, double n = double m -> n = mn, m:natdouble n = double m -> n = mn:natdouble n = double 0 -> n = 0n, m':natIHm':double n = double m' -> n = m'double n = double (S m') -> n = S m'n:natdouble n = double 0 -> n = 0n:natdouble n = 0 -> n = 0n:nateq:double n = 0n = 0n:natE:n = 0eq:double 0 = 00 = 0n, n':natE:n = S n'eq:double (S n') = 0S n' = 0reflexivity.n:natE:n = 0eq:double 0 = 00 = 0discriminate eq.n, n':natE:n = S n'eq:double (S n') = 0S n' = 0n, m':natIHm':double n = double m' -> n = m'double n = double (S m') -> n = S m'n, m':natIHm':double n = double m' -> n = m'eq:double n = double (S m')n = S m'n, m':natE:n = 0IHm':double 0 = double m' -> 0 = m'eq:double 0 = double (S m')0 = S m'n, m', n':natE:n = S n'IHm':double (S n') = double m' -> S n' = m'eq:double (S n') = double (S m')S n' = S m'discriminate eq.n, m':natE:n = 0IHm':double 0 = double m' -> 0 = m'eq:double 0 = double (S m')0 = S m'n, m', n':natE:n = S n'IHm':double (S n') = double m' -> S n' = m'eq:double (S n') = double (S m')S n' = S m'(* Stuck again here, just like before. *) Abort.n, m', n':natE:n = S n'IHm':double (S n') = double m' -> S n' = m'eq:double (S n') = double (S m')n' = m'
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.
forall n m : nat, double n = double m -> n = mforall n m : nat, double n = double m -> n = m(* [n] and [m] are both in the context *)n, m:natdouble n = double m -> n = m(* Now [n] is back in the goal and we can do induction on [m] and get a sufficiently general IH. *)m:natforall n : nat, double n = double m -> n = mforall n : nat, double n = double 0 -> n = 0m':natIHm':forall n : nat, double n = double m' -> n = m'forall n : nat, double n = double (S m') -> n = S m'forall n : nat, double n = double 0 -> n = 0forall n : nat, double n = 0 -> n = 0n:nateq:double n = 0n = 0n:natE:n = 0eq:double 0 = 00 = 0n, n':natE:n = S n'eq:double (S n') = 0S n' = 0reflexivity.n:natE:n = 0eq:double 0 = 00 = 0discriminate eq.n, n':natE:n = S n'eq:double (S n') = 0S n' = 0m':natIHm':forall n : nat, double n = double m' -> n = m'forall n : nat, double n = double (S m') -> n = S m'm':natIHm':forall n0 : nat, double n0 = double m' -> n0 = m'n:nateq:double n = double (S m')n = S m'm':natIHm':forall n0 : nat, double n0 = double m' -> n0 = m'n:natE:n = 0eq:double 0 = double (S m')0 = S m'm':natIHm':forall n0 : nat, double n0 = double m' -> n0 = m'n, n':natE:n = S n'eq:double (S n') = double (S m')S n' = S m'discriminate eq.m':natIHm':forall n0 : nat, double n0 = double m' -> n0 = m'n:natE:n = 0eq:double 0 = double (S m')0 = S m'm':natIHm':forall n0 : nat, double n0 = double m' -> n0 = m'n, n':natE:n = S n'eq:double (S n') = double (S m')S n' = S m'm':natIHm':forall n0 : nat, double n0 = double m' -> n0 = m'n, n':natE:n = S n'eq:double (S n') = double (S m')n' = m'm':natIHm':forall n0 : nat, double n0 = double m' -> n0 = m'n, n':natE:n = S n'eq:double (S n') = double (S m')double n' = double m'apply goal. Qed.m':natIHm':forall n0 : nat, double n0 = double m' -> n0 = m'n, n':natE:n = S n'goal:double n' = double m'double n' = double m'
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'.
Before we close this section and move on to some exercises,
let's digress briefly and use eqb_true to prove a similar
property of identifiers that we'll need in later chapters:
forall x y : id, eqb_id x y = true -> x = yforall x y : id, eqb_id x y = true -> x = ym, n:nateqb_id (Id m) (Id n) = true -> Id m = Id nm, n:nat(m =? n) = true -> Id m = Id nm, n:natH:(m =? n) = trueId m = Id nm, n:natH:(m =? n) = truem = nm, n:natH:(m =? n) = trueH':m = nId m = Id nm, n:natH:(m =? n) = truem = napply H.m, n:natH:(m =? n) = true(m =? n) = truem, n:natH:(m =? n) = trueH':m = nId m = Id nreflexivity. Qed.m, n:natH:(m =? n) = trueH':m = nId n = Id n
forall (n : nat) (X : Type) (l : list X), length l = n -> nth_error l n = None(* FILL IN HERE *) Admitted.forall (n : nat) (X : Type) (l : list X), length l = n -> nth_error l n = None
☐
(* ################################################################# *)
It sometimes happens that we need to manually unfold a name that
has been introduced by a Definition so that we can manipulate
its right-hand side. For example, if we define...
Definition square n := n * n.
... and try to prove a simple fact about square...
forall n m : nat, square (n * m) = square n * square mforall n m : nat, square (n * m) = square n * square mn, m:natsquare (n * m) = square n * square mn, m:natsquare (n * m) = square n * square m
... we appear to be 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:
n, m:natn * m * (n * m) = n * n * (m * m)
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 it is not hard
to finish the proof.
n, m:natn * m * n * m = n * n * (m * m)n, m:natn * m * n = n * n * mn, m:natH:n * m * n = n * n * mn * m * n * m = n * n * (m * m)n, m:natn * m * n = n * n * mapply mult_assoc.n, m:natn * (n * m) = n * n * mn, m:natH:n * m * n = n * n * mn * m * n * m = n * n * (m * m)n, m:natH:n * m * n = n * n * mn * n * m * m = n * n * (m * m)reflexivity. Qed.n, m:natH:n * m * n = n * n * mn * n * m * m = n * n * m * m
At this point, some 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...
Definition foo (x: nat) := 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.
forall m : nat, foo m + 1 = foo (m + 1) + 1forall m : nat, foo m + 1 = foo (m + 1) + 1m:natfoo m + 1 = foo (m + 1) + 1reflexivity. Qed.m:nat6 = 6
However, this automatic unfolding is somewhat conservative. For
example, if we define a slightly more complicated function
involving a pattern match...
Definition bar x :=
match x with
| O => 5
| S _ => 5
end.
...then the analogous proof will get stuck:
forall m : nat, bar m + 1 = bar (m + 1) + 1forall m : nat, bar m + 1 = bar (m + 1) + 1m:natbar m + 1 = bar (m + 1) + 1Abort.m:natbar m + 1 = bar (m + 1) + 1
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 cannot itself 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.
forall m : nat, bar m + 1 = bar (m + 1) + 1forall m : nat, bar m + 1 = bar (m + 1) + 1m:natbar m + 1 = bar (m + 1) + 1m:natE:m = 0bar 0 + 1 = bar (0 + 1) + 1m, n:natE:m = S nbar (S n) + 1 = bar (S n + 1) + 1m:natE:m = 0bar 0 + 1 = bar (0 + 1) + 1reflexivity.m:natE:m = 06 = 6m, n:natE:m = S nbar (S n) + 1 = bar (S n + 1) + 1reflexivity. Qed.m, n:natE:m = S n6 = 6
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.
forall m : nat, bar m + 1 = bar (m + 1) + 1forall m : nat, bar m + 1 = bar (m + 1) + 1m:natbar m + 1 = bar (m + 1) + 1m:natmatch m with | 0 | _ => 5 end + 1 = match m + 1 with | 0 | _ => 5 end + 1
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.
m:natE:m = 05 + 1 = match 0 + 1 with | 0 | _ => 5 end + 1m, n:natE:m = S n5 + 1 = match S n + 1 with | 0 | _ => 5 end + 1reflexivity.m:natE:m = 05 + 1 = match 0 + 1 with | 0 | _ => 5 end + 1reflexivity. Qed. (* ################################################################# *)m, n:natE:m = S n5 + 1 = match S n + 1 with | 0 | _ => 5 end + 1
We have seen many examples where destruct is used to
perform case analysis of the value of some variable. But
sometimes we need to reason by cases on the result of some
expression. We can also do this with destruct.
Here are some examples:
Definition sillyfun (n : nat) : bool := if n =? 3 then false else if n =? 5 then false else false.forall n : nat, sillyfun n = falseforall n : nat, sillyfun n = falsen:natsillyfun n = falsen:nat(if n =? 3 then false else if n =? 5 then false else false) = falsen:natE1:(n =? 3) = truefalse = falsen:natE1:(n =? 3) = false(if n =? 5 then false else false) = falsereflexivity.n:natE1:(n =? 3) = truefalse = falsen:natE1:(n =? 3) = false(if n =? 5 then false else false) = falsen:natE1:(n =? 3) = falseE2:(n =? 5) = truefalse = falsen:natE1:(n =? 3) = falseE2:(n =? 5) = falsefalse = falsereflexivity.n:natE1:(n =? 3) = falseE2:(n =? 5) = truefalse = falsereflexivity. Qed.n:natE1:(n =? 3) = falseE2:(n =? 5) = falsefalse = false
After unfolding sillyfun in the above proof, we find that
we are stuck on if (n =? 3) then ... else .... But either
n is equal to 3 or it isn't, so we can use destruct (eqb
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, standard, optional (combine_split)
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.
Prove that split and combine are inverses in the following
sense:
forall (X Y : Type) (l : list (X * Y)) (l1 : list X) (l2 : list Y), split l = (l1, l2) -> combine l1 l2 = l(* FILL IN HERE *) Admitted.forall (X Y : Type) (l : list (X * Y)) (l1 : list X) (l2 : list Y), split l = (l1, l2) -> combine l1 l2 = l
☐
The eqn: part of the destruct tactic is optional: We've chosen
to include it most of the time, just for the sake of
documentation, but many Coq proofs omit it.
When destructing compound expressions, however, the information
recorded by the eqn: can actually be critical: if we leave it
out, then destruct can sometimes erase information we need to
complete a proof.
For example, suppose we define a function sillyfun1 like
this:
Definition sillyfun1 (n : nat) : bool :=
if n =? 3 then true
else if 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. If we start the proof like this (with no eqn: on the
destruct)...
forall n : nat, sillyfun1 n = true -> oddb n = trueforall n : nat, sillyfun1 n = true -> oddb n = truen:nateq:sillyfun1 n = trueoddb n = truen:nateq:(if n =? 3 then true else if n =? 5 then true else false) = trueoddb n = true(* stuck... *) Abort.n:nateq:true = trueoddb n = truen:nateq:(if n =? 5 then true else false) = trueoddb n = true
... then we are 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 quite brutal --
in this case, it thows away every occurrence of 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 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 want here is to substitute away all existing occurences of
n =? 3, but at the same time add an equation to the context that
records which case we are in. This is precisely what the eqn:
qualifier does.
forall n : nat, sillyfun1 n = true -> oddb n = trueforall n : nat, sillyfun1 n = true -> oddb n = truen:nateq:sillyfun1 n = trueoddb n = truen:nateq:(if n =? 3 then true else if n =? 5 then true else false) = trueoddb n = true(* 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. *)n:natHeqe3:(n =? 3) = trueeq:true = trueoddb n = truen:natHeqe3:(n =? 3) = falseeq:(if n =? 5 then true else false) = trueoddb n = truen:natHeqe3:(n =? 3) = trueeq:true = trueoddb n = truen:natHeqe3:n = 3eq:true = trueoddb n = truereflexivity.n:natHeqe3:n = 3eq:true = trueoddb 3 = true(* 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, allowing us to finish the proof. *)n:natHeqe3:(n =? 3) = falseeq:(if n =? 5 then true else false) = trueoddb n = truen:natHeqe3:(n =? 3) = falseHeqe5:(n =? 5) = trueeq:true = trueoddb n = truen:natHeqe3:(n =? 3) = falseHeqe5:(n =? 5) = falseeq:false = trueoddb n = truen:natHeqe3:(n =? 3) = falseHeqe5:(n =? 5) = trueeq:true = trueoddb n = truen:natHeqe3:(n =? 3) = falseHeqe5:n = 5eq:true = trueoddb n = truereflexivity.n:natHeqe3:(n =? 3) = falseHeqe5:n = 5eq:true = trueoddb 5 = truediscriminate eq. Qed.n:natHeqe3:(n =? 3) = falseHeqe5:(n =? 5) = falseeq:false = trueoddb n = true
forall (f : bool -> bool) (b : bool), f (f (f b)) = f b(* FILL IN HERE *) Admitted.forall (f : bool -> bool) (b : bool), f (f (f b)) = f b
☐
(* ################################################################# *)
We've now seen many of Coq's most fundamental tactics. We'll
introduce a few more in the coming chapters, and later on we'll
see some more powerful automation tactics that make Coq help us
with low-level details. But basically we've got what we need to
get work done.
Here are the ones we've seen:
- 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
- injection: reason by injectivity on equalities
between values of inductively defined types
- discriminate: reason by disjointness of constructors on
equalities between values of inductively defined types
- 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
(* ################################################################# *)
forall n m : nat, (n =? m) = (m =? n)(* FILL IN HERE *) Admitted.forall n m : nat, (n =? m) = (m =? n)
☐
Exercise: 3 stars, advanced, optional (eqb_sym_informal)
(* FILL IN HERE
[] *)
forall n m p : nat, (n =? m) = true -> (m =? p) = true -> (n =? p) = true(* FILL IN HERE *) Admitted.forall n m p : nat, (n =? m) = true -> (m =? p) = true -> (n =? p) = true
☐
Exercise: 3 stars, advanced (split_combine)
Admitted.Propsplit_combine_statement(* FILL IN HERE *) Admitted. (* Do not modify the following line: *) Definition manual_grade_for_split_combine : option (nat*string) := None.split_combine_statement
☐
Exercise: 3 stars, advanced (filter_exercise)
forall (X : Type) (test : X -> bool) (x : X) (l lf : list X), filter test l = x :: lf -> test x = true(* FILL IN HERE *) Admitted.forall (X : Type) (test : X -> bool) (x : X) (l lf : list X), filter test l = x :: lf -> test x = true
☐
Exercise: 4 stars, advanced, recommended (forall_exists_challenge)
Admitted.forallb:forall X0 : Type, (X0 -> bool) -> list X0 -> boolX:Typetest:X -> booll:list Xboolforallb oddb [1; 3; 5; 7; 9] = trueAdmitted.forallb oddb [1; 3; 5; 7; 9] = trueforallb negb [false; false] = trueAdmitted.forallb negb [false; false] = trueforallb evenb [0; 2; 4; 5] = falseAdmitted.forallb evenb [0; 2; 4; 5] = falseforallb (eqb 5) [ ] = trueAdmitted.forallb (eqb 5) [ ] = trueAdmitted.existsb:forall X0 : Type, (X0 -> bool) -> list X0 -> boolX:Typetest:X -> booll:list Xboolexistsb (eqb 5) [0; 2; 3; 6] = falseAdmitted.existsb (eqb 5) [0; 2; 3; 6] = falseexistsb (andb true) [true; true; false] = trueAdmitted.existsb (andb true) [true; true; false] = trueexistsb oddb [1; 0; 0; 0; 0; 3] = trueAdmitted.existsb oddb [1; 0; 0; 0; 0; 3] = trueexistsb evenb [ ] = falseAdmitted.existsb evenb [ ] = falseAdmitted.X:Typetest:X -> booll:list Xboolforall (X : Type) (test : X -> bool) (l : list X), existsb test l = existsb' test lAdmitted.forall (X : Type) (test : X -> bool) (l : list X), existsb test l = existsb' test l
☐
(* Wed Jan 9 12:02:44 EST 2019 *)