Skip to content

Commit

Permalink
feat: GCDMonoid (Associates α) instance (#11618)
Browse files Browse the repository at this point in the history
Co-authored-by: Eric Wieser <wieser.eric@gmail.com>
  • Loading branch information
2 people authored and atarnoam committed Apr 16, 2024
1 parent d791774 commit 231a045
Show file tree
Hide file tree
Showing 2 changed files with 39 additions and 0 deletions.
3 changes: 3 additions & 0 deletions Mathlib/Algebra/Associated.lean
Original file line number Diff line number Diff line change
Expand Up @@ -819,6 +819,9 @@ theorem quot_out [Monoid α] (a : Associates α) : Associates.mk (Quot.out a) =
rw [← quot_mk_eq_mk, Quot.out_eq]
#align associates.quot_out Associates.quot_outₓ

theorem mk_quot_out [Monoid α] (a : α) : Quot.out (Associates.mk a) ~ᵤ a := by
rw [← Associates.mk_eq_mk_iff_associated, Associates.quot_out]

theorem forall_associated [Monoid α] {p : Associates α → Prop} :
(∀ a, p a) ↔ ∀ a, p (Associates.mk a) :=
Iff.intro (fun h _ => h _) fun h a => Quotient.inductionOn a h
Expand Down
36 changes: 36 additions & 0 deletions Mathlib/Algebra/GCDMonoid/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -1420,3 +1420,39 @@ theorem normalize_eq_one {a : G₀} (h0 : a ≠ 0) : normalize a = 1 := by simp
#align comm_group_with_zero.normalize_eq_one CommGroupWithZero.normalize_eq_one

end CommGroupWithZero

theorem Associated.gcd [CancelCommMonoidWithZero α] [GCDMonoid α]
{a₁ a₂ b₁ b₂ : α} (ha : Associated a₁ a₂) (hb : Associated b₁ b₂) :
Associated (gcd a₁ b₁) (gcd a₂ b₂) :=
associated_of_dvd_dvd (gcd_dvd_gcd ha.dvd hb.dvd) (gcd_dvd_gcd ha.symm.dvd hb.symm.dvd)

theorem Associated.lcm [CancelCommMonoidWithZero α] [GCDMonoid α]
{a₁ a₂ b₁ b₂ : α} (ha : Associated a₁ a₂) (hb : Associated b₁ b₂) :
Associated (lcm a₁ b₁) (lcm a₂ b₂) :=
associated_of_dvd_dvd (lcm_dvd_lcm ha.dvd hb.dvd) (lcm_dvd_lcm ha.symm.dvd hb.symm.dvd)

namespace Associates

variable [CancelCommMonoidWithZero α] [GCDMonoid α]

instance instGCDMonoid : GCDMonoid (Associates α) where
gcd := Quotient.map₂' gcd fun a₁ a₂ (ha : Associated _ _) b₁ b₂ (hb : Associated _ _) => ha.gcd hb
lcm := Quotient.map₂' lcm fun a₁ a₂ (ha : Associated _ _) b₁ b₂ (hb : Associated _ _) => ha.lcm hb
gcd_dvd_left := by rintro ⟨a⟩ ⟨b⟩; exact mk_le_mk_of_dvd (gcd_dvd_left _ _)
gcd_dvd_right := by rintro ⟨a⟩ ⟨b⟩; exact mk_le_mk_of_dvd (gcd_dvd_right _ _)
dvd_gcd := by
rintro ⟨a⟩ ⟨b⟩ ⟨c⟩ hac hbc
exact mk_le_mk_of_dvd (dvd_gcd (dvd_of_mk_le_mk hac) (dvd_of_mk_le_mk hbc))
gcd_mul_lcm := by
rintro ⟨a⟩ ⟨b⟩
rw [associated_iff_eq]
exact Quotient.sound <| gcd_mul_lcm _ _
lcm_zero_left := by rintro ⟨a⟩; exact congr_arg Associates.mk <| lcm_zero_left _
lcm_zero_right := by rintro ⟨a⟩; exact congr_arg Associates.mk <| lcm_zero_right _

theorem gcd_mk_mk {a b : α} : gcd (Associates.mk a) (Associates.mk b) = Associates.mk (gcd a b) :=
rfl
theorem lcm_mk_mk {a b : α} : lcm (Associates.mk a) (Associates.mk b) = Associates.mk (lcm a b) :=
rfl

end Associates

0 comments on commit 231a045

Please sign in to comment.