@@ -22,6 +22,36 @@ structure Int8 where
22
22
-/
23
23
toUInt8 : UInt8
24
24
25
+ /--
26
+ The type of signed 16-bit integers. This type has special support in the
27
+ compiler to make it actually 16 bits rather than wrapping a `Nat`.
28
+ -/
29
+ structure Int16 where
30
+ /--
31
+ Obtain the `UInt16` that is 2's complement equivalent to the `Int16`.
32
+ -/
33
+ toUInt16 : UInt16
34
+
35
+ /--
36
+ The type of signed 32-bit integers. This type has special support in the
37
+ compiler to make it actually 32 bits rather than wrapping a `Nat`.
38
+ -/
39
+ structure Int32 where
40
+ /--
41
+ Obtain the `UInt32` that is 2's complement equivalent to the `Int32`.
42
+ -/
43
+ toUInt32 : UInt32
44
+
45
+ /--
46
+ The type of signed 64-bit integers. This type has special support in the
47
+ compiler to make it actually 64 bits rather than wrapping a `Nat`.
48
+ -/
49
+ structure Int64 where
50
+ /--
51
+ Obtain the `UInt64` that is 2's complement equivalent to the `Int64`.
52
+ -/
53
+ toUInt64 : UInt64
54
+
25
55
/-- The size of type `Int8`, that is, `2^8 = 256`. -/
26
56
abbrev Int8.size : Nat := 256
27
57
@@ -32,12 +62,16 @@ Obtain the `BitVec` that contains the 2's complement representation of the `Int8
32
62
33
63
@[extern "lean_int8_of_int"]
34
64
def Int8.ofInt (i : @& Int) : Int8 := ⟨⟨BitVec.ofInt 8 i⟩⟩
35
- @[extern "lean_int8_of_int "]
65
+ @[extern "lean_int8_of_nat "]
36
66
def Int8.ofNat (n : @& Nat) : Int8 := ⟨⟨BitVec.ofNat 8 n⟩⟩
37
67
abbrev Int.toInt8 := Int8.ofInt
38
68
abbrev Nat.toInt8 := Int8.ofNat
39
69
@[extern "lean_int8_to_int"]
40
70
def Int8.toInt (i : Int8) : Int := i.toBitVec.toInt
71
+ /--
72
+ This function has the same behavior as `Int.toNat` for negative numbers.
73
+ If you want to obtain the 2's complement representation use `toBitVec`.
74
+ -/
41
75
@[inline] def Int8.toNat (i : Int8) : Nat := i.toInt.toNat
42
76
@[extern "lean_int8_neg"]
43
77
def Int8.neg (i : Int8) : Int8 := ⟨⟨-i.toBitVec⟩⟩
@@ -58,17 +92,17 @@ def Int8.mul (a b : Int8) : Int8 := ⟨⟨a.toBitVec * b.toBitVec⟩⟩
58
92
@[extern "lean_int8_div"]
59
93
def Int8.div (a b : Int8) : Int8 := ⟨⟨BitVec.sdiv a.toBitVec b.toBitVec⟩⟩
60
94
@[extern "lean_int8_mod"]
61
- def Int8.mod (a b : Int8) : Int8 := ⟨⟨BitVec.smod a.toBitVec b.toBitVec⟩⟩
95
+ def Int8.mod (a b : Int8) : Int8 := ⟨⟨BitVec.srem a.toBitVec b.toBitVec⟩⟩
62
96
@[extern "lean_int8_land"]
63
97
def Int8.land (a b : Int8) : Int8 := ⟨⟨a.toBitVec &&& b.toBitVec⟩⟩
64
98
@[extern "lean_int8_lor"]
65
99
def Int8.lor (a b : Int8) : Int8 := ⟨⟨a.toBitVec ||| b.toBitVec⟩⟩
66
100
@[extern "lean_int8_xor"]
67
101
def Int8.xor (a b : Int8) : Int8 := ⟨⟨a.toBitVec ^^^ b.toBitVec⟩⟩
68
102
@[extern "lean_int8_shift_left"]
69
- def Int8.shiftLeft (a b : Int8) : Int8 := ⟨⟨a.toBitVec <<< (mod b 8 ).toBitVec ⟩⟩
103
+ def Int8.shiftLeft (a b : Int8) : Int8 := ⟨⟨a.toBitVec <<< (b.toBitVec.smod 8 )⟩⟩
70
104
@[extern "lean_int8_shift_right"]
71
- def Int8.shiftRight (a b : Int8) : Int8 := ⟨⟨BitVec.sshiftRight' a.toBitVec (mod b 8 ).toBitVec ⟩⟩
105
+ def Int8.shiftRight (a b : Int8) : Int8 := ⟨⟨BitVec.sshiftRight' a.toBitVec (b.toBitVec.smod 8 )⟩⟩
72
106
@[extern "lean_int8_complement"]
73
107
def Int8.complement (a : Int8) : Int8 := ⟨⟨~~~a.toBitVec⟩⟩
74
108
@@ -114,3 +148,318 @@ instance (a b : Int8) : Decidable (a < b) := Int8.decLt a b
114
148
instance (a b : Int8) : Decidable (a ≤ b) := Int8.decLe a b
115
149
instance : Max Int8 := maxOfLe
116
150
instance : Min Int8 := minOfLe
151
+
152
+ /-- The size of type `Int16`, that is, `2^16 = 65536`. -/
153
+ abbrev Int16.size : Nat := 65536
154
+
155
+ /--
156
+ Obtain the `BitVec` that contains the 2's complement representation of the `Int16`.
157
+ -/
158
+ @[inline] def Int16.toBitVec (x : Int16) : BitVec 16 := x.toUInt16.toBitVec
159
+
160
+ @[extern "lean_int16_of_int"]
161
+ def Int16.ofInt (i : @& Int) : Int16 := ⟨⟨BitVec.ofInt 16 i⟩⟩
162
+ @[extern "lean_int16_of_nat"]
163
+ def Int16.ofNat (n : @& Nat) : Int16 := ⟨⟨BitVec.ofNat 16 n⟩⟩
164
+ abbrev Int.toInt16 := Int16.ofInt
165
+ abbrev Nat.toInt16 := Int16.ofNat
166
+ @[extern "lean_int16_to_int"]
167
+ def Int16.toInt (i : Int16) : Int := i.toBitVec.toInt
168
+ /--
169
+ This function has the same behavior as `Int.toNat` for negative numbers.
170
+ If you want to obtain the 2's complement representation use `toBitVec`.
171
+ -/
172
+ @[inline] def Int16.toNat (i : Int16) : Nat := i.toInt.toNat
173
+ @[extern "lean_int16_to_int8"]
174
+ def Int16.toInt8 (a : Int16) : Int8 := ⟨⟨a.toBitVec.signExtend 8 ⟩⟩
175
+ @[extern "lean_int8_to_int16"]
176
+ def Int8.toInt16 (a : Int8) : Int16 := ⟨⟨a.toBitVec.signExtend 16 ⟩⟩
177
+ @[extern "lean_int16_neg"]
178
+ def Int16.neg (i : Int16) : Int16 := ⟨⟨-i.toBitVec⟩⟩
179
+
180
+ instance : ToString Int16 where
181
+ toString i := toString i.toInt
182
+
183
+ instance : OfNat Int16 n := ⟨Int16.ofNat n⟩
184
+ instance : Neg Int16 where
185
+ neg := Int16.neg
186
+
187
+ @[extern "lean_int16_add"]
188
+ def Int16.add (a b : Int16) : Int16 := ⟨⟨a.toBitVec + b.toBitVec⟩⟩
189
+ @[extern "lean_int16_sub"]
190
+ def Int16.sub (a b : Int16) : Int16 := ⟨⟨a.toBitVec - b.toBitVec⟩⟩
191
+ @[extern "lean_int16_mul"]
192
+ def Int16.mul (a b : Int16) : Int16 := ⟨⟨a.toBitVec * b.toBitVec⟩⟩
193
+ @[extern "lean_int16_div"]
194
+ def Int16.div (a b : Int16) : Int16 := ⟨⟨BitVec.sdiv a.toBitVec b.toBitVec⟩⟩
195
+ @[extern "lean_int16_mod"]
196
+ def Int16.mod (a b : Int16) : Int16 := ⟨⟨BitVec.srem a.toBitVec b.toBitVec⟩⟩
197
+ @[extern "lean_int16_land"]
198
+ def Int16.land (a b : Int16) : Int16 := ⟨⟨a.toBitVec &&& b.toBitVec⟩⟩
199
+ @[extern "lean_int16_lor"]
200
+ def Int16.lor (a b : Int16) : Int16 := ⟨⟨a.toBitVec ||| b.toBitVec⟩⟩
201
+ @[extern "lean_int16_xor"]
202
+ def Int16.xor (a b : Int16) : Int16 := ⟨⟨a.toBitVec ^^^ b.toBitVec⟩⟩
203
+ @[extern "lean_int16_shift_left"]
204
+ def Int16.shiftLeft (a b : Int16) : Int16 := ⟨⟨a.toBitVec <<< (b.toBitVec.smod 16 )⟩⟩
205
+ @[extern "lean_int16_shift_right"]
206
+ def Int16.shiftRight (a b : Int16) : Int16 := ⟨⟨BitVec.sshiftRight' a.toBitVec (b.toBitVec.smod 16 )⟩⟩
207
+ @[extern "lean_int16_complement"]
208
+ def Int16.complement (a : Int16) : Int16 := ⟨⟨~~~a.toBitVec⟩⟩
209
+
210
+ @[extern "lean_int16_dec_eq"]
211
+ def Int16.decEq (a b : Int16) : Decidable (a = b) :=
212
+ match a, b with
213
+ | ⟨n⟩, ⟨m⟩ =>
214
+ if h : n = m then
215
+ isTrue <| h ▸ rfl
216
+ else
217
+ isFalse (fun h' => Int16.noConfusion h' (fun h' => absurd h' h))
218
+
219
+ def Int16.lt (a b : Int16) : Prop := a.toBitVec.slt b.toBitVec
220
+ def Int16.le (a b : Int16) : Prop := a.toBitVec.sle b.toBitVec
221
+
222
+ instance : Inhabited Int16 where
223
+ default := 0
224
+
225
+ instance : Add Int16 := ⟨Int16.add⟩
226
+ instance : Sub Int16 := ⟨Int16.sub⟩
227
+ instance : Mul Int16 := ⟨Int16.mul⟩
228
+ instance : Mod Int16 := ⟨Int16.mod⟩
229
+ instance : Div Int16 := ⟨Int16.div⟩
230
+ instance : LT Int16 := ⟨Int16.lt⟩
231
+ instance : LE Int16 := ⟨Int16.le⟩
232
+ instance : Complement Int16 := ⟨Int16.complement⟩
233
+ instance : AndOp Int16 := ⟨Int16.land⟩
234
+ instance : OrOp Int16 := ⟨Int16.lor⟩
235
+ instance : Xor Int16 := ⟨Int16.xor⟩
236
+ instance : ShiftLeft Int16 := ⟨Int16.shiftLeft⟩
237
+ instance : ShiftRight Int16 := ⟨Int16.shiftRight⟩
238
+ instance : DecidableEq Int16 := Int16.decEq
239
+
240
+ @[extern "lean_int16_dec_lt"]
241
+ def Int16.decLt (a b : Int16) : Decidable (a < b) :=
242
+ inferInstanceAs (Decidable (a.toBitVec.slt b.toBitVec))
243
+
244
+ @[extern "lean_int16_dec_le"]
245
+ def Int16.decLe (a b : Int16) : Decidable (a ≤ b) :=
246
+ inferInstanceAs (Decidable (a.toBitVec.sle b.toBitVec))
247
+
248
+ instance (a b : Int16) : Decidable (a < b) := Int16.decLt a b
249
+ instance (a b : Int16) : Decidable (a ≤ b) := Int16.decLe a b
250
+ instance : Max Int16 := maxOfLe
251
+ instance : Min Int16 := minOfLe
252
+
253
+ /-- The size of type `Int32`, that is, `2^32 = 4294967296`. -/
254
+ abbrev Int32.size : Nat := 4294967296
255
+
256
+ /--
257
+ Obtain the `BitVec` that contains the 2's complement representation of the `Int32`.
258
+ -/
259
+ @[inline] def Int32.toBitVec (x : Int32) : BitVec 32 := x.toUInt32.toBitVec
260
+
261
+ @[extern "lean_int32_of_int"]
262
+ def Int32.ofInt (i : @& Int) : Int32 := ⟨⟨BitVec.ofInt 32 i⟩⟩
263
+ @[extern "lean_int32_of_nat"]
264
+ def Int32.ofNat (n : @& Nat) : Int32 := ⟨⟨BitVec.ofNat 32 n⟩⟩
265
+ abbrev Int.toInt32 := Int32.ofInt
266
+ abbrev Nat.toInt32 := Int32.ofNat
267
+ @[extern "lean_int32_to_int"]
268
+ def Int32.toInt (i : Int32) : Int := i.toBitVec.toInt
269
+ /--
270
+ This function has the same behavior as `Int.toNat` for negative numbers.
271
+ If you want to obtain the 2's complement representation use `toBitVec`.
272
+ -/
273
+ @[inline] def Int32.toNat (i : Int32) : Nat := i.toInt.toNat
274
+ @[extern "lean_int32_to_int8"]
275
+ def Int32.toInt8 (a : Int32) : Int8 := ⟨⟨a.toBitVec.signExtend 8 ⟩⟩
276
+ @[extern "lean_int32_to_int16"]
277
+ def Int32.toInt16 (a : Int32) : Int16 := ⟨⟨a.toBitVec.signExtend 16 ⟩⟩
278
+ @[extern "lean_int8_to_int32"]
279
+ def Int8.toInt32 (a : Int8) : Int32 := ⟨⟨a.toBitVec.signExtend 32 ⟩⟩
280
+ @[extern "lean_int16_to_int32"]
281
+ def Int16.toInt32 (a : Int16) : Int32 := ⟨⟨a.toBitVec.signExtend 32 ⟩⟩
282
+ @[extern "lean_int32_neg"]
283
+ def Int32.neg (i : Int32) : Int32 := ⟨⟨-i.toBitVec⟩⟩
284
+
285
+ instance : ToString Int32 where
286
+ toString i := toString i.toInt
287
+
288
+ instance : OfNat Int32 n := ⟨Int32.ofNat n⟩
289
+ instance : Neg Int32 where
290
+ neg := Int32.neg
291
+
292
+ @[extern "lean_int32_add"]
293
+ def Int32.add (a b : Int32) : Int32 := ⟨⟨a.toBitVec + b.toBitVec⟩⟩
294
+ @[extern "lean_int32_sub"]
295
+ def Int32.sub (a b : Int32) : Int32 := ⟨⟨a.toBitVec - b.toBitVec⟩⟩
296
+ @[extern "lean_int32_mul"]
297
+ def Int32.mul (a b : Int32) : Int32 := ⟨⟨a.toBitVec * b.toBitVec⟩⟩
298
+ @[extern "lean_int32_div"]
299
+ def Int32.div (a b : Int32) : Int32 := ⟨⟨BitVec.sdiv a.toBitVec b.toBitVec⟩⟩
300
+ @[extern "lean_int32_mod"]
301
+ def Int32.mod (a b : Int32) : Int32 := ⟨⟨BitVec.srem a.toBitVec b.toBitVec⟩⟩
302
+ @[extern "lean_int32_land"]
303
+ def Int32.land (a b : Int32) : Int32 := ⟨⟨a.toBitVec &&& b.toBitVec⟩⟩
304
+ @[extern "lean_int32_lor"]
305
+ def Int32.lor (a b : Int32) : Int32 := ⟨⟨a.toBitVec ||| b.toBitVec⟩⟩
306
+ @[extern "lean_int32_xor"]
307
+ def Int32.xor (a b : Int32) : Int32 := ⟨⟨a.toBitVec ^^^ b.toBitVec⟩⟩
308
+ @[extern "lean_int32_shift_left"]
309
+ def Int32.shiftLeft (a b : Int32) : Int32 := ⟨⟨a.toBitVec <<< (b.toBitVec.smod 32 )⟩⟩
310
+ @[extern "lean_int32_shift_right"]
311
+ def Int32.shiftRight (a b : Int32) : Int32 := ⟨⟨BitVec.sshiftRight' a.toBitVec (b.toBitVec.smod 32 )⟩⟩
312
+ @[extern "lean_int32_complement"]
313
+ def Int32.complement (a : Int32) : Int32 := ⟨⟨~~~a.toBitVec⟩⟩
314
+
315
+ @[extern "lean_int32_dec_eq"]
316
+ def Int32.decEq (a b : Int32) : Decidable (a = b) :=
317
+ match a, b with
318
+ | ⟨n⟩, ⟨m⟩ =>
319
+ if h : n = m then
320
+ isTrue <| h ▸ rfl
321
+ else
322
+ isFalse (fun h' => Int32.noConfusion h' (fun h' => absurd h' h))
323
+
324
+ def Int32.lt (a b : Int32) : Prop := a.toBitVec.slt b.toBitVec
325
+ def Int32.le (a b : Int32) : Prop := a.toBitVec.sle b.toBitVec
326
+
327
+ instance : Inhabited Int32 where
328
+ default := 0
329
+
330
+ instance : Add Int32 := ⟨Int32.add⟩
331
+ instance : Sub Int32 := ⟨Int32.sub⟩
332
+ instance : Mul Int32 := ⟨Int32.mul⟩
333
+ instance : Mod Int32 := ⟨Int32.mod⟩
334
+ instance : Div Int32 := ⟨Int32.div⟩
335
+ instance : LT Int32 := ⟨Int32.lt⟩
336
+ instance : LE Int32 := ⟨Int32.le⟩
337
+ instance : Complement Int32 := ⟨Int32.complement⟩
338
+ instance : AndOp Int32 := ⟨Int32.land⟩
339
+ instance : OrOp Int32 := ⟨Int32.lor⟩
340
+ instance : Xor Int32 := ⟨Int32.xor⟩
341
+ instance : ShiftLeft Int32 := ⟨Int32.shiftLeft⟩
342
+ instance : ShiftRight Int32 := ⟨Int32.shiftRight⟩
343
+ instance : DecidableEq Int32 := Int32.decEq
344
+
345
+ @[extern "lean_int32_dec_lt"]
346
+ def Int32.decLt (a b : Int32) : Decidable (a < b) :=
347
+ inferInstanceAs (Decidable (a.toBitVec.slt b.toBitVec))
348
+
349
+ @[extern "lean_int32_dec_le"]
350
+ def Int32.decLe (a b : Int32) : Decidable (a ≤ b) :=
351
+ inferInstanceAs (Decidable (a.toBitVec.sle b.toBitVec))
352
+
353
+ instance (a b : Int32) : Decidable (a < b) := Int32.decLt a b
354
+ instance (a b : Int32) : Decidable (a ≤ b) := Int32.decLe a b
355
+ instance : Max Int32 := maxOfLe
356
+ instance : Min Int32 := minOfLe
357
+
358
+ /-- The size of type `Int64`, that is, `2^64 = 18446744073709551616`. -/
359
+ abbrev Int64.size : Nat := 18446744073709551616
360
+
361
+ /--
362
+ Obtain the `BitVec` that contains the 2's complement representation of the `Int64`.
363
+ -/
364
+ @[inline] def Int64.toBitVec (x : Int64) : BitVec 64 := x.toUInt64.toBitVec
365
+
366
+ @[extern "lean_int64_of_int"]
367
+ def Int64.ofInt (i : @& Int) : Int64 := ⟨⟨BitVec.ofInt 64 i⟩⟩
368
+ @[extern "lean_int64_of_nat"]
369
+ def Int64.ofNat (n : @& Nat) : Int64 := ⟨⟨BitVec.ofNat 64 n⟩⟩
370
+ abbrev Int.toInt64 := Int64.ofInt
371
+ abbrev Nat.toInt64 := Int64.ofNat
372
+ @[extern "lean_int64_to_int_sint"]
373
+ def Int64.toInt (i : Int64) : Int := i.toBitVec.toInt
374
+ /--
375
+ This function has the same behavior as `Int.toNat` for negative numbers.
376
+ If you want to obtain the 2's complement representation use `toBitVec`.
377
+ -/
378
+ @[inline] def Int64.toNat (i : Int64) : Nat := i.toInt.toNat
379
+ @[extern "lean_int64_to_int8"]
380
+ def Int64.toInt8 (a : Int64) : Int8 := ⟨⟨a.toBitVec.signExtend 8 ⟩⟩
381
+ @[extern "lean_int64_to_int16"]
382
+ def Int64.toInt16 (a : Int64) : Int16 := ⟨⟨a.toBitVec.signExtend 16 ⟩⟩
383
+ @[extern "lean_int64_to_int32"]
384
+ def Int64.toInt32 (a : Int64) : Int32 := ⟨⟨a.toBitVec.signExtend 32 ⟩⟩
385
+ @[extern "lean_int8_to_int64"]
386
+ def Int8.toInt64 (a : Int8) : Int64 := ⟨⟨a.toBitVec.signExtend 64 ⟩⟩
387
+ @[extern "lean_int16_to_int64"]
388
+ def Int16.toInt64 (a : Int16) : Int64 := ⟨⟨a.toBitVec.signExtend 64 ⟩⟩
389
+ @[extern "lean_int32_to_int64"]
390
+ def Int32.toInt64 (a : Int32) : Int64 := ⟨⟨a.toBitVec.signExtend 64 ⟩⟩
391
+ @[extern "lean_int64_neg"]
392
+ def Int64.neg (i : Int64) : Int64 := ⟨⟨-i.toBitVec⟩⟩
393
+
394
+ instance : ToString Int64 where
395
+ toString i := toString i.toInt
396
+
397
+ instance : OfNat Int64 n := ⟨Int64.ofNat n⟩
398
+ instance : Neg Int64 where
399
+ neg := Int64.neg
400
+
401
+ @[extern "lean_int64_add"]
402
+ def Int64.add (a b : Int64) : Int64 := ⟨⟨a.toBitVec + b.toBitVec⟩⟩
403
+ @[extern "lean_int64_sub"]
404
+ def Int64.sub (a b : Int64) : Int64 := ⟨⟨a.toBitVec - b.toBitVec⟩⟩
405
+ @[extern "lean_int64_mul"]
406
+ def Int64.mul (a b : Int64) : Int64 := ⟨⟨a.toBitVec * b.toBitVec⟩⟩
407
+ @[extern "lean_int64_div"]
408
+ def Int64.div (a b : Int64) : Int64 := ⟨⟨BitVec.sdiv a.toBitVec b.toBitVec⟩⟩
409
+ @[extern "lean_int64_mod"]
410
+ def Int64.mod (a b : Int64) : Int64 := ⟨⟨BitVec.srem a.toBitVec b.toBitVec⟩⟩
411
+ @[extern "lean_int64_land"]
412
+ def Int64.land (a b : Int64) : Int64 := ⟨⟨a.toBitVec &&& b.toBitVec⟩⟩
413
+ @[extern "lean_int64_lor"]
414
+ def Int64.lor (a b : Int64) : Int64 := ⟨⟨a.toBitVec ||| b.toBitVec⟩⟩
415
+ @[extern "lean_int64_xor"]
416
+ def Int64.xor (a b : Int64) : Int64 := ⟨⟨a.toBitVec ^^^ b.toBitVec⟩⟩
417
+ @[extern "lean_int64_shift_left"]
418
+ def Int64.shiftLeft (a b : Int64) : Int64 := ⟨⟨a.toBitVec <<< (b.toBitVec.smod 64 )⟩⟩
419
+ @[extern "lean_int64_shift_right"]
420
+ def Int64.shiftRight (a b : Int64) : Int64 := ⟨⟨BitVec.sshiftRight' a.toBitVec (b.toBitVec.smod 64 )⟩⟩
421
+ @[extern "lean_int64_complement"]
422
+ def Int64.complement (a : Int64) : Int64 := ⟨⟨~~~a.toBitVec⟩⟩
423
+
424
+ @[extern "lean_int64_dec_eq"]
425
+ def Int64.decEq (a b : Int64) : Decidable (a = b) :=
426
+ match a, b with
427
+ | ⟨n⟩, ⟨m⟩ =>
428
+ if h : n = m then
429
+ isTrue <| h ▸ rfl
430
+ else
431
+ isFalse (fun h' => Int64.noConfusion h' (fun h' => absurd h' h))
432
+
433
+ def Int64.lt (a b : Int64) : Prop := a.toBitVec.slt b.toBitVec
434
+ def Int64.le (a b : Int64) : Prop := a.toBitVec.sle b.toBitVec
435
+
436
+ instance : Inhabited Int64 where
437
+ default := 0
438
+
439
+ instance : Add Int64 := ⟨Int64.add⟩
440
+ instance : Sub Int64 := ⟨Int64.sub⟩
441
+ instance : Mul Int64 := ⟨Int64.mul⟩
442
+ instance : Mod Int64 := ⟨Int64.mod⟩
443
+ instance : Div Int64 := ⟨Int64.div⟩
444
+ instance : LT Int64 := ⟨Int64.lt⟩
445
+ instance : LE Int64 := ⟨Int64.le⟩
446
+ instance : Complement Int64 := ⟨Int64.complement⟩
447
+ instance : AndOp Int64 := ⟨Int64.land⟩
448
+ instance : OrOp Int64 := ⟨Int64.lor⟩
449
+ instance : Xor Int64 := ⟨Int64.xor⟩
450
+ instance : ShiftLeft Int64 := ⟨Int64.shiftLeft⟩
451
+ instance : ShiftRight Int64 := ⟨Int64.shiftRight⟩
452
+ instance : DecidableEq Int64 := Int64.decEq
453
+
454
+ @[extern "lean_int64_dec_lt"]
455
+ def Int64.decLt (a b : Int64) : Decidable (a < b) :=
456
+ inferInstanceAs (Decidable (a.toBitVec.slt b.toBitVec))
457
+
458
+ @[extern "lean_int64_dec_le"]
459
+ def Int64.decLe (a b : Int64) : Decidable (a ≤ b) :=
460
+ inferInstanceAs (Decidable (a.toBitVec.sle b.toBitVec))
461
+
462
+ instance (a b : Int64) : Decidable (a < b) := Int64.decLt a b
463
+ instance (a b : Int64) : Decidable (a ≤ b) := Int64.decLe a b
464
+ instance : Max Int64 := maxOfLe
465
+ instance : Min Int64 := minOfLe
0 commit comments