1 2 3 4 5 6 7 8 9 10 11 12 13 14 15 16 17
|
Require Import Uint63.
Set Implicit Arguments.
Open Scope uint63_scope.
Check (eq_refl : 3 - 2 = 1).
Check (eq_refl 1 <: 3 - 2 = 1).
Check (eq_refl 1 <<: 3 - 2 = 1).
Definition compute1 := Eval compute in 3 - 2.
Check (eq_refl compute1 : 1 = 1).
Check (eq_refl : 0 - 1 = 9223372036854775807).
Check (eq_refl 9223372036854775807 <: 0 - 1 = 9223372036854775807).
Check (eq_refl 9223372036854775807 <<: 0 - 1 = 9223372036854775807).
Definition compute2 := Eval compute in 0 - 1.
Check (eq_refl compute2 : 9223372036854775807 = 9223372036854775807).
|