[ fix #1579 ] Nat hack for comparison operators (#1580)

Co-authored-by: Guillaume ALLAIS <guillaume.allais@ens-lyon.org>
This commit is contained in:
Tim Engler 2021-06-28 23:48:08 +08:00 committed by GitHub
parent 8d7d13dd41
commit 05f28724ed
No known key found for this signature in database
GPG Key ID: 4AEE18F83AFDEB23
8 changed files with 81 additions and 31 deletions

View File

@ -89,6 +89,7 @@ public export
interface Eq ty => Ord ty where
constructor MkOrd
compare : ty -> ty -> Ordering
compare x y = if x < y then LT else if x == y then EQ else GT
(<) : ty -> ty -> Bool
(<) x y = compare x y == LT
@ -129,8 +130,6 @@ Ord Bool where
public export
Ord Int where
compare x y = if x < y then LT else if x == y then EQ else GT
(<) x y = intToBool (prim__lt_Int x y)
(<=) x y = intToBool (prim__lte_Int x y)
(>) x y = intToBool (prim__gt_Int x y)
@ -138,17 +137,18 @@ Ord Int where
public export
Ord Integer where
compare x y = if x < y then LT else if x == y then EQ else GT
(<) x y = intToBool (prim__lt_Integer x y)
(<=) x y = intToBool (prim__lte_Integer x y)
(>) x y = intToBool (prim__gt_Integer x y)
(>=) x y = intToBool (prim__gte_Integer x y)
-- Used for the nat hack
public export
compareInteger : (x, y : Integer) -> Ordering
compareInteger = compare
public export
Ord Bits8 where
compare x y = if x < y then LT else if x == y then EQ else GT
(<) x y = intToBool (prim__lt_Bits8 x y)
(<=) x y = intToBool (prim__lte_Bits8 x y)
(>) x y = intToBool (prim__gt_Bits8 x y)
@ -156,8 +156,6 @@ Ord Bits8 where
public export
Ord Bits16 where
compare x y = if x < y then LT else if x == y then EQ else GT
(<) x y = intToBool (prim__lt_Bits16 x y)
(<=) x y = intToBool (prim__lte_Bits16 x y)
(>) x y = intToBool (prim__gt_Bits16 x y)
@ -165,8 +163,6 @@ Ord Bits16 where
public export
Ord Bits32 where
compare x y = if x < y then LT else if x == y then EQ else GT
(<) x y = intToBool (prim__lt_Bits32 x y)
(<=) x y = intToBool (prim__lte_Bits32 x y)
(>) x y = intToBool (prim__gt_Bits32 x y)
@ -174,8 +170,6 @@ Ord Bits32 where
public export
Ord Bits64 where
compare x y = if x < y then LT else if x == y then EQ else GT
(<) x y = intToBool (prim__lt_Bits64 x y)
(<=) x y = intToBool (prim__lte_Bits64 x y)
(>) x y = intToBool (prim__gt_Bits64 x y)
@ -183,8 +177,6 @@ Ord Bits64 where
public export
Ord Double where
compare x y = if x < y then LT else if x == y then EQ else GT
(<) x y = intToBool (prim__lt_Double x y)
(<=) x y = intToBool (prim__lte_Double x y)
(>) x y = intToBool (prim__gt_Double x y)

View File

@ -69,18 +69,28 @@ Num Nat where
fromInteger x = integerToNat x
-- used for nat hack
public export
equalNat : (m, n : Nat) -> Bool
equalNat Z Z = True
equalNat (S j) (S k) = equalNat j k
equalNat _ _ = False
public export
Eq Nat where
Z == Z = True
S j == S k = j == k
_ == _ = False
(==) = equalNat
-- used for nat hack
public export
compareNat : (m, n : Nat) -> Ordering
compareNat Z Z = EQ
compareNat Z (S k) = LT
compareNat (S k) Z = GT
compareNat (S j) (S k) = compareNat j k
public export
Ord Nat where
compare Z Z = EQ
compare Z (S k) = LT
compare (S k) Z = GT
compare (S j) (S k) = compare j k
compare = compareNat
public export
natToInteger : Nat -> Integer

View File

@ -203,6 +203,10 @@ natHack =
, MagicCRef (NS typesNS (UN "mult")) 2
(\ fc, fc', [m,n] => CApp fc (CRef fc' (UN "prim__mul_Integer")) [m, n])
, MagicCRef (NS typesNS (UN "minus")) 2 magic__natMinus
, MagicCRef (NS typesNS (UN "equalNat")) 2
(\ fc, fc', [m,n] => CApp fc (CRef fc' (UN "prim__eq_Integer")) [m, n])
, MagicCRef (NS typesNS (UN "compareNat")) 2
(\ fc, fc', [m,n] => CApp fc (CRef fc' (NS eqOrdNS (UN "compareInteger"))) [m, n])
]
-- get all transformation from %builtin pragmas

