mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-11-29 03:24:02 +03:00
a972778eab
They don't all pass yet, for minor reasons. Coming shortly... Unfortunately the startup overhead for chez is really noticeable here!
16 lines
377 B
Idris
16 lines
377 B
Idris
data MyMaybe : (0 x : Type) -> Type where
|
|
MyNothing : MyMaybe a
|
|
MyJust : a -> MyMaybe a
|
|
|
|
-- Should fail since type argument is deleted
|
|
nameOf : Type -> String
|
|
nameOf (MyMaybe Bool) = "MyMaybe Bool"
|
|
nameOf (MyMaybe Int) = "MyMaybe Int"
|
|
nameOf _ = "Unknown"
|
|
|
|
main : IO ()
|
|
main = do
|
|
putStrLn (nameOf (MyMaybe Bool))
|
|
putStrLn (nameOf (MyMaybe Int))
|
|
putStrLn (nameOf Int)
|