Idris2/tests/idris2/basic032/Idiom.idr
MarcelineVQ 961a28ee64 fix idiom brackets to account for IAlternative
Things like (,) () aren't straightforward IVar's but are IAlternative's
which present options about how the term should resolve. [| |] was not
accounting for this.
2020-09-24 15:27:40 +01:00

15 lines
313 B
Idris

export
record Core t where
constructor MkCore
runCore : IO (Either String t)
pure : a -> Core a
pure x = MkCore (pure (pure x))
export
(<*>) : Core (a -> b) -> Core a -> Core b
(<*>) (MkCore f) (MkCore a) = MkCore [| f `Prelude.(<*>)` a |]
addm : Maybe Int -> Maybe Int -> Maybe Int
addm x y = [| x + y |]