Idris2/tests/idris2/basic068/Issue2138.idr
Balazs Komuves 6cc20a9974
[ fix #2138] Add an %unhide pragma (#2181)
* add %unhide pragma

* add a test case

* clean up

* more consistent English usage (+fix some typos)

* add a warning for unhiding not-already-hidden names

* move `unhide` (and `hide`) to the bottom of the source file to avoid having to forward-declare `recordWarning`
2021-12-11 18:03:36 -08:00

42 lines
527 B
Idris

namespace Foo
public export
data T
= A Int
| B T
export
Show T where
show (A n) = "[A " ++ show n ++ "]"
show (B t) = "[B " ++ show t ++ "]"
namespace Bar
%hide Foo.T
public export
data T
= A Int
| B T
| C Bool
export
Show T where
show (A n) = "{A " ++ show n ++ "}"
show (B t) = "{B " ++ show t ++ "}"
show (C b) = "{C " ++ show b ++ "}"
%unhide Foo.T
foo : Foo.T
foo = B (A 5)
bar : Bar.T
bar = B (A 6)
main : IO ()
main = do
printLn foo
printLn bar