File tree 10 files changed +93
-3
lines changed
tests/idris2/reflection/reflection027
10 files changed +93
-3
lines changed Original file line number Diff line number Diff line change 6
6
7
7
### Compiler changes
8
8
9
+ * Elaborator script's expression under the ` %runElab ` is now typechecked in the context
10
+ of quantity ` 0 ` , because it is present and works only at compile time
11
+
9
12
### Library changes
10
13
11
14
#### Prelude
Original file line number Diff line number Diff line change @@ -29,3 +29,9 @@ public export
29
29
interface Preorder a => Top a where
30
30
top : a
31
31
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
Original file line number Diff line number Diff line change @@ -72,6 +72,14 @@ Top ZeroOneOmega where
72
72
topAbs {x = Rig1 } = Refl
73
73
topAbs {x = RigW } = Refl
74
74
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
+
75
83
-- --------------------------------------
76
84
77
85
rigPlusAssociative : (x, y, z : ZeroOneOmega) ->
Original file line number Diff line number Diff line change @@ -353,8 +353,8 @@ checkRunElab rig elabinfo nest env fc reqExt script exp
353
353
throw (GenericMsg fc " %language ElabReflection not enabled" )
354
354
let n = NS reflectionNS (UN $ Basic " Elab" )
355
355
elabtt <- appCon fc defs n [expected]
356
- (stm, sty ) <- runDelays (const True ) $
357
- check rig elabinfo nest env script (Just (gnf env elabtt))
356
+ (stm, _ ) <- runDelays (const True ) $
357
+ check bot elabinfo nest env script (Just (gnf env elabtt))
358
358
solveConstraints inTerm Normal
359
359
defs <- get Ctxt -- checking might have resolved some holes
360
360
ntm <- elabScript rig fc nest env
Original file line number Diff line number Diff line change @@ -38,5 +38,6 @@ processRunElab eopts nest env fc tm
38
38
unit <- getCon fc defs (builtin " Unit" )
39
39
exp <- appCon fc defs n [unit]
40
40
41
- stm <- checkTerm tidx InExpr eopts nest env tm (gnf env exp )
41
+ e <- newRef EST $ initEStateSub tidx env SubRefl
42
+ (stm, _ ) <- check bot (initElabInfo InExpr ) nest env tm $ Just $ gnf env exp
42
43
ignore $ elabScript top fc nest env ! (nfOpts withAll defs env stm) Nothing
Original file line number Diff line number Diff line change
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
Original file line number Diff line number Diff line change
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}
Original file line number Diff line number Diff line change
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
Original file line number Diff line number Diff line change
1
+ 1/1: Building RunElab0 (RunElab0.idr)
2
+ 1/1: Building NoEscape (NoEscape.idr)
3
+ 1/1: Building NoEscapePar (NoEscapePar.idr)
Original file line number Diff line number Diff line change
1
+ . ../../../testutils.sh
2
+
3
+ check RunElab0.idr
4
+ check NoEscape.idr
5
+ check NoEscapePar.idr
You can’t perform that action at this time.
0 commit comments