Idris2/tests/idris2/reflection010/Name.idr
MarcelineVQ e03096188a add more Name reflections
broaden what Names can be reflected and refied
I did not add the Names I wasn't sure how to test but have put placeholders
that produce clearer error messages.
2020-09-20 00:54:49 -07:00

41 lines
886 B
Idris

import Language.Reflection
%language ElabReflection
%logging 1
-- This test is for checking changes to how Names are reflected and reified.
-- It currently tests Names that refer to nested functions and Names that refer
-- to nested cases.
-- Please add future tests for Names here if they would fit in. There's plenty
-- of room.
data Identity a = MkIdentity a
nested : Identity Int
nested = MkIdentity foo
where
foo : Int
foo = 12
-- a pattern matching lambda is really a case
cased : Identity Int -> Int
cased = \(MkIdentity x) => x
test : Elab ()
test = do
n <- quote nested
logTerm "" 1 "nested" n
MkIdentity n' <- check {expected=Identity Int} n
logMsg "" 1 $ show (n' == 12)
c <- quote cased
logTerm "" 1 "cased" c
c' <- check {expected=Identity Int -> Int} c
logMsg "" 1 $ show (c' (MkIdentity 10))
pure ()
%runElab test