Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit bf9bbbc

Browse files
committed
feat(field_theory/ratfunc): The numerator and denominator of a rational function are coprime (#18652)
Also make more arguments to `gcd_ne_zero_of_left`/`gcd_ne_zero_of_right` implicit. Co-authored-by: Bhavik Mehta <bhavikmehta8@gmail.com>
1 parent 88474d1 commit bf9bbbc

File tree

3 files changed

+21
-5
lines changed

3 files changed

+21
-5
lines changed

src/algebra/euclidean_domain/basic.lean

Lines changed: 3 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -84,6 +84,9 @@ begin
8484
rw [mul_div_cancel_left _ hz, mul_left_comm, mul_div_cancel_left _ hz]
8585
end
8686

87+
protected lemma mul_div_cancel' {a b : R} (hb : b ≠ 0) (hab : b ∣ a) : b * (a / b) = a :=
88+
by rw [←mul_div_assoc _ hab, mul_div_cancel_left _ hb]
89+
8790
@[simp, priority 900] -- This generalizes `int.div_one`, see note [simp-normal form]
8891
lemma div_one (p : R) : p / 1 = p :=
8992
(euclidean_domain.eq_div_of_mul_eq_left (one_ne_zero' R) (mul_one p)).symm

src/field_theory/ratfunc.lean

Lines changed: 8 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -1030,6 +1030,14 @@ x.induction_on (λ p q hq, begin
10301030
exact inv_ne_zero (polynomial.leading_coeff_ne_zero.mpr q_div_ne_zero) },
10311031
end)
10321032

1033+
lemma is_coprime_num_denom (x : ratfunc K) : is_coprime x.num x.denom :=
1034+
begin
1035+
induction x using ratfunc.induction_on with p q hq,
1036+
rw [num_div, denom_div _ hq],
1037+
exact (is_coprime_mul_unit_left ((leading_coeff_ne_zero.2 $ right_div_gcd_ne_zero
1038+
hq).is_unit.inv.map C) _ _).2 (is_coprime_div_gcd_div_gcd hq),
1039+
end
1040+
10331041
@[simp] lemma num_eq_zero_iff {x : ratfunc K} : num x = 0 ↔ x = 0 :=
10341042
⟨λ h, by rw [← num_div_denom x, h, ring_hom.map_zero, zero_div],
10351043
λ h, h.symm ▸ num_zero⟩

src/ring_theory/euclidean_domain.lean

Lines changed: 10 additions & 5 deletions
Original file line numberDiff line numberDiff line change
@@ -29,14 +29,12 @@ open euclidean_domain set ideal
2929

3030
section gcd_monoid
3131

32-
variables {R : Type*} [euclidean_domain R] [gcd_monoid R]
32+
variables {R : Type*} [euclidean_domain R] [gcd_monoid R] {p q : R}
3333

34-
lemma gcd_ne_zero_of_left (p q : R) (hp : p ≠ 0) :
35-
gcd_monoid.gcd p q ≠ 0 :=
34+
lemma gcd_ne_zero_of_left (hp : p ≠ 0) : gcd_monoid.gcd p q ≠ 0 :=
3635
λ h, hp $ eq_zero_of_zero_dvd (h ▸ gcd_dvd_left p q)
3736

38-
lemma gcd_ne_zero_of_right (p q : R) (hp : q ≠ 0) :
39-
gcd_monoid.gcd p q ≠ 0 :=
37+
lemma gcd_ne_zero_of_right (hp : q ≠ 0) : gcd_monoid.gcd p q ≠ 0 :=
4038
λ h, hp $ eq_zero_of_zero_dvd (h ▸ gcd_dvd_right p q)
4139

4240
lemma left_div_gcd_ne_zero {p q : R} (hp : p ≠ 0) :
@@ -57,6 +55,13 @@ begin
5755
exact r0,
5856
end
5957

58+
lemma is_coprime_div_gcd_div_gcd (hq : q ≠ 0) :
59+
is_coprime (p / gcd_monoid.gcd p q) (q / gcd_monoid.gcd p q) :=
60+
(gcd_is_unit_iff _ _).1 $ is_unit_gcd_of_eq_mul_gcd
61+
(euclidean_domain.mul_div_cancel' (gcd_ne_zero_of_right hq) $ gcd_dvd_left _ _).symm
62+
(euclidean_domain.mul_div_cancel' (gcd_ne_zero_of_right hq) $ gcd_dvd_right _ _).symm $
63+
gcd_ne_zero_of_right hq
64+
6065
end gcd_monoid
6166

6267
namespace euclidean_domain

0 commit comments

Comments
 (0)