Skip to content

Commit 3bc6212

Browse files
committed
refactor: use Is*CancelMulZero (#1137)
This is a Lean 4 version of leanprover-community/mathlib3#17963
1 parent 4b0358a commit 3bc6212

File tree

6 files changed

+67
-99
lines changed

6 files changed

+67
-99
lines changed

Mathlib/Algebra/Associated.lean

Lines changed: 14 additions & 28 deletions
Original file line numberDiff line numberDiff line change
@@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Johannes Hölzl, Jens Wagemaker
55
66
! This file was ported from Lean 3 source module algebra.associated
7-
! leanprover-community/mathlib commit dcf2250875895376a142faeeac5eabff32c48655
7+
! leanprover-community/mathlib commit 2f3994e1b117b1e1da49bcfb67334f33460c3ce4
88
! Please do not edit these lines, except to modify the commit id
99
! if you have ported upstream changes.
1010
-/
@@ -1099,28 +1099,19 @@ instance : PartialOrder (Associates α) where
10991099
instance : OrderedCommMonoid (Associates α) where
11001100
mul_le_mul_left := fun a _ ⟨d, hd⟩ c => hd.symm ▸ mul_assoc c a d ▸ le_mul_right
11011101

1102+
instance : CancelCommMonoidWithZero (Associates α) :=
1103+
{ (by infer_instance : CommMonoidWithZero (Associates α)) with
1104+
mul_left_cancel_of_ne_zero := by
1105+
rintro ⟨a⟩ ⟨b⟩ ⟨c⟩ ha h
1106+
rcases Quotient.exact' h with ⟨u, hu⟩
1107+
have hu : a * (b * ↑u) = a * c := by rwa [← mul_assoc]
1108+
exact Quotient.sound' ⟨u, mul_left_cancel₀ (mk_ne_zero.1 ha) hu⟩ }
1109+
11021110
instance : NoZeroDivisors (Associates α) :=
1103-
by
1104-
intro x y
1105-
exact
1106-
Quotient.inductionOn₂ x y <| fun a b h =>
1107-
have : a * b = 0 := (associated_zero_iff_eq_zero _).1 (Quotient.exact h)
1108-
have : a = 0 ∨ b = 0 := mul_eq_zero.1 this
1109-
this.imp (fun h => by rw [h]; rfl) fun h => by rw [h]; rfl⟩
1110-
1111-
theorem eq_of_mul_eq_mul_left : ∀ a b c : Associates α, a ≠ 0 → a * b = a * c → b = c := by
1112-
rintro ⟨a⟩ ⟨b⟩ ⟨c⟩ ha h
1113-
rcases Quotient.exact' h with ⟨u, hu⟩
1114-
have hu : a * (b * ↑u) = a * c := by rwa [← mul_assoc]
1115-
exact Quotient.sound' ⟨u, mul_left_cancel₀ (mk_ne_zero.1 ha) hu⟩
1116-
#align associates.eq_of_mul_eq_mul_left Associates.eq_of_mul_eq_mul_left
1117-
1118-
theorem eq_of_mul_eq_mul_right : ∀ a b c : Associates α, b ≠ 0 → a * b = c * b → a = c :=
1119-
fun a b c bne0 => mul_comm b a ▸ mul_comm b c ▸ eq_of_mul_eq_mul_left b a c bne0
1120-
#align associates.eq_of_mul_eq_mul_right Associates.eq_of_mul_eq_mul_right
1111+
by infer_instance
11211112

11221113
theorem le_of_mul_le_mul_left (a b c : Associates α) (ha : a ≠ 0) : a * b ≤ a * c → b ≤ c
1123-
| ⟨d, hd⟩ => ⟨d, eq_of_mul_eq_mul_left a _ _ ha <| by rwa [← mul_assoc]⟩
1114+
| ⟨d, hd⟩ => ⟨d, mul_left_cancel₀ ha <| by rwa [← mul_assoc]⟩
11241115
#align associates.le_of_mul_le_mul_left Associates.le_of_mul_le_mul_left
11251116

11261117
theorem one_or_eq_of_le_of_prime : ∀ p m : Associates α, Prime p → m ≤ p → m = 1 ∨ m = p
@@ -1150,15 +1141,10 @@ theorem one_or_eq_of_le_of_prime : ∀ p m : Associates α, Prime p → m ≤ p
11501141
exact Or.inl <| bot_unique <| Associates.le_of_mul_le_mul_left d m 1 ‹d ≠ 0› this
11511142
#align associates.one_or_eq_of_le_of_prime Associates.one_or_eq_of_le_of_prime
11521143

1153-
instance : CancelCommMonoidWithZero (Associates α) :=
1154-
{ (inferInstance : CommMonoidWithZero (Associates α)) with
1155-
mul_left_cancel_of_ne_zero := by apply eq_of_mul_eq_mul_left
1156-
mul_right_cancel_of_ne_zero := by apply eq_of_mul_eq_mul_right }
1157-
11581144
instance : CanonicallyOrderedMonoid (Associates α) where
1159-
exists_mul_of_le := by intro a b h; exact h
1160-
le_self_mul := fun a b => ⟨b, rfl⟩
1161-
bot_le := by apply one_le
1145+
exists_mul_of_le := fun h => h
1146+
le_self_mul := fun _ b => ⟨b, rfl⟩
1147+
bot_le := fun _ => one_le
11621148

11631149
theorem dvdNotUnit_iff_lt {a b : Associates α} : DvdNotUnit a b ↔ a < b :=
11641150
dvd_and_not_dvd_iff.symm

Mathlib/Algebra/GroupWithZero/Basic.lean

Lines changed: 3 additions & 9 deletions
Original file line numberDiff line numberDiff line change
@@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Johan Commelin
55
66
! This file was ported from Lean 3 source module algebra.group_with_zero.basic
7-
! leanprover-community/mathlib commit a148d797a1094ab554ad4183a4ad6f130358ef64
7+
! leanprover-community/mathlib commit 2f3994e1b117b1e1da49bcfb67334f33460c3ce4
88
! Please do not edit these lines, except to modify the commit id
99
! if you have ported upstream changes.
1010
-/
@@ -149,14 +149,8 @@ variable [CancelMonoidWithZero M₀] {a b c : M₀}
149149

150150
-- see Note [lower instance priority]
151151
instance (priority := 10) CancelMonoidWithZero.toNoZeroDivisors : NoZeroDivisors M₀ :=
152-
fun {a b} ab0 => by
153-
by_cases a = 0
154-
· left
155-
exact h
156-
157-
right
158-
apply CancelMonoidWithZero.mul_left_cancel_of_ne_zero h
159-
rw [ab0, mul_zero]⟩
152+
fun ab0 => or_iff_not_imp_left.mpr <| fun ha => mul_left_cancel₀ ha <|
153+
ab0.trans (mul_zero _).symm⟩
160154
#align cancel_monoid_with_zero.to_no_zero_divisors CancelMonoidWithZero.toNoZeroDivisors
161155

162156
theorem mul_left_inj' (hc : c ≠ 0) : a * c = b * c ↔ a = b :=

Mathlib/Algebra/GroupWithZero/Defs.lean

Lines changed: 40 additions & 47 deletions
Original file line numberDiff line numberDiff line change
@@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Johan Commelin
55
66
! 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
88
! Please do not edit these lines, except to modify the commit id
99
! if you have ported upstream changes.
1010
-/
@@ -56,11 +56,35 @@ class IsLeftCancelMulZero (M₀ : Type u) [Mul M₀] [Zero M₀] : Prop where
5656
/-- Multiplicatin by a nonzero element is left cancellative. -/
5757
protected mul_left_cancel_of_ne_zero : ∀ {a b c : M₀}, a ≠ 0 → a * b = a * c → b = c
5858

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+
5971
/-- A mixin for right cancellative multiplication by nonzero elements. -/
6072
class IsRightCancelMulZero (M₀ : Type u) [Mul M₀] [Zero M₀] : Prop where
6173
/-- Multiplicatin by a nonzero element is right cancellative. -/
6274
protected mul_right_cancel_of_ne_zero : ∀ {a b c : M₀}, b ≠ 0 → a * b = c * b → a = c
6375

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+
6488
/-- A mixin for cancellative multiplication by nonzero elements. -/
6589
class IsCancelMulZero (M₀ : Type u) [Mul M₀] [Zero M₀]
6690
extends IsLeftCancelMulZero M₀, IsRightCancelMulZero M₀ : Prop
@@ -88,75 +112,44 @@ class MonoidWithZero (M₀ : Type u) extends Monoid M₀, MulZeroOneClass M₀,
88112

89113
/-- A type `M` is a `CancelMonoidWithZero` if it is a monoid with zero element, `0` is left
90114
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₀
121116

122117
/-- A type `M` is a commutative “monoid with zero” if it is a commutative monoid with zero
123118
element, and `0` is left and right absorbing. -/
124119
class CommMonoidWithZero (M₀ : Type _) extends CommMonoid M₀, MonoidWithZero M₀
125120

126-
namespace CommMonoidWithZero
121+
section CommSemigroup
127122

128-
variable [CommMonoidWithZero M₀]
123+
variable [CommSemigroup M₀] [Zero M₀]
129124

130125
lemma IsLeftCancelMulZero.to_IsRightCancelMulZero [IsLeftCancelMulZero M₀] :
131126
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 _ _)) }
136129

137130
lemma IsRightCancelMulZero.to_IsLeftCancelMulZero [IsRightCancelMulZero M₀] :
138131
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 _ _)) }
143134

144135
lemma IsLeftCancelMulZero.to_IsCancelMulZero [IsLeftCancelMulZero M₀] :
145136
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 }
148138

149139
lemma IsRightCancelMulZero.to_IsCancelMulZero [IsRightCancelMulZero M₀] :
150140
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 }
153142

154-
end CommMonoidWithZero
143+
end CommSemigroup
155144

156145
/-- A type `M` is a `CancelCommMonoidWithZero` if it is a commutative monoid with zero element,
157146
`0` is left and right absorbing,
158147
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 }
160153

161154
/-- A type `G₀` is a “group with zero” if it is a monoid with zero element (distinct from `1`)
162155
such that every nonzero element is invertible.

Mathlib/Algebra/Ring/Regular.lean

Lines changed: 4 additions & 6 deletions
Original file line numberDiff line numberDiff line change
@@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Jeremy Avigad, Leonardo de Moura, Floris van Doorn, Yury Kudryashov, Neil Strickland
55
66
! This file was ported from Lean 3 source module algebra.ring.regular
7-
! leanprover-community/mathlib commit 70d50ecfd4900dd6d328da39ab7ebd516abe4025
7+
! leanprover-community/mathlib commit 2f3994e1b117b1e1da49bcfb67334f33460c3ce4
88
! Please do not edit these lines, except to modify the commit id
99
! if you have ported upstream changes.
1010
-/
@@ -67,24 +67,22 @@ Note this is not an instance as it forms a typeclass loop. -/
6767
@[reducible]
6868
def NoZeroDivisors.toCancelCommMonoidWithZero [CommRing α] [NoZeroDivisors α] :
6969
CancelCommMonoidWithZero α :=
70-
{ NoZeroDivisors.toCancelMonoidWithZero, (by infer_instance : CommMonoidWithZero α) with }
70+
{ NoZeroDivisors.toCancelMonoidWithZero, ‹CommRing α› with }
7171
#align no_zero_divisors.to_cancel_comm_monoid_with_zero NoZeroDivisors.toCancelCommMonoidWithZero
7272

7373
section IsDomain
7474

7575
-- see Note [lower instance priority]
7676
instance (priority := 100) IsDomain.toCancelMonoidWithZero [Semiring α] [IsDomain α] :
7777
CancelMonoidWithZero α :=
78-
{ mul_left_cancel_of_ne_zero := IsLeftCancelMulZero.mul_left_cancel_of_ne_zero
79-
mul_right_cancel_of_ne_zero := IsRightCancelMulZero.mul_right_cancel_of_ne_zero }
78+
{ }
8079
#align is_domain.to_cancel_monoid_with_zero IsDomain.toCancelMonoidWithZero
8180

8281
variable [CommSemiring α] [IsDomain α]
8382

8483
-- see Note [lower instance priority]
8584
instance (priority := 100) IsDomain.toCancelCommMonoidWithZero : CancelCommMonoidWithZero α :=
86-
{ mul_left_cancel_of_ne_zero := IsLeftCancelMulZero.mul_left_cancel_of_ne_zero
87-
mul_right_cancel_of_ne_zero := IsRightCancelMulZero.mul_right_cancel_of_ne_zero }
85+
{ mul_left_cancel_of_ne_zero := IsLeftCancelMulZero.mul_left_cancel_of_ne_zero }
8886
#align is_domain.to_cancel_comm_monoid_with_zero IsDomain.toCancelCommMonoidWithZero
8987

9088
end IsDomain

Mathlib/Data/Nat/Basic.lean

Lines changed: 2 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Floris van Doorn, Leonardo de Moura, Jeremy Avigad, Mario Carneiro
55
66
! This file was ported from Lean 3 source module data.nat.basic
7-
! leanprover-community/mathlib commit 70d50ecfd4900dd6d328da39ab7ebd516abe4025
7+
! leanprover-community/mathlib commit 2f3994e1b117b1e1da49bcfb67334f33460c3ce4
88
! Please do not edit these lines, except to modify the commit id
99
! if you have ported upstream changes.
1010
-/
@@ -116,9 +116,7 @@ protected theorem nsmul_eq_mul (m n : ℕ) : m • n = m * n :=
116116
instance cancelCommMonoidWithZero : CancelCommMonoidWithZero ℕ :=
117117
{ (inferInstance : CommMonoidWithZero ℕ) with
118118
mul_left_cancel_of_ne_zero :=
119-
fun {_ _ _} h1 h2 => Nat.eq_of_mul_eq_mul_left (Nat.pos_of_ne_zero h1) h2,
120-
mul_right_cancel_of_ne_zero :=
121-
fun {_ _ _} h1 h2 => Nat.eq_of_mul_eq_mul_right (Nat.pos_of_ne_zero h1) h2 }
119+
fun h1 h2 => Nat.eq_of_mul_eq_mul_left (Nat.pos_of_ne_zero h1) h2 }
122120
#align nat.cancel_comm_monoid_with_zero Nat.cancelCommMonoidWithZero
123121

124122
attribute [simp]

Mathlib/Data/Nat/Choose/Basic.lean

Lines changed: 4 additions & 5 deletions
Original file line numberDiff line numberDiff line change
@@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Chris Hughes, Bhavik Mehta, Stuart Presnell
55
66
! This file was ported from Lean 3 source module data.nat.choose.basic
7-
! leanprover-community/mathlib commit d012cd09a9b256d870751284dd6a29882b0be105
7+
! leanprover-community/mathlib commit 2f3994e1b117b1e1da49bcfb67334f33460c3ce4
88
! Please do not edit these lines, except to modify the commit id
99
! if you have ported upstream changes.
1010
-/
@@ -144,10 +144,9 @@ theorem choose_mul_factorial_mul_factorial : ∀ {n k}, k ≤ n → choose n k *
144144
#align nat.choose_mul_factorial_mul_factorial Nat.choose_mul_factorial_mul_factorial
145145

146146
theorem choose_mul {n k s : ℕ} (hkn : k ≤ n) (hsk : s ≤ k) :
147-
n.choose k * k.choose s = n.choose s * (n - s).choose (k - s) := by
148-
have h : 0 < (n - k)! * (k - s)! * s ! :=
149-
mul_pos (mul_pos (factorial_pos _) (factorial_pos _)) (factorial_pos _)
150-
refine' eq_of_mul_eq_mul_right h _
147+
n.choose k * k.choose s = n.choose s * (n - s).choose (k - s) :=
148+
have h : (n - k)! * (k - s)! * s ! ≠ 0:= by apply_rules [factorial_ne_zero, mul_ne_zero]
149+
mul_right_cancel₀ h <|
151150
calc
152151
n.choose k * k.choose s * ((n - k)! * (k - s)! * s !) =
153152
n.choose k * (k.choose s * s ! * (k - s)!) * (n - k)! :=

0 commit comments

Comments
 (0)