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 : 2 * 3 = 6).
Check (eq_refl 6 <: 2 * 3 = 6).
Check (eq_refl 6 <<: 2 * 3 = 6).
Definition compute1 := Eval compute in 2 * 3.
Check (eq_refl compute1 : 6 = 6).
Check (eq_refl : 9223372036854775807 * 2 = 9223372036854775806).
Check (eq_refl 9223372036854775806 <: 9223372036854775807 * 2 = 9223372036854775806).
Check (eq_refl 9223372036854775806 <<: 9223372036854775807 * 2 = 9223372036854775806).
Definition compute2 := Eval compute in 9223372036854775807 * 2.
Check (eq_refl compute2 : 9223372036854775806 = 9223372036854775806).
|