Skip to content

Commit 2d3a97f

Browse files
committed
[ elab ] Make %runElab expressions to have unrestricted quantity
1 parent 2f6ec76 commit 2d3a97f

File tree

10 files changed

+113
-2
lines changed

10 files changed

+113
-2
lines changed

src/Algebra/Preorder.idr

+6
Original file line numberDiff line numberDiff line change
@@ -29,3 +29,9 @@ public export
2929
interface Preorder a => Top a where
3030
top : a
3131
topAbs : {x : a} -> x <= top = True
32+
33+
||| The least bound of a bounded lattice
34+
public export
35+
interface Preorder a => Bot a where
36+
bot : a
37+
botAbs : {x : a} -> bot <= x = True

src/Algebra/ZeroOneOmega.idr

+8
Original file line numberDiff line numberDiff line change
@@ -70,6 +70,14 @@ Top ZeroOneOmega where
7070
topAbs {x = Rig1} = Refl
7171
topAbs {x = RigW} = Refl
7272

73+
||| The bottm value of a lattice
74+
export
75+
Bot ZeroOneOmega where
76+
bot = Rig0
77+
botAbs {x = Rig0} = Refl
78+
botAbs {x = Rig1} = Refl
79+
botAbs {x = RigW} = Refl
80+
7381
----------------------------------------
7482

7583
rigPlusAssociative : (x, y, z : ZeroOneOmega) ->

src/TTImp/Elab/RunElab.idr

+1-1
Original file line numberDiff line numberDiff line change
@@ -236,7 +236,7 @@ checkRunElab rig elabinfo nest env fc script exp
236236
let ttn = reflectiontt "TT"
237237
elabtt <- appCon fc defs n [expected]
238238
(stm, sty) <- runDelays (const True) $
239-
check rig elabinfo nest env script (Just (gnf env elabtt))
239+
check bot elabinfo nest env script (Just (gnf env elabtt))
240240
defs <- get Ctxt -- checking might have resolved some holes
241241
ntm <- elabScript fc nest env
242242
!(nfOpts withAll defs env stm) (Just (gnf env expected))

src/TTImp/ProcessRunElab.idr

+1-1
Original file line numberDiff line numberDiff line change
@@ -36,5 +36,5 @@ processRunElab eopts nest env fc tm
3636
unit <- getCon fc defs (builtin "Unit")
3737
exp <- appCon fc defs n [unit]
3838

39-
stm <- checkTerm tidx InExpr eopts nest env tm (gnf env exp)
39+
stm <- checkTerm tidx InType eopts nest env tm (gnf env exp)
4040
ignore $ elabScript fc nest env !(nfOpts withAll defs env stm) Nothing

tests/Main.idr

+1
Original file line numberDiff line numberDiff line change
@@ -214,6 +214,7 @@ idrisTests = MkTestPool "Misc" [] Nothing
214214
"reflection001", "reflection002", "reflection003", "reflection004",
215215
"reflection005", "reflection006", "reflection007", "reflection008",
216216
"reflection009", "reflection010", "reflection011", "reflection012",
217+
"reflection013",
217218
-- The 'with' rule
218219
"with001", "with002", "with004", "with005", "with006",
219220
-- with-disambiguation
+14
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,14 @@
1+
module NoEscape
2+
3+
import Language.Reflection
4+
5+
%language ElabReflection
6+
7+
0 n : Nat
8+
n = 3
9+
10+
0 elabScript : Elab Nat
11+
elabScript = pure n
12+
13+
m : Nat
14+
m = %runElab elabScript
+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 NoEscapePar
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}
+13
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,13 @@
1+
module RunElab0
2+
3+
import Language.Reflection
4+
5+
%language ElabReflection
6+
7+
0 elabScript : Elab Unit
8+
elabScript = pure ()
9+
10+
x : Unit
11+
x = %runElab elabScript
12+
13+
%runElab elabScript

tests/idris2/reflection013/expected

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

tests/idris2/reflection013/run

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

0 commit comments

Comments
 (0)