MoreCoqMore About Coq
Require Export Poly.
This chapter introduces several more proof strategies and tactics that,
together, allow us to prove theorems about the functional
programs we have been writing. In particular, we'll reason about
functions that work with natural numbers and lists.
In particular, we will see:
- how to use auxiliary lemmas, in both forwards and backwards reasoning
- how to reason about data constructors, which are injective and disjoint
- how to create a strong induction hypotheses (and when strengthening is required)
- 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.
(* At this point, we could finish with
"rewrite → eq2. reflexivity." as we have
done several times above. But 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.
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.
You may find it instructive to experiment with this proof
and see if there is a way to complete it using just rewrite
instead of apply.
Typically, when we use apply H, the statement H will
begin with a ∀ binding 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.
Theorem silly_ex :
(∀n, evenb n = true → oddb (S n) = true) →
evenb 3 = true →
oddb 4 = true.
Proof.
(* FILL IN HERE *) Admitted.
☐
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.
Theorem silly3_firsttry : ∀(n : nat),
true = beq_nat n 5 →
beq_nat (S (S n)) 7 = true.
Proof.
intros n H.
simpl.
(* Here we cannot use apply directly *)
Abort.
In this case we can use the symmetry tactic, which switches the
left and right sides of an equality in the goal.
Theorem silly3 : ∀(n : nat),
true = beq_nat n 5 →
beq_nat (S (S n)) 7 = true.
Proof.
intros n H.
symmetry.
simpl. (* Actually, this simpl is unnecessary, since
apply will perform simplification first. *)
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 SearchAbout is your friend.Theorem rev_exercise1 : ∀(l l' : list nat),
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. Are there situations where both can usefully be applied? (* FILL IN HERE *)☐
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.
Since this is a common pattern, we might
abstract 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.
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.
(* 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. *)
apply trans_eq with (m:=[c;d]). apply eq1. apply eq2. Qed.
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 clear 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 other facts:
| O : nat
| S : nat → nat.
- The constructor S is injective. That is, the only way we can
have S n = S m is if 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 instructs Coq to "invert" this
equality to extract the information it contains about these terms:
- If c and d are the same constructor, then we know, by the
injectivity of this constructor, that a1 = b1, a2 = b2,
etc.; 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. That is, a false assumption has crept into the context, and this means that any goal whatsoever is provable! In this case, inversion H marks the current goal as completed and pops it off the goal stack.
Theorem eq_add_S : ∀(n m : nat),
S n = S m →
n = m.
Proof.
intros n m eq. inversion eq. reflexivity. Qed.
Theorem silly4 : ∀(n m : nat),
[n] = [m] →
n = m.
Proof.
intros n o eq. inversion eq. reflexivity. Qed.
As a convenience, the inversion tactic can also
destruct equalities between complex values, binding
multiple variables as it goes.
Theorem silly5 : ∀(n m o : nat),
[n;m] = [o;o] →
[n] = [m].
Proof.
intros n m o eq. inversion eq. reflexivity. Qed.
Example sillyex1 : ∀(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.
☐
Theorem silly6 : ∀(n : nat),
S n = O →
2 + 2 = 5.
Proof.
intros n contra. inversion contra. Qed.
Theorem silly7 : ∀(n m : nat),
false = true →
[n] = [m].
Proof.
intros n m contra. inversion contra. Qed.
Example sillyex2 : ∀(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 :: l = [] →
y :: l = z :: j →
x = z.
Proof.
(* FILL IN HERE *) Admitted.
☐
While the injectivity of constructors allows us to reason
∀ (n m : nat), S n = S m → n = m, the reverse direction of
the implication is an instance of a more general fact about
constructors and functions, which we will often find useful:
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.
Exercise: 2 stars, optional (practice)
A couple more nontrivial but not-too-complicated proofs to work together in class, or for you to work as exercises.Theorem beq_nat_0_l : ∀n,
beq_nat 0 n = true → n = 0.
Proof.
(* FILL IN HERE *) Admitted.
Theorem beq_nat_0_r : ∀n,
beq_nat n 0 = true → n = 0.
Proof.
(* FILL IN HERE *) Admitted.
☐
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.
Similarly, the tactic 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
gives us 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.
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, Coq tends to favor backward reasoning, but in some
situations the forward style can be easier to use or to think
about.
Exercise: 3 stars (plus_n_n_injective)
Practice using "in" variants in this exercise.Theorem plus_n_n_injective : ∀n m,
n + n = m + m →
n = m.
Proof.
intros n. induction n as [| n'].
(* Hint: use the plus_n_Sm lemma *)
(* 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 little bit delicate: if we
begin it 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...
Theorem double_injective_FAILED : ∀n m,
double n = double m →
n = m.
Proof.
intros n m. induction n as [| n'].
Case "n = O". simpl. intros eq. destruct m as [| m'].
SCase "m = O". reflexivity.
SCase "m = S m'". inversion eq.
Case "n = S n'". intros eq. destruct m as [| m'].
SCase "m = O". inversion eq.
SCase "m = S m'". apply f_equal.
(* Here we are stuck. 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 this 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 that
the proposition
holds for all n 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 with proving R! (If we
tried to prove R from Q, we would say 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 at this point.)
To summarize: Trying to carry out this proof by induction on n
when m is already in the context doesn't work because we are
trying to prove a relation involving every n but just a
single m.
The good 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'].
Case "n = O". simpl. intros m eq. destruct m as [| m'].
SCase "m = O". reflexivity.
SCase "m = S m'". inversion eq.
Case "n = S n'".
(* Notice that both the goal and the induction
hypothesis have changed: 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 choose a particular m and introduce the
assumption that double n = double m. Since we
are doing a case analysis on n, we need a case
analysis on m to keep the two "in sync." *)
destruct m as [| m'].
SCase "m = O".
(* The 0 case is trivial *)
inversion eq.
SCase "m = S m'".
apply f_equal.
(* At this point, since we are in the second
branch of the destruct m, the m' mentioned
in the context at this point is actually the
predecessor of the one 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
m' that we are talking about right now (this
instantiation is performed automatically by
apply), then IHn' gives us exactly what we
need to finish the proof. *)
apply IHn'. inversion eq. reflexivity. Qed.
What this teaches us is that we need to be careful about using
induction to try to prove something too specific: If we're proving
a property of n and m by induction on n, we may need to
leave m generic.
The proof of this theorem (left as an exercise) has to be treated similarly:
Exercise: 2 stars (beq_nat_true)
Theorem beq_nat_true : ∀n m,
beq_nat n m = true → n = m.
Proof.
(* FILL IN HERE *) Admitted.
beq_nat n m = true → n = m.
Proof.
(* FILL IN HERE *) Admitted.
☐
Exercise: 2 stars, advanced (beq_nat_true_informal)
Give a careful informal proof of beq_nat_true, being as explicit as possible about quantifiers.(* FILL IN HERE *)
☐
The strategy of doing fewer intros before an induction doesn't
always work directly; sometimes a little rearrangement of
quantified variables is needed. Suppose, for example, that we
wanted to prove double_injective by induction on m instead of
n.
Theorem double_injective_take2_FAILED : ∀n m,
double n = double m →
n = m.
Proof.
intros n m. induction m as [| m'].
Case "m = O". simpl. intros eq. destruct n as [| n'].
SCase "n = O". reflexivity.
SCase "n = S n'". inversion eq.
Case "m = S m'". intros eq. destruct n as [| n'].
SCase "n = O". inversion eq.
SCase "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
will work, but it's not nice: We don't want to have to mangle the
statements of lemmas to fit the needs of a particular strategy for
proving them — we want to state them in the most clear and
natural way.
What we can do instead is to first introduce all the
quantified variables and then re-generalize one or more of
them, taking them 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'].
Case "m = O". simpl. intros n eq. destruct n as [| n'].
SCase "n = O". reflexivity.
SCase "n = S n'". inversion eq.
Case "m = S m'". intros n eq. destruct n as [| n'].
SCase "n = O". inversion eq.
SCase "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.
Here's another illustration of inversion and using an
appropriately general induction hypothesis. This is a slightly
roundabout way of stating a fact that we have already proved
above. The extra equalities force us to do a little more
equational reasoning and exercise some of the tactics we've seen
recently.
- First, suppose m = 0, and suppose n is a number such
that double n = double m. We must show that n = 0.
- Otherwise, 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'.
Theorem length_snoc' : ∀(X : Type) (v : X)
(l : list X) (n : nat),
length l = n →
length (snoc l v) = S n.
Proof.
intros X v l. induction l as [| v' l'].
Case "l = []".
intros n eq. rewrite ← eq. reflexivity.
Case "l = v' :: l'".
intros n eq. simpl. destruct n as [| n'].
SCase "n = 0". inversion eq.
SCase "n = S n'".
apply f_equal. apply IHl'. inversion eq. reflexivity. Qed.
It might be tempting to start proving the above theorem
by introducing n and eq at the outset. However, this leads
to an induction hypothesis that is not strong enough. Compare
the above to the following (aborted) attempt:
Theorem length_snoc_bad : ∀(X : Type) (v : X)
(l : list X) (n : nat),
length l = n →
length (snoc l v) = S n.
Proof.
intros X v l n eq. induction l as [| v' l'].
Case "l = []".
rewrite ← eq. reflexivity.
Case "l = v' :: l'".
simpl. destruct n as [| n'].
SCase "n = 0". inversion eq.
SCase "n = S n'".
apply f_equal. Abort. (* apply IHl'. *) (* The IH doesn't apply! *)
As in the double examples, the problem is that by
introducing n before doing induction on l, the induction
hypothesis is specialized to one particular natural number, namely
n. In the induction case, however, we need to be able to use
the induction hypothesis on some other natural number n'.
Retaining the more general form of the induction hypothesis thus
gives us more flexibility.
In general, a good rule of thumb is to make the induction hypothesis
as general as possible.
Exercise: 3 stars (gen_dep_practice)
Prove this by induction on l.Theorem index_after_last: ∀(n : nat) (X : Type) (l : list X),
length l = n →
index n l = None.
Proof.
(* FILL IN HERE *) Admitted.
☐
Theorem: For all sets X, lists l : list X, and numbers
n, if length l = n then index n l = None.
Proof:
(* FILL IN HERE *)
☐
Exercise: 3 stars, advanced, optional (index_after_last_informal)
Write an informal proof corresponding to your Coq proof of index_after_last:☐
Exercise: 3 stars, optional (gen_dep_practice_more)
Prove this by induction on l.Theorem length_snoc''' : ∀(n : nat) (X : Type)
(v : X) (l : list X),
length l = n →
length (snoc l v) = S n.
Proof.
(* FILL IN HERE *) Admitted.
☐
Exercise: 3 stars, optional (app_length_cons)
Prove this by induction on l1, without using app_length from Lists.Theorem app_length_cons : ∀(X : Type) (l1 l2 : list X)
(x : X) (n : nat),
length (l1 ++ (x :: l2)) = n →
S (length (l1 ++ l2)) = n.
Proof.
(* FILL IN HERE *) Admitted.
☐
Exercise: 4 stars, optional (app_length_twice)
Prove this by induction on l, without using app_length.Theorem app_length_twice : ∀(X:Type) (n:nat) (l:list X),
length l = n →
length (l ++ l) = n + n.
Proof.
(* FILL IN HERE *) Admitted.
☐
Exercise: 3 stars, optional (double_induction)
Prove the following principle of induction over two naturals.Theorem double_induction: ∀(P : nat → nat → Prop),
P 0 0 →
(∀m, P m 0 → P (S m) 0) →
(∀n, P 0 n → P 0 (S n)) →
(∀m n, P m n → P (S m) (S n)) →
∀m n, P m n.
Proof.
(* FILL IN HERE *) Admitted.
☐
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).
Case "beq_nat n 3 = true". reflexivity.
Case "beq_nat n 3 = false". destruct (beq_nat n 5).
SCase "beq_nat n 5 = true". reflexivity.
SCase "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 .... Well,
either n is equal to 3 or it isn't, so we 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: 1 star (override_shadow)
Theorem override_shadow : ∀(X:Type) x1 x2 k1 k2 (f : nat→X),
(override (override f k1 x2) k1 x1) k2 = (override f k1 x1) k2.
Proof.
(* FILL IN HERE *) Admitted.
(override (override f k1 x2) k1 x1) k2 = (override f k1 x1) k2.
Proof.
(* FILL IN HERE *) Admitted.
Theorem combine_split : ∀X Y (l : list (X × Y)) l1 l2,
split l = (l1, l2) →
combine l1 l2 = l.
Proof.
(* FILL IN HERE *) Admitted.
☐
Sometimes, doing a destruct on a compound expression (a
non-variable) will erase information we need to complete a proof. For example, suppose
we define a function sillyfun1 like this:
Definition sillyfun1 (n : nat) : bool :=
if beq_nat n 3 then true
else if beq_nat n 5 then true
else false.
And suppose that we want to convince Coq of the rather
obvious observation 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.
We get stuck at this point because the context does not
contain enough information to prove the goal! The problem is that
the substitution peformed 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, in this branch of the case
analysis, beq_nat n 3 = true, 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 (with whatever
name 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. *)
Case "e3 = true". apply beq_nat_true in Heqe3.
rewrite → Heqe3. reflexivity.
Case "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.
SCase "e5 = true".
apply beq_nat_true in Heqe5.
rewrite → Heqe5. reflexivity.
SCase "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.
Theorem override_same : ∀(X:Type) x1 k1 k2 (f : nat→X),
f k1 = x1 →
(override f k1 x1) k2 = f k2.
Proof.
(* FILL IN HERE *) Admitted.
f k1 = x1 →
(override f k1 x1) k2 = f k2.
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 (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
Theorem beq_nat_sym : ∀(n m : nat),
beq_nat n m = beq_nat m n.
Proof.
(* FILL IN HERE *) Admitted.
beq_nat n m = beq_nat m n.
Proof.
(* FILL IN HERE *) Admitted.
☐
Theorem: For any nats n m, beq_nat n m = beq_nat m n.
Proof:
(* FILL IN HERE *)
☐
Exercise: 3 stars, advanced, optional (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.
☐
Complete the definition of split_combine_statement below with a
property that states that split is the inverse of
combine. Then, prove that the property holds. (Be sure to leave
your induction hypothesis general by not doing intros on more
things than necessary. Hint: what property do you need of l1
and l2 for split combine l1 l2 = (l1,l2) to be true?)
Exercise: 3 stars, advanced (split_combine)
We have just proven 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 :=
(* FILL IN HERE *) admit.
Theorem split_combine : split_combine_statement.
Proof.
(* FILL IN HERE *) Admitted.
Theorem override_permute : ∀(X:Type) x1 x2 k1 k2 k3 (f : nat→X),
beq_nat k2 k1 = false →
(override (override f k2 x2) k1 x1) k3 = (override (override f k1 x1) k2 x2) k3.
Proof.
(* FILL IN HERE *) Admitted.
beq_nat k2 k1 = false →
(override (override f k2 x2) k1 x1) k3 = (override (override f k1 x1) k2 x2) k3.
Proof.
(* FILL IN HERE *) Admitted.
☐
Exercise: 3 stars, advanced (filter_exercise)
This one is a bit challenging. Pay attention to the form of your IH.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.
☐
Prove that existsb' and existsb have the same behavior.
Exercise: 4 stars, advanced (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 *)
☐
(* $Date: 2014-10-13 22:45:24 -0400 (Mon, 13 Oct 2014) $ *)