coq-club AT inria.fr
Subject: The Coq mailing list
List archive
- From: Yves Bertot <Yves.Bertot AT sophia.inria.fr>
- To: roconnor AT theorem.ca
- Cc: fr�d�ric BESSON <fbesson AT irisa.fr>, coq-club AT pauillac.inria.fr
- Subject: Re: [Coq-Club] Shortening proofs that the square root of 2 is non rational.
- Date: Wed, 18 Apr 2007 09:56:11 +0200
- List-archive: <http://pauillac.inria.fr/pipermail/coq-club/>
roconnor AT theorem.ca
wrote:
On Mon, 16 Apr 2007, frédéric BESSON wrote:Here is a solution that will please both Frédéric and Russell, but can probably be made better.
Theorem sqrt2_not_rational : ~exists n, exists p, n ^2 = 2* p^2 /\ n <> 0.
Now that we have standard rationals (QArith), I think the statement should read as follows.
Theorem sqrt2_not_rational : ~exists (x:Q), x^2 = 2.
The proof still fits in 40 lines. Notice I use more "asserts" than Frédéric,
hoping that this makes the script almost readable without figuring out what Coq does.
It's a pity micromega is not part of standard Coq.
Yves
===================
Require Import ZArith Zwf Micromegatac QArith.
Open Scope Z_scope.
Lemma Zabs_square : forall x, (Zabs x)^2 = x^2.
Proof.
intros ; case (Zabs_dec x) ; intros ; micromega.
Qed.
Hint Resolve Zabs_pos Zabs_square.
Lemma integer_statement : ~exists n, exists p, n^2 = 2*p^2 /\ n <> 0.
Proof.
intros [n [p [Heq Hnz]]]; pose (n' := Zabs n); pose (p':=Zabs p).
assert (facts : 0 <= Zabs n /\ 0 <= Zabs p /\ Zabs n^2=n^2
/\ Zabs p^2 = p^2) by auto.
assert (H : (0 < n' /\ 0 <= p' /\ n' ^2 = 2* p' ^2))
by (destruct facts as [Hf1 [Hf2 [Hf3 Hf4]]]; unfold n', p'; micromega).
generalize p' H; elim n' using (well_founded_ind (Zwf_well_founded 0)); clear.
intros n IHn p [Hn [Hp Heq]].
assert (Hzwf : Zwf 0 (2*p-n) n) by (unfold Zwf; zrelax; micromega).
assert (Hdecr : 0 < 2*p-n /\ 0 <= n-p /\ (2*p-n)^2=2*(n-p)^2)
by (zrelax; micromega).
apply (IHn (2*p-n) Hzwf (n-p) Hdecr).
Qed.
Open Scope Q_scope.
Theorem sqrt2_not_rational : ~exists x:Q, x^2==2#1.
Proof.
unfold Qeq; intros [x]; simpl (Qden (2#1)); rewrite Zmult_1_r.
intros HQeq.
assert (Heq : (Qnum x ^ 2 = 2 * ' Qden x ^ 2)%Z) by exact HQeq.
assert (Hnx : (Qnum x <> 0)%Z)
by (intros Hx; simpl in HQeq; rewrite Hx in HQeq; discriminate HQeq).
apply integer_statement; exists (Qnum x); exists (' Qden x); auto.
Qed.
- [Coq-Club] Shortening proofs that the square root of 2 is non rational., Yves Bertot
- [Coq-Club] Elliptic curve in Coq, Thery Laurent
- Re: [Coq-Club] Shortening proofs that the square root of 2 is non rational.,
frédéric BESSON
- Re: [Coq-Club] Shortening proofs that the square root of 2 is non rational.,
roconnor
- Re: [Coq-Club] Shortening proofs that the square root of 2 is non rational., Yves Bertot
- Re: [Coq-Club] Shortening proofs that the square root of 2 is non rational.,
roconnor
Archive powered by MhonArc 2.6.16.