@@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
4
4
Authors: Johan Commelin
5
5
6
6
! This file was ported from Lean 3 source module algebra.group_with_zero.defs
7
- ! leanprover-community/mathlib commit 2aa04f651209dc8f37b9937a8c4c20c79571ac52
7
+ ! leanprover-community/mathlib commit 2f3994e1b117b1e1da49bcfb67334f33460c3ce4
8
8
! Please do not edit these lines, except to modify the commit id
9
9
! if you have ported upstream changes.
10
10
-/
@@ -56,11 +56,35 @@ class IsLeftCancelMulZero (M₀ : Type u) [Mul M₀] [Zero M₀] : Prop where
56
56
/-- Multiplicatin by a nonzero element is left cancellative. -/
57
57
protected mul_left_cancel_of_ne_zero : ∀ {a b c : M₀}, a ≠ 0 → a * b = a * c → b = c
58
58
59
+ section IsLeftCancelMulZero
60
+
61
+ variable [Mul M₀] [Zero M₀] [IsLeftCancelMulZero M₀] {a b c : M₀}
62
+
63
+ theorem mul_left_cancel₀ (ha : a ≠ 0 ) (h : a * b = a * c) : b = c :=
64
+ IsLeftCancelMulZero.mul_left_cancel_of_ne_zero ha h
65
+
66
+ theorem mul_right_injective₀ (ha : a ≠ 0 ) : Function.Injective ((· * ·) a) :=
67
+ fun _ _ => mul_left_cancel₀ ha
68
+
69
+ end IsLeftCancelMulZero
70
+
59
71
/-- A mixin for right cancellative multiplication by nonzero elements. -/
60
72
class IsRightCancelMulZero (M₀ : Type u) [Mul M₀] [Zero M₀] : Prop where
61
73
/-- Multiplicatin by a nonzero element is right cancellative. -/
62
74
protected mul_right_cancel_of_ne_zero : ∀ {a b c : M₀}, b ≠ 0 → a * b = c * b → a = c
63
75
76
+ section IsRightCancelMulZero
77
+
78
+ variable [Mul M₀] [Zero M₀] [IsRightCancelMulZero M₀] {a b c : M₀}
79
+
80
+ theorem mul_right_cancel₀ (hb : b ≠ 0 ) (h : a * b = c * b) : a = c :=
81
+ IsRightCancelMulZero.mul_right_cancel_of_ne_zero hb h
82
+
83
+ theorem mul_left_injective₀ (hb : b ≠ 0 ) : Function.Injective fun a => a * b :=
84
+ fun _ _ => mul_right_cancel₀ hb
85
+
86
+ end IsRightCancelMulZero
87
+
64
88
/-- A mixin for cancellative multiplication by nonzero elements. -/
65
89
class IsCancelMulZero (M₀ : Type u) [Mul M₀] [Zero M₀]
66
90
extends IsLeftCancelMulZero M₀, IsRightCancelMulZero M₀ : Prop
@@ -88,75 +112,44 @@ class MonoidWithZero (M₀ : Type u) extends Monoid M₀, MulZeroOneClass M₀,
88
112
89
113
/-- A type `M` is a `CancelMonoidWithZero` if it is a monoid with zero element, `0` is left
90
114
and right absorbing, and left/right multiplication by a non-zero element is injective. -/
91
- class CancelMonoidWithZero (M₀ : Type _) extends MonoidWithZero M₀ where
92
- /-- Left multiplication by a non-zero element is injective. -/
93
- protected mul_left_cancel_of_ne_zero : ∀ {a b c : M₀}, a ≠ 0 → a * b = a * c → b = c
94
- /-- Right multiplication by a non-zero element is injective. -/
95
- protected mul_right_cancel_of_ne_zero : ∀ {a b c : M₀}, b ≠ 0 → a * b = c * b → a = c
96
-
97
- section CancelMonoidWithZero
98
-
99
- variable [CancelMonoidWithZero M₀] {a b c : M₀}
100
-
101
- theorem mul_left_cancel₀ (ha : a ≠ 0 ) (h : a * b = a * c) : b = c :=
102
- CancelMonoidWithZero.mul_left_cancel_of_ne_zero ha h
103
-
104
- theorem mul_right_cancel₀ (hb : b ≠ 0 ) (h : a * b = c * b) : a = c :=
105
- CancelMonoidWithZero.mul_right_cancel_of_ne_zero hb h
106
-
107
- theorem mul_right_injective₀ (ha : a ≠ 0 ) : Function.Injective ((· * ·) a) :=
108
- fun _ _ => mul_left_cancel₀ ha
109
-
110
- theorem mul_left_injective₀ (hb : b ≠ 0 ) : Function.Injective fun a => a * b :=
111
- fun _ _ => mul_right_cancel₀ hb
112
-
113
- /-- A `CancelMonoidWithZero` satisfies `IsCancelMulZero`. -/
114
- instance (priority := 100 ) CancelMonoidWithZero.to_IsCancelMulZero : IsCancelMulZero M₀ :=
115
- { mul_left_cancel_of_ne_zero := fun ha h ↦
116
- CancelMonoidWithZero.mul_left_cancel_of_ne_zero ha h
117
- mul_right_cancel_of_ne_zero := fun hb h ↦
118
- CancelMonoidWithZero.mul_right_cancel_of_ne_zero hb h, }
119
-
120
- end CancelMonoidWithZero
115
+ class CancelMonoidWithZero (M₀ : Type _) extends MonoidWithZero M₀, IsCancelMulZero M₀
121
116
122
117
/-- A type `M` is a commutative “monoid with zero” if it is a commutative monoid with zero
123
118
element, and `0` is left and right absorbing. -/
124
119
class CommMonoidWithZero (M₀ : Type _) extends CommMonoid M₀, MonoidWithZero M₀
125
120
126
- namespace CommMonoidWithZero
121
+ section CommSemigroup
127
122
128
- variable [CommMonoidWithZero M₀]
123
+ variable [CommSemigroup M₀] [Zero M₀]
129
124
130
125
lemma IsLeftCancelMulZero.to_IsRightCancelMulZero [IsLeftCancelMulZero M₀] :
131
126
IsRightCancelMulZero M₀ :=
132
- { mul_right_cancel_of_ne_zero := by
133
- intros a b c ha h
134
- rw [mul_comm, mul_comm c] at h
135
- exact IsLeftCancelMulZero.mul_left_cancel_of_ne_zero ha h }
127
+ { mul_right_cancel_of_ne_zero :=
128
+ fun hb h => mul_left_cancel₀ hb <| (mul_comm _ _).trans (h.trans (mul_comm _ _)) }
136
129
137
130
lemma IsRightCancelMulZero.to_IsLeftCancelMulZero [IsRightCancelMulZero M₀] :
138
131
IsLeftCancelMulZero M₀ :=
139
- { mul_left_cancel_of_ne_zero := by
140
- intros a b c ha h
141
- rw [mul_comm a, mul_comm a c] at h
142
- exact IsRightCancelMulZero.mul_right_cancel_of_ne_zero ha h }
132
+ { mul_left_cancel_of_ne_zero :=
133
+ fun hb h => mul_right_cancel₀ hb <| (mul_comm _ _).trans (h.trans (mul_comm _ _)) }
143
134
144
135
lemma IsLeftCancelMulZero.to_IsCancelMulZero [IsLeftCancelMulZero M₀] :
145
136
IsCancelMulZero M₀ :=
146
- { mul_right_cancel_of_ne_zero := fun ha h ↦
147
- IsLeftCancelMulZero.to_IsRightCancelMulZero.mul_right_cancel_of_ne_zero ha h }
137
+ { IsLeftCancelMulZero.to_IsRightCancelMulZero with }
148
138
149
139
lemma IsRightCancelMulZero.to_IsCancelMulZero [IsRightCancelMulZero M₀] :
150
140
IsCancelMulZero M₀ :=
151
- { mul_left_cancel_of_ne_zero := fun ha h ↦
152
- IsRightCancelMulZero.to_IsLeftCancelMulZero.mul_left_cancel_of_ne_zero ha h }
141
+ { IsRightCancelMulZero.to_IsLeftCancelMulZero with }
153
142
154
- end CommMonoidWithZero
143
+ end CommSemigroup
155
144
156
145
/-- A type `M` is a `CancelCommMonoidWithZero` if it is a commutative monoid with zero element,
157
146
`0` is left and right absorbing,
158
147
and left/right multiplication by a non-zero element is injective. -/
159
- class CancelCommMonoidWithZero (M₀ : Type _) extends CommMonoidWithZero M₀, CancelMonoidWithZero M₀
148
+ class CancelCommMonoidWithZero (M₀ : Type _) extends CommMonoidWithZero M₀, IsLeftCancelMulZero M₀
149
+
150
+ instance (priority := 100 ) CancelCommMonoidWithZero.to_CancelMonoidWithZero
151
+ [CancelCommMonoidWithZero M₀] : CancelMonoidWithZero M₀ :=
152
+ { IsLeftCancelMulZero.to_IsCancelMulZero with }
160
153
161
154
/-- A type `G₀` is a “group with zero” if it is a monoid with zero element (distinct from `1`)
162
155
such that every nonzero element is invertible.
0 commit comments