Skip to content

Commit

Permalink
feat: The numerator and denominator of a rational function are coprime
Browse files Browse the repository at this point in the history
  • Loading branch information
YaelDillies committed Jun 16, 2023
1 parent f1adfdf commit a4bcc12
Show file tree
Hide file tree
Showing 3 changed files with 27 additions and 6 deletions.
6 changes: 5 additions & 1 deletion Mathlib/Algebra/EuclideanDomain/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Louis Carlin, Mario Carneiro
! This file was ported from Lean 3 source module algebra.euclidean_domain.basic
! leanprover-community/mathlib commit e8638a0fcaf73e4500469f368ef9494e495099b3
! leanprover-community/mathlib commit bf9bbbcf0c1c1ead18280b0d010e417b10abb1b6

Check notice on line 7 in Mathlib/Algebra/EuclideanDomain/Basic.lean

View workflow job for this annotation

GitHub Actions / Add annotations

Synchronization

See https://leanprover-community.github.io/mathlib-port-status/file/algebra/euclidean_domain/basic?range=e8638a0fcaf73e4500469f368ef9494e495099b3..bf9bbbcf0c1c1ead18280b0d010e417b10abb1b6
! Please do not edit these lines, except to modify the commit id
! if you have ported upstream changes.
-/
Expand Down Expand Up @@ -107,6 +107,10 @@ theorem mul_div_assoc (x : R) {y z : R} (h : z ∣ y) : x * y / z = x * (y / z)
rw [mul_div_cancel_left _ hz, mul_left_comm, mul_div_cancel_left _ hz]
#align euclidean_domain.mul_div_assoc EuclideanDomain.mul_div_assoc

protected theorem mul_div_cancel' {a b : R} (hb : b ≠ 0) (hab : b ∣ a) : b * (a / b) = a := by
rw [← mul_div_assoc _ hab, mul_div_cancel_left _ hb]
#align euclidean_domain.mul_div_cancel' EuclideanDomain.mul_div_cancel'

-- This generalizes `Int.div_one`, see note [simp-normal form]
@[simp]
theorem div_one (p : R) : p / 1 = p :=
Expand Down
10 changes: 9 additions & 1 deletion Mathlib/FieldTheory/RatFunc.lean
Original file line number Diff line number Diff line change
Expand Up @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Anne Baanen
! This file was ported from Lean 3 source module field_theory.ratfunc
! leanprover-community/mathlib commit 67237461d7cbf410706a6a06f4d40cbb594c32dc
! leanprover-community/mathlib commit bf9bbbcf0c1c1ead18280b0d010e417b10abb1b6

Check notice on line 7 in Mathlib/FieldTheory/RatFunc.lean

View workflow job for this annotation

GitHub Actions / Add annotations

Synchronization

See https://leanprover-community.github.io/mathlib-port-status/file/field_theory/ratfunc?range=67237461d7cbf410706a6a06f4d40cbb594c32dc..bf9bbbcf0c1c1ead18280b0d010e417b10abb1b6
! Please do not edit these lines, except to modify the commit id
! if you have ported upstream changes.
-/
Expand Down Expand Up @@ -1248,6 +1248,14 @@ theorem num_div_denom (x : RatFunc K) : algebraMap _ _ (num x) / algebraMap _ _
exact inv_ne_zero (Polynomial.leadingCoeff_ne_zero.mpr q_div_ne_zero)
#align ratfunc.num_div_denom RatFunc.num_div_denom

theorem isCoprime_num_denom (x : RatFunc K) : IsCoprime x.num x.den := by

Check failure on line 1251 in Mathlib/FieldTheory/RatFunc.lean

View workflow job for this annotation

GitHub Actions / Build

invalid field 'den', the environment does not contain 'RatFunc.den'
induction' x using RatFunc.induction_on with p q hq
rw [num_div, denom_div _ hq]

Check failure on line 1253 in Mathlib/FieldTheory/RatFunc.lean

View workflow job for this annotation

GitHub Actions / Build

