mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-12-25 04:33:45 +03:00
[ fix ] issue 2537 (#2538)
Co-authored-by: Zoe Stafford <zoepolarsax@gmail.com>
This commit is contained in:
parent
7db20d38a3
commit
5a962929e7
@ -446,22 +446,23 @@ mutual
|
||||
-- I'm (edwinb) keeping it visible here because I plan to put it back in
|
||||
-- more or less this form once case inlining works better and the whole thing
|
||||
-- works in a nice principled way.
|
||||
-- if noworld -- just substitute the scrutinee into
|
||||
-- -- the RHS
|
||||
-- then
|
||||
if noworld -- just substitute the scrutinee into
|
||||
-- the RHS
|
||||
then
|
||||
let env : SubstCEnv args vars
|
||||
= mkSubst 0 scr pos args in
|
||||
pure $ Just (substs env !(toCExpTree n sc))
|
||||
-- else -- let bind the scrutinee, and substitute the
|
||||
-- -- name into the RHS
|
||||
-- let env : SubstCEnv args (MN "eff" 0 :: vars)
|
||||
-- = mkSubst 0 (CLocal fc First) pos args in
|
||||
-- do sc' <- toCExpTree n sc
|
||||
-- let scope = insertNames {outer=args}
|
||||
-- {inner=vars}
|
||||
-- [MN "eff" 0] sc'
|
||||
-- pure $ Just (CLet fc (MN "eff" 0) False scr
|
||||
-- (substs env scope))
|
||||
else -- let bind the scrutinee, and substitute the
|
||||
-- name into the RHS
|
||||
let env : SubstCEnv args (MN "eff" 0 :: vars)
|
||||
= mkSubst 0 (CLocal fc First) pos args in
|
||||
do sc' <- toCExpTree n sc
|
||||
let scope = insertNames {outer=args}
|
||||
{inner=vars}
|
||||
{ns = [MN "eff" 0]}
|
||||
(mkSizeOf _) (mkSizeOf _) sc'
|
||||
pure $ Just (CLet fc (MN "eff" 0) False scr
|
||||
(substs env scope))
|
||||
_ => pure Nothing -- there's a normal match to do
|
||||
where
|
||||
mkSubst : Nat -> CExp vs ->
|
||||
|
@ -123,13 +123,13 @@ NS_UN : Namespace -> String -> Name
|
||||
NS_UN ns un = NS ns (UN $ Basic un)
|
||||
|
||||
argError : Ref State InterpState => Stack -> Vect h Object -> Core a
|
||||
argError stk obj = interpError stk $ "Unexpected arguments: " ++ foldMap ((" " ++) . showType) obj
|
||||
argError stk obj = interpError stk $ "Unexpected arguments: " ++ foldMap ((" " ++) . showDepth 1) obj
|
||||
|
||||
unit : Object
|
||||
unit = Const (I 0)
|
||||
|
||||
ioRes : Object -> Object
|
||||
ioRes obj = Constructor (Left 0) [Const WorldVal, obj]
|
||||
ioRes obj = obj -- ioRes is a newtype -- Constructor (Left 0) [Const WorldVal, obj]
|
||||
|
||||
-- TODO: add more?
|
||||
knownForeign : NameMap (ar ** (Ref State InterpState => Stack -> Vect ar Object -> Core Object))
|
||||
@ -140,17 +140,27 @@ knownForeign = fromList
|
||||
, (NS_UN ioNS "prim__putStr", (2 ** prim_putStr))
|
||||
]
|
||||
where
|
||||
-- %MkWorld should not be matched on
|
||||
-- however a value of type %World should only be %MkWorld or and erased value
|
||||
world : Ref State InterpState => Stack -> Object -> Core ()
|
||||
world stk Null = pure ()
|
||||
world stk (Const WorldVal) = pure ()
|
||||
world stk o = interpError stk $ "expected %MkWorld or Null, got \{show o}"
|
||||
|
||||
prim_putChar : Ref State InterpState => Stack -> Vect 2 Object -> Core Object
|
||||
prim_putChar _ [Const (Ch c), Const WorldVal] = ioRes unit <$ coreLift_ (putChar c)
|
||||
prim_putChar stk [Const (Ch c), w] = world stk w *> (ioRes unit <$ coreLift_ (putChar c))
|
||||
prim_putChar stk as = argError stk as
|
||||
|
||||
prim_getChar : Ref State InterpState => Stack -> Vect 1 Object -> Core Object
|
||||
prim_getChar _ [Const WorldVal] = ioRes . Const . Ch <$> coreLift getChar
|
||||
prim_getChar stk [w] = world stk w *> (ioRes . Const . Ch <$> coreLift getChar)
|
||||
prim_getChar stk as = argError stk as
|
||||
|
||||
prim_getStr : Ref State InterpState => Stack -> Vect 1 Object -> Core Object
|
||||
prim_getStr _ [Const WorldVal] = Const . Str <$> coreLift getLine
|
||||
prim_getStr stk [w] = world stk w *> (ioRes . Const . Str <$> coreLift getLine)
|
||||
prim_getStr stk as = argError stk as
|
||||
|
||||
prim_putStr : Ref State InterpState => Stack -> Vect 2 Object -> Core Object
|
||||
prim_putStr _ [Const (Str s), Const WorldVal] = ioRes unit <$ coreLift_ (putStr s)
|
||||
prim_putStr stk [Const (Str s), w] = world stk w *> (ioRes unit <$ coreLift_ (putStr s))
|
||||
prim_putStr stk as = argError stk as
|
||||
|
||||
knownExtern : NameMap (ar ** (Ref State InterpState => Stack -> Vect ar Object -> Core Object))
|
||||
|
@ -257,6 +257,8 @@ idrisTests = MkTestPool "Misc" [] Nothing
|
||||
"with003",
|
||||
-- pretty printing
|
||||
"pretty001", "pretty002",
|
||||
-- PrimIO
|
||||
"primloop",
|
||||
-- golden file testing
|
||||
"golden001",
|
||||
-- quantifiers
|
||||
|
@ -6,6 +6,6 @@ Main.plus = [{arg:N}, {arg:N}]: (%case !{arg:N} [(%constcase 0 !{arg:N})] Just (
|
||||
Main.main = [{ext:N}]: (Main.plus [1, 2])
|
||||
Builtin.believe_me = [{arg:N}]: (believe_me [___, ___, !{arg:N}])
|
||||
Prelude.Types.prim__integerToNat = [{arg:N}]: (%case (%case (<=Integer [0, !{arg:N}]) [(%constcase 0 0)] Just 1) [(%constcase 1 (Builtin.believe_me [!{arg:N}])), (%constcase 0 0)] Nothing)
|
||||
PrimIO.unsafePerformIO = [{arg:N}]: (PrimIO.unsafeCreateWorld [(%lam w (!{arg:N} [!w]))])
|
||||
PrimIO.unsafePerformIO = [{arg:N}]: (%let {eff:N} !{arg:N} (PrimIO.unsafeCreateWorld [(%lam w (%let {eff:N} (!{eff:N} [!w]) !{eff:N}))]))
|
||||
PrimIO.unsafeCreateWorld = [{arg:N}]: (!{arg:N} [%MkWorld])
|
||||
|
||||
|
12
tests/idris2/primloop/PrimLoop.idr
Normal file
12
tests/idris2/primloop/PrimLoop.idr
Normal file
@ -0,0 +1,12 @@
|
||||
module PrimLoop
|
||||
|
||||
%default total
|
||||
|
||||
loop : Nat -> PrimIO ()
|
||||
loop Z w = MkIORes () w
|
||||
loop (S k) w =
|
||||
let MkIORes () w2 := toPrim (printLn (S k)) w
|
||||
in loop k w2
|
||||
|
||||
main : IO ()
|
||||
main = fromPrim $ loop 10
|
12
tests/idris2/primloop/expected
Normal file
12
tests/idris2/primloop/expected
Normal file
@ -0,0 +1,12 @@
|
||||
1/1: Building PrimLoop (PrimLoop.idr)
|
||||
PrimLoop> 10
|
||||
9
|
||||
8
|
||||
7
|
||||
6
|
||||
5
|
||||
4
|
||||
3
|
||||
2
|
||||
1
|
||||
PrimLoop> Bye for now!
|
2
tests/idris2/primloop/input
Normal file
2
tests/idris2/primloop/input
Normal file
@ -0,0 +1,2 @@
|
||||
:exec main
|
||||
:q
|
3
tests/idris2/primloop/run
Normal file
3
tests/idris2/primloop/run
Normal file
@ -0,0 +1,3 @@
|
||||
rm -rf build
|
||||
|
||||
$1 --no-color --console-width 0 --no-banner PrimLoop.idr < input
|
@ -1,2 +1,2 @@
|
||||
allocated AnyPtr freed
|
||||
allocated (Ptr t) freed
|
||||
allocated AnyPtr freed
|
||||
|
Loading…
Reference in New Issue
Block a user