Refactor even/odd stuff

It's now broken down into multiple lemmas, and my weird tactics are
introduced along the way.

I also changed some style things later on.
This commit is contained in:
Jarvis Carroll 2025-11-22 02:03:51 +11:00
parent 4d8d3e5f81
commit 48136669a9

380
pascal.v
View File

@ -1,3 +1,6 @@
Require Import Ring.
Require Import ArithRing.
(* Coq is a powerful language, that can search for proofs automatically, and
then verify them automatically. This results in proofs like *)
@ -31,97 +34,30 @@ Qed.
A number is even if it is equal to some other number, doubled. *)
Definition Even (n: nat): Prop :=
exists k: nat, 2 * k = n.
exists k: nat, n = 2 * k.
(* And a number is odd if it is one more than that. *)
Definition Odd (n: nat): Prop :=
exists k: nat, 2 * k + 1 = n.
exists k: nat, n = 2 * k + 1.
(* Now let's get into the proof. We are going to try to use automation more
tastefully this time, to only do mundane manipulations. Let's import the
'Ring' tactics, which is where Coq defines how to automatically simplify
expressions built out of + and * operations. *)
(* Let's try proving a number is even. We'll do zero, since then we have a base
case for induction later. *)
Require Import Ring.
Require Import ArithRing.
(* Now let's get into the proof. This proof is chosen because it is a good
demonstration of a lot of the features, strengths, and weaknesses of Coq,
all in one place. It will be easier to follow than an 'auto' one-liner, but
it still features a lot of the hidden state that is typical of Coq proofs.
Feel free to read through by yourself, or to step through with the IDE to
see what that hidden state actually looks like on each line, or once you get
the point, that hidden state is hard to follow, feel free to move on to
where we try and address exactly that, by making the state more explicit. *)
Theorem even_or_odd: forall n: nat,
Even n \/ Odd n.
Lemma zero_is_even: Even 0.
Proof.
intro n.
(* This is an inductive proof. We prove 0 is even or odd, then 1 is even or
odd, then 2 is even or odd, etc. *)
induction n.
(* First prove 0 is even or odd... Let's go with even! *)
*left.
(* We have to choose k so that 2*k = 0. *)
exists 0.
(* Now we have to actually prove that 2*k = 0. In DTT this simplifies to
0 = 0, so we can use reflexivity on 0. *)
reflexivity.
(* Now the inductive case. If n is even or odd, we want to show that n+1 is
even or odd. Check which case we are in first. *)
*destruct IHn as [IsEven | IsOdd].
(* The first case is that n is even, in which case n+1 will be odd. *)
+right.
(* So, we know 2*k = n, and we want to show 2*k' + 1 = n + 1... Clearly
k' = k. *)
destruct IsEven as [k Eq].
exists k.
(* Now Coq gives us the goal 2*k + 1 = n + 1... Let's substitute n in. *)
rewrite <- Eq.
(* Now the goal is 2*k + 1 = 1 + 2 * k. This is clearly true by the
algebraic properties of + and *. We use the 'ring' tactic, even though
natural numbers are only a semiring. The logic still works! *)
ring.
(* That handles the even case. Now, if n is odd, we want to show that n+1 is
even. *)
+left.
(* n is odd, i.e. n = 2*k+1, and we want to show n+1 is even, i.e.
n+1 = 2*k'. Choose k' = k + 1. *)
destruct IsOdd as [k Eq].
exists (k + 1).
(* The goal is 2*(k+1) = 1+n. Substitute n in. *)
rewrite <- Eq.
(* Now the goal is 2*(k+1) = 1 + (2*k + 1)... Let Coq do the algebra. *)
ring.
(* We need to find a k so that 2 * k = 0. Let's try.... 0 *)
exists 0.
(* Now to show that 2 * 0 = 0, we let Coq beta/eta reduce it to 0 = 0, and
then fall back to 'reflexivity', the defining property of Leibniz equality. *)
reflexivity.
Qed.
(* So we see that more explicit control flow sets us up for quite nice proofs,
where we can see that the goal follows by simple algebra, algebra which Coq
is prepared to do for us.
Agda and Idris programmers will recognize that Even and Odd could have been
jerry-rigged so that the `ring` tactics weren't required at all, and while
learning that sort of trick is immensely productive for understanding DTT,
it is not very useful for understanding number theory.
The problem remains, however, that we had to write comments to track what
the state of the program was going to be. That is confusing, even with
comments, and who knows if the comments are up to date! They aren't checked,
despite us being in a language specifically designed to check things like
that. Let's make some tactics that let us declare what algebraic expression
we think we are working with, so that our proofs become just a little more
self-documenting.
A very common pattern when proving relations in real math proofs, is to
start with one side of the relation, and to manipulate it bit by bit, until
it matches the other side that we wanted. As long as each manipulation
proves a transitive relation, all those transitive steps can be composed
together into a bigger proof. Let's make some tactics that let us do that,
when our goal is some relation L = R, or Rel L R more generally. *)
(* It is annoying that we have to write all of these comments to explain what
the state of the proof assistant is *going* to be. Plus, these comments can
get out of sync with the program, which seems absurd in a language
specifically designed to check such things. Let's make some new tactics to
let us represent this information better. *)
(* This first tactic lets us start such a string of manipulations, by declaring
what the LHS of our goal is, effectively just a comment, but checked. *)
@ -147,10 +83,8 @@ Ltac want expr :=
fail "Goal is not a relation. Goal:" Other
end.
(* Next, all algebraic manipulations in type theory must be shown to follow
from the lambda calculus rules of beta/eta reduction, or similar, so let's
make a tactic that lets us manipulate our LHS using whatever beta/eta rules
Coq supports. *)
(* Next, all manipulations are ultimately built out of beta/eta equivalence,
so let's make a tactic to check these equivalences for us. *)
Ltac thatis expr :=
lazymatch goal with
| |- ?Rel ?L ?R =>
@ -160,24 +94,9 @@ Ltac thatis expr :=
fail "Goal is not a relation. Goal:" Other
end.
(* If beta/eta isn't enough, then we can often specify an intermediate
expression, and get Coq to step there using ring algebra. *)
Ltac rearrange expr :=
lazymatch goal with
| |- ?Rel ?L ?R =>
transitivity expr; [> ring | ]
| |- ?Other =>
fail "Goal is not a relation. Goal:" Other
end.
(* Finally, if we manipulate L using thatis or rewrite tactics, then we will
eventually get to the point where it matches R exactly, at which point one
normally uses the 'reflexivity' tactic, but that is confusing, if we had to
do manipulations precisely because our goal was not reflexive, and further,
reflexivity might sneak more beta/eta manipulations in, making our
annotations irrelevant. This tactic lets us check that L and R match
exactly, no beta/eta required, and only then uses reflexivity to finish the
proof. *)
(* Finally, if we want to be sure that our above annotations are correct, this
tactic will check that the LHS and RHS are not just equivalent, but already
identical. *)
Ltac done :=
lazymatch goal with
| |- ?Rel ?L ?L =>
@ -188,6 +107,63 @@ Ltac done :=
fail "Goal is not a relation. Goal:" Other
end.
(* Let's try again at that simple zero example, but with annotations. *)
Lemma zero_is_even_explicit: Even 0.
Proof.
(* We still need to pick a value for k. *)
exists 0.
(* Now let's fill our goal a little more explicitly. *)
want (2 * 0).
have 0.
(* So that is our RHS and LHS... How are we ever going to connect them?
Well let's compute our LHS a little! *)
thatis (2 * 0).
(* And that matches our goal! Look! *)
done.
Qed.
(* Probably overkill for that case, and frankly unintuitive when all the
computation happens on the RHS. Oh well, let's try something more involved. *)
Lemma even_implies_odd: forall n: nat,
Even n -> Odd (S n).
Proof.
intros n IsEven.
destruct IsEven as [k Eq].
(* n is 2 * k, and we need to show that n+1 = 2*k' + 1, I wonder what k' can
be? *)
exists k.
want (2 * k + 1).
have (S n).
(* Now let's do some real tactical proving! The equation that proves n is
even can now be used to manipulate our term. *)
rewrite Eq.
(* Now let's see what that did to our LHS. *)
have (S (2 * k)).
(* This isn't quite our LHS. Let's dig into some natural number identities. *)
rewrite (plus_n_O (2 * k)) at 1.
have (S (2 * k + 0)).
rewrite (plus_n_Sm (2 * k) 0).
have (2 * k + 1).
done.
Qed.
(* Rewriting n into 2 * k with these tactics was nice, but the identities at
the end sucked. Let's bring back a little more automation; just enough to do
algebra, but without hiding any important control flow. *)
Ltac rearrange expr :=
lazymatch goal with
| |- ?Rel ?L ?R =>
transitivity expr; [> ring | ]
| |- ?Other =>
fail "Goal is not a relation. Goal:" Other
end.
(* Before we try using this one, let's also define some things that will help
us annotate these Even/Odd/exists manipulations. *)
(* If our goal isn't a relation yet, we might still want to declare it. *)
Ltac havegoal expr :=
lazymatch goal with
@ -197,8 +173,9 @@ Ltac havegoal expr :=
fail "Expected goal to be" expr "but it was" Other
end.
(* When manipulating dysjunctions, like the Even/Odd case coming up, it can
be useful to restate what our current hypotheses are. *)
(* The place where we learn new equations is often quite far from where we
use them. This tactic lets us state in-context what hypothesis we are
using, and defer choosing a name for it until then as well. *)
Ltac weknow newid P :=
lazymatch goal with
| proof: P |- ?Goal =>
@ -210,52 +187,64 @@ Ltac weknow newid P :=
Tactic Notation "weknow" ident(newid) ":" constr(P) :=
weknow newid P.
(* Now let's try and prove our theorem again, but using these self-documenting
tactics. *)
Theorem even_or_odd_chained: forall n: nat,
Even n \/ Odd n.
(* We're going to need these for odd_implies_even, as that case is harder! *)
Lemma odd_implies_even: forall n: nat,
Odd n -> Even (S n).
Proof.
induction n.
*havegoal (Even 0 \/ Odd 0).
left; havegoal (Even 0). (* Choose (Even 0) as a new goal. *)
exists 0; havegoal (2 * 0 = 0).
have (2 * 0).
thatis 0.
done.
*weknow IndCase: (Even n \/ Odd n).
destruct IndCase.
+havegoal (Even (S n) \/ Odd (S n)).
weknow IsEven: (Even n).
right; havegoal (Odd (S n)).
destruct IsEven as [k Eq].
weknow Eq: (2 * k = n).
exists k; havegoal (2 * k + 1 = S n).
apply eq_sym. (* Flip the goal, to work right to left. *)
have (S n).
rewrite <- Eq.
have (S (2 * k)).
rearrange (2 * k + 1).
done.
+havegoal (Even (S n) \/ Odd (S n)).
weknow IsOdd: (Odd n).
left; havegoal (Even (S n)).
destruct IsOdd as [k Eq].
exists (S k).
havegoal (2 * S k = S n).
apply eq_sym.
have (S n).
rewrite <- Eq.
have (S (2 * k + 1)).
rearrange (2 * S k).
done.
intros.
weknow IsOdd: (Odd n).
destruct IsOdd as [k].
havegoal (Even (S n)).
weknow OddEq: (n = 2 * k + 1).
(* From now on let's do the goal manipulation, and then justify it afterwards
by annotating what our goal is now. *)
exists (S k); havegoal (S n = 2 * S k).
have (S n).
(* Let's also combine rewrites with the new expression, if they'll fit on one line. *)
rewrite OddEq; have (S (2 * k + 1)).
(* Now let's write the algebra we would write on paper, and let Coq justify it. *)
rearrange (2 * S k).
done.
Qed.
(* That was much longer, but now the proof guides you through its own state,
and if you edit this file so that any of the annotations are wrong, the
proof no longer completes! In general our proofs are going to be pretty long
from now on, but hopefully they are much more pleasant to read.
(* You might see where this is going now. We have all the pieces to build an
inductive proof! *)
Now, finally, let's get on with some number theory! First let's look at
Theorem even_or_odd: forall n: nat,
Even n \/ Odd n.
Proof.
intro n.
induction n.
*havegoal (Even 0 \/ Odd 0).
left; havegoal (Even 0).
exact zero_is_even.
*havegoal (Even (S n) \/ Odd (S n)).
weknow Ind: (Even n \/ Odd n).
destruct Ind.
+weknow IsEven: (Even n).
right; havegoal (Odd (S n)).
exact (even_implies_odd n IsEven).
+weknow IsOdd: (Odd n).
left; havegoal (Even (S n)).
exact (odd_implies_even n IsOdd).
Qed.
(* Wasn't that a fun and pleasant adventure?
So we see that more explicit control flow sets us up for quite nice proofs,
where we can see that the goal follows by simple algebra, algebra which Coq
is prepared to do for us.
Agda and Idris programmers will recognize that Even and Odd could have been
jerry-rigged so that the `ring` tactics weren't required at all, and while
we will be exploiting that sort of trick to simplify things when they get
hard, but for this simple example, having suboptimal definitions was useful
as it motivated us to learn how to make Coq do the boring parts for us. *)
(* Now, finally, let's get on with some number theory! First let's look at
Pascal's triangle, since understanding that turns out to be essential to
understanding how to find points on elliptic curves.
@ -281,17 +270,20 @@ Fixpoint partial_sum (seq: nat -> nat) (count: nat): nat :=
end.
(* But Triangular numbers aren't the sum of 0 or more numbers, they are the sum
of one or more numbers! So let's also define a 'sequence' variation. *)
of one or more numbers! So let's also define a 'series' variation. *)
Definition sequence (seq: nat -> nat) (index: nat): nat :=
Definition series (seq: nat -> nat) (index: nat): nat :=
partial_sum seq (1 + index).
Fixpoint pascal (r: nat): nat -> nat :=
match r with
| 0 => fun _ => 1
| S r' => sequence (pascal r')
| S r' => series (pascal r')
end.
(* Now the first 'row' of our function is all 1 by definition, but it is useful
to show that the first 'column' is all 1 as well. *)
Lemma pascal_0: forall r: nat,
pascal r 0 = 1.
Proof.
@ -301,13 +293,16 @@ Proof.
thatis 1.
done.
*have (pascal (S r) 0).
thatis (sequence (pascal r) 0).
thatis (series (pascal r) 0).
thatis (pascal r 0).
want 1.
weknow Ind: (pascal r 0 = 1).
exact Ind.
Qed.
(* Next let's define some combinatoric functions, and see if we can't prove how
they relate to the recursive formula above. *)
Fixpoint factorial (n : nat): nat :=
match n with
| S n' => n * factorial n'
@ -320,41 +315,9 @@ Fixpoint prod_up (k r: nat): nat :=
| S k' => prod_up k' (1 + r) * r
end.
Lemma prod_up_down: forall k r: nat,
prod_up (1 + k) r = (k + r) * prod_up k r.
Proof.
intro k.
induction k; intro r.
*cbn.
havegoal (r + 0 = r * 1).
ring.
*have (prod_up (1 + S k) r).
thatis (prod_up (1 + k) (1 + r) * r).
weknow Ind: (forall r': nat, prod_up (1 + k) r' = (k + r') * prod_up k r').
rewrite (Ind (1 + r)).
have ((k + (1 + r)) * prod_up k (1 + r) * r).
rearrange ((S k + r) * (prod_up k (1 + r) * r)).
thatis ((S k + r) * prod_up (S k) r).
done.
Qed.
Lemma factorial_up: forall n: nat,
factorial n = prod_up n 1.
Proof.
intro n.
induction n.
*reflexivity.
*have (factorial (S n)).
thatis ((S n) * factorial n).
weknow Ind: (factorial n = prod_up n 1).
rewrite Ind.
have ((S n) * prod_up n 1).
rearrange ((n + 1) * prod_up n 1).
rewrite <- (prod_up_down n 1).
have (prod_up (1 + n) 1).
thatis (prod_up (S n) 1).
done.
Qed.
(* Often this permutation prod_up is notated as factorial (k + r) / factorial k,
but we don't have division in peano arithmetic, so let's show this relation
through multiplication instead. *)
Lemma prod_up_ratio: forall k r: nat,
prod_up k (1 + r) * factorial r = factorial (k + r).
@ -376,6 +339,45 @@ Proof.
done.
Qed.
(* If we let r = 0 then we even show that the factorial is a special case of
this product function. *)
Lemma factorial_up: forall n: nat,
prod_up n 1 = factorial n.
Proof.
intro n.
have (prod_up n 1).
rearrange (prod_up n 1 * 1).
thatis (prod_up n (1 + 0) * factorial 0).
rewrite (prod_up_ratio n 0); have (factorial (n + 0)).
rewrite <- (plus_n_O); have (factorial n).
done.
Qed.
(* Another trick that is implicit in pen-and-paper notation, that we have to be
explicit about in code, is the fact that we can pull terms off of either
side of this product, to do algebraic manipulations with. *)
Lemma prod_up_down: forall k r: nat,
prod_up (1 + k) r = (k + r) * prod_up k r.
Proof.
intro k.
induction k; intro r.
*cbn.
havegoal (r + 0 = r * 1).
ring.
*have (prod_up (1 + S k) r).
thatis (prod_up (1 + k) (1 + r) * r).
weknow Ind: (forall r': nat, prod_up (1 + k) r' = (k + r') * prod_up k r').
rewrite (Ind (1 + r)).
have ((k + (1 + r)) * prod_up k (1 + r) * r).
rearrange ((S k + r) * (prod_up k (1 + r) * r)).
thatis ((S k + r) * prod_up (S k) r).
done.
Qed.
(* Now let's get to the theorem! I haven't broken this one up at all. Turn your
font up or something. *)
Theorem pascal_choose: forall r k: nat,
pascal r k * factorial k = prod_up k (S r).
Proof.
@ -383,7 +385,7 @@ Proof.
induction r; intro k.
*cbn.
ring_simplify.
exact (factorial_up k).
exact (eq_sym (factorial_up k)).
*induction k.
**cbn.
ring_simplify.
@ -416,3 +418,5 @@ Proof.
done.
Qed.