Skip to content

Commit 62617a5

Browse files
committed
[ fix #119 ] Throw an error when an unknown name appears in generated Haskell code
1 parent 6a9885a commit 62617a5

File tree

10 files changed

+90
-20
lines changed

10 files changed

+90
-20
lines changed

src/Agda2Hs/Compile/ClassInstance.hs

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -333,7 +333,7 @@ resolveStringName s = do
333333
lookupDefaultImplementations :: QName -> [Hs.Name ()] -> C [Definition]
334334
lookupDefaultImplementations recName fields = do
335335
let modName = qnameToMName recName
336-
isField f _ = (`elem` fields) . unQual <$> compileQName f
336+
isField f _ = (`elem` fields) <$> compileName (qnameName f)
337337
findDefinitions isField modName
338338

339339
classMemberNames :: Definition -> C [Hs.Name ()]

src/Agda2Hs/Compile/Name.hs

Lines changed: 26 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -31,8 +31,10 @@ import qualified Agda.Syntax.Common.Pretty as P
3131
import Agda.TypeChecking.Datatypes ( isDataOrRecordType )
3232
import Agda.TypeChecking.Pretty
3333
import Agda.TypeChecking.Records ( isRecordConstructor )
34+
import Agda.TypeChecking.Warnings ( warning )
3435

3536
import qualified Agda.Utils.List1 as List1
37+
import Agda.Utils.Monad
3638
import Agda.Utils.Maybe ( isJust, isNothing, whenJust, fromMaybe, caseMaybeM )
3739
import Agda.Utils.Monad ( whenM )
3840

@@ -102,11 +104,28 @@ compileQName f
102104
Just (r, def) | not (_recNamedCon def) -> r -- use record name for unnamed constructors
103105
_ -> f
104106
hf0 <- compileName (qnameName f)
105-
(hf, mimpBuiltin) <- fromMaybe (hf0, Nothing) <$> isSpecialName f
107+
special <- isSpecialName f
108+
let (hf, mimpBuiltin) = fromMaybe (hf0, Nothing) special
109+
106110
parent <- parentName f
107111
par <- traverse (compileName . qnameName) parent
108112
let mod0 = qnameModule $ fromMaybe f parent
109113
mod <- compileModuleName mod0
114+
115+
existsInHaskell <- orM
116+
[ pure $ isJust special
117+
, pure $ isPrimModule mod
118+
, hasCompilePragma f
119+
, isClassFunction f
120+
, isWhereFunction f
121+
, maybe (pure False) hasCompilePragma parent
122+
]
123+
124+
unless existsInHaskell $ do
125+
reportSDoc "agda2hs.name" 20 $ text "DOES NOT EXIST IN HASKELL"
126+
typeError $ CustomBackendError "agda2hs" $ P.text $
127+
"Symbol " ++ Hs.prettyPrint hf ++ " is missing a COMPILE pragma or rewrite rule"
128+
110129
currMod <- hsTopLevelModuleName <$> asks currModule
111130
let skipModule = mod == currMod
112131
|| isJust mimpBuiltin
@@ -185,14 +204,19 @@ compileQName f
185204

186205
mkImport mod qual par hf maybeIsType
187206
-- make sure the Prelude is properly qualified
188-
| any (`isPrefixOf` pp mod) primModules
207+
| isPrimModule mod
189208
= if isQualified qual then
190209
let mod' = hsModuleName "Prelude"
191210
in (mod', Just (Import mod' qual Nothing hf maybeIsType))
192211
else (mod, Nothing)
193212
| otherwise
194213
= (mod, Just (Import mod qual par hf maybeIsType))
195214

215+
isWhereFunction :: QName -> C Bool
216+
isWhereFunction f = do
217+
whereMods <- asks whereModules
218+
return $ any (qnameModule f `isLeChildModuleOf`) whereMods
219+
196220
hsTopLevelModuleName :: TopLevelModuleName -> Hs.ModuleName ()
197221
hsTopLevelModuleName = hsModuleName . intercalate "." . map unpack
198222
. List1.toList . moduleNameParts

