[ fix ] issue 2537 (#2538)

Co-authored-by: Zoe Stafford <zoepolarsax@gmail.com>
This commit is contained in:
Stefan Höck 2022-07-07 08:54:52 +00:00 committed by GitHub
parent 7db20d38a3
commit 5a962929e7
No known key found for this signature in database
GPG Key ID: 4AEE18F83AFDEB23
9 changed files with 63 additions and 21 deletions

View File

@ -446,22 +446,23 @@ mutual
-- I'm (edwinb) keeping it visible here because I plan to put it back in -- 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 -- more or less this form once case inlining works better and the whole thing
-- works in a nice principled way. -- works in a nice principled way.
-- if noworld -- just substitute the scrutinee into if noworld -- just substitute the scrutinee into
-- -- the RHS -- the RHS
-- then then
let env : SubstCEnv args vars let env : SubstCEnv args vars
= mkSubst 0 scr pos args in = mkSubst 0 scr pos args in
pure $ Just (substs env !(toCExpTree n sc)) pure $ Just (substs env !(toCExpTree n sc))
-- else -- let bind the scrutinee, and substitute the else -- let bind the scrutinee, and substitute the
-- -- name into the RHS -- name into the RHS
-- let env : SubstCEnv args (MN "eff" 0 :: vars) let env : SubstCEnv args (MN "eff" 0 :: vars)
-- = mkSubst 0 (CLocal fc First) pos args in = mkSubst 0 (CLocal fc First) pos args in
-- do sc' <- toCExpTree n sc do sc' <- toCExpTree n sc
-- let scope = insertNames {outer=args} let scope = insertNames {outer=args}
-- {inner=vars} {inner=vars}
-- [MN "eff" 0] sc' {ns = [MN "eff" 0]}
-- pure $ Just (CLet fc (MN "eff" 0) False scr (mkSizeOf _) (mkSizeOf _) sc'
-- (substs env scope)) pure $ Just (CLet fc (MN "eff" 0) False scr
(substs env scope))
_ => pure Nothing -- there's a normal match to do _ => pure Nothing -- there's a normal match to do
where where
mkSubst : Nat -> CExp vs -> mkSubst : Nat -> CExp vs ->

View File

@ -123,13 +123,13 @@ NS_UN : Namespace -> String -> Name
NS_UN ns un = NS ns (UN $ Basic un) NS_UN ns un = NS ns (UN $ Basic un)
argError : Ref State InterpState => Stack -> Vect h Object -> Core a 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 : Object
unit = Const (I 0) unit = Const (I 0)
ioRes : Object -> Object 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? -- TODO: add more?
knownForeign : NameMap (ar ** (Ref State InterpState => Stack -> Vect ar Object -> Core Object)) 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)) , (NS_UN ioNS "prim__putStr", (2 ** prim_putStr))
] ]
where 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 : 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_putChar stk as = argError stk as
prim_getChar : Ref State InterpState => Stack -> Vect 1 Object -> Core Object 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_getChar stk as = argError stk as
prim_getStr : Ref State InterpState => Stack -> Vect 1 Object -> Core Object 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_getStr stk as = argError stk as
prim_putStr : Ref State InterpState => Stack -> Vect 2 Object -> Core Object 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 prim_putStr stk as = argError stk as
knownExtern : NameMap (ar ** (Ref State InterpState => Stack -> Vect ar Object -> Core Object)) knownExtern : NameMap (ar ** (Ref State InterpState => Stack -> Vect ar Object -> Core Object))

View File

@ -257,6 +257,8 @@ idrisTests = MkTestPool "Misc" [] Nothing
"with003", "with003",
-- pretty printing -- pretty printing
"pretty001", "pretty002", "pretty001", "pretty002",
-- PrimIO
"primloop",
-- golden file testing -- golden file testing
"golden001", "golden001",
-- quantifiers -- quantifiers

View File

@ -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]) Main.main = [{ext:N}]: (Main.plus [1, 2])
Builtin.believe_me = [{arg:N}]: (believe_me [___, ___, !{arg:N}]) 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) 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]) PrimIO.unsafeCreateWorld = [{arg:N}]: (!{arg:N} [%MkWorld])

View 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

View 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!

View File

@ -0,0 +1,2 @@
:exec main
:q

View File

@ -0,0 +1,3 @@
rm -rf build
$1 --no-color --console-width 0 --no-banner PrimLoop.idr < input

View File

@ -1,2 +1,2 @@
allocated AnyPtr freed
allocated (Ptr t) freed allocated (Ptr t) freed
allocated AnyPtr freed