Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit a9902d5

Browse files
committed
feat(algebra/divisibility): generalise basic facts to semigroups (#12325)
Co-authored-by: Scott Morrison <scott.morrison@gmail.com>
1 parent cc9de07 commit a9902d5

File tree

3 files changed

+38
-24
lines changed

3 files changed

+38
-24
lines changed

src/algebra/divisibility.lean

Lines changed: 33 additions & 21 deletions
Original file line numberDiff line numberDiff line change
@@ -31,14 +31,14 @@ divisibility, divides
3131

3232
variables {α : Type*}
3333

34-
section monoid
34+
section semigroup
3535

36-
variables [monoid α] {a b c : α}
36+
variables [semigroup α] {a b c : α}
3737

3838
/-- There are two possible conventions for divisibility, which coincide in a `comm_monoid`.
3939
This matches the convention for ordinals. -/
4040
@[priority 100]
41-
instance monoid_has_dvd : has_dvd α :=
41+
instance semigroup_has_dvd : has_dvd α :=
4242
has_dvd.mk (λ a b, ∃ c, b = a * c)
4343

4444
-- TODO: this used to not have `c` explicit, but that seems to be important
@@ -53,12 +53,6 @@ theorem exists_eq_mul_right_of_dvd (h : a ∣ b) : ∃ c, b = a * c := h
5353
theorem dvd.elim {P : Prop} {a b : α} (H₁ : a ∣ b) (H₂ : ∀ c, b = a * c → P) : P :=
5454
exists.elim H₁ H₂
5555

56-
@[refl, simp] theorem dvd_refl (a : α) : a ∣ a :=
57-
dvd.intro 1 (mul_one _)
58-
59-
lemma dvd_rfl {a : α} : a ∣ a :=
60-
dvd_refl a
61-
6256
local attribute [simp] mul_assoc mul_comm mul_left_comm
6357

6458
@[trans] theorem dvd_trans (h₁ : a ∣ b) (h₂ : b ∣ c) : a ∣ c :=
@@ -69,8 +63,6 @@ end
6963

7064
alias dvd_trans ← has_dvd.dvd.trans
7165

72-
theorem one_dvd (a : α) : 1 ∣ a := dvd.intro a (one_mul _)
73-
7466
@[simp] theorem dvd_mul_right (a b : α) : a ∣ a * b := dvd.intro b rfl
7567

7668
theorem dvd_mul_of_dvd_left (h : a ∣ b) (c : α) : a ∣ b * c :=
@@ -93,11 +85,25 @@ f.to_mul_hom.map_dvd
9385

9486
end map_dvd
9587

88+
end semigroup
89+
90+
section monoid
91+
92+
variables [monoid α]
93+
94+
@[refl, simp] theorem dvd_refl (a : α) : a ∣ a :=
95+
dvd.intro 1 (mul_one _)
96+
97+
lemma dvd_rfl {a : α} : a ∣ a :=
98+
dvd_refl a
99+
100+
theorem one_dvd (a : α) : 1 ∣ a := dvd.intro a (one_mul _)
101+
96102
end monoid
97103

98-
section comm_monoid
104+
section comm_semigroup
99105

100-
variables [comm_monoid α] {a b c : α}
106+
variables [comm_semigroup α] {a b c : α}
101107

102108
theorem dvd.intro_left (c : α) (h : c * a = b) : a ∣ b :=
103109
dvd.intro _ (begin rewrite mul_comm at h, apply h end)
@@ -125,32 +131,38 @@ local attribute [simp] mul_assoc mul_comm mul_left_comm
125131
theorem mul_dvd_mul : ∀ {a b c d : α}, a ∣ b → c ∣ d → a * c ∣ b * d
126132
| a ._ c ._ ⟨e, rfl⟩ ⟨f, rfl⟩ := ⟨e * f, by simp⟩
127133

134+
theorem dvd_of_mul_left_dvd (h : a * b ∣ c) : b ∣ c :=
135+
dvd.elim h (λ d ceq, dvd.intro (a * d) (by simp [ceq]))
136+
137+
end comm_semigroup
138+
139+
section comm_monoid
140+
141+
variables [comm_monoid α] {a b : α}
142+
128143
theorem mul_dvd_mul_left (a : α) {b c : α} (h : b ∣ c) : a * b ∣ a * c :=
129144
mul_dvd_mul (dvd_refl a) h
130145

131146
theorem mul_dvd_mul_right (h : a ∣ b) (c : α) : a * c ∣ b * c :=
132147
mul_dvd_mul h (dvd_refl c)
133148

134-
theorem dvd_of_mul_left_dvd (h : a * b ∣ c) : b ∣ c :=
135-
dvd.elim h (λ d ceq, dvd.intro (a * d) (by simp [ceq]))
136-
137149
end comm_monoid
138150

139-
section monoid_with_zero
151+
section semigroup_with_zero
140152

141-
variables [monoid_with_zero α] {a : α}
153+
variables [semigroup_with_zero α] {a : α}
142154

143155
theorem eq_zero_of_zero_dvd (h : 0 ∣ a) : a = 0 :=
144156
dvd.elim h (assume c, assume H' : a = 0 * c, eq.trans H' (zero_mul c))
145157

146-
/-- Given an element `a` of a commutative monoid with zero, there exists another element whose
158+
/-- Given an element `a` of a commutative semigroup with zero, there exists another element whose
147159
product with zero equals `a` iff `a` equals zero. -/
148160
@[simp] lemma zero_dvd_iff : 0 ∣ a ↔ a = 0 :=
149-
⟨eq_zero_of_zero_dvd, λ h, by rw h⟩
161+
⟨eq_zero_of_zero_dvd, λ h, by { rw h, use 0, simp, }
150162

151163
@[simp] theorem dvd_zero (a : α) : a ∣ 0 := dvd.intro 0 (by simp)
152164

153-
end monoid_with_zero
165+
end semigroup_with_zero
154166

155167
/-- Given two elements `b`, `c` of a `cancel_monoid_with_zero` and a nonzero element `a`,
156168
`a*b` divides `a*c` iff `b` divides `c`. -/

src/algebra/gcd_monoid/basic.lean

Lines changed: 4 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -557,9 +557,11 @@ begin
557557
mul_dvd_mul_iff_left h1, mul_dvd_mul_iff_right h2, and_comm] }
558558
end
559559

560-
lemma dvd_lcm_left [gcd_monoid α] (a b : α) : a ∣ lcm a b := (lcm_dvd_iff.1 dvd_rfl).1
560+
lemma dvd_lcm_left [gcd_monoid α] (a b : α) : a ∣ lcm a b :=
561+
(lcm_dvd_iff.1 (dvd_refl (lcm a b))).1
561562

562-
lemma dvd_lcm_right [gcd_monoid α] (a b : α) : b ∣ lcm a b := (lcm_dvd_iff.1 dvd_rfl).2
563+
lemma dvd_lcm_right [gcd_monoid α] (a b : α) : b ∣ lcm a b :=
564+
(lcm_dvd_iff.1 (dvd_refl (lcm a b))).2
563565

564566
lemma lcm_dvd [gcd_monoid α] {a b c : α} (hab : a ∣ b) (hcb : c ∣ b) : lcm a c ∣ b :=
565567
lcm_dvd_iff.2 ⟨hab, hcb⟩

src/number_theory/padics/ring_homs.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -378,7 +378,7 @@ begin
378378
{ rw zero_pow hc0,
379379
simp only [sub_zero, zmod.cast_zero, mul_zero],
380380
rw unit_coeff_spec hc',
381-
exact (dvd_pow_self _ hc0.ne').mul_left _ }
381+
exact (dvd_pow_self (p : ℤ_[p]) hc0.ne').mul_left _, },
382382
end
383383

384384
attribute [irreducible] appr

0 commit comments

Comments
 (0)