src/Agda2Hs/Compile/Record.hs

Lines changed: 9 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -16,7 +16,7 @@ import Agda.Syntax.Common ( Arg(unArg), defaultArg )
1616
import Agda.Syntax.Internal
1717
import Agda.Syntax.Common.Pretty ( prettyShow )
1818

19-
import Agda.TypeChecking.Pretty ( ($$), (<+>), text, vcat, prettyTCM )
19+
import Agda.TypeChecking.Pretty ( ($$), (<+>), text, vcat, prettyTCM, pretty, prettyList_ )
2020
import Agda.TypeChecking.Substitute ( TelV(TelV), Apply(apply) )
2121
import Agda.TypeChecking.Telescope
2222

@@ -38,7 +38,10 @@ withMinRecord :: QName -> C a -> C a
3838
withMinRecord m = local $ \ e -> e { minRecordName = Just (qnameToMName m) }
3939

4040
compileMinRecord :: [Hs.Name ()] -> QName -> C MinRecord
41-
compileMinRecord fieldNames m = do
41+
compileMinRecord fieldNames m = withMinRecord m $ do
42+
reportSDoc "agda2hs.record.min" 20 $
43+
text "Compiling minimal record" <+> pretty m <+>
44+
text "with field names" <+> prettyList_ (map (text . pp) fieldNames)
4245
rdef <- getConstInfo m
4346
definedFields <- classMemberNames rdef
4447
let Record{recPars = npars, recTel = tel} = theDef rdef
@@ -48,9 +51,12 @@ compileMinRecord fieldNames m = do
4851
-- We can't simply compileFun here for two reasons:
4952
-- * it has an explicit dictionary argument
5053
-- * it's using the fields and definitions from the minimal record and not the parent record
51-
compiled <- withMinRecord m $ addContext (defaultDom rtype) $ compileLocal $
54+
compiled <- addContext (defaultDom rtype) $ compileLocal $
5255
fmap concat $ traverse (compileFun False) defaults
5356
let declMap = Map.fromList [ (definedName c, def) | def@(Hs.FunBind _ (c : _)) <- compiled ]
57+
reportSDoc "agda2hs.record.min" 20 $
58+
text "Done compiling minimal record" <+> pretty m <+>
59+
text "defined fields: " <+> prettyList_ (map (text . pp) definedFields)
5460
return (definedFields, declMap)
5561

5662

src/Agda2Hs/Compile/Types.hs

Lines changed: 2 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -22,6 +22,7 @@ import qualified Language.Haskell.Exts.Comments as Hs
2222
import Agda.Compiler.Backend
2323
import Agda.Syntax.Position ( Range )
2424
import Agda.Syntax.TopLevelModuleName ( TopLevelModuleName )
25+
import Agda.TypeChecking.Warnings ( MonadWarning )
2526
import Agda.Utils.Null
2627
import Agda.Utils.Impossible
2728

@@ -174,6 +175,7 @@ instance MonadStConcreteNames C where
174175
runStConcreteNames m = rwsT $ \r s -> runStConcreteNames $ StateT $ \ns -> do
175176
((x, ns'), s', w) <- runRWST (runStateT m ns) r s
176177
pure ((x, s', w), ns')
178+
instance MonadWarning C where
177179
instance IsString a => IsString (C a) where fromString = pure . fromString
178180
instance PureTCM C where
179181
instance Null a => Null (C a) where

src/Agda2Hs/Compile/Utils.hs

Lines changed: 15 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -9,6 +9,7 @@ import Control.Monad.State ( put, modify )
99

1010
import Data.Maybe ( isJust )
1111
import qualified Data.Map as M
12+
import Data.List ( isPrefixOf )
1213

1314
import qualified Language.Haskell.Exts as Hs
1415

@@ -60,6 +61,8 @@ primModules =
6061
, "Haskell.Law"
6162
]
6263

64+
isPrimModule :: Hs.ModuleName () -> Bool
65+
isPrimModule mod = any (`isPrefixOf` pp mod) primModules
6366

6467
concatUnzip :: [([a], [b])] -> ([a], [b])
6568
concatUnzip = (concat *** concat) . unzip
@@ -152,6 +155,18 @@ isErasedBaseType x = orM
152155
]
153156
where b = El __DUMMY_SORT__ x
154157

158+
hasCompilePragma :: QName -> C Bool
159+
hasCompilePragma q = processPragma q <&> \case
160+
NoPragma{} -> False
161+
InlinePragma{} -> True
162+
DefaultPragma{} -> True
163+
ClassPragma{} -> True
164+
ExistingClassPragma{} -> True
165+
UnboxPragma{} -> True
166+
TransparentPragma{} -> True
167+
NewTypePragma{} -> True
168+
DerivePragma{} -> True
169+
155170
-- Exploits the fact that the name of the record type and the name of the record module are the
156171
-- same, including the unique name ids.
157172
isClassFunction :: QName -> C Bool

test/Delay.agda

Lines changed: 11 additions & 6 deletions
Original file line numberDiff line numberDiff line change
@@ -8,16 +8,21 @@ open import Haskell.Extra.Delay
88
open import Agda.Builtin.Size
99

1010
postulate
11-
div : Int Int Int
12-
mod : Int Int Int
11+
div' : Int Int Int
12+
mod' : Int Int Int
1313

14-
even : Int Bool
15-
even x = mod x 2 == 0
14+
{-# COMPILE AGDA2HS div' #-}
15+
{-# COMPILE AGDA2HS mod' #-}
16+
17+
even' : Int Bool
18+
even' x = mod' x 2 == 0
19+
20+
{-# COMPILE AGDA2HS even' #-}
1621

1722
collatz : {@0 j} Int Delay Int j
18-
collatz x =
23+
collatz x =
1924
if x == 0 then now 0
20-
else if even x then later (λ where .force collatz (div x 2))
25+
else if even' x then later (λ where .force collatz (div' x 2))
2126
else later λ where .force collatz (3 * x + 1)
2227

2328
{-# COMPILE AGDA2HS collatz #-}

test/Fail/Issue119.agda

Lines changed: 13 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,13 @@
1+
module Fail.Issue119 where
2+
3+
open import Haskell.Prelude
4+
5+
aaa : Int
6+
aaa = 21
7+
8+
-- Oops, forgot compile pragma for aaa
9+
10+
bbb : Int
11+
bbb = aaa + aaa
12+
13+
{-# COMPILE AGDA2HS bbb #-}

test/TypeOperatorImport.agda

Lines changed: 1 addition & 7 deletions
Original file line numberDiff line numberDiff line change
@@ -2,15 +2,9 @@ module TypeOperatorImport where
22

33
{-# FOREIGN AGDA2HS {-# LANGUAGE TypeOperators #-} #-}
44

5-
open import Agda.Builtin.Unit
6-
open import Agda.Builtin.Bool
7-
open import Haskell.Prelude using (_∘_)
5+
open import Haskell.Prelude hiding (_<_)
86
open import TypeOperatorExport
97

10-
not : Bool Bool
11-
not true = false
12-
not false = true
13-
148
test1 : ⊤ < Bool
159
test1 = tt
1610
{-# COMPILE AGDA2HS test1 #-}

test/golden/Delay.hs

Lines changed: 10 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -1,7 +1,16 @@
11
module Delay where
22

3+
div' :: Int -> Int -> Int
4+
div' = error "postulate: Int -> Int -> Int"
5+
6+
mod' :: Int -> Int -> Int
7+
mod' = error "postulate: Int -> Int -> Int"
8+
9+
even' :: Int -> Bool
10+
even' x = mod' x 2 == 0
11+
312
collatz :: Int -> Int
413
collatz x
514
= if x == 0 then 0 else
6-
if even x then collatz (div x 2) else collatz (3 * x + 1)
15+
if even' x then collatz (div' x 2) else collatz (3 * x + 1)
716

test/golden/Issue119.err

Lines changed: 2 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,2 @@
1+
test/Fail/Issue119.agda:10,1-4
2+
agda2hs: Symbol aaa is missing a COMPILE pragma or rewrite rule

0 commit comments

Comments
 (0)