Skip to content
New issue

Have a question about this project? Sign up for a free GitHub account to open an issue and contact its maintainers and the community.

By clicking “Sign up for GitHub”, you agree to our terms of service and privacy statement. We’ll occasionally send you account related emails.

Already on GitHub? Sign in to your account

[Merged by Bors] - refactor(algebra/associated): change several instances from [integral_domain] to [comm_cancel_monoid_with_zero] #3744

Closed
wants to merge 5 commits into from
Closed
Show file tree
Hide file tree
Changes from 2 commits
Commits
File filter

Filter by extension

Filter by extension

Conversations
Failed to load comments.
Jump to
Jump to file
Failed to load files.
Diff view
Diff view
37 changes: 16 additions & 21 deletions src/algebra/associated.lean
Original file line number Diff line number Diff line change
Expand Up @@ -34,14 +34,15 @@ theorem is_unit_int {n : ℤ} : is_unit n ↔ n.nat_abs = 1 :=
lemma is_unit_of_dvd_one [comm_monoid α] : ∀a ∣ 1, is_unit (a:α)
| a ⟨b, eq⟩ := ⟨units.mk_of_mul_eq_one a b eq.symm, rfl⟩

lemma dvd_and_not_dvd_iff [integral_domain α] {x y : α} :
lemma dvd_and_not_dvd_iff [comm_cancel_monoid_with_zero α] {x y : α} :
x ∣ y ∧ ¬y ∣ x ↔ x ≠ 0 ∧ ∃ d : α, ¬ is_unit d ∧ y = x * d :=
⟨λ ⟨⟨d, hd⟩, hyx⟩, ⟨λ hx0, by simpa [hx0] using hyx, ⟨d,
mt is_unit_iff_dvd_one.1 (λ ⟨e, he⟩, hyx ⟨e, by rw [hd, mul_assoc, ← he, mul_one]⟩), hd⟩⟩,
λ ⟨hx0, d, hdu, hdx⟩, ⟨⟨d, hdx⟩, λ ⟨e, he⟩, hdu (is_unit_of_dvd_one _
⟨e, mul_left_cancel' hx0 $ by conv {to_lhs, rw [he, hdx]};simp [mul_assoc]⟩)⟩⟩

lemma pow_dvd_pow_iff [integral_domain α] {x : α} {n m : ℕ} (h0 : x ≠ 0) (h1 : ¬ is_unit x) :
lemma pow_dvd_pow_iff [comm_cancel_monoid_with_zero α] [no_zero_divisors α]
awainverse marked this conversation as resolved.
Show resolved Hide resolved
{x : α} {n m : ℕ} (h0 : x ≠ 0) (h1 : ¬ is_unit x) :
x ^ n ∣ x ^ m ↔ n ≤ m :=
begin
split,
Expand Down Expand Up @@ -150,7 +151,7 @@ begin
exact H _ o.1 _ o.2 h.symm
end

lemma irreducible_of_prime [integral_domain α] {p : α} (hp : prime p) : irreducible p :=
lemma irreducible_of_prime [comm_cancel_monoid_with_zero α] {p : α} (hp : prime p) : irreducible p :=
⟨hp.not_unit, λ a b hab,
(show a * b ∣ a ∨ a * b ∣ b, from hab ▸ hp.div_or_div (hab ▸ (dvd_refl _))).elim
(λ ⟨x, hx⟩, or.inr (is_unit_iff_dvd_one.2
Expand All @@ -160,9 +161,9 @@ lemma irreducible_of_prime [integral_domain α] {p : α} (hp : prime p) : irredu
⟨x, mul_right_cancel' (show b ≠ 0, from λ h, by simp [*, prime] at *)
$ by conv {to_lhs, rw hx}; simp [mul_comm, mul_assoc, mul_left_comm]⟩))⟩

lemma succ_dvd_or_succ_dvd_of_succ_sum_dvd_mul [integral_domain α] {p : α} (hp : prime p) {a b : α}
{k l : ℕ} : p ^ k ∣ a → p ^ l ∣ b → p ^ ((k + l) + 1) ∣ a * b →
p ^ (k + 1) ∣ a ∨ p ^ (l + 1) ∣ b :=
lemma succ_dvd_or_succ_dvd_of_succ_sum_dvd_mul [comm_cancel_monoid_with_zero α] [no_zero_divisors α]
awainverse marked this conversation as resolved.
Show resolved Hide resolved
{p : α} (hp : prime p) {a b : α} {k l : ℕ} :
p ^ k ∣ a → p ^ l ∣ b → p ^ ((k + l) + 1) ∣ a * b → p ^ (k + 1) ∣ a ∨ p ^ (l + 1) ∣ b :=
λ ⟨x, hx⟩ ⟨y, hy⟩ ⟨z, hz⟩,
have h : p ^ (k + l) * (x * y) = p ^ (k + l) * (p * z),
by simpa [mul_comm, _root_.pow_add, hx, hy, mul_assoc, mul_left_comm] using hz,
Expand Down Expand Up @@ -255,7 +256,7 @@ end
theorem dvd_dvd_iff_associated [cancel_monoid_with_zero α] {a b : α} : a ∣ b ∧ b ∣ a ↔ a ~ᵤ b :=
⟨λ ⟨h1, h2⟩, associated_of_dvd_dvd h1 h2, dvd_dvd_of_associated⟩

lemma exists_associated_mem_of_dvd_prod [integral_domain α] {p : α}
lemma exists_associated_mem_of_dvd_prod [comm_cancel_monoid_with_zero α] {p : α}
(hp : prime p) {s : multiset α} : (∀ r ∈ s, prime r) → p ∣ s.prod → ∃ q ∈ s, p ~ᵤ q :=
multiset.induction_on s (by simp [mt is_unit_iff_dvd_one.2 hp.not_unit])
(λ a s ih hs hps, begin
Expand Down Expand Up @@ -310,7 +311,7 @@ lemma irreducible_iff_of_associated [comm_semiring α] {p q : α} (h : p ~ᵤ q)
irreducible p ↔ irreducible q :=
⟨irreducible_of_associated h, irreducible_of_associated h.symm⟩

lemma associated_mul_left_cancel [integral_domain α] {a b c d : α}
lemma associated_mul_left_cancel [comm_cancel_monoid_with_zero α] {a b c d : α}
(h : a * b ~ᵤ c * d) (h₁ : a ~ᵤ c) (ha : a ≠ 0) : b ~ᵤ d :=
let ⟨u, hu⟩ := h in let ⟨v, hv⟩ := associated.symm h₁ in
⟨u * (v : units α), mul_left_cancel' ha
Expand All @@ -319,7 +320,7 @@ let ⟨u, hu⟩ := h in let ⟨v, hv⟩ := associated.symm h₁ in
simp [hv.symm, mul_assoc, mul_comm, mul_left_comm]
end⟩

lemma associated_mul_right_cancel [integral_domain α] {a b c d : α} :
lemma associated_mul_right_cancel [comm_cancel_monoid_with_zero α] {a b c d : α} :
a * b ~ᵤ c * d → b ~ᵤ d → b ≠ 0 → a ~ᵤ c :=
by rw [mul_comm a, mul_comm c]; exact associated_mul_left_cancel

Expand Down Expand Up @@ -474,12 +475,6 @@ have (0 : α) ~ᵤ 1, from quotient.exact h,
have (0 : α) = 1, from ((associated_zero_iff_eq_zero 1).1 this.symm).symm,
zero_ne_one this⟩⟩

end comm_monoid_with_zero

section comm_semiring

variables [comm_semiring α]

theorem dvd_of_mk_le_mk {a b : α} : associates.mk a ≤ associates.mk b → a ∣ b
| ⟨c', hc'⟩ := (quotient.induction_on c' $ assume c hc,
let ⟨d, hd⟩ := (quotient.exact hc).symm in
Expand Down Expand Up @@ -533,10 +528,10 @@ begin
rw [mk_mul_mk, mk_le_mk_iff_dvd_iff, mk_le_mk_iff_dvd_iff, mk_le_mk_iff_dvd_iff]
end

end comm_semiring
end comm_monoid_with_zero

section integral_domain
variable [integral_domain α]
section comm_cancel_monoid_with_zero
variable [comm_cancel_monoid_with_zero α]

instance : partial_order (associates α) :=
{ le_antisymm := λ a' b', quotient.induction_on₂ a' b' (λ a b hab hba,
Expand All @@ -553,14 +548,14 @@ instance : order_top (associates α) :=
le_top := assume a, ⟨0, (mul_zero a).symm⟩,
.. associates.partial_order }

instance : no_zero_divisors (associates α) :=
instance [no_zero_divisors α] : no_zero_divisors (associates α) :=
awainverse marked this conversation as resolved.
Show resolved Hide resolved
⟨λ x y,
(quotient.induction_on₂ x y $ assume a b h,
have a * b = 0, from (associated_zero_iff_eq_zero _).1 (quotient.exact h),
have a = 0 ∨ b = 0, from mul_eq_zero.1 this,
this.imp (assume h, h.symm ▸ rfl) (assume h, h.symm ▸ rfl))⟩

theorem prod_eq_zero_iff {s : multiset (associates α)} :
theorem prod_eq_zero_iff [no_zero_divisors α] [nontrivial α] {s : multiset (associates α)} :
awainverse marked this conversation as resolved.
Show resolved Hide resolved
s.prod = 0 ↔ (0 : associates α) ∈ s :=
multiset.induction_on s (by simp) $
assume a s, by simp [mul_eq_zero, @eq_comm _ 0 a] {contextual := tt}
Expand Down Expand Up @@ -610,6 +605,6 @@ match h m d (le_refl _) with
or.inl $ bot_unique $ associates.le_of_mul_le_mul_left d m 1 ‹d ≠ 0› this
end

end integral_domain
end comm_cancel_monoid_with_zero

end associates
8 changes: 7 additions & 1 deletion src/algebra/divisibility.lean
Original file line number Diff line number Diff line change
Expand Up @@ -130,12 +130,18 @@ dvd.elim h (assume c, assume H' : a = 0 * c, eq.trans H' (zero_mul c))

end monoid_with_zero

/-- Given two elements `b`, `c` of an integral domain and a nonzero element `a`,
/-- Given two elements `b`, `c` of a `cancel_monoid_with_zero` and a nonzero element `a`,
`a*b` divides `a*c` iff `b` divides `c`. -/
theorem mul_dvd_mul_iff_left [cancel_monoid_with_zero α] {a b c : α}
(ha : a ≠ 0) : a * b ∣ a * c ↔ b ∣ c :=
exists_congr $ λ d, by rw [mul_assoc, mul_right_inj' ha]

/-- Given two elements `a`, `b` of a commutative `cancel_monoid_with_zero` and a nonzero
element `c`, `a*c` divides `b*c` iff `a` divides `b`. -/
theorem mul_dvd_mul_iff_right [comm_cancel_monoid_with_zero α] {a b c : α} (hc : c ≠ 0) :
a * c ∣ b * c ↔ a ∣ b :=
exists_congr $ λ d, by rw [mul_right_comm, mul_left_inj' hc]

/-!
### Units in various monoids
-/
Expand Down
29 changes: 29 additions & 0 deletions src/algebra/group_with_zero.lean
Original file line number Diff line number Diff line change
Expand Up @@ -212,6 +212,12 @@ element, and `0` is left and right absorbing. -/
@[protect_proj]
class comm_monoid_with_zero (M₀ : Type*) extends comm_monoid M₀, monoid_with_zero M₀.

/-- A type `M` is a `comm_cancel_monoid_with_zero` if it is a commutative monoid with zero element,
`0` is left and right absorbing,
and left/right multiplication by a non-zero element is injective. -/
@[protect_proj] class comm_cancel_monoid_with_zero (M₀ : Type*) extends
comm_monoid_with_zero M₀, cancel_monoid_with_zero M₀.

/-- A type `G₀` is a “group with zero” if it is a monoid with zero element (distinct from `1`)
such that every nonzero element is invertible.
The type is required to come with an “inverse” function, and the inverse of `0` must be `0`.
Expand Down Expand Up @@ -371,6 +377,21 @@ classical.by_contradiction $ λ ha, h₁ $ mul_right_cancel' ha $ h₂.symm ▸

end cancel_monoid_with_zero

section comm_cancel_monoid_with_zero

variables [cancel_monoid_with_zero M₀] {a b c : M₀}
awainverse marked this conversation as resolved.
Show resolved Hide resolved

section prio
set_option default_priority 10 -- see Note [default priority]

instance comm_cancel_monoid_with_zero.no_zero_divisors : no_zero_divisors M₀ :=
⟨λ a b ab0, by { by_cases a = 0, {left, exact h}, right,
awainverse marked this conversation as resolved.
Show resolved Hide resolved
apply cancel_monoid_with_zero.mul_left_cancel_of_ne_zero h, rw [ab0, mul_zero], }⟩

end prio

end comm_cancel_monoid_with_zero

section group_with_zero
variables [group_with_zero G₀]

Expand Down Expand Up @@ -714,6 +735,14 @@ end group_with_zero
section comm_group_with_zero -- comm
variables [comm_group_with_zero G₀] {a b c : G₀}

section prio
set_option default_priority 10 -- see Note [default priority]

instance comm_group_with_zero.comm_cancel_monoid_with_zero : comm_cancel_monoid_with_zero G₀ :=
{ .. group_with_zero.cancel_monoid_with_zero, ..comm_group_with_zero.to_comm_monoid_with_zero G₀ }
awainverse marked this conversation as resolved.
Show resolved Hide resolved

end prio

/-- Pullback a `comm_group_with_zero` class along an injective function. -/
protected def function.injective.comm_group_with_zero [has_zero G₀'] [has_mul G₀'] [has_one G₀']
[has_inv G₀'] (f : G₀' → G₀) (hf : injective f) (zero : f 0 = 0) (one : f 1 = 1)
Expand Down
8 changes: 3 additions & 5 deletions src/algebra/ring/basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -732,6 +732,9 @@ class integral_domain (α : Type u) extends comm_ring α, domain α
section integral_domain
variables [integral_domain α] {a b c d e : α}

instance integral_domain.to_comm_cancel_monoid_with_zero : comm_cancel_monoid_with_zero α :=
{ ..comm_semiring.to_comm_monoid_with_zero, ..domain.to_cancel_monoid_with_zero }

/-- Pullback an `integral_domain` instance along an injective function. -/
protected def function.injective.integral_domain [has_zero β] [has_one β] [has_add β]
[has_mul β] [has_neg β] (f : β → α) (hf : injective f) (zero : f 0 = 0) (one : f 1 = 1)
Expand All @@ -747,11 +750,6 @@ by rw [← sub_eq_zero, mul_self_sub_mul_self, mul_eq_zero, or_comm, sub_eq_zero
lemma mul_self_eq_one_iff {a : α} : a * a = 1 ↔ a = 1 ∨ a = -1 :=
by rw [← mul_self_eq_mul_self_iff, one_mul]

/-- Given two elements a, b of an integral domain and a nonzero element c, a*c divides b*c iff
a divides b. -/
theorem mul_dvd_mul_iff_right (hc : c ≠ 0) : a * c ∣ b * c ↔ a ∣ b :=
exists_congr $ λ d, by rw [mul_right_comm, mul_left_inj' hc]

/-- In the unit group of an integral domain, a unit is its own inverse iff the unit is one or
one's additive inverse. -/
lemma units.inv_eq_self_iff (u : units α) : u⁻¹ = u ↔ u = 1 ∨ u = -1 :=
Expand Down