Skip to content

Commit deee51f

Browse files
committed
[ elab ] Make %runElab expressions to have unrestricted quantity
1 parent d78d0bf commit deee51f

File tree

10 files changed

+89
-3
lines changed

10 files changed

+89
-3
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
@@ -72,6 +72,14 @@ Top ZeroOneOmega where
7272
topAbs {x = Rig1} = Refl
7373
topAbs {x = RigW} = Refl
7474

75+
||| The bottom value of a lattice
76+
export
77+
Bot ZeroOneOmega where
78+
bot = Rig0
79+
botAbs {x = Rig0} = Refl
80+
botAbs {x = Rig1} = Refl
81+
botAbs {x = RigW} = Refl
82+
7583
----------------------------------------
7684

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

src/TTImp/Elab/RunElab.idr

+1-1
Original file line numberDiff line numberDiff line change
@@ -249,7 +249,7 @@ checkRunElab rig elabinfo nest env fc script exp
249249
let n = NS reflectionNS (UN $ Basic "Elab")
250250
elabtt <- appCon fc defs n [expected]
251251
(stm, sty) <- runDelays (const True) $
252-
check rig elabinfo nest env script (Just (gnf env elabtt))
252+
check bot elabinfo nest env script (Just (gnf env elabtt))
253253
solveConstraints inTerm Normal
254254
defs <- get Ctxt -- checking might have resolved some holes
255255
ntm <- elabScript rig fc nest env

src/TTImp/ProcessRunElab.idr

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

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

tests/Main.idr

+1-1
Original file line numberDiff line numberDiff line change
@@ -241,7 +241,7 @@ idrisTestsReflection = MkTestPool "Quotation and Reflection" [] Nothing
241241
"reflection005", "reflection006", "reflection007", "reflection008",
242242
"reflection009", "reflection010", "reflection011", "reflection012",
243243
"reflection013", "reflection014", "reflection015", "reflection016",
244-
"reflection017"]
244+
"reflection017", "reflection018"]
245245

246246
idrisTestsWith : TestPool
247247
idrisTestsWith = MkTestPool "With abstraction" [] Nothing
+16
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,16 @@
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+
failing "n is not accessible in this context"
14+
15+
m : Nat
16+
m = %runElab elabScript
+35
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,35 @@
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+
failing "x is not accessible in this context"
12+
13+
esc : (0 _ : a) -> a
14+
esc = %runElab escScr
15+
16+
escd : (0 _ : a) -> a
17+
18+
0 escd' : (0 _ : a) -> a
19+
20+
escDecl : Name -> Elab Unit
21+
escDecl name = declare [
22+
IDef EmptyFC name [
23+
PatClause EmptyFC
24+
-- lhs
25+
(IApp EmptyFC (IVar EmptyFC name) (IBindVar EmptyFC "x"))
26+
-- rhs
27+
`(x)
28+
]
29+
]
30+
31+
%runElab escDecl `{escd'}
32+
33+
failing "x is not accessible in this context"
34+
35+
%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/reflection018/expected

+3
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,3 @@
1+
1/1: Building RunElab0 (RunElab0.idr)
2+
1/1: Building NoEscape (NoEscape.idr)
3+
1/1: Building NoEscapePar (NoEscapePar.idr)

tests/idris2/reflection018/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)