@@ -18,16 +18,16 @@ open Function OrderDual Set
18
18
19
19
universe u
20
20
21
- variable {α β K : Type *}
21
+ variable {K L : Type *}
22
22
23
23
section DivisionSemiring
24
24
25
- variable [DivisionSemiring α ] {a b c d : α }
25
+ variable [DivisionSemiring K ] {a b c d : K }
26
26
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]
28
28
29
29
@[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 :=
31
31
(add_div _ _ _).symm
32
32
33
33
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 :
49
49
1 / a * (a + b) * (1 / b) = 1 / a + 1 / b := by
50
50
simpa only [one_div] using (inv_add_inv' ha hb).symm
51
51
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 :=
53
53
(eq_div_iff_mul_eq hc).2 <| by rw [right_distrib, div_mul_cancel₀ _ hc]
54
54
55
55
@[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
57
57
rw [add_div, mul_div_cancel_right₀ _ hc]
58
58
59
59
@[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
61
61
rwa [add_comm, add_div', add_comm]
62
62
63
63
protected theorem Commute.div_add_div (hbc : Commute b c) (hbd : Commute b d) (hb : b ≠ 0 )
@@ -167,9 +167,9 @@ end DivisionRing
167
167
168
168
section Semifield
169
169
170
- variable [Semifield α ] {a b c d : α }
170
+ variable [Semifield K ] {a b c d : K }
171
171
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 ) :
173
173
a / b + c / d = (a * d + b * c) / (b * d) :=
174
174
(Commute.all b _).div_add_div (Commute.all _ _) hb hd
175
175
@@ -211,7 +211,7 @@ end Field
211
211
212
212
namespace RingHom
213
213
214
- protected theorem injective [DivisionRing α ] [Semiring β ] [Nontrivial β ] (f : α →+* β ) :
214
+ protected theorem injective [DivisionRing K ] [Semiring L ] [Nontrivial L ] (f : K →+* L ) :
215
215
Injective f :=
216
216
(injective_iff_map_eq_zero f).2 fun _ ↦ (map_eq_zero f).1
217
217
@@ -242,18 +242,18 @@ noncomputable abbrev Field.ofIsUnitOrEqZero [CommRing R] (h : ∀ a : R, IsUnit
242
242
end NoncomputableDefs
243
243
244
244
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 : K → L ) (hf : Injective f)
248
248
249
249
/-- Pullback a `DivisionSemiring` along an injective function. -/
250
250
-- 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 )
252
252
(add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y)
253
253
(inv : ∀ x, f x⁻¹ = (f x)⁻¹) (div : ∀ x y, f (x / y) = f x / f y)
254
254
(nsmul : ∀ (n : ℕ) (x), f (n • x) = n • f x) (nnqsmul : ∀ (q : ℚ≥0 ) (x), f (q • x) = q • f x)
255
255
(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
257
257
toSemiring := hf.semiring f zero one add mul nsmul npow natCast
258
258
__ := hf.groupWithZero f zero one mul inv div npow zpow
259
259
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 :
262
262
263
263
/-- Pullback a `DivisionSemiring` along an injective function. -/
264
264
-- 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 )
266
266
(add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y)
267
267
(neg : ∀ x, f (-x) = -f x) (sub : ∀ x y, f (x - y) = f x - f y) (inv : ∀ x, f x⁻¹ = (f x)⁻¹)
268
268
(div : ∀ x y, f (x / y) = f x / f y)
269
269
(nsmul : ∀ (n : ℕ) (x), f (n • x) = n • f x) (zsmul : ∀ (n : ℤ) (x), f (n • x) = n • f x)
270
270
(nnqsmul : ∀ (q : ℚ≥0 ) (x), f (q • x) = q • f x) (qsmul : ∀ (q : ℚ) (x), f (q • x) = q • f x)
271
271
(npow : ∀ (x) (n : ℕ), f (x ^ n) = f x ^ n) (zpow : ∀ (x) (n : ℤ), f (x ^ n) = f x ^ n)
272
272
(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
274
274
toRing := hf.ring f zero one add mul neg sub nsmul zsmul npow natCast intCast
275
275
__ := hf.groupWithZero f zero one mul inv div npow zpow
276
276
__ := 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)
280
280
281
281
/-- Pullback a `Field` along an injective function. -/
282
282
-- 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 )
284
284
(add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y)
285
285
(inv : ∀ x, f x⁻¹ = (f x)⁻¹) (div : ∀ x y, f (x / y) = f x / f y)
286
286
(nsmul : ∀ (n : ℕ) (x), f (n • x) = n • f x) (nnqsmul : ∀ (q : ℚ≥0 ) (x), f (q • x) = q • f x)
287
287
(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
289
289
toCommSemiring := hf.commSemiring f zero one add mul nsmul npow natCast
290
290
__ := hf.commGroupWithZero f zero one mul inv div npow zpow
291
291
__ := hf.divisionSemiring f zero one add mul inv div nsmul nnqsmul npow zpow natCast nnratCast
292
292
293
293
/-- Pullback a `Field` along an injective function. -/
294
294
-- 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 )
296
296
(add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y)
297
297
(neg : ∀ x, f (-x) = -f x) (sub : ∀ x y, f (x - y) = f x - f y) (inv : ∀ x, f x⁻¹ = (f x)⁻¹)
298
298
(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)
301
301
(npow : ∀ (x) (n : ℕ), f (x ^ n) = f x ^ n) (zpow : ∀ (x) (n : ℤ), f (x ^ n) = f x ^ n)
302
302
(natCast : ∀ n : ℕ, f n = n) (intCast : ∀ n : ℤ, f n = n) (nnratCast : ∀ q : ℚ≥0 , f q = q)
303
303
(ratCast : ∀ q : ℚ, f q = q) :
304
- Field α where
304
+ Field K where
305
305
toCommRing := hf.commRing f zero one add mul neg sub nsmul zsmul npow natCast intCast
306
306
__ := hf.divisionRing f zero one add mul neg sub inv div nsmul zsmul nnqsmul qsmul npow zpow
307
307
natCast intCast nnratCast ratCast
@@ -312,30 +312,30 @@ end Function.Injective
312
312
313
313
namespace OrderDual
314
314
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ᵒᵈ := ‹_›
320
320
321
321
end OrderDual
322
322
323
- @[simp] lemma toDual_ratCast [RatCast α ] (n : ℚ) : toDual (n : α ) = n := rfl
323
+ @[simp] lemma toDual_ratCast [RatCast K ] (n : ℚ) : toDual (n : K ) = n := rfl
324
324
325
- @[simp] lemma ofDual_ratCast [RatCast α ] (n : ℚ) : (ofDual n : α ) = n := rfl
325
+ @[simp] lemma ofDual_ratCast [RatCast K ] (n : ℚ) : (ofDual n : K ) = n := rfl
326
326
327
327
/-! ### Lexicographic order -/
328
328
329
329
namespace Lex
330
330
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 ) := ‹_›
336
336
337
337
end Lex
338
338
339
- @[simp] lemma toLex_ratCast [RatCast α ] (n : ℚ) : toLex (n : α ) = n := rfl
339
+ @[simp] lemma toLex_ratCast [RatCast K ] (n : ℚ) : toLex (n : K ) = n := rfl
340
340
341
- @[simp] lemma ofLex_ratCast [RatCast α ] (n : ℚ) : (ofLex n : α ) = n := rfl
341
+ @[simp] lemma ofLex_ratCast [RatCast K ] (n : ℚ) : (ofLex n : K ) = n := rfl
0 commit comments