Skip to content

Commit

Permalink
[ elab ] Make %runElab expressions to have unrestricted quantity
Browse files Browse the repository at this point in the history
  • Loading branch information
buzden committed Nov 8, 2021
1 parent 2f6ec76 commit 3ae7ce0
Show file tree
Hide file tree
Showing 10 changed files with 113 additions and 2 deletions.
6 changes: 6 additions & 0 deletions src/Algebra/Preorder.idr
Original file line number Diff line number Diff line change
Expand Up @@ -29,3 +29,9 @@ public export
interface Preorder a => Top a where
top : a
topAbs : {x : a} -> x <= top = True

||| The least bound of a bounded lattice
public export
interface Preorder a => Bot a where
bot : a
botAbs : {x : a} -> bot <= x = True
8 changes: 8 additions & 0 deletions src/Algebra/ZeroOneOmega.idr
Original file line number Diff line number Diff line change
Expand Up @@ -70,6 +70,14 @@ Top ZeroOneOmega where
topAbs {x = Rig1} = Refl
topAbs {x = RigW} = Refl

||| The bottm value of a lattice
export
Bot ZeroOneOmega where
bot = Rig0
botAbs {x = Rig0} = Refl
botAbs {x = Rig1} = Refl
botAbs {x = RigW} = Refl

----------------------------------------

rigPlusAssociative : (x, y, z : ZeroOneOmega) ->
Expand Down
2 changes: 1 addition & 1 deletion src/TTImp/Elab/RunElab.idr
Original file line number Diff line number Diff line change
Expand Up @@ -236,7 +236,7 @@ checkRunElab rig elabinfo nest env fc script exp
let ttn = reflectiontt "TT"
elabtt <- appCon fc defs n [expected]
(stm, sty) <- runDelays (const True) $
check rig elabinfo nest env script (Just (gnf env elabtt))
check bot elabinfo nest env script (Just (gnf env elabtt))
defs <- get Ctxt -- checking might have resolved some holes
ntm <- elabScript fc nest env
!(nfOpts withAll defs env stm) (Just (gnf env expected))
Expand Down
2 changes: 1 addition & 1 deletion src/TTImp/ProcessRunElab.idr
Original file line number Diff line number Diff line change
Expand Up @@ -36,5 +36,5 @@ processRunElab eopts nest env fc tm
unit <- getCon fc defs (builtin "Unit")
exp <- appCon fc defs n [unit]

stm <- checkTerm tidx InExpr eopts nest env tm (gnf env exp)
stm <- checkTerm tidx InType eopts nest env tm (gnf env exp)
ignore $ elabScript fc nest env !(nfOpts withAll defs env stm) Nothing
1 change: 1 addition & 0 deletions tests/Main.idr
Original file line number Diff line number Diff line change
Expand Up @@ -214,6 +214,7 @@ idrisTests = MkTestPool "Misc" [] Nothing
"reflection001", "reflection002", "reflection003", "reflection004",
"reflection005", "reflection006", "reflection007", "reflection008",
"reflection009", "reflection010", "reflection011", "reflection012",
"reflection013",
-- The 'with' rule
"with001", "with002", "with004", "with005", "with006",
-- with-disambiguation
Expand Down
14 changes: 14 additions & 0 deletions tests/idris2/reflection013/NoEscape.idr
Original file line number Diff line number Diff line change
@@ -0,0 +1,14 @@
module NoEscape

import Language.Reflection

%language ElabReflection

0 n : Nat
n = 3

0 elabScript : Elab Nat
elabScript = pure n

m : Nat
m = %runElab elabScript
31 changes: 31 additions & 0 deletions tests/idris2/reflection013/NoEscapePar.idr
Original file line number Diff line number Diff line change
@@ -0,0 +1,31 @@
||| Check that we cannot implement function illegally escaping zero quantity using elaboration reflection
module NoEscapePar

import Language.Reflection

%language ElabReflection

escScr : Elab $ (0 _ : a) -> a
escScr = check $ ILam EmptyFC M0 ExplicitArg (Just `{x}) `(a) `(x)

esc : (0 _ : a) -> a
esc = %runElab escScr

escd : (0 _ : a) -> a

0 escd' : (0 _ : a) -> a

escDecl : Name -> Elab Unit
escDecl name = declare [
IDef EmptyFC name [
PatClause EmptyFC
-- lhs
(IApp EmptyFC (IVar EmptyFC name) (IBindVar EmptyFC "x"))
-- rhs
`(x)
]
]

%runElab escDecl `{escd'}

%runElab escDecl `{escd}
13 changes: 13 additions & 0 deletions tests/idris2/reflection013/RunElab0.idr
Original file line number Diff line number Diff line change
@@ -0,0 +1,13 @@
module RunElab0

import Language.Reflection

%language ElabReflection

0 elabScript : Elab Unit
elabScript = pure ()

x : Unit
x = %runElab elabScript

%runElab elabScript
33 changes: 33 additions & 0 deletions tests/idris2/reflection013/expected
Original file line number Diff line number Diff line change
@@ -0,0 +1,33 @@
1/1: Building RunElab0 (RunElab0.idr)
1/1: Building NoEscape (NoEscape.idr)
Error: While processing right hand side of m. NoEscape.n is not accessible in this context.

NoEscape:11:19--11:20
07 | 0 n : Nat
08 | n = 3
09 |
10 | 0 elabScript : Elab Nat
11 | elabScript = pure n
^

1/1: Building NoEscapePar (NoEscapePar.idr)
Error: While processing right hand side of esc. x is not accessible in this context.

NoEscape:9:65--9:66
5 |
6 | %language ElabReflection
7 |
8 | escScr : Elab $ (0 _ : a) -> a
9 | escScr = check $ ILam EmptyFC M0 ExplicitArg (Just `{x}) `(a) `(x)
^

Error: While processing right hand side of escd. x is not accessible in this context.

NoEscape:25:24--25:25
21 | PatClause EmptyFC
22 | -- lhs
23 | (IApp EmptyFC (IVar EmptyFC name) (IBindVar EmptyFC "x"))
24 | -- rhs
25 | `(x)
^

5 changes: 5 additions & 0 deletions tests/idris2/reflection013/run
Original file line number Diff line number Diff line change
@@ -0,0 +1,5 @@
rm -rf build

$1 --no-color --console-width 0 --check RunElab0.idr
$1 --no-color --console-width 0 --check NoEscape.idr
$1 --no-color --console-width 0 --check NoEscapePar.idr

0 comments on commit 3ae7ce0

Please sign in to comment.