@@ -89,11 +89,14 @@ theorem equivTuple_apply {R : Type*} (c₁ c₂ : R) (x : ℍ[R,c₁,c₂]) :
89
89
@[simp]
90
90
theorem mk.eta {R : Type *} {c₁ c₂} (a : ℍ[R,c₁,c₂]) : mk a.1 a.2 a.3 a.4 = a := rfl
91
91
92
- variable {S T R : Type *} [CommRing R] {c₁ c₂ : R} (r x y z : R) (a b c : ℍ[R,c₁,c₂])
92
+ variable {S T R : Type *} {c₁ c₂ : R} (r x y z : R) (a b c : ℍ[R,c₁,c₂])
93
93
94
94
instance [Subsingleton R] : Subsingleton ℍ[R, c₁, c₂] := (equivTuple c₁ c₂).subsingleton
95
95
instance [Nontrivial R] : Nontrivial ℍ[R, c₁, c₂] := (equivTuple c₁ c₂).surjective.nontrivial
96
96
97
+ section Zero
98
+ variable [Zero R]
99
+
97
100
/-- The imaginary part of a quaternion. -/
98
101
def im (x : ℍ[R,c₁,c₂]) : ℍ[R,c₁,c₂] :=
99
102
⟨0 , x.imI, x.imJ, x.imK⟩
@@ -159,6 +162,9 @@ theorem coe_zero : ((0 : R) : ℍ[R,c₁,c₂]) = 0 := rfl
159
162
160
163
instance : Inhabited ℍ[R,c₁,c₂] := ⟨0 ⟩
161
164
165
+ section One
166
+ variable [One R]
167
+
162
168
-- Porting note: removed `simps`, added simp lemmas manually
163
169
instance : One ℍ[R,c₁,c₂] := ⟨⟨1 , 0 , 0 , 0 ⟩⟩
164
170
@@ -175,6 +181,11 @@ instance : One ℍ[R,c₁,c₂] := ⟨⟨1, 0, 0, 0⟩⟩
175
181
@[simp, norm_cast]
176
182
theorem coe_one : ((1 : R) : ℍ[R,c₁,c₂]) = 1 := rfl
177
183
184
+ end One
185
+ end Zero
186
+ section Add
187
+ variable [Add R]
188
+
178
189
-- Porting note: removed `simps`, added simp lemmas manually
179
190
instance : Add ℍ[R,c₁,c₂] :=
180
191
⟨fun a b => ⟨a.1 + b.1 , a.2 + b.2 , a.3 + b.3 , a.4 + b.4 ⟩⟩
@@ -187,17 +198,27 @@ instance : Add ℍ[R,c₁,c₂] :=
187
198
188
199
@[simp] theorem add_imK : (a + b).imK = a.imK + b.imK := rfl
189
200
190
- @[simp] theorem add_im : (a + b).im = a.im + b.im :=
191
- QuaternionAlgebra.ext (zero_add _).symm rfl rfl rfl
192
-
193
201
@[simp]
194
202
theorem mk_add_mk (a₁ a₂ a₃ a₄ b₁ b₂ b₃ b₄ : R) :
195
203
(mk a₁ a₂ a₃ a₄ : ℍ[R,c₁,c₂]) + mk b₁ b₂ b₃ b₄ = mk (a₁ + b₁) (a₂ + b₂) (a₃ + b₃) (a₄ + b₄) :=
196
204
rfl
197
205
206
+ end Add
207
+
208
+ section AddZeroClass
209
+ variable [AddZeroClass R]
210
+
211
+ @[simp] theorem add_im : (a + b).im = a.im + b.im :=
212
+ QuaternionAlgebra.ext (zero_add _).symm rfl rfl rfl
213
+
198
214
@[simp, norm_cast]
199
215
theorem coe_add : ((x + y : R) : ℍ[R,c₁,c₂]) = x + y := by ext <;> simp
200
216
217
+ end AddZeroClass
218
+
219
+ section Neg
220
+ variable [Neg R]
221
+
201
222
-- Porting note: removed `simps`, added simp lemmas manually
202
223
instance : Neg ℍ[R,c₁,c₂] := ⟨fun a => ⟨-a.1 , -a.2 , -a.3 , -a.4 ⟩⟩
203
224
@@ -209,13 +230,18 @@ instance : Neg ℍ[R,c₁,c₂] := ⟨fun a => ⟨-a.1, -a.2, -a.3, -a.4⟩⟩
209
230
210
231
@[simp] theorem neg_imK : (-a).imK = -a.imK := rfl
211
232
212
- @[simp] theorem neg_im : (-a).im = -a.im :=
213
- QuaternionAlgebra.ext neg_zero.symm rfl rfl rfl
214
-
215
233
@[simp]
216
234
theorem neg_mk (a₁ a₂ a₃ a₄ : R) : -(mk a₁ a₂ a₃ a₄ : ℍ[R,c₁,c₂]) = ⟨-a₁, -a₂, -a₃, -a₄⟩ :=
217
235
rfl
218
236
237
+ end Neg
238
+
239
+ section AddGroup
240
+ variable [AddGroup R]
241
+
242
+ @[simp] theorem neg_im : (-a).im = -a.im :=
243
+ QuaternionAlgebra.ext neg_zero.symm rfl rfl rfl
244
+
219
245
@[simp, norm_cast]
220
246
theorem coe_neg : ((-x : R) : ℍ[R,c₁,c₂]) = -x := by ext <;> simp
221
247
@@ -254,6 +280,11 @@ theorem sub_self_im : a - a.im = a.re :=
254
280
theorem sub_self_re : a - a.re = a.im :=
255
281
QuaternionAlgebra.ext (sub_self _) (sub_zero _) (sub_zero _) (sub_zero _)
256
282
283
+ end AddGroup
284
+
285
+ section Ring
286
+ variable [Ring R]
287
+
257
288
/-- Multiplication is given by
258
289
259
290
* `1 * x = x * 1 = x`;
@@ -290,11 +321,11 @@ theorem mk_mul_mk (a₁ a₂ a₃ a₄ b₁ b₂ b₃ b₄ : R) :
290
321
a₁ * b₃ + c₁ * a₂ * b₄ + a₃ * b₁ - c₁ * a₄ * b₂, a₁ * b₄ + a₂ * b₃ - a₃ * b₂ + a₄ * b₁⟩ :=
291
322
rfl
292
323
293
- section
324
+ end Ring
325
+ section SMul
294
326
295
327
variable [SMul S R] [SMul T R] (s : S)
296
328
297
- -- Porting note: Lean 4 auto drops the unused `[Ring R]` argument
298
329
instance : SMul S ℍ[R,c₁,c₂] where smul s a := ⟨s • a.1 , s • a.2 , s • a.3 , s • a.4 ⟩
299
330
300
331
instance [SMul S T] [IsScalarTower S T R] : IsScalarTower S T ℍ[R,c₁,c₂] where
@@ -311,25 +342,28 @@ instance [SMulCommClass S T R] : SMulCommClass S T ℍ[R,c₁,c₂] where
311
342
312
343
@[simp] theorem smul_imK : (s • a).imK = s • a.imK := rfl
313
344
314
- @[simp] theorem smul_im {S} [SMulZeroClass S R] (s : S) : (s • a).im = s • a.im :=
345
+ @[simp] theorem smul_im {S} [CommRing R] [ SMulZeroClass S R] (s : S) : (s • a).im = s • a.im :=
315
346
QuaternionAlgebra.ext (smul_zero s).symm rfl rfl rfl
316
347
317
348
@[simp]
318
349
theorem smul_mk (re im_i im_j im_k : R) :
319
350
s • (⟨re, im_i, im_j, im_k⟩ : ℍ[R,c₁,c₂]) = ⟨s • re, s • im_i, s • im_j, s • im_k⟩ :=
320
351
rfl
321
352
322
- end
353
+ end SMul
323
354
324
355
@[simp, norm_cast]
325
- theorem coe_smul [SMulZeroClass S R] (s : S) (r : R) :
356
+ theorem coe_smul [Zero R] [ SMulZeroClass S R] (s : S) (r : R) :
326
357
(↑(s • r) : ℍ[R,c₁,c₂]) = s • (r : ℍ[R,c₁,c₂]) :=
327
358
QuaternionAlgebra.ext rfl (smul_zero s).symm (smul_zero s).symm (smul_zero s).symm
328
359
329
- instance : AddCommGroup ℍ[R,c₁,c₂] :=
360
+ instance [AddCommGroup R] : AddCommGroup ℍ[R,c₁,c₂] :=
330
361
(equivProd c₁ c₂).injective.addCommGroup _ rfl (fun _ _ ↦ rfl) (fun _ ↦ rfl) (fun _ _ ↦ rfl)
331
362
(fun _ _ ↦ rfl) (fun _ _ ↦ rfl)
332
363
364
+ section AddCommGroupWithOne
365
+ variable [AddCommGroupWithOne R]
366
+
333
367
instance : AddCommGroupWithOne ℍ[R,c₁,c₂] where
334
368
natCast n := ((n : R) : ℍ[R,c₁,c₂])
335
369
natCast_zero := by simp
@@ -424,6 +458,11 @@ theorem coe_intCast (z : ℤ) : ↑(z : R) = (z : ℍ[R,c₁,c₂]) :=
424
458
@[deprecated (since := "2024-04-17")]
425
459
alias coe_int_cast := coe_intCast
426
460
461
+ end AddCommGroupWithOne
462
+
463
+ -- For the remainder of the file we assume `CommRing R`.
464
+ variable [CommRing R]
465
+
427
466
instance instRing : Ring ℍ[R,c₁,c₂] where
428
467
__ := inferInstanceAs (AddCommGroupWithOne ℍ[R,c₁,c₂])
429
468
left_distrib _ _ _ := by ext <;> simp <;> ring
0 commit comments