tactic 'rewrite' failed, did not find instance of the pattern in the target expression
exact (isCoprime_mul_unit_left
((leading_coeff_ne_zero.2 <| right_div_gcd_ne_zero hq).IsUnit.inv.map C) _ _).2
(isCoprime_div_gcd_div_gcd hq)
#align ratfunc.is_coprime_num_denom RatFunc.isCoprime_num_denom

@[simp]
theorem num_eq_zero_iff {x : RatFunc K} : num x = 0 ↔ x = 0 :=
fun h => by rw [← num_div_denom x, h, RingHom.map_zero, zero_div], fun h => h.symm ▸ num_zero⟩
Expand Down
17 changes: 13 additions & 4 deletions Mathlib/RingTheory/EuclideanDomain.lean
Original file line number Diff line number Diff line change
Expand Up @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Mario Carneiro, Chris Hughes
! This file was ported from Lean 3 source module ring_theory.euclidean_domain
! leanprover-community/mathlib commit 70fd9563a21e7b963887c9360bd29b2393e6225a
! leanprover-community/mathlib commit bf9bbbcf0c1c1ead18280b0d010e417b10abb1b6

Check notice on line 7 in Mathlib/RingTheory/EuclideanDomain.lean

View workflow job for this annotation

GitHub Actions / Add annotations

Synchronization

See https://leanprover-community.github.io/mathlib-port-status/file/ring_theory/euclidean_domain?range=70fd9563a21e7b963887c9360bd29b2393e6225a..bf9bbbcf0c1c1ead18280b0d010e417b10abb1b6
! Please do not edit these lines, except to modify the commit id
! if you have ported upstream changes.
-/
Expand Down Expand Up @@ -34,13 +34,13 @@ open EuclideanDomain Set Ideal

section GCDMonoid

variable {R : Type _} [EuclideanDomain R] [GCDMonoid R]
variable {R : Type _} [EuclideanDomain R] [GCDMonoid R] {p q : R}

theorem gcd_ne_zero_of_left (p q : R) (hp : p ≠ 0) : GCDMonoid.gcd p q ≠ 0 := fun h =>
theorem gcd_ne_zero_of_left (hp : p ≠ 0) : GCDMonoid.gcd p q ≠ 0 := fun h =>
hp <| eq_zero_of_zero_dvd (h ▸ gcd_dvd_left p q)
#align gcd_ne_zero_of_left gcd_ne_zero_of_left

theorem gcd_ne_zero_of_right (p q : R) (hp : q ≠ 0) : GCDMonoid.gcd p q ≠ 0 := fun h =>
theorem gcd_ne_zero_of_right (hp : q ≠ 0) : GCDMonoid.gcd p q ≠ 0 := fun h =>
hp <| eq_zero_of_zero_dvd (h ▸ gcd_dvd_right p q)
#align gcd_ne_zero_of_right gcd_ne_zero_of_right

Expand All @@ -60,6 +60,15 @@ theorem right_div_gcd_ne_zero {p q : R} (hq : q ≠ 0) : q / GCDMonoid.gcd p q
exact r0
#align right_div_gcd_ne_zero right_div_gcd_ne_zero

theorem isCoprime_div_gcd_div_gcd (hq : q ≠ 0) :
IsCoprime (p / GCDMonoid.gcd p q) (q / GCDMonoid.gcd p q) :=
(gcd_isUnit_iff _ _).1 <|
isUnit_gcd_of_eq_mul_gcd
(EuclideanDomain.mul_div_cancel' (gcd_ne_zero_of_right hq) <| gcd_dvd_left _ _).symm
(EuclideanDomain.mul_div_cancel' (gcd_ne_zero_of_right hq) <| gcd_dvd_right _ _).symm <|
gcd_ne_zero_of_right hq
#align is_coprime_div_gcd_div_gcd isCoprime_div_gcd_div_gcd

end GCDMonoid

namespace EuclideanDomain
Expand Down

0 comments on commit a4bcc12

Please sign in to comment.