mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-12-20 10:02:03 +03:00
19 lines
346 B
Idris
19 lines
346 B
Idris
module Debug.Trace
|
|
|
|
import Prelude
|
|
import PrimIO
|
|
|
|
%default total
|
|
|
|
export
|
|
trace : (msg : String) -> (result : a) -> a
|
|
trace x val = unsafePerformIO (do putStrLn x; pure val)
|
|
|
|
export %inline
|
|
traceValBy : (msgF : a -> String) -> (result : a) -> a
|
|
traceValBy f v = trace (f v) v
|
|
|
|
export %inline
|
|
traceVal : Show a => a -> a
|
|
traceVal = traceValBy show
|