Skip to content

Commit afa2032

Browse files
chore: move Associates.quot_out earlier (#8484)
And don't require commutativity while we're here.
1 parent b9a6d49 commit afa2032

File tree

2 files changed

+5
-4
lines changed

2 files changed

+5
-4
lines changed

Mathlib/Algebra/Associated.lean

Lines changed: 5 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -779,6 +779,11 @@ theorem quot_mk_eq_mk [Monoid α] (a : α) : Quot.mk Setoid.r a = Associates.mk
779779
rfl
780780
#align associates.quot_mk_eq_mk Associates.quot_mk_eq_mk
781781

782+
@[simp]
783+
theorem quot_out [Monoid α] (a : Associates α) : Associates.mk (Quot.out a) = a := by
784+
rw [← quot_mk_eq_mk, Quot.out_eq]
785+
#align associates.quot_out Associates.quot_outₓ
786+
782787
theorem forall_associated [Monoid α] {p : Associates α → Prop} :
783788
(∀ a, p a) ↔ ∀ a, p (Associates.mk a) :=
784789
Iff.intro (fun h _ => h _) fun h a => Quotient.inductionOn a h

Mathlib/RingTheory/UniqueFactorizationDomain.lean

Lines changed: 0 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -1920,10 +1920,6 @@ section
19201920

19211921
open Associates UniqueFactorizationMonoid
19221922

1923-
theorem Associates.quot_out {α : Type*} [CommMonoid α] (a : Associates α) :
1924-
Associates.mk (Quot.out a) = a := by rw [← quot_mk_eq_mk, Quot.out_eq]
1925-
#align associates.quot_out Associates.quot_out
1926-
19271923
/-- `toGCDMonoid` constructs a GCD monoid out of a unique factorization domain. -/
19281924
noncomputable def UniqueFactorizationMonoid.toGCDMonoid (α : Type*) [CancelCommMonoidWithZero α]
19291925
[UniqueFactorizationMonoid α] [DecidableEq (Associates α)] [DecidableEq α] : GCDMonoid α where

0 commit comments

Comments
 (0)