Skip to content

Commit 76f3b3b

Browse files
committed
[ elab ] Add test of no escaping the zero quantity
1 parent 3854eec commit 76f3b3b

File tree

4 files changed

+56
-0
lines changed

4 files changed

+56
-0
lines changed

tests/Main.idr

+1
Original file line numberDiff line numberDiff line change
@@ -212,6 +212,7 @@ idrisTests = MkTestPool "Misc" [] Nothing
212212
"reflection001", "reflection002", "reflection003", "reflection004",
213213
"reflection005", "reflection006", "reflection007", "reflection008",
214214
"reflection009", "reflection010", "reflection011", "reflection012",
215+
"reflection013",
215216
-- The 'with' rule
216217
"with001", "with002", "with004", "with005", "with006",
217218
-- with-disambiguation
+31
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,31 @@
1+
||| Check that we cannot implement function illegally escaping zero quantity using elaboration reflection
2+
module NoEscape
3+
4+
import Language.Reflection
5+
6+
%language ElabReflection
7+
8+
escScr : Elab $ (0 _ : a) -> a
9+
escScr = check $ ILam EmptyFC M0 ExplicitArg (Just `{x}) `(a) `(x)
10+
11+
esc : (0 _ : a) -> a
12+
esc = %runElab escScr
13+
14+
escd : (0 _ : a) -> a
15+
16+
0 escd' : (0 _ : a) -> a
17+
18+
escDecl : Name -> Elab Unit
19+
escDecl name = declare [
20+
IDef EmptyFC name [
21+
PatClause EmptyFC
22+
-- lhs
23+
(IApp EmptyFC (IVar EmptyFC name) (IBindVar EmptyFC "x"))
24+
-- rhs
25+
`(x)
26+
]
27+
]
28+
29+
%runElab escDecl `{escd'}
30+
31+
%runElab escDecl `{escd}

tests/idris2/reflection013/expected

+21
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,21 @@
1+
1/1: Building NoEscape (NoEscape.idr)
2+
Error: While processing right hand side of esc. x is not accessible in this context.
3+
4+
NoEscape:9:65--9:66
5+
5 |
6+
6 | %language ElabReflection
7+
7 |
8+
8 | escScr : Elab $ (0 _ : a) -> a
9+
9 | escScr = check $ ILam EmptyFC M0 ExplicitArg (Just `{x}) `(a) `(x)
10+
^
11+
12+
Error: While processing right hand side of escd. x is not accessible in this context.
13+
14+
NoEscape:25:24--25:25
15+
21 | PatClause EmptyFC
16+
22 | -- lhs
17+
23 | (IApp EmptyFC (IVar EmptyFC name) (IBindVar EmptyFC "x"))
18+
24 | -- rhs
19+
25 | `(x)
20+
^
21+

tests/idris2/reflection013/run

+3
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,3 @@
1+
rm -rf build
2+
3+
$1 --no-color --console-width 0 --check NoEscape.idr

0 commit comments

Comments
 (0)