Skip to content

Commit b4dda71

Browse files
committed
chore(*): Use non-greek variable names here and there (#17035)
... for monoids, rings, and fields.
1 parent 5d4874c commit b4dda71

File tree

3 files changed

+64
-64
lines changed

3 files changed

+64
-64
lines changed

Mathlib/Algebra/Associated/OrderedCommMonoid.lean

Lines changed: 4 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -23,14 +23,14 @@ Then we show that the quotient type `Associates` is a monoid
2323
and prove basic properties of this quotient.
2424
-/
2525

26-
variable {α : Type*} [CancelCommMonoidWithZero α]
26+
variable {M : Type*} [CancelCommMonoidWithZero M]
2727

2828
namespace Associates
2929

30-
instance instOrderedCommMonoid : OrderedCommMonoid (Associates α) where
31-
mul_le_mul_left := fun a _ ⟨d, hd⟩ c => hd.symm ▸ mul_assoc c a d ▸ le_mul_right (α := α)
30+
instance instOrderedCommMonoid : OrderedCommMonoid (Associates M) where
31+
mul_le_mul_left := fun a _ ⟨d, hd⟩ c => hd.symm ▸ mul_assoc c a d ▸ le_mul_right
3232

33-
instance : CanonicallyOrderedCommMonoid (Associates α) where
33+
instance : CanonicallyOrderedCommMonoid (Associates M) where
3434
exists_mul_of_le h := h
3535
le_self_mul _ b := ⟨b, rfl⟩
3636
bot_le _ := one_le

Mathlib/Algebra/Field/Basic.lean

Lines changed: 35 additions & 35 deletions
Original file line numberDiff line numberDiff line change
@@ -18,16 +18,16 @@ open Function OrderDual Set
1818

1919
universe u
2020

21-
variable {α β K : Type*}
21+
variable {K L : Type*}
2222

2323
section DivisionSemiring
2424

25-
variable [DivisionSemiring α] {a b c d : α}
25+
variable [DivisionSemiring K] {a b c d : K}
2626

27-
theorem add_div (a b c : α) : (a + b) / c = a / c + b / c := by simp_rw [div_eq_mul_inv, add_mul]
27+
theorem add_div (a b c : K) : (a + b) / c = a / c + b / c := by simp_rw [div_eq_mul_inv, add_mul]
2828

2929
@[field_simps]
30-
theorem div_add_div_same (a b c : α) : a / c + b / c = (a + b) / c :=
30+
theorem div_add_div_same (a b c : K) : a / c + b / c = (a + b) / c :=
3131
(add_div _ _ _).symm
3232

3333
theorem same_add_div (h : b ≠ 0) : (b + a) / b = 1 + a / b := by rw [← div_self h, add_div]
@@ -49,15 +49,15 @@ theorem one_div_mul_add_mul_one_div_eq_one_div_add_one_div (ha : a ≠ 0) (hb :
4949
1 / a * (a + b) * (1 / b) = 1 / a + 1 / b := by
5050
simpa only [one_div] using (inv_add_inv' ha hb).symm
5151

52-
theorem add_div_eq_mul_add_div (a b : α) (hc : c ≠ 0) : a + b / c = (a * c + b) / c :=
52+
theorem add_div_eq_mul_add_div (a b : K) (hc : c ≠ 0) : a + b / c = (a * c + b) / c :=
5353
(eq_div_iff_mul_eq hc).2 <| by rw [right_distrib, div_mul_cancel₀ _ hc]
5454

5555
@[field_simps]
56-
theorem add_div' (a b c : α) (hc : c ≠ 0) : b + a / c = (b * c + a) / c := by
56+
theorem add_div' (a b c : K) (hc : c ≠ 0) : b + a / c = (b * c + a) / c := by
5757
rw [add_div, mul_div_cancel_right₀ _ hc]
5858

5959
@[field_simps]
60-
theorem div_add' (a b c : α) (hc : c ≠ 0) : a / c + b = (a + b * c) / c := by
60+
theorem div_add' (a b c : K) (hc : c ≠ 0) : a / c + b = (a + b * c) / c := by
6161
rwa [add_comm, add_div', add_comm]
6262

6363
protected theorem Commute.div_add_div (hbc : Commute b c) (hbd : Commute b d) (hb : b ≠ 0)
@@ -167,9 +167,9 @@ end DivisionRing
167167

168168
section Semifield
169169

170-
variable [Semifield α] {a b c d : α}
170+
variable [Semifield K] {a b c d : K}
171171

172-
theorem div_add_div (a : α) (c : α) (hb : b ≠ 0) (hd : d ≠ 0) :
172+
theorem div_add_div (a : K) (c : K) (hb : b ≠ 0) (hd : d ≠ 0) :
173173
a / b + c / d = (a * d + b * c) / (b * d) :=
174174
(Commute.all b _).div_add_div (Commute.all _ _) hb hd
175175

@@ -211,7 +211,7 @@ end Field
211211

212212
namespace RingHom
213213

214-
protected theorem injective [DivisionRing α] [Semiring β] [Nontrivial β] (f : α →+* β) :
214+
protected theorem injective [DivisionRing K] [Semiring L] [Nontrivial L] (f : K →+* L) :
215215
Injective f :=
216216
(injective_iff_map_eq_zero f).2 fun _ ↦ (map_eq_zero f).1
217217

@@ -242,18 +242,18 @@ noncomputable abbrev Field.ofIsUnitOrEqZero [CommRing R] (h : ∀ a : R, IsUnit
242242
end NoncomputableDefs
243243

244244
namespace Function.Injective
245-
variable [Zero α] [Add α] [Neg α] [Sub α] [One α] [Mul α] [Inv α] [Div α] [SMul ℕ α] [SMul ℤ α]
246-
[SMul ℚ≥0 α] [SMul ℚ α] [Pow α ℕ] [Pow α ℤ] [NatCast α] [IntCast α] [NNRatCast α] [RatCast α]
247-
(f : αβ) (hf : Injective f)
245+
variable [Zero K] [Add K] [Neg K] [Sub K] [One K] [Mul K] [Inv K] [Div K] [SMul ℕ K] [SMul ℤ K]
246+
[SMul ℚ≥0 K] [SMul ℚ K] [Pow K ℕ] [Pow K ℤ] [NatCast K] [IntCast K] [NNRatCast K] [RatCast K]
247+
(f : KL) (hf : Injective f)
248248

249249
/-- Pullback a `DivisionSemiring` along an injective function. -/
250250
-- See note [reducible non-instances]
251-
protected abbrev divisionSemiring [DivisionSemiring β] (zero : f 0 = 0) (one : f 1 = 1)
251+
protected abbrev divisionSemiring [DivisionSemiring L] (zero : f 0 = 0) (one : f 1 = 1)
252252
(add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y)
253253
(inv : ∀ x, f x⁻¹ = (f x)⁻¹) (div : ∀ x y, f (x / y) = f x / f y)
254254
(nsmul : ∀ (n : ℕ) (x), f (n • x) = n • f x) (nnqsmul : ∀ (q : ℚ≥0) (x), f (q • x) = q • f x)
255255
(npow : ∀ (x) (n : ℕ), f (x ^ n) = f x ^ n) (zpow : ∀ (x) (n : ℤ), f (x ^ n) = f x ^ n)
256-
(natCast : ∀ n : ℕ, f n = n) (nnratCast : ∀ q : ℚ≥0, f q = q) : DivisionSemiring α where
256+
(natCast : ∀ n : ℕ, f n = n) (nnratCast : ∀ q : ℚ≥0, f q = q) : DivisionSemiring K where
257257
toSemiring := hf.semiring f zero one add mul nsmul npow natCast
258258
__ := hf.groupWithZero f zero one mul inv div npow zpow
259259
nnratCast_def q := hf <| by rw [nnratCast, NNRat.cast_def, div, natCast, natCast]
@@ -262,15 +262,15 @@ protected abbrev divisionSemiring [DivisionSemiring β] (zero : f 0 = 0) (one :
262262

263263
/-- Pullback a `DivisionSemiring` along an injective function. -/
264264
-- See note [reducible non-instances]
265-
protected abbrev divisionRing [DivisionRing β] (zero : f 0 = 0) (one : f 1 = 1)
265+
protected abbrev divisionRing [DivisionRing L] (zero : f 0 = 0) (one : f 1 = 1)
266266
(add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y)
267267
(neg : ∀ x, f (-x) = -f x) (sub : ∀ x y, f (x - y) = f x - f y) (inv : ∀ x, f x⁻¹ = (f x)⁻¹)
268268
(div : ∀ x y, f (x / y) = f x / f y)
269269
(nsmul : ∀ (n : ℕ) (x), f (n • x) = n • f x) (zsmul : ∀ (n : ℤ) (x), f (n • x) = n • f x)
270270
(nnqsmul : ∀ (q : ℚ≥0) (x), f (q • x) = q • f x) (qsmul : ∀ (q : ℚ) (x), f (q • x) = q • f x)
271271
(npow : ∀ (x) (n : ℕ), f (x ^ n) = f x ^ n) (zpow : ∀ (x) (n : ℤ), f (x ^ n) = f x ^ n)
272272
(natCast : ∀ n : ℕ, f n = n) (intCast : ∀ n : ℤ, f n = n) (nnratCast : ∀ q : ℚ≥0, f q = q)
273-
(ratCast : ∀ q : ℚ, f q = q) : DivisionRing α where
273+
(ratCast : ∀ q : ℚ, f q = q) : DivisionRing K where
274274
toRing := hf.ring f zero one add mul neg sub nsmul zsmul npow natCast intCast
275275
__ := hf.groupWithZero f zero one mul inv div npow zpow
276276
__ := hf.divisionSemiring f zero one add mul inv div nsmul nnqsmul npow zpow natCast nnratCast
@@ -280,19 +280,19 @@ protected abbrev divisionRing [DivisionRing β] (zero : f 0 = 0) (one : f 1 = 1)
280280

281281
/-- Pullback a `Field` along an injective function. -/
282282
-- See note [reducible non-instances]
283-
protected abbrev semifield [Semifield β] (zero : f 0 = 0) (one : f 1 = 1)
283+
protected abbrev semifield [Semifield L] (zero : f 0 = 0) (one : f 1 = 1)
284284
(add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y)
285285
(inv : ∀ x, f x⁻¹ = (f x)⁻¹) (div : ∀ x y, f (x / y) = f x / f y)
286286
(nsmul : ∀ (n : ℕ) (x), f (n • x) = n • f x) (nnqsmul : ∀ (q : ℚ≥0) (x), f (q • x) = q • f x)
287287
(npow : ∀ (x) (n : ℕ), f (x ^ n) = f x ^ n) (zpow : ∀ (x) (n : ℤ), f (x ^ n) = f x ^ n)
288-
(natCast : ∀ n : ℕ, f n = n) (nnratCast : ∀ q : ℚ≥0, f q = q) : Semifield α where
288+
(natCast : ∀ n : ℕ, f n = n) (nnratCast : ∀ q : ℚ≥0, f q = q) : Semifield K where
289289
toCommSemiring := hf.commSemiring f zero one add mul nsmul npow natCast
290290
__ := hf.commGroupWithZero f zero one mul inv div npow zpow
291291
__ := hf.divisionSemiring f zero one add mul inv div nsmul nnqsmul npow zpow natCast nnratCast
292292

293293
/-- Pullback a `Field` along an injective function. -/
294294
-- See note [reducible non-instances]
295-
protected abbrev field [Field β] (zero : f 0 = 0) (one : f 1 = 1)
295+
protected abbrev field [Field L] (zero : f 0 = 0) (one : f 1 = 1)
296296
(add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y)
297297
(neg : ∀ x, f (-x) = -f x) (sub : ∀ x y, f (x - y) = f x - f y) (inv : ∀ x, f x⁻¹ = (f x)⁻¹)
298298
(div : ∀ x y, f (x / y) = f x / f y)
@@ -301,7 +301,7 @@ protected abbrev field [Field β] (zero : f 0 = 0) (one : f 1 = 1)
301301
(npow : ∀ (x) (n : ℕ), f (x ^ n) = f x ^ n) (zpow : ∀ (x) (n : ℤ), f (x ^ n) = f x ^ n)
302302
(natCast : ∀ n : ℕ, f n = n) (intCast : ∀ n : ℤ, f n = n) (nnratCast : ∀ q : ℚ≥0, f q = q)
303303
(ratCast : ∀ q : ℚ, f q = q) :
304-
Field α where
304+
Field K where
305305
toCommRing := hf.commRing f zero one add mul neg sub nsmul zsmul npow natCast intCast
306306
__ := hf.divisionRing f zero one add mul neg sub inv div nsmul zsmul nnqsmul qsmul npow zpow
307307
natCast intCast nnratCast ratCast
@@ -312,30 +312,30 @@ end Function.Injective
312312

313313
namespace OrderDual
314314

315-
instance instRatCast [RatCast α] : RatCast αᵒᵈ := ‹_›
316-
instance instDivisionSemiring [DivisionSemiring α] : DivisionSemiring αᵒᵈ := ‹_›
317-
instance instDivisionRing [DivisionRing α] : DivisionRing αᵒᵈ := ‹_›
318-
instance instSemifield [Semifield α] : Semifield αᵒᵈ := ‹_›
319-
instance instField [Field α] : Field αᵒᵈ := ‹_›
315+
instance instRatCast [RatCast K] : RatCast Kᵒᵈ := ‹_›
316+
instance instDivisionSemiring [DivisionSemiring K] : DivisionSemiring Kᵒᵈ := ‹_›
317+
instance instDivisionRing [DivisionRing K] : DivisionRing Kᵒᵈ := ‹_›
318+
instance instSemifield [Semifield K] : Semifield Kᵒᵈ := ‹_›
319+
instance instField [Field K] : Field Kᵒᵈ := ‹_›
320320

321321
end OrderDual
322322

323-
@[simp] lemma toDual_ratCast [RatCast α] (n : ℚ) : toDual (n : α) = n := rfl
323+
@[simp] lemma toDual_ratCast [RatCast K] (n : ℚ) : toDual (n : K) = n := rfl
324324

325-
@[simp] lemma ofDual_ratCast [RatCast α] (n : ℚ) : (ofDual n : α) = n := rfl
325+
@[simp] lemma ofDual_ratCast [RatCast K] (n : ℚ) : (ofDual n : K) = n := rfl
326326

327327
/-! ### Lexicographic order -/
328328

329329
namespace Lex
330330

331-
instance instRatCast [RatCast α] : RatCast (Lex α) := ‹_›
332-
instance instDivisionSemiring [DivisionSemiring α] : DivisionSemiring (Lex α) := ‹_›
333-
instance instDivisionRing [DivisionRing α] : DivisionRing (Lex α) := ‹_›
334-
instance instSemifield [Semifield α] : Semifield (Lex α) := ‹_›
335-
instance instField [Field α] : Field (Lex α) := ‹_›
331+
instance instRatCast [RatCast K] : RatCast (Lex K) := ‹_›
332+
instance instDivisionSemiring [DivisionSemiring K] : DivisionSemiring (Lex K) := ‹_›
333+
instance instDivisionRing [DivisionRing K] : DivisionRing (Lex K) := ‹_›
334+
instance instSemifield [Semifield K] : Semifield (Lex K) := ‹_›
335+
instance instField [Field K] : Field (Lex K) := ‹_›
336336

337337
end Lex
338338

339-
@[simp] lemma toLex_ratCast [RatCast α] (n : ℚ) : toLex (n : α) = n := rfl
339+
@[simp] lemma toLex_ratCast [RatCast K] (n : ℚ) : toLex (n : K) = n := rfl
340340

341-
@[simp] lemma ofLex_ratCast [RatCast α] (n : ℚ) : (ofLex n : α) = n := rfl
341+
@[simp] lemma ofLex_ratCast [RatCast K] (n : ℚ) : (ofLex n : K) = n := rfl

Mathlib/Algebra/Field/Defs.lean

Lines changed: 25 additions & 25 deletions
Original file line numberDiff line numberDiff line change
@@ -54,7 +54,7 @@ open Function Set
5454

5555
universe u
5656

57-
variable {α β K : Type*}
57+
variable {K : Type*}
5858

5959
/-- The default definition of the coercion `ℚ≥0 → K` for a division semiring `K`.
6060
@@ -81,23 +81,23 @@ itself). See also note [forgetful inheritance].
8181
8282
If the division semiring has positive characteristic `p`, our division by zero convention forces
8383
`nnratCast (1 / p) = 1 / 0 = 0`. -/
84-
class DivisionSemiring (α : Type*) extends Semiring α, GroupWithZero α, NNRatCast α where
84+
class DivisionSemiring (K : Type*) extends Semiring K, GroupWithZero K, NNRatCast K where
8585
protected nnratCast := NNRat.castRec
8686
/-- However `NNRat.cast` is defined, it must be propositionally equal to `a / b`.
8787
8888
Do not use this lemma directly. Use `NNRat.cast_def` instead. -/
89-
protected nnratCast_def (q : ℚ≥0) : (NNRat.cast q : α) = q.num / q.den := by intros; rfl
89+
protected nnratCast_def (q : ℚ≥0) : (NNRat.cast q : K) = q.num / q.den := by intros; rfl
9090
/-- Scalar multiplication by a nonnegative rational number.
9191
9292
Unless there is a risk of a `Module ℚ≥0 _` instance diamond, write `nnqsmul := _`. This will set
9393
`nnqsmul` to `(NNRat.cast · * ·)` thanks to unification in the default proof of `nnqsmul_def`.
9494
9595
Do not use directly. Instead use the `•` notation. -/
96-
protected nnqsmul : ℚ≥0αα
96+
protected nnqsmul : ℚ≥0KK
9797
/-- However `qsmul` is defined, it must be propositionally equal to multiplication by `Rat.cast`.
9898
9999
Do not use this lemma directly. Use `NNRat.smul_def` instead. -/
100-
protected nnqsmul_def (q : ℚ≥0) (a : α) : nnqsmul q a = NNRat.cast q * a := by intros; rfl
100+
protected nnqsmul_def (q : ℚ≥0) (a : K) : nnqsmul q a = NNRat.cast q * a := by intros; rfl
101101

102102
/-- A `DivisionRing` is a `Ring` with multiplicative inverses for nonzero elements.
103103
@@ -109,48 +109,48 @@ See also note [forgetful inheritance]. Similarly, there are maps `nnratCast ℚ
109109
110110
If the division ring has positive characteristic `p`, our division by zero convention forces
111111
`ratCast (1 / p) = 1 / 0 = 0`. -/
112-
class DivisionRing (α : Type*)
113-
extends Ring α, DivInvMonoid α, Nontrivial α, NNRatCast α, RatCast α where
112+
class DivisionRing (K : Type*)
113+
extends Ring K, DivInvMonoid K, Nontrivial K, NNRatCast K, RatCast K where
114114
/-- For a nonzero `a`, `a⁻¹` is a right multiplicative inverse. -/
115-
protected mul_inv_cancel : ∀ (a : α), a ≠ 0 → a * a⁻¹ = 1
115+
protected mul_inv_cancel : ∀ (a : K), a ≠ 0 → a * a⁻¹ = 1
116116
/-- The inverse of `0` is `0` by convention. -/
117-
protected inv_zero : (0 : α)⁻¹ = 0
117+
protected inv_zero : (0 : K)⁻¹ = 0
118118
protected nnratCast := NNRat.castRec
119119
/-- However `NNRat.cast` is defined, it must be equal to `a / b`.
120120
121121
Do not use this lemma directly. Use `NNRat.cast_def` instead. -/
122-
protected nnratCast_def (q : ℚ≥0) : (NNRat.cast q : α) = q.num / q.den := by intros; rfl
122+
protected nnratCast_def (q : ℚ≥0) : (NNRat.cast q : K) = q.num / q.den := by intros; rfl
123123
/-- Scalar multiplication by a nonnegative rational number.
124124
125125
Unless there is a risk of a `Module ℚ≥0 _` instance diamond, write `nnqsmul := _`. This will set
126126
`nnqsmul` to `(NNRat.cast · * ·)` thanks to unification in the default proof of `nnqsmul_def`.
127127
128128
Do not use directly. Instead use the `•` notation. -/
129-
protected nnqsmul : ℚ≥0αα
129+
protected nnqsmul : ℚ≥0KK
130130
/-- However `qsmul` is defined, it must be propositionally equal to multiplication by `Rat.cast`.
131131
132132
Do not use this lemma directly. Use `NNRat.smul_def` instead. -/
133-
protected nnqsmul_def (q : ℚ≥0) (a : α) : nnqsmul q a = NNRat.cast q * a := by intros; rfl
133+
protected nnqsmul_def (q : ℚ≥0) (a : K) : nnqsmul q a = NNRat.cast q * a := by intros; rfl
134134
protected ratCast := Rat.castRec
135135
/-- However `Rat.cast q` is defined, it must be propositionally equal to `q.num / q.den`.
136136
137137
Do not use this lemma directly. Use `Rat.cast_def` instead. -/
138-
protected ratCast_def (q : ℚ) : (Rat.cast q : α) = q.num / q.den := by intros; rfl
138+
protected ratCast_def (q : ℚ) : (Rat.cast q : K) = q.num / q.den := by intros; rfl
139139
/-- Scalar multiplication by a rational number.
140140
141141
Unless there is a risk of a `Module ℚ _` instance diamond, write `qsmul := _`. This will set
142142
`qsmul` to `(Rat.cast · * ·)` thanks to unification in the default proof of `qsmul_def`.
143143
144144
Do not use directly. Instead use the `•` notation. -/
145-
protected qsmul : ℚ → αα
145+
protected qsmul : ℚ → KK
146146
/-- However `qsmul` is defined, it must be propositionally equal to multiplication by `Rat.cast`.
147147
148148
Do not use this lemma directly. Use `Rat.cast_def` instead. -/
149-
protected qsmul_def (a : ℚ) (x : α) : qsmul a x = Rat.cast a * x := by intros; rfl
149+
protected qsmul_def (a : ℚ) (x : K) : qsmul a x = Rat.cast a * x := by intros; rfl
150150

151151
-- see Note [lower instance priority]
152-
instance (priority := 100) DivisionRing.toDivisionSemiring [DivisionRing α] : DivisionSemiring α :=
153-
{ ‹DivisionRing αwith }
152+
instance (priority := 100) DivisionRing.toDivisionSemiring [DivisionRing K] : DivisionSemiring K :=
153+
{ ‹DivisionRing Kwith }
154154

155155
/-- A `Semifield` is a `CommSemiring` with multiplicative inverses for nonzero elements.
156156
@@ -161,7 +161,7 @@ itself). See also note [forgetful inheritance].
161161
162162
If the semifield has positive characteristic `p`, our division by zero convention forces
163163
`nnratCast (1 / p) = 1 / 0 = 0`. -/
164-
class Semifield (α : Type*) extends CommSemiring α, DivisionSemiring α, CommGroupWithZero α
164+
class Semifield (K : Type*) extends CommSemiring K, DivisionSemiring K, CommGroupWithZero K
165165

166166
/-- A `Field` is a `CommRing` with multiplicative inverses for nonzero elements.
167167
@@ -175,19 +175,19 @@ If the field has positive characteristic `p`, our division by zero convention fo
175175
class Field (K : Type u) extends CommRing K, DivisionRing K
176176

177177
-- see Note [lower instance priority]
178-
instance (priority := 100) Field.toSemifield [Field α] : Semifield α := { ‹Field αwith }
178+
instance (priority := 100) Field.toSemifield [Field K] : Semifield K := { ‹Field Kwith }
179179

180180
namespace NNRat
181-
variable [DivisionSemiring α]
181+
variable [DivisionSemiring K]
182182

183-
instance (priority := 100) smulDivisionSemiring : SMul ℚ≥0 α := ⟨DivisionSemiring.nnqsmul⟩
183+
instance (priority := 100) smulDivisionSemiring : SMul ℚ≥0 K := ⟨DivisionSemiring.nnqsmul⟩
184184

185-
lemma cast_def (q : ℚ≥0) : (q : α) = q.num / q.den := DivisionSemiring.nnratCast_def _
186-
lemma smul_def (q : ℚ≥0) (a : α) : q • a = q * a := DivisionSemiring.nnqsmul_def q a
185+
lemma cast_def (q : ℚ≥0) : (q : K) = q.num / q.den := DivisionSemiring.nnratCast_def _
186+
lemma smul_def (q : ℚ≥0) (a : K) : q • a = q * a := DivisionSemiring.nnqsmul_def q a
187187

188-
variable (α)
188+
variable (K)
189189

190-
@[simp] lemma smul_one_eq_cast (q : ℚ≥0) : q • (1 : α) = q := by rw [NNRat.smul_def, mul_one]
190+
@[simp] lemma smul_one_eq_cast (q : ℚ≥0) : q • (1 : K) = q := by rw [NNRat.smul_def, mul_one]
191191

192192
@[deprecated (since := "2024-05-03")] alias smul_one_eq_coe := smul_one_eq_cast
193193

0 commit comments

Comments
 (0)