Free standing case blocks return ()

Fixes #116. This is the solution Idris 1 took, and while it is a special
case, the syntax does make it explicit (in a way) that the result of the
case is unused - if you mean something other than (), you must now say
so!
This commit is contained in:
Edwin Brady 2020-06-07 22:47:41 +01:00
parent 6898965a56
commit d47d495744
2 changed files with 8 additions and 1 deletions

View File

@ -392,10 +392,14 @@ mutual
expandDo side ps topfc (DoExp fc tm :: rest)
= do tm' <- desugar side ps tm
rest' <- expandDo side ps topfc rest
-- A free standing 'case' block must return ()
let ty = case tm' of
ICase _ _ _ _ => IVar fc (UN "Unit")
_ => Implicit fc False
gam <- get Ctxt
pure $ IApp fc (IApp fc (IVar fc (UN ">>=")) tm')
(ILam fc top Explicit Nothing
(Implicit fc False) rest')
ty rest')
expandDo side ps topfc (DoBind fc n tm :: rest)
= do tm' <- desugar side ps tm
rest' <- expandDo side ps topfc rest

View File

@ -222,6 +222,9 @@ runTest opts testPath
| Left err => do print err
pure False
let result = normalize out == normalize exp
-- The issue #116 that made this necessary is fixed, but
-- please resist putting 'result' here until it's also
-- fixed in Idris2-boot!
if normalize out == normalize exp
then putStrLn "success"
else do