Skip to content

Commit

Permalink
fix: make EuclideanDomain.gcdMonoid computable (#6076)
Browse files Browse the repository at this point in the history
This of course still works non-computably in the presence of `Classical.decEq`, but the caller now has the option.
  • Loading branch information
eric-wieser committed Jul 29, 2023
1 parent 91917cb commit e6e9a9e
Showing 1 changed file with 7 additions and 8 deletions.
15 changes: 7 additions & 8 deletions Mathlib/RingTheory/EuclideanDomain.lean
Expand Up @@ -23,9 +23,7 @@ euclidean domain
-/


noncomputable section

open Classical
section

open EuclideanDomain Set Ideal

Expand Down Expand Up @@ -71,7 +69,8 @@ end GCDMonoid
namespace EuclideanDomain

/-- Create a `GCDMonoid` whose `GCDMonoid.gcd` matches `EuclideanDomain.gcd`. -/
def gcdMonoid (R) [EuclideanDomain R] : GCDMonoid R where
-- porting note: added `DecidableEq R`
def gcdMonoid (R) [EuclideanDomain R] [DecidableEq R] : GCDMonoid R where
gcd := gcd
lcm := lcm
gcd_dvd_left := gcd_dvd_left
Expand All @@ -84,26 +83,26 @@ def gcdMonoid (R) [EuclideanDomain R] : GCDMonoid R where

variable {α : Type _} [EuclideanDomain α] [DecidableEq α]

theorem span_gcd {α} [EuclideanDomain α] (x y : α) :
theorem span_gcd (x y : α) :
span ({gcd x y} : Set α) = span ({x, y} : Set α) :=
letI := EuclideanDomain.gcdMonoid α
_root_.span_gcd x y
#align euclidean_domain.span_gcd EuclideanDomain.span_gcd

theorem gcd_isUnit_iff {α} [EuclideanDomain α] {x y : α} : IsUnit (gcd x y) ↔ IsCoprime x y :=
theorem gcd_isUnit_iff {x y : α} : IsUnit (gcd x y) ↔ IsCoprime x y :=
letI := EuclideanDomain.gcdMonoid α
_root_.gcd_isUnit_iff x y
#align euclidean_domain.gcd_is_unit_iff EuclideanDomain.gcd_isUnit_iff

-- this should be proved for UFDs surely?
theorem isCoprime_of_dvd {α} [EuclideanDomain α] {x y : α} (nonzero : ¬(x = 0 ∧ y = 0))
theorem isCoprime_of_dvd {x y : α} (nonzero : ¬(x = 0 ∧ y = 0))
(H : ∀ z ∈ nonunits α, z ≠ 0 → z ∣ x → ¬z ∣ y) : IsCoprime x y :=
letI := EuclideanDomain.gcdMonoid α
_root_.isCoprime_of_dvd x y nonzero H
#align euclidean_domain.is_coprime_of_dvd EuclideanDomain.isCoprime_of_dvd

-- this should be proved for UFDs surely?
theorem dvd_or_coprime {α} [EuclideanDomain α] (x y : α) (h : Irreducible x) :
theorem dvd_or_coprime (x y : α) (h : Irreducible x) :
x ∣ y ∨ IsCoprime x y :=
letI := EuclideanDomain.gcdMonoid α
_root_.dvd_or_coprime x y h
Expand Down

0 comments on commit e6e9a9e

Please sign in to comment.