LogicLogic in Coq
(* $Date: 2012-07-22 18:36:58 -0400 (Sun, 22 Jul 2012) $ *)
Require Export "Prop".
Coq's built-in logic is extremely small: only Inductive
definitions, universal quantification (∀), and
implication (→) are primitive, while all the other familiar
logical connectives — conjunction, disjunction, negation,
existential quantification, even equality — can be defined using
just these.
Quantification and Implication
Definition funny_prop1 :=
∀n, ∀(E : beautiful n), beautiful (n+3).
If we had a proof term inhabiting this proposition, it would
be a function with two arguments: a number n and some evidence
that n is beautiful. But the name E for this evidence is not
used in the rest of the statement of funny_prop1, so it's a bit
silly to bother making up a name. We could write it like this
instead, using the dummy identifier _ in place of a real
name:
Definition funny_prop1' :=
∀n, ∀(_ : beautiful n), beautiful (n+3).
Or, equivalently, we can write it in more familiar notation:
Definition funny_prop1'' :=
∀n, beautiful n → beautiful (n+3).
This illustrates that "P → Q" is just syntactic sugar for
"∀ (_:P), Q".
Conjunction
Inductive and (P Q : Prop) : Prop :=
conj : P → Q → (and P Q).
Note that, like the definition of ev in the previous
chapter, this definition is parameterized; however, in this case,
the parameters are themselves propositions, rather than numbers.
The intuition behind this definition is simple: to
construct evidence for and P Q, we must provide evidence
for P and evidence for Q. More precisely:
Since we'll be using conjunction a lot, let's introduce a more
familiar-looking infix notation for it.
- conj p q can be taken as evidence for and P Q if p
is evidence for P and q is evidence for Q; and
- this is the only way to give evidence for and P Q — that is, if someone gives us evidence for and P Q, we know it must have the form conj p q, where p is evidence for P and q is evidence for Q.
Notation "P ∧ Q" := (and P Q) : type_scope.
(The type_scope annotation tells Coq that this notation
will be appearing in propositions, not values.)
Consider the "type" of the constructor conj:
Check conj.
(* ===> forall P Q : Prop, P -> Q -> P /\ Q *)
Notice that it takes 4 inputs — namely the propositions P
and Q and evidence for P and Q — and returns as output the
evidence of P ∧ Q.
Besides the elegance of building everything up from a tiny
foundation, what's nice about defining conjunction this way is
that we can prove statements involving conjunction using the
tactics that we already know. For example, if the goal statement
is a conjuction, we can prove it by applying the single
constructor conj, which (as can be seen from the type of conj)
solves the current goal and leaves the two parts of the
conjunction as subgoals to be proved separately.
Theorem and_example :
(beautiful 0) ∧ (beautiful 3).
Proof.
apply conj.
(* Case "left". *) apply b_0.
(* Case "right". *) apply b_3. Qed.
Let's take a look at the proof object for the above theorem.
Print and_example.
(* ===> conj (beautiful 0) (beautiful 3) b_0 b_3
: beautiful 0 /\ beautiful 3 *)
Note that the proof is of the form
Just for convenience, we can use the tactic split as a shorthand for
apply conj.
conj (beautiful 0) (beautiful 3)
(...pf of beautiful 3...) (...pf of beautiful 3...)
as you'd expect, given the type of conj.
(...pf of beautiful 3...) (...pf of beautiful 3...)
Theorem and_example' :
(ev 0) ∧ (ev 4).
Proof.
split.
Case "left". apply ev_0.
Case "right". apply ev_SS. apply ev_SS. apply ev_0. Qed.
Conversely, the inversion tactic can be used to take a
conjunction hypothesis in the context, calculate what evidence
must have been used to build it, and add variables representing
this evidence to the proof context.
Theorem proj1 : ∀P Q : Prop,
P ∧ Q → P.
Proof.
intros P Q H.
inversion H as [HP HQ].
apply HP. Qed.
Theorem proj2 : ∀P Q : Prop,
P ∧ Q → Q.
Proof.
(* FILL IN HERE *) Admitted.
P ∧ Q → Q.
Proof.
(* FILL IN HERE *) Admitted.
☐
Theorem and_commut : ∀P Q : Prop,
P ∧ Q → Q ∧ P.
Proof.
(* WORKED IN CLASS *)
intros P Q H.
inversion H as [HP HQ].
split.
(* Case "left". *) apply HQ.
(* Case "right".*) apply HP. Qed.
Once again, we have commented out the Case tactics to make the
proof object for this theorem easy to understand. Examining it
shows that all that is really happening is taking apart a record
containing evidence for P and Q and rebuilding it in the
opposite order:
Print and_commut.
(* ===>
and_commut =
fun (P Q : Prop) (H : P /\ Q) =>
let H0 := match H with
| conj HP HQ => conj Q P HQ HP
end
in H0
: forall P Q : Prop, P /\ Q -> Q /\ P *)
Exercise: 2 stars (and_assoc)
In the following proof, notice how the nested pattern in the inversion breaks the hypothesis H : P ∧ (Q ∧ R) down into HP: P, HQ : Q, and HR : R. Finish the proof from there:Theorem and_assoc : ∀P Q R : Prop,
P ∧ (Q ∧ R) → (P ∧ Q) ∧ R.
Proof.
intros P Q R H.
inversion H as [HP [HQ HR]].
(* FILL IN HERE *) Admitted.
☐
Exercise: 2 stars, recommended (even__ev)
Now we can prove the other direction of the equivalence of even and ev, which we left hanging in chapter Prop. Notice that the left-hand conjunct here is the statement we are actually interested in; the right-hand conjunct is needed in order to make the induction hypothesis strong enough that we can carry out the reasoning in the inductive step. (To see why this is needed, try proving the left conjunct by itself and observe where things get stuck.)Theorem even__ev : ∀n : nat,
(even n → ev n) ∧ (even (S n) → ev (S n)).
Proof.
(* Hint: Use induction on n. *)
(* FILL IN HERE *) Admitted.
☐
Exercise: 2 stars, optional (conj_fact)
Construct a proof object demonstrating the following proposition.Definition conj_fact : ∀P Q R, P ∧ Q → Q ∧ R → P ∧ R :=
(* FILL IN HERE *) admit.
☐
Definition iff (P Q : Prop) := (P → Q) ∧ (Q → P).
Notation "P ↔ Q" := (iff P Q)
(at level 95, no associativity) : type_scope.
Theorem iff_implies : ∀P Q : Prop,
(P ↔ Q) → P → Q.
Proof.
intros P Q H.
inversion H as [HAB HBA]. apply HAB. Qed.
Theorem iff_sym : ∀P Q : Prop,
(P ↔ Q) → (Q ↔ P).
Proof.
(* WORKED IN CLASS *)
intros P Q H.
inversion H as [HAB HBA].
split.
Case "→". apply HBA.
Case "←". apply HAB. Qed.
Exercise: 1 star, optional (iff_properties)
Using the above proof that ↔ is symmetric (iff_sym) as a guide, prove that it is also reflexive and transitive.Theorem iff_refl : ∀P : Prop,
P ↔ P.
Proof.
(* FILL IN HERE *) Admitted.
Theorem iff_trans : ∀P Q R : Prop,
(P ↔ Q) → (Q ↔ R) → (P ↔ R).
Proof.
(* FILL IN HERE *) Admitted.
Hint: If you have an iff hypothesis in the context, you can use
inversion to break it into two separate implications. (Think
about why this works.) ☐
We have seen that the families of propositions beautiful and
gorgeous actually characterize the same set of numbers.
Prove that beautiful n ↔ gorgeous n for all n. Just for
fun, write your proof as an explicit proof object, rather than
using tactics. (Hint: if you make use of previously defined
theorems, you should only need a single line!)
Exercise: 2 stars, optional (beautiful_iff_gorgeous)
Definition beautiful_iff_gorgeous :
∀n, beautiful n ↔ gorgeous n :=
(* FILL IN HERE *) admit.
☐
Some of Coq's tactics treat iff statements specially, thus
avoiding the need for some low-level manipulation when reasoning
with them. In particular, rewrite can be used with iff
statements, not just equalities.
Inductive or (P Q : Prop) : Prop :=
| or_introl : P → or P Q
| or_intror : Q → or P Q.
Notation "P ∨ Q" := (or P Q) : type_scope.
Consider the "type" of the constructor or_introl:
Check or_introl.
(* ===> forall P Q : Prop, P -> P \/ Q *)
It takes 3 inputs, namely the propositions P, Q and
evidence of P, and returns, as output, the evidence of P ∨ Q.
Next, look at the type of or_intror:
Check or_intror.
(* ===> forall P Q : Prop, Q -> P \/ Q *)
It is like or_introl but it requires evidence of Q
instead of evidence of P.
Intuitively, there are two ways of giving evidence for P ∨ Q:
Since P ∨ Q has two constructors, doing inversion on a
hypothesis of type P ∨ Q yields two subgoals.
- give evidence for P (and say that it is P you are giving
evidence for — this is the function of the or_introl
constructor), or
- give evidence for Q, tagged with the or_intror constructor.
Theorem or_commut : ∀P Q : Prop,
P ∨ Q → Q ∨ P.
Proof.
intros P Q H.
inversion H as [HP | HQ].
Case "left". apply or_intror. apply HP.
Case "right". apply or_introl. apply HQ. Qed.
From here on, we'll use the shorthand tactics left and right
in place of apply or_introl and apply or_intror.
Theorem or_commut' : ∀P Q : Prop,
P ∨ Q → Q ∨ P.
Proof.
intros P Q H.
inversion H as [HP | HQ].
Case "left". right. apply HP.
Case "right". left. apply HQ. Qed.
Exercise: 2 stars, optional (or_commut'')
Try to write down an explicit proof object for or_commut (without using Print to peek at the ones we already defined!).(* FILL IN HERE *)
☐
Theorem or_distributes_over_and_1 : ∀P Q R : Prop,
P ∨ (Q ∧ R) → (P ∨ Q) ∧ (P ∨ R).
Proof.
intros P Q R. intros H. inversion H as [HP | [HQ HR]].
Case "left". split.
SCase "left". left. apply HP.
SCase "right". left. apply HP.
Case "right". split.
SCase "left". right. apply HQ.
SCase "right". right. apply HR. Qed.
Theorem or_distributes_over_and_2 : ∀P Q R : Prop,
(P ∨ Q) ∧ (P ∨ R) → P ∨ (Q ∧ R).
Proof.
(* FILL IN HERE *) Admitted.
(P ∨ Q) ∧ (P ∨ R) → P ∨ (Q ∧ R).
Proof.
(* FILL IN HERE *) Admitted.
Theorem or_distributes_over_and : ∀P Q R : Prop,
P ∨ (Q ∧ R) ↔ (P ∨ Q) ∧ (P ∨ R).
Proof.
(* FILL IN HERE *) Admitted.
P ∨ (Q ∧ R) ↔ (P ∨ Q) ∧ (P ∨ R).
Proof.
(* FILL IN HERE *) Admitted.
☐
Relating ∧ and ∨ with andb and orb
Theorem andb_true__and : ∀b c,
andb b c = true → b = true ∧ c = true.
Proof.
(* WORKED IN CLASS *)
intros b c H.
destruct b.
Case "b = true". destruct c.
SCase "c = true". apply conj. reflexivity. reflexivity.
SCase "c = false". inversion H.
Case "b = false". inversion H. Qed.
Theorem and__andb_true : ∀b c,
b = true ∧ c = true → andb b c = true.
Proof.
(* WORKED IN CLASS *)
intros b c H.
inversion H.
rewrite H0. rewrite H1. reflexivity. Qed.
Theorem andb_false : ∀b c,
andb b c = false → b = false ∨ c = false.
Proof.
(* FILL IN HERE *) Admitted.
Theorem orb_true : ∀b c,
orb b c = true → b = true ∨ c = true.
Proof.
(* FILL IN HERE *) Admitted.
Theorem orb_false : ∀b c,
orb b c = false → b = false ∧ c = false.
Proof.
(* FILL IN HERE *) Admitted.
andb b c = false → b = false ∨ c = false.
Proof.
(* FILL IN HERE *) Admitted.
Theorem orb_true : ∀b c,
orb b c = true → b = true ∨ c = true.
Proof.
(* FILL IN HERE *) Admitted.
Theorem orb_false : ∀b c,
orb b c = false → b = false ∧ c = false.
Proof.
(* FILL IN HERE *) Admitted.
☐
Falsehood
Inductive False : Prop := .
Intuition: False is a proposition for which there is no way
to give evidence.
Exercise: 1 star (False_ind_principle)
Can you predict the induction principle for falsehood?(* Check False_ind. *)
☐
Since False has no constructors, inverting an assumption
of type False always yields zero subgoals, allowing us to
immediately prove any goal.
Theorem False_implies_nonsense :
False → 2 + 2 = 5.
Proof.
intros contra.
inversion contra. Qed.
How does this work? The inversion tactic breaks contra into
each of its possible cases, and yields a subgoal for each case.
As contra is evidence for False, it has no possible cases,
hence, there are no possible subgoals and the proof is done.
Conversely, the only way to prove False is if there is already
something nonsensical or contradictory in the context:
Theorem nonsense_implies_False :
2 + 2 = 5 → False.
Proof.
intros contra.
inversion contra. Qed.
Actually, since the proof of False_implies_nonsense
doesn't actually have anything to do with the specific nonsensical
thing being proved; it can easily be generalized to work for an
arbitrary P:
Theorem ex_falso_quodlibet : ∀(P:Prop),
False → P.
Proof.
(* WORKED IN CLASS *)
intros P contra.
inversion contra. Qed.
The Latin ex falso quodlibet means, literally, "from
falsehood follows whatever you please." This theorem is also
known as the principle of explosion.
Truth
Exercise: 2 stars, optional (True_induction)
Define True as another inductively defined proposition. What induction principle will Coq generate for your definition? (The intution is that True should be a proposition for which it is trivial to give evidence. Alternatively, you may find it easiest to start with the induction principle and work backwards to the inductive definition.)(* FILL IN HERE *)
☐
However, unlike False, which we'll use extensively, True is
just a theoretical curiosity: it is trivial (and therefore
uninteresting) to prove as a goal, and it carries no useful
information as a hypothesis.
Definition not (P:Prop) := P → False.
The intuition is that, if P is not true, then anything at
all (even False) follows from assuming P.
Notation "~ x" := (not x) : type_scope.
Check not.
(* ===> Prop -> Prop *)
It takes a little practice to get used to working with
negation in Coq. Even though you can see perfectly well why
something is true, it can be a little hard at first to get things
into the right configuration so that Coq can see it! Here are
proofs of a few familiar facts about negation to get you warmed
up.
Theorem not_False :
~ False.
Proof.
unfold not. intros H. inversion H. Qed.
Theorem contradiction_implies_anything : ∀P Q : Prop,
(P ∧ ~P) → Q.
Proof.
(* WORKED IN CLASS *)
intros P Q H. inversion H as [HP HNA]. unfold not in HNA.
apply HNA in HP. inversion HP. Qed.
Theorem double_neg : ∀P : Prop,
P → ~~P.
Proof.
(* WORKED IN CLASS *)
intros P H. unfold not. intros G. apply G. apply H. Qed.
Exercise: 2 stars, recommended (double_neg_inf)
Write an informal proof of double_neg:☐
Exercise: 2 stars, recommended (contrapositive)
Theorem contrapositive : ∀P Q : Prop,
(P → Q) → (~Q → ~P).
Proof.
(* FILL IN HERE *) Admitted.
(P → Q) → (~Q → ~P).
Proof.
(* FILL IN HERE *) Admitted.
Theorem not_both_true_and_false : ∀P : Prop,
~ (P ∧ ~P).
Proof.
(* FILL IN HERE *) Admitted.
~ (P ∧ ~P).
Proof.
(* FILL IN HERE *) Admitted.
☐
Exercise: 1 star (informal_not_PNP)
Write an informal proof (in English) of the proposition ∀ P : Prop, ~(P ∧ ~P).(* FILL IN HERE *)
☐
Theorem five_not_even :
~ ev 5.
Proof.
(* WORKED IN CLASS *)
unfold not. intros Hev5. inversion Hev5 as [|n Hev3 Heqn].
inversion Hev3 as [|n' Hev1 Heqn']. inversion Hev1. Qed.
Exercise: 1 star (ev_not_ev_S)
Theorem five_not_even confirms the unsurprising fact that five is not an even number. Prove this more interesting fact:Theorem ev_not_ev_S : ∀n,
ev n → ~ ev (S n).
Proof.
unfold not. intros n H. induction H. (* not n! *)
(* FILL IN HERE *) Admitted.
☐
Note that some theorems that are true in classical logic are not
provable in Coq's (constructive) logic. E.g., let's look at how
this proof gets stuck...
Theorem classic_double_neg : ∀P : Prop,
~~P → P.
Proof.
(* WORKED IN CLASS *)
intros P H. unfold not in H.
(* But now what? There is no way to "invent" evidence for P. *)
Admitted.
Exercise: 5 stars, optional (classical_axioms)
For those who like a challenge, here is an exercise taken from the Coq'Art book (p. 123). The following five statements are often considered as characterizations of classical logic (as opposed to constructive logic, which is what is "built in" to Coq). We can't prove them in Coq, but we can consistently add any one of them as an unproven axiom if we wish to work in classical logic. Prove that these five propositions are equivalent.Definition peirce := ∀P Q: Prop,
((P→Q)→P)→P.
Definition classic := ∀P:Prop,
~~P → P.
Definition excluded_middle := ∀P:Prop,
P ∨ ~P.
Definition de_morgan_not_and_not := ∀P Q:Prop,
~(~P/\~Q) → P∨Q.
Definition implies_to_or := ∀P Q:Prop,
(P→Q) → (~P∨Q).
(* FILL IN HERE *)
☐
Notation "x <> y" := (~ (x = y)) : type_scope.
Since inequality involves a negation, it again requires
a little practice to be able to work with it fluently. Here
is one very useful trick. If you are trying to prove a goal
that is nonsensical (e.g., the goal state is false = true),
apply the lemma ex_falso_quodlibet to change the goal to
False. This makes it easier to use assumptions of the form
~P that are available in the context — in particular,
assumptions of the form x<>y.
Theorem not_false_then_true : ∀b : bool,
b <> false → b = true.
Proof.
intros b H. destruct b.
Case "b = true". reflexivity.
Case "b = false".
unfold not in H.
apply ex_falso_quodlibet.
apply H. reflexivity. Qed.
Theorem not_eq_beq_false : ∀n n' : nat,
n <> n' →
beq_nat n n' = false.
Proof.
(* FILL IN HERE *) Admitted.
n <> n' →
beq_nat n n' = false.
Proof.
(* FILL IN HERE *) Admitted.
Theorem beq_false_not_eq : ∀n m,
false = beq_nat n m → n <> m.
Proof.
(* FILL IN HERE *) Admitted.
false = beq_nat n m → n <> m.
Proof.
(* FILL IN HERE *) Admitted.
☐
Existential Quantification
Inductive ex (X:Type) (P : X→Prop) : Prop :=
ex_intro : ∀(witness:X), P witness → ex X P.
That is, ex is a family of propositions indexed by a type X
and a property P over X. In order to give evidence for the
assertion "there exists an x for which the property P holds"
we must actually name a witness — a specific value x — and
then give evidence for P x, i.e., evidence that x has the
property P.
For example, consider this existentially quantified proposition:
Definition some_nat_is_even : Prop :=
ex nat ev.
To prove this proposition, we need to choose a particular number
as witness — say, 4 — and give some evidence that that number is
even.
Definition snie : some_nat_is_even :=
ex_intro _ ev 4 (ev_SS 2 (ev_SS 0 ev_0)).
Coq's notation definition facility can be used to introduce
more familiar notation for writing existentially quantified
propositions, exactly parallel to the built-in syntax for
universally quantified propositions. Instead of writing ex nat
ev to express the proposition that there exists some number that
is even, for example, we can write ∃ x:nat, ev x. (It is
not necessary to understand exactly how the Notation definition
works.)
Notation "'exists' x , p" := (ex _ (fun x => p))
(at level 200, x ident, right associativity) : type_scope.
Notation "'exists' x : X , p" := (ex _ (fun x:X => p))
(at level 200, x ident, right associativity) : type_scope.
We can use the same set of tactics as always for
manipulating existentials. For example, if to prove an
existential, we apply the constructor ex_intro. Since the
premise of ex_intro involves a variable (witness) that does
not appear in its conclusion, we need to explicitly give its value
when we use apply.
Example exists_example_1 : ∃n, n + (n * n) = 6.
Proof.
apply ex_intro with (witness:=2).
reflexivity. Qed.
Note, again, that we have to explicitly give the witness.
Or, instead of writing apply ex_intro with (witness:=e) all the
time, we can use the convenient shorthand ∃ e, which means
the same thing.
Example exists_example_1' : ∃n,
n + (n * n) = 6.
Proof.
∃2.
reflexivity. Qed.
Conversely, if we have an existential hypothesis in the
context, we can eliminate it with inversion. Note the use
of the as... pattern to name the variable that Coq
introduces to name the witness value and get evidence that
the hypothesis holds for the witness. (If we don't
explicitly choose one, Coq will just call it witness, which
makes proofs confusing.)
Theorem exists_example_2 : ∀n,
(∃m, n = 4 + m) →
(∃o, n = 2 + o).
Proof.
intros n H.
inversion H as [m Hm].
∃(2 + m).
apply Hm. Qed.
Exercise: 1 star, optional (english_exists)
In English, what does the proposition
ex nat (fun n => beautiful (S n))
mean?
(* FILL IN HERE *)
Complete the definition of the following proof object:
Definition p : ex nat (fun n => beautiful (S n)) :=
(* FILL IN HERE *) admit.
☐
Exercise: 1 star (dist_not_exists)
Prove that "P holds for all x" and "there is no x for which P does not hold" are equivalent assertions.Theorem dist_not_exists : ∀(X:Type) (P : X → Prop),
(∀x, P x) → ~ (∃x, ~ P x).
Proof.
(* FILL IN HERE *) Admitted.
☐
Exercise: 3 stars, optional (not_exists_dist)
The other direction requires the classical "law of the excluded middle":Theorem not_exists_dist :
excluded_middle →
∀(X:Type) (P : X → Prop),
~ (∃x, ~ P x) → (∀x, P x).
Proof.
(* FILL IN HERE *) Admitted.
☐
Exercise: 2 stars (dist_exists_or)
Prove that existential quantification distributes over disjunction.Theorem dist_exists_or : ∀(X:Type) (P Q : X → Prop),
(∃x, P x ∨ Q x) ↔ (∃x, P x) ∨ (∃x, Q x).
Proof.
(* FILL IN HERE *) Admitted.
☐
(* Print dist_exists_or. *)
Equality
Module MyEquality.
Inductive eq (X:Type) : X → X → Prop :=
refl_equal : ∀x, eq X x x.
Standard infix notation (using Coq's type argument synthesis):
Notation "x = y" := (eq _ x y)
(at level 70, no associativity) : type_scope.
This is a bit subtle. The way to think about it is that, given a
set X, it defines a family of propositions "x is equal to
y," indexed by pairs of values (x and y) from X. There is
just one way of constructing evidence for members of this family:
applying the constructor refl_equal to a type X and a value x
: X yields evidence that x is equal to x.
Here is a slightly different definition — the one that actually
appears in the Coq standard library.
Inductive eq' (X:Type) (x:X) : X → Prop :=
refl_equal' : eq' X x x.
Notation "x =' y" := (eq' _ x y)
(at level 70, no associativity) : type_scope.
Exercise: 3 stars, optional (two_defs_of_eq_coincide)
Verify that the two definitions of equality are equivalent.Theorem two_defs_of_eq_coincide : ∀(X:Type) (x y : X),
x = y ↔ x =' y.
Proof.
(* FILL IN HERE *) Admitted.
☐
The advantage of the second definition is that the induction
principle that Coq derives for it is precisely the familiar
principle of Leibniz equality: what we mean when we say "x and
y are equal" is that every property on P that is true of x
is also true of y.
Check eq'_ind.
(* ===>
forall (X : Type) (x : X) (P : X -> Prop),
P x -> forall y : X, x =' y -> P y
===> (i.e., after a little reorganization)
forall (X : Type) (x : X) forall y : X,
x =' y ->
forall P : X -> Prop, P x -> P y *)
One important consideration remains. Clearly, we can use
refl_equal to construct evidence that, for example, 2 = 2.
Can we also use it to construct evidence that 1 + 1 = 2? Yes:
indeed, it is the very same piece of evidence! The reason is that
Coq treats as "the same" any two terms that are convertible
according to a simple set of computation rules. These rules,
which are similar to those used by Eval simpl, include
evaluation of function application, inlining of definitions, and
simplification of matches.
In tactic-based proofs of equality, the conversion rules are
normally hidden in uses of simpl (either explicit or implicit in
other tactics such as reflexivity). But you can see them
directly at work in the following explicit proof objects:
Definition four : 2 + 2 = 1 + 3 :=
refl_equal nat 4.
Definition singleton : ∀(X:Set) (x:X), []++[x] = x::[] :=
fun (X:Set) (x:X) => refl_equal (list X) [x].
End MyEquality.
Inversion, Again
- takes a hypothesis H whose type P is inductively defined,
and
- for each constructor C in P's definition,
- generates a new subgoal in which we assume H was
built with C,
- adds the arguments (premises) of C to the context of
the subgoal as extra hypotheses,
- matches the conclusion (result type) of C against the
current goal and calculates a set of equalities that must
hold in order for C to be applicable,
- adds these equalities to the context (and, for convenience,
rewrites them in the goal), and
- if the equalities are not satisfiable (e.g., they involve things like S n = O), immediately solves the subgoal.
- generates a new subgoal in which we assume H was
built with C,
Relations as Propositions
Module LeFirstTry.
We've already seen an inductive definition of one
fundamental relation: equality. Another useful one is the "less
than or equal to" relation on numbers:
The following definition should be fairly intuitive. It
says that there are two ways to give evidence that one number is
less than or equal to another: either observe that they are the
same number, or give evidence that the first is less than or equal
to the predecessor of the second.
Inductive le : nat → nat → Prop :=
| le_n : ∀n, le n n
| le_S : ∀n m, (le n m) → (le n (S m)).
End LeFirstTry.
This is a reasonable definition of the <= relation, but we
can streamline it a little by observing that the left-hand
argument n is the same everywhere in the definition, so we can
actually make it a "general parameter" to the whole definition,
rather than an argument to each constructor. This is similar to
what we did in our second definition of the eq relation,
above.
Inductive le (n:nat) : nat → Prop :=
| le_n : le n n
| le_S : ∀m, (le n m) → (le n (S m)).
Notation "m <= n" := (le m n).
The second one is better, even though it looks less symmetric.
Why? Because it gives us a simpler induction principle.
(The same was true of our second version of eq.)
Check le_ind.
(* ===> forall (n : nat) (P : nat -> Prop),
P n ->
(forall m : nat, n <= m -> P m -> P (S m)) ->
forall n0 : nat, n <= n0 -> P n0 *)
By contrast, the induction principle that Coq calculates for the
first definition has a lot of extra quantifiers, which makes it
messier to work with when proving things by induction. Here is
the induction principle for the first le:
(* le_ind :
forall P : nat -> nat -> Prop,
(forall n : nat, P n n) ->
(forall n m : nat, le n m -> P n m -> P n (S m)) ->
forall n n0 : nat, le n n0 -> P n n0 *)
Proofs of facts about <= using the constructors le_n and
le_S follow the same patterns as proofs about properties, like
ev in chapter Prop. We can apply the constructors to prove <=
goals (e.g., to show that 3<=3 or 3<=6), and we can use
tactics like inversion to extract information from <=
hypotheses in the context (e.g., to prove that ~(2 <= 1).)
Here are some sanity checks on the definition. (Notice that,
although these are the same kind of simple "unit tests" as we gave
for the testing functions we wrote in the first few lectures, we
must construct their proofs explicitly — simpl and
reflexivity don't do the job, because the proofs aren't just a
matter of simplifying computations.)
Theorem test_le1 :
3 <= 3.
Proof.
(* WORKED IN CLASS *)
apply le_n. Qed.
Theorem test_le2 :
3 <= 6.
Proof.
(* WORKED IN CLASS *)
apply le_S. apply le_S. apply le_S. apply le_n. Qed.
Theorem test_le3 :
~ (2 <= 1).
Proof.
(* WORKED IN CLASS *)
intros H. inversion H. inversion H1. Qed.
The "strictly less than" relation n < m can now be defined
in terms of le.
Definition lt (n m:nat) := le (S n) m.
Notation "m < n" := (lt m n).
Here are a few more simple relations on numbers:
Inductive square_of : nat → nat → Prop :=
sq : ∀n:nat, square_of n (n * n).
Inductive next_nat (n:nat) : nat → Prop :=
| nn : next_nat n (S n).
Inductive next_even (n:nat) : nat → Prop :=
| ne_1 : ev (S n) → next_even n (S n)
| ne_2 : ev (S (S n)) → next_even n (S (S n)).
Exercise: 2 stars, recommended (total_relation)
Define an inductive binary relation total_relation that holds between every pair of natural numbers.(* FILL IN HERE *)
☐
Exercise: 2 stars (empty_relation)
Define an inductive binary relation empty_relation (on numbers) that never holds.(* FILL IN HERE *)
Module R.
We can define three-place relations, four-place relations,
etc., in just the same way as binary relations. For example,
consider the following three-place relation on numbers:
Inductive R : nat → nat → nat → Prop :=
| c1 : R 0 0 0
| c2 : ∀m n o, R m n o → R (S m) n (S o)
| c3 : ∀m n o, R m n o → R m (S n) (S o)
| c4 : ∀m n o, R (S m) (S n) (S (S o)) → R m n o
| c5 : ∀m n o, R m n o → R n m o.
- Which of the following propositions are provable?
- R 1 1 2
- R 2 2 6
- If we dropped constructor c5 from the definition of R,
would the set of provable propositions change? Briefly (1
sentence) explain your answer.
- If we dropped constructor c4 from the definition of R, would the set of provable propositions change? Briefly (1 sentence) explain your answer.
☐
Exercise: 3 stars, optional (R_fact)
State and prove an equivalent characterization of the relation R. That is, if R m n o is true, what can we say about m, n, and o, and vice versa?(* FILL IN HERE *)
☐
End R.
Exercise: 3 stars, recommended (all_forallb)
Inductively define a property all of lists, parameterized by a type X and a property P : X → Prop, such that all X P l asserts that P is true for every element of the list l.Inductive all (X : Type) (P : X → Prop) : list X → Prop :=
(* FILL IN HERE *)
.
Recall the function forallb, from the exercise
forall_exists_challenge in chapter Poly:
Fixpoint forallb {X : Type} (test : X → bool) (l : list X) : bool :=
match l with
| [] => true
| x :: l' => andb (test x) (forallb test l')
end.
Using the property all, write down a specification for forallb,
and prove that it satisfies the specification. Try to make your
specification as precise as possible.
Are there any important properties of the function forallb which
are not captured by your specification?
(* FILL IN HERE *)
☐
Suppose we have a set X, a function test: X→bool, and a list
l of type list X. Suppose further that l is an "in-order
merge" of two lists, l1 and l2, such that every item in l1
satisfies test and no item in l2 satisfies test. Then filter
test l = l1.
A list l is an "in-order merge" of l1 and l2 if it contains
all the same elements as l1 and l2, in the same order as l1
and l2, but possibly interleaved. For example,
Exercise: 4 stars, optional (filter_challenge)
One of the main purposes of Coq is to prove that programs match their specifications. To this end, let's prove that our definition of filter matches a specification. Here is the specification, written out informally in English.
[1,4,6,2,3]
is an in-order merge of
[1,6,2]
and
[4,3].
Your job is to translate this specification into a Coq theorem and
prove it. (Hint: You'll need to begin by defining what it means
for one list to be a merge of two others. Do this with an
inductive relation, not a Fixpoint.)
(* FILL IN HERE *)
☐
Exercise: 5 stars, optional (filter_challenge_2)
A different way to formally characterize the behavior of filter goes like this: Among all subsequences of l with the property that test evaluates to true on all their members, filter test l is the longest. Express this claim formally and prove it.(* FILL IN HERE *)
Inductive appears_in {X:Type} (a:X) : list X → Prop :=
| ai_here : ∀l, appears_in a (a::l)
| ai_later : ∀b l, appears_in a l → appears_in a (b::l).
...gives us a precise way of saying that a value a appears at
least once as a member of a list l.
Here's a pair of warm-ups about appears_in.
Lemma appears_in_app : ∀{X:Type} (xs ys : list X) (x:X),
appears_in x (xs ++ ys) → appears_in x xs ∨ appears_in x ys.
Proof.
(* FILL IN HERE *) Admitted.
Lemma app_appears_in : ∀{X:Type} (xs ys : list X) (x:X),
appears_in x xs ∨ appears_in x ys → appears_in x (xs ++ ys).
Proof.
(* FILL IN HERE *) Admitted.
Now use appears_in to define a proposition disjoint X l1 l2,
which should be provable exactly when l1 and l2 are
lists (with elements of type X) that have no elements in common.
(* FILL IN HERE *)
Next, use appears_in to define an inductive proposition
no_repeats X l, which should be provable exactly when l is a
list (with elements of type X) where every member is different
from every other. For example, no_repeats nat [1,2,3,4] and
no_repeats bool [] should be provable, while no_repeats nat
[1,2,1] and no_repeats bool [true,true] should not be.
(* FILL IN HERE *)
Finally, state and prove one or more interesting theorems relating
disjoint, no_repeats and ++ (list append).
(* FILL IN HERE *)
☐
Digression: More Facts about <= and <
Exercise: 2 stars, optional (le_exercises)
Theorem O_le_n : ∀n,
0 <= n.
Proof.
(* FILL IN HERE *) Admitted.
Theorem n_le_m__Sn_le_Sm : ∀n m,
n <= m → S n <= S m.
Proof.
(* FILL IN HERE *) Admitted.
Theorem Sn_le_Sm__n_le_m : ∀n m,
S n <= S m → n <= m.
Proof.
intros n m. generalize dependent n. induction m.
(* FILL IN HERE *) Admitted.
Theorem le_plus_l : ∀a b,
a <= a + b.
Proof.
(* FILL IN HERE *) Admitted.
Theorem plus_lt : ∀n1 n2 m,
n1 + n2 < m →
n1 < m ∧ n2 < m.
Proof.
(* FILL IN HERE *) Admitted.
Theorem lt_S : ∀n m,
n < m →
n < S m.
Proof.
(* FILL IN HERE *) Admitted.
Theorem ble_nat_true : ∀n m,
ble_nat n m = true → n <= m.
Proof.
(* FILL IN HERE *) Admitted.
Theorem ble_nat_n_Sn_false : ∀n m,
ble_nat n (S m) = false →
ble_nat n m = false.
Proof.
(* FILL IN HERE *) Admitted.
Theorem ble_nat_false : ∀n m,
ble_nat n m = false → ~(n <= m).
Proof.
(* Hint: Do the right induction! *)
(* FILL IN HERE *) Admitted.
0 <= n.
Proof.
(* FILL IN HERE *) Admitted.
Theorem n_le_m__Sn_le_Sm : ∀n m,
n <= m → S n <= S m.
Proof.
(* FILL IN HERE *) Admitted.
Theorem Sn_le_Sm__n_le_m : ∀n m,
S n <= S m → n <= m.
Proof.
intros n m. generalize dependent n. induction m.
(* FILL IN HERE *) Admitted.
Theorem le_plus_l : ∀a b,
a <= a + b.
Proof.
(* FILL IN HERE *) Admitted.
Theorem plus_lt : ∀n1 n2 m,
n1 + n2 < m →
n1 < m ∧ n2 < m.
Proof.
(* FILL IN HERE *) Admitted.
Theorem lt_S : ∀n m,
n < m →
n < S m.
Proof.
(* FILL IN HERE *) Admitted.
Theorem ble_nat_true : ∀n m,
ble_nat n m = true → n <= m.
Proof.
(* FILL IN HERE *) Admitted.
Theorem ble_nat_n_Sn_false : ∀n m,
ble_nat n (S m) = false →
ble_nat n m = false.
Proof.
(* FILL IN HERE *) Admitted.
Theorem ble_nat_false : ∀n m,
ble_nat n m = false → ~(n <= m).
Proof.
(* Hint: Do the right induction! *)
(* FILL IN HERE *) Admitted.
☐
We say that a list of numbers "stutters" if it repeats the same
number consecutively. The predicate "nostutter mylist" means
that mylist does not stutter. Formulate an inductive definition
for nostutter. (This is different from the no_repeats
predicate in the exercise above; the sequence 1,4,1 repeats but
does not stutter.)
Exercise: 3 stars, recommended (nostutter)
Formulating inductive definitions of predicates is an important skill you'll need in this course. Try to solve this exercise without any help at all (except from your study group partner, if you have one).Inductive nostutter: list nat → Prop :=
(* FILL IN HERE *)
.
Make sure each of these tests succeeds, but you are free
to change the proof if the given one doesn't work for you.
Your definition might be different from mine and still correct,
in which case the examples might need a different proof.
The suggested proofs for the examples (in comments) use a number
of tactics we haven't talked about, to try to make them robust
with respect to different possible ways of defining nostutter.
You should be able to just uncomment and use them as-is, but if
you prefer you can also prove each example with more basic
tactics.
Example test_nostutter_1: nostutter [3,1,4,1,5,6].
(* FILL IN HERE *) Admitted.
(*
Proof. repeat constructor; apply beq_false_not_eq; auto. Qed.
*)
Example test_nostutter_2: nostutter [].
(* FILL IN HERE *) Admitted.
(*
Proof. repeat constructor; apply beq_false_not_eq; auto. Qed.
*)
Example test_nostutter_3: nostutter [5].
(* FILL IN HERE *) Admitted.
(*
Proof. repeat constructor; apply beq_false_not_eq; auto. Qed.
*)
Example test_nostutter_4: not (nostutter [3,1,1,4]).
(* FILL IN HERE *) Admitted.
(*
Proof. intro.
repeat match goal with
h: nostutter _ |- _ => inversion h; clear h; subst
end.
contradiction H1; auto. Qed.
*)
☐
First a pair of useful lemmas... (we already proved this for lists
of naturals, but not for arbitrary lists.)
Exercise: 4 stars, optional (pigeonhole principle)
The "pigeonhole principle" states a basic fact about counting: if you distribute more than n items into n pigeonholes, some pigeonhole must contain at least two items. As is often the case, this apparently trivial fact about numbers requires non-trivial machinery to prove, but we now have enough...Lemma app_length : ∀{X:Type} (l1 l2 : list X),
length (l1 ++ l2) = length l1 + length l2.
Proof.
(* FILL IN HERE *) Admitted.
Lemma appears_in_app_split : ∀{X:Type} (x:X) (l:list X),
appears_in x l →
∃l1, ∃l2, l = l1 ++ (x::l2).
Proof.
(* FILL IN HERE *) Admitted.
Now define a predicate repeats (analogous to no_repeats in the
exercise above), such that repeats X l asserts that l contains
at least one repeated element (of type X).
Inductive repeats {X:Type} : list X → Prop :=
(* FILL IN HERE *)
.
Now here's a way to formalize the pigeonhole principle. List l2
represents a list of pigeonhole labels, and list l1 represents an
assignment of items to labels: if there are more items than labels,
at least two items must have the same label. You will almost
certainly need to use the excluded_middle hypothesis.
Theorem pigeonhole_principle: ∀{X:Type} (l1 l2:list X),
excluded_middle →
(∀x, appears_in x l1 → appears_in x l2) →
length l2 < length l1 →
repeats l1.
Proof. intros X l1. induction l1.
(* FILL IN HERE *) Admitted.
☐
Informal Proofs
Informal Proofs by Induction
Induction Over an Inductively Defined Set
- Theorem: <Universally quantified proposition of the form
"For all n:S, P(n)," where S is some inductively defined
set.>
- Suppose n = c a1 ... ak, where <...and here we state
the IH for each of the a's that has type S, if any>.
We must show <...and here we restate P(c a1 ... ak)>.
- <other cases similarly...> ☐
- Suppose n = c a1 ... ak, where <...and here we state
the IH for each of the a's that has type S, if any>.
We must show <...and here we restate P(c a1 ... ak)>.
- Theorem: For all sets X, lists l : list X, and numbers
n, if length l = n then index (S n) l = None.
- Suppose l = []. We must show, for all numbers n,
that, if length [] = n, then index (S n) [] =
None.
- Suppose l = x :: l' for some x and l', where
length l' = n' implies index (S n') l' = None, for
any number n'. We must show, for all n, that, if
length (x::l') = n then index (S n) (x::l') =
None.
length l = length (x::l') = S (length l'),it suffices to show thatindex (S (length l')) l' = None.But this follows directly from the induction hypothesis, picking n' to be length l'. ☐
- Suppose l = []. We must show, for all numbers n,
that, if length [] = n, then index (S n) [] =
None.
Induction Over an Inductively Defined Proposition
- Theorem: <Proposition of the form "Q → P," where Q is
some inductively defined proposition (more generally,
"For all x y z, Q x y z → P x y z")>
- Suppose the final rule used to show Q is c. Then
<...and here we state the types of all of the a's
together with any equalities that follow from the
definition of the constructor and the IH for each of
the a's that has type Q, if there are any>. We must
show <...and here we restate P>.
- <other cases similarly...> ☐
- Suppose the final rule used to show Q is c. Then
<...and here we state the types of all of the a's
together with any equalities that follow from the
definition of the constructor and the IH for each of
the a's that has type Q, if there are any>. We must
show <...and here we restate P>.
- Theorem: The <= relation is transitive — i.e., for all
numbers n, m, and o, if n <= m and m <= o, then
n <= o.
- Suppose the final rule used to show m <= o is
le_n. Then m = o and we must show that n <= m,
which is immediate by hypothesis.
- Suppose the final rule used to show m <= o is
le_S. Then o = S o' for some o' with m <= o'.
We must show that n <= S o'.
By induction hypothesis, n <= o'.
- Suppose the final rule used to show m <= o is
le_n. Then m = o and we must show that n <= m,
which is immediate by hypothesis.
Induction Principles for ∧ and ∨
Exercise: 1 star, optional (and_ind_principle)
See if you can predict the induction principle for conjunction.(* Check and_ind. *)
☐
Exercise: 1 star, optional (or_ind_principle)
See if you can predict the induction principle for disjunction.(* Check or_ind. *)
☐
Check and_ind.
From the inductive definition of the proposition and P Q
Inductive and (P Q : Prop) : Prop :=
conj : P → Q → (and P Q).
we might expect Coq to generate this induction principle
conj : P → Q → (and P Q).
and_ind_max :
∀ (P Q : Prop) (P0 : P ∧ Q → Prop),
(∀ (a : P) (b : Q), P0 (conj P Q a b)) →
∀ a : P ∧ Q, P0 a
but actually it generates this simpler and more useful one:
∀ (P Q : Prop) (P0 : P ∧ Q → Prop),
(∀ (a : P) (b : Q), P0 (conj P Q a b)) →
∀ a : P ∧ Q, P0 a
and_ind :
∀ P Q P0 : Prop,
(P → Q → P0) →
P ∧ Q → P0
In the same way, when given the inductive definition of or P Q
∀ P Q P0 : Prop,
(P → Q → P0) →
P ∧ Q → P0
Inductive or (P Q : Prop) : Prop :=
| or_introl : P → or P Q
| or_intror : Q → or P Q.
instead of the "maximal induction principle"
| or_introl : P → or P Q
| or_intror : Q → or P Q.
or_ind_max :
∀ (P Q : Prop) (P0 : P ∨ Q → Prop),
(∀ a : P, P0 (or_introl P Q a)) →
(∀ b : Q, P0 (or_intror P Q b)) →
∀ o : P ∨ Q, P0 o
what Coq actually generates is this:
∀ (P Q : Prop) (P0 : P ∨ Q → Prop),
(∀ a : P, P0 (or_introl P Q a)) →
(∀ b : Q, P0 (or_intror P Q b)) →
∀ o : P ∨ Q, P0 o
or_ind :
∀ P Q P0 : Prop,
(P → P0) →
(Q → P0) →
P ∨ Q → P0
∀ P Q P0 : Prop,
(P → P0) →
(Q → P0) →
P ∨ Q → P0
Explicit Proof Objects for Induction
(* Check nat_ind. *)
(* ===>
nat_ind : forall P : nat -> Prop,
P 0%nat ->
(forall n : nat, P n -> P (S n)) ->
forall n : nat, P n *)
There's nothing magic about this induction lemma: it's just
another Coq lemma that requires a proof. Coq generates the proof
automatically too...
Print nat_ind.
(* ===> (after some manual tidying)
nat_ind =
fun (P : nat -> Type)
(f : P 0)
(f0 : forall n : nat, P n -> P (S n)) =>
fix F (n : nat) : P n :=
match n as n0 return (P n0) with
| 0 => f
| S n0 => f0 n0 (F n0)
end.
*)
We can read this as follows:
Suppose we have evidence f that P holds on 0, and
evidence f0 that ∀ n:nat, P n → P (S n).
Then we can prove that P holds of an arbitrary nat n via
a recursive function F (here defined using the expression
form Fix rather than by a top-level Fixpoint
declaration). F pattern matches on n:
Aside to those interested in functional programming: You may
notice that the match in F requires an annotation as n0
return (P n0) to help Coq's typechecker realize that the two arms
of the match actually return the same type (namely P n). This
is essentially like matching over a GADT (generalized algebraic
datatype) in Haskell. In fact, F has a dependent type: its
result type depends on its argument; GADT's can be used to
describe simple dependent types like this.
We can adapt this approach to proving nat_ind to help prove
non-standard induction principles too. Recall our desire to
prove that
∀ n : nat, even n → ev n.
Attempts to do this by standard induction on n fail, because the
induction principle only lets us proceed when we can prove that
even n → even (S n) — which is of course never provable. What
we did earlier in this chapter was a bit of a hack:
Theorem even__ev : ∀ n : nat,
(even n → ev n) ∧ (even (S n) → ev (S n)).
We can make a much better proof by defining and proving a
non-standard induction principle that goes "by twos":
- If it finds 0, F uses f to show that P n holds.
- If it finds S n0, F applies itself recursively on n0 to obtain evidence that P n0 holds; then it applies f0 on that evidence to show that P (S n) holds.
Definition nat_ind2 :
∀(P : nat → Prop),
P 0 →
P 1 →
(∀n : nat, P n → P (S(S n))) →
∀n : nat , P n :=
fun P => fun P0 => fun P1 => fun PSS =>
fix f (n:nat) := match n return P n with
0 => P0
| 1 => P1
| S (S n') => PSS n' (f n')
end.
Once you get the hang of it, it is entirely straightforward to
give an explicit proof term for induction principles like this.
Proving this as a lemma using tactics is much less intuitive (try
it!).
The induction ... using tactic gives a convenient way to
specify a non-standard induction principle like this.
Lemma even__ev' : ∀n, even n → ev n.
Proof.
intros.
induction n as [ | |n'] using nat_ind2.
Case "even 0".
apply ev_0.
Case "even 1".
inversion H.
Case "even (S(S n'))".
apply ev_SS.
apply IHn'. unfold even. unfold even in H. simpl in H. apply H.
Qed.
The Coq Trusted Computing Base
- Since Coq types can themselves be expressions, the checker must
normalize these (by using the conversion rules) before
comparing them.
- The checker must make sure that match expressions are
exhaustive. That is, there must be an arm for every possible
constructor. To see why, consider the following alleged proof
object:
Definition or_bogus : ∀ P Q, P ∨ Q → P :=All the types here match correctly, but the match only considers one of the possible constructors for or. Coq's exhaustiveness check will reject this definition.
fun (P Q : Prop) (A : P ∨ Q) =>
match A with
| or_introl H => H
end. - The checker must make sure that each fix expression
terminates. It does this using a syntactic check to make sure
that each recursive call is on a subexpression of the original
argument. To see why this is essential, consider this alleged
proof:
Definition nat_false : ∀ (n:nat), False :=Again, this is perfectly well-typed, but (fortunately) Coq will reject it.
fix f (n:nat) : False := f n.