View File

@ -6,3 +6,9 @@ doMult x y = x `mult` y
doMinus : Nat -> Nat -> Nat
doMinus x y = x `minus` y
doEqual : (m, n : Nat) -> Bool
doEqual m n = m `equalNat` n
doCompare : (m, n : Nat) -> Ordering
doCompare m n = m `compareNat` n

View File

@ -38,4 +38,30 @@ Refers to: [Prelude.Types.mult]
Refers to (runtime): []
Flags: [covering]
Size change: Prelude.Types.mult: [Just (0, Same), Just (1, Same)]
Main> Main.doCompare ==> [{arg:0}, {arg:1}];
Compile time tree: [0] (Prelude.Types.compareNat {arg:0}[0] {arg:1}[1])
Run time tree: Error: ""
RigW
Erasable args: []
Detaggable arg types: []
Specialise args: []
Inferrable args: []
Compiled: [{arg:0}, {arg:1}]: (Prelude.EqOrd.compareInteger [!{arg:0}, !{arg:1}])
Refers to: [Prelude.Types.compareNat]
Refers to (runtime): [Prelude.EqOrd.compareInteger]
Flags: [covering]
Size change: Prelude.Types.compareNat: [Just (0, Same), Just (1, Same)]
Main> Main.doEqual ==> [{arg:0}, {arg:1}];
Compile time tree: [0] (Prelude.Types.equalNat {arg:0}[0] {arg:1}[1])
Run time tree: Error: ""
RigW
Erasable args: []
Detaggable arg types: []
Specialise args: []
Inferrable args: []
Compiled: [{arg:0}, {arg:1}]: (==Integer [!{arg:0}, !{arg:1}])
Refers to: [Prelude.Types.equalNat]
Refers to (runtime): []
Flags: [covering]
Size change: Prelude.Types.equalNat: [Just (0, Same), Just (1, Same)]
Main> Bye for now!

View File

@ -1,4 +1,6 @@
:di doPlus
:di doMinus
:di doMult
:di doCompare
:di doEqual
:q

View File

@ -6,7 +6,16 @@ Eq Foo where
C == C = True
_ == _ = False
Ord Foo where
A < B = True
B < C = True
_ < _ = False
interface Read a where
readPrefix : String -> Maybe (a, String)
read : String -> Maybe a
read str = case readPrefix str of
Just (a, "") => pure a
Nothing => Nothing
Read Foo where
read "A" = A
read "B" = B
read "C" = C

View File

@ -1,9 +1,10 @@
1/1: Building Methods (Methods.idr)
Error: Missing methods in Ord: compare
Error: Missing methods in Read: readPrefix
Methods:9:1--12:17
09 | Ord Foo where
10 | A < B = True
11 | B < C = True
12 | _ < _ = False
Methods:17:1--21:15
17 | Read Foo where
18 |
19 | read "A" = A
20 | read "B" = B
21 | read "C" = C