@@ -31,14 +31,14 @@ divisibility, divides
31
31
32
32
variables {α : Type *}
33
33
34
- section monoid
34
+ section semigroup
35
35
36
- variables [monoid α] {a b c : α}
36
+ variables [semigroup α] {a b c : α}
37
37
38
38
/-- There are two possible conventions for divisibility, which coincide in a `comm_monoid`.
39
39
This matches the convention for ordinals. -/
40
40
@[priority 100 ]
41
- instance monoid_has_dvd : has_dvd α :=
41
+ instance semigroup_has_dvd : has_dvd α :=
42
42
has_dvd.mk (λ a b, ∃ c, b = a * c)
43
43
44
44
-- 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
53
53
theorem dvd.elim {P : Prop } {a b : α} (H₁ : a ∣ b) (H₂ : ∀ c, b = a * c → P) : P :=
54
54
exists.elim H₁ H₂
55
55
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
-
62
56
local attribute [simp] mul_assoc mul_comm mul_left_comm
63
57
64
58
@[trans] theorem dvd_trans (h₁ : a ∣ b) (h₂ : b ∣ c) : a ∣ c :=
69
63
70
64
alias dvd_trans ← has_dvd.dvd.trans
71
65
72
- theorem one_dvd (a : α) : 1 ∣ a := dvd.intro a (one_mul _)
73
-
74
66
@[simp] theorem dvd_mul_right (a b : α) : a ∣ a * b := dvd.intro b rfl
75
67
76
68
theorem dvd_mul_of_dvd_left (h : a ∣ b) (c : α) : a ∣ b * c :=
@@ -93,11 +85,25 @@ f.to_mul_hom.map_dvd
93
85
94
86
end map_dvd
95
87
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
+
96
102
end monoid
97
103
98
- section comm_monoid
104
+ section comm_semigroup
99
105
100
- variables [comm_monoid α] {a b c : α}
106
+ variables [comm_semigroup α] {a b c : α}
101
107
102
108
theorem dvd.intro_left (c : α) (h : c * a = b) : a ∣ b :=
103
109
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
125
131
theorem mul_dvd_mul : ∀ {a b c d : α}, a ∣ b → c ∣ d → a * c ∣ b * d
126
132
| a ._ c ._ ⟨e, rfl⟩ ⟨f, rfl⟩ := ⟨e * f, by simp⟩
127
133
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
+
128
143
theorem mul_dvd_mul_left (a : α) {b c : α} (h : b ∣ c) : a * b ∣ a * c :=
129
144
mul_dvd_mul (dvd_refl a) h
130
145
131
146
theorem mul_dvd_mul_right (h : a ∣ b) (c : α) : a * c ∣ b * c :=
132
147
mul_dvd_mul h (dvd_refl c)
133
148
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
149
end comm_monoid
138
150
139
- section monoid_with_zero
151
+ section semigroup_with_zero
140
152
141
- variables [monoid_with_zero α] {a : α}
153
+ variables [semigroup_with_zero α] {a : α}
142
154
143
155
theorem eq_zero_of_zero_dvd (h : 0 ∣ a) : a = 0 :=
144
156
dvd.elim h (assume c, assume H' : a = 0 * c, eq.trans H' (zero_mul c))
145
157
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
147
159
product with zero equals `a` iff `a` equals zero. -/
148
160
@[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, } ⟩
150
162
151
163
@[simp] theorem dvd_zero (a : α) : a ∣ 0 := dvd.intro 0 (by simp)
152
164
153
- end monoid_with_zero
165
+ end semigroup_with_zero
154
166
155
167
/-- Given two elements `b`, `c` of a `cancel_monoid_with_zero` and a nonzero element `a`,
156
168
`a*b` divides `a*c` iff `b` divides `c`. -/
0 commit comments