2015-06-27 09:13:36 +00:00
|
|
|
|
/-
|
|
|
|
|
Copyright (c) 2015 William Peterson. All rights reserved.
|
|
|
|
|
Released under Apache 2.0 license as described in the file LICENSE.
|
|
|
|
|
Authors: William Peterson, Jeremy Avigad
|
|
|
|
|
|
|
|
|
|
Extended gcd, Bezout's theorem, chinese remainder theorem.
|
|
|
|
|
-/
|
2015-07-04 16:49:14 +00:00
|
|
|
|
import data.nat.div data.int .primes
|
2015-07-04 12:56:55 +00:00
|
|
|
|
|
|
|
|
|
/- Bezout's theorem -/
|
|
|
|
|
|
|
|
|
|
section Bezout
|
|
|
|
|
|
2015-06-27 09:13:36 +00:00
|
|
|
|
open nat int
|
|
|
|
|
open eq.ops well_founded decidable prod
|
2015-10-12 03:35:45 +00:00
|
|
|
|
open algebra
|
2015-06-27 09:13:36 +00:00
|
|
|
|
|
|
|
|
|
private definition pair_nat.lt : ℕ × ℕ → ℕ × ℕ → Prop := measure pr₂
|
|
|
|
|
private definition pair_nat.lt.wf : well_founded pair_nat.lt := intro_k (measure.wf pr₂) 20
|
|
|
|
|
local attribute pair_nat.lt.wf [instance]
|
|
|
|
|
local infixl `≺`:50 := pair_nat.lt
|
|
|
|
|
|
2015-10-29 19:36:26 +00:00
|
|
|
|
private definition gcd.lt.dec (x y₁ : ℕ) : (succ y₁, x % succ y₁) ≺ (x, succ y₁) :=
|
2015-06-27 09:13:36 +00:00
|
|
|
|
!nat.mod_lt (succ_pos y₁)
|
|
|
|
|
|
|
|
|
|
private definition egcd_rec_f (z : ℤ) : ℤ → ℤ → ℤ × ℤ := λ s t, (t, s - t * z)
|
|
|
|
|
|
|
|
|
|
definition egcd.F : Π (p₁ : ℕ × ℕ), (Π p₂ : ℕ × ℕ, p₂ ≺ p₁ → ℤ × ℤ) → ℤ × ℤ
|
|
|
|
|
| (x, y) := nat.cases_on y
|
|
|
|
|
(λ f, (1, 0) )
|
|
|
|
|
(λ y₁ (f : Π p₂, p₂ ≺ (x, succ y₁) → ℤ × ℤ),
|
2015-10-29 19:36:26 +00:00
|
|
|
|
let bz := f (succ y₁, x % succ y₁) !gcd.lt.dec in
|
|
|
|
|
prod.cases_on bz (egcd_rec_f (x / succ y₁)))
|
2015-06-27 09:13:36 +00:00
|
|
|
|
|
|
|
|
|
definition egcd (x y : ℕ) := fix egcd.F (pair x y)
|
|
|
|
|
|
|
|
|
|
theorem egcd_zero (x : ℕ) : egcd x 0 = (1, 0) :=
|
|
|
|
|
well_founded.fix_eq egcd.F (x, 0)
|
|
|
|
|
|
|
|
|
|
theorem egcd_succ (x y : ℕ) :
|
2015-10-29 19:36:26 +00:00
|
|
|
|
egcd x (succ y) = prod.cases_on (egcd (succ y) (x % succ y)) (egcd_rec_f (x / succ y)) :=
|
2015-06-27 09:13:36 +00:00
|
|
|
|
well_founded.fix_eq egcd.F (x, succ y)
|
|
|
|
|
|
|
|
|
|
theorem egcd_of_pos (x : ℕ) {y : ℕ} (ypos : y > 0) :
|
2015-10-29 19:36:26 +00:00
|
|
|
|
let erec := egcd y (x % y), u := pr₁ erec, v := pr₂ erec in
|
|
|
|
|
egcd x y = (v, u - v * (x / y)) :=
|
2015-10-09 01:35:37 +00:00
|
|
|
|
obtain (y' : nat) (yeq : y = succ y'), from exists_eq_succ_of_pos ypos,
|
|
|
|
|
begin
|
|
|
|
|
rewrite [yeq, egcd_succ, -prod.eta (egcd _ _)],
|
|
|
|
|
esimp, unfold egcd_rec_f,
|
|
|
|
|
rewrite [of_nat_div]
|
|
|
|
|
end
|
2015-06-27 09:13:36 +00:00
|
|
|
|
|
|
|
|
|
theorem egcd_prop (x y : ℕ) : (pr₁ (egcd x y)) * x + (pr₂ (egcd x y)) * y = gcd x y :=
|
|
|
|
|
gcd.induction x y
|
2015-10-09 01:35:37 +00:00
|
|
|
|
(take m, by krewrite [egcd_zero, algebra.mul_zero, algebra.one_mul])
|
2015-06-27 09:13:36 +00:00
|
|
|
|
(take m n,
|
|
|
|
|
assume npos : 0 < n,
|
|
|
|
|
assume IH,
|
|
|
|
|
begin
|
|
|
|
|
let H := egcd_of_pos m npos, esimp at H,
|
2015-10-09 01:35:37 +00:00
|
|
|
|
rewrite H,
|
|
|
|
|
esimp,
|
|
|
|
|
rewrite [gcd_rec, -IH],
|
|
|
|
|
rewrite [algebra.add.comm],
|
|
|
|
|
rewrite [-of_nat_mod],
|
2015-10-29 19:36:26 +00:00
|
|
|
|
rewrite [int.mod_def],
|
2015-10-09 01:35:37 +00:00
|
|
|
|
rewrite [+algebra.mul_sub_right_distrib],
|
2015-10-22 20:36:45 +00:00
|
|
|
|
rewrite [+algebra.mul_sub_left_distrib, *left_distrib],
|
2015-10-29 19:36:26 +00:00
|
|
|
|
rewrite [*sub_eq_add_neg, {pr₂ (egcd n (m % n)) * of_nat m + - _}algebra.add.comm],
|
|
|
|
|
rewrite [-algebra.add.assoc ,algebra.mul.assoc]
|
2015-06-27 09:13:36 +00:00
|
|
|
|
end)
|
|
|
|
|
|
2015-07-04 12:56:55 +00:00
|
|
|
|
theorem Bezout_aux (x y : ℕ) : ∃ a b : ℤ, a * x + b * y = gcd x y :=
|
2015-06-27 09:13:36 +00:00
|
|
|
|
exists.intro _ (exists.intro _ (egcd_prop x y))
|
2015-07-04 12:56:55 +00:00
|
|
|
|
|
|
|
|
|
theorem Bezout (x y : ℤ) : ∃ a b : ℤ, a * x + b * y = gcd x y :=
|
|
|
|
|
obtain a' b' (H : a' * nat_abs x + b' * nat_abs y = gcd x y), from !Bezout_aux,
|
|
|
|
|
begin
|
|
|
|
|
existsi (a' * sign x),
|
|
|
|
|
existsi (b' * sign y),
|
2015-10-22 20:36:45 +00:00
|
|
|
|
rewrite [*mul.assoc, -*abs_eq_sign_mul, -*of_nat_nat_abs],
|
2015-07-04 12:56:55 +00:00
|
|
|
|
apply H
|
|
|
|
|
end
|
|
|
|
|
end Bezout
|
|
|
|
|
|
2015-07-08 01:32:23 +00:00
|
|
|
|
/-
|
|
|
|
|
A sample application of Bezout's theorem, namely, an alternative proof that irreducible
|
|
|
|
|
implies prime (dvd_or_dvd_of_prime_of_dvd_mul).
|
|
|
|
|
-/
|
|
|
|
|
|
2015-07-04 12:56:55 +00:00
|
|
|
|
namespace nat
|
2015-10-12 03:35:45 +00:00
|
|
|
|
open int algebra
|
2015-07-04 12:56:55 +00:00
|
|
|
|
|
2015-07-08 01:32:23 +00:00
|
|
|
|
example {p x y : ℕ} (pp : prime p) (H : p ∣ x * y) : p ∣ x ∨ p ∣ y :=
|
2015-07-04 12:56:55 +00:00
|
|
|
|
decidable.by_cases
|
2015-07-24 01:15:24 +00:00
|
|
|
|
(suppose p ∣ x, or.inl this)
|
|
|
|
|
(suppose ¬ p ∣ x,
|
|
|
|
|
have cpx : coprime p x, from coprime_of_prime_of_not_dvd pp this,
|
2015-10-09 01:35:37 +00:00
|
|
|
|
obtain (a b : ℤ) (Hab : a * p + b * x = gcd p x), from Bezout_aux p x,
|
2015-07-24 01:15:24 +00:00
|
|
|
|
assert a * p * y + b * x * y = y,
|
2015-10-22 20:36:45 +00:00
|
|
|
|
by rewrite [-right_distrib, Hab, ↑coprime at cpx, cpx, int.one_mul],
|
2015-07-24 01:15:24 +00:00
|
|
|
|
have p ∣ y,
|
2015-07-04 12:56:55 +00:00
|
|
|
|
begin
|
|
|
|
|
apply dvd_of_of_nat_dvd_of_nat,
|
2015-07-24 01:15:24 +00:00
|
|
|
|
rewrite [-this],
|
2015-10-09 01:35:37 +00:00
|
|
|
|
apply @dvd_add,
|
|
|
|
|
{apply dvd_mul_of_dvd_left,
|
|
|
|
|
apply dvd_mul_of_dvd_right,
|
|
|
|
|
apply dvd.refl},
|
2015-10-22 20:36:45 +00:00
|
|
|
|
{rewrite mul.assoc,
|
2015-10-09 01:35:37 +00:00
|
|
|
|
apply dvd_mul_of_dvd_right,
|
2015-07-04 12:56:55 +00:00
|
|
|
|
apply of_nat_dvd_of_nat_of_dvd H}
|
|
|
|
|
end,
|
2015-07-24 01:15:24 +00:00
|
|
|
|
or.inr this)
|
2015-07-04 12:56:55 +00:00
|
|
|
|
|
|
|
|
|
end nat
|