Skip to content

Commit

Permalink
[ fix #324 ] Keep track of instance-only imports
Browse files Browse the repository at this point in the history
  • Loading branch information
jespercockx committed Dec 17, 2024
1 parent beb8afb commit 84c0c36
Show file tree
Hide file tree
Showing 15 changed files with 120 additions and 59 deletions.
1 change: 1 addition & 0 deletions src/Agda2Hs/Compile/ClassInstance.hs
Original file line number Diff line number Diff line change
Expand Up @@ -36,6 +36,7 @@ import Agda.Utils.Impossible ( __IMPOSSIBLE__ )
import Agda2Hs.AgdaUtils
import Agda2Hs.Compile.Function
import Agda2Hs.Compile.Name
import Agda2Hs.Compile.Term
import Agda2Hs.Compile.Type
import Agda2Hs.Compile.Types
import Agda2Hs.Compile.Utils
Expand Down
34 changes: 20 additions & 14 deletions src/Agda2Hs/Compile/Imports.hs
Original file line number Diff line number Diff line change
Expand Up @@ -25,10 +25,12 @@ type ImportDeclMap = Map (Hs.ModuleName (), Qualifier) ImportSpecMap

compileImports :: String -> Imports -> TCM [Hs.ImportDecl ()]
compileImports top is0 = do
let is = filter (not . (top ==) . Hs.prettyPrint . importModule) is0
let is = filter ((top /=) . Hs.prettyPrint . importModule) is0
checkClashingImports is
let imps = Map.toList $ groupModules is
return $ map (uncurry $ uncurry makeImportDecl) imps
reportSLn "agda2hs.import" 10 $ "All imported modules: " ++ show (map (pp . fst . fst) imps)
let decls = map (uncurry $ uncurry makeImportDecl) imps
return decls
where
mergeChildren :: ImportSpecMap -> ImportSpecMap -> ImportSpecMap
mergeChildren = Map.unionWith Set.union
Expand All @@ -38,18 +40,20 @@ compileImports top is0 = do
makeSingle (Just p) q = Map.singleton p $ Set.singleton q

groupModules :: [Import] -> ImportDeclMap
groupModules = foldr
(\(Import mod as p q ns) -> Map.insertWith mergeChildren (mod,as)
(makeSingle (parentNN p) (NamespacedName ns q)))
Map.empty
where
parentNN :: Maybe (Hs.Name ()) -> Maybe NamespacedName
parentNN (Just name@(Hs.Symbol _ _)) = Just $ NamespacedName (Hs.TypeNamespace ()) name
-- ^ for parents, if they are operators, we assume they are type operators
-- but actually, this will get lost anyway because of the structure of ImportSpec
-- the point is that there should not be two tuples with the same name and diffenrent namespaces
parentNN (Just name) = Just $ NamespacedName (Hs.NoNamespace ()) name
parentNN Nothing = Nothing
groupModules = flip foldr Map.empty $ \case
(Import mod as p q ns) ->
Map.insertWith mergeChildren (mod,as)
(makeSingle (parentNN p) (NamespacedName ns q))
(ImportInstances mod) ->
Map.insertWith mergeChildren (mod,Unqualified) Map.empty
where
parentNN :: Maybe (Hs.Name ()) -> Maybe NamespacedName
parentNN (Just name@(Hs.Symbol _ _)) = Just $ NamespacedName (Hs.TypeNamespace ()) name
-- ^ for parents, if they are operators, we assume they are type operators
-- but actually, this will get lost anyway because of the structure of ImportSpec
-- the point is that there should not be two tuples with the same name and diffenrent namespaces
parentNN (Just name) = Just $ NamespacedName (Hs.NoNamespace ()) name
parentNN Nothing = Nothing

-- TODO: avoid having to do this by having a CName instead of a
-- Name in the Import datatype
Expand Down Expand Up @@ -81,11 +85,13 @@ compileImports top is0 = do
[] -> checkClashingImports is
where
isClashing (Import _ as' p' q' _) = (as' == as) && (p' /= p) && (q' == q)
isClashing ImportInstances{} = False
mkErrorMsg (Import _ _ p' q' _) =
"Clashing import: " ++ pp q ++ " (both from "
++ prettyShow (pp <$> p) ++ " and "
++ prettyShow (pp <$> p') ++ ")"
-- TODO: no range information as we only have Haskell names at this point
checkClashingImports (ImportInstances mod : is) = checkClashingImports is


-- | Generate a prelude import considering prelude config options (hiding, implicit, etc).
Expand Down
7 changes: 7 additions & 0 deletions src/Agda2Hs/Compile/Name.hs
Original file line number Diff line number Diff line change
Expand Up @@ -230,3 +230,10 @@ compileModuleName m = do
text "Top-level module name for" <+> prettyTCM m <+>
text "is" <+> text (pp tlm)
return tlm

importInstance :: QName -> C ()
importInstance f = do
mod <- compileModuleName $ qnameModule f
unless (isPrimModule mod) $ do
reportSLn "agda2hs.import" 20 $ "Importing instances from " ++ pp mod
tellImport $ ImportInstances mod
49 changes: 47 additions & 2 deletions src/Agda2Hs/Compile/Term.hs
Original file line number Diff line number Diff line change
Expand Up @@ -2,6 +2,7 @@ module Agda2Hs.Compile.Term where

import Control.Arrow ( (>>>), (&&&), second )
import Control.Monad ( unless, zipWithM )
import Control.Monad.Except
import Control.Monad.Reader

import Data.Foldable ( toList )
Expand All @@ -19,11 +20,16 @@ import Agda.Syntax.Common
import Agda.Syntax.Literal
import Agda.Syntax.Internal

import Agda.TypeChecking.CheckInternal ( infer )
import Agda.TypeChecking.Constraints ( noConstraints )
import Agda.TypeChecking.Conversion ( equalTerm )
import Agda.TypeChecking.InstanceArguments ( findInstance )
import Agda.TypeChecking.MetaVars ( newInstanceMeta )
import Agda.TypeChecking.Monad
import Agda.TypeChecking.Pretty
import Agda.TypeChecking.Records ( shouldBeProjectible, isRecordType, recordFieldNames )
import Agda.TypeChecking.Datatypes ( getConType )
import Agda.TypeChecking.Reduce ( unfoldDefinitionStep, instantiate )
import Agda.TypeChecking.Reduce ( unfoldDefinitionStep, instantiate, reduce )
import Agda.TypeChecking.Substitute
import Agda.TypeChecking.Telescope ( telView, mustBePi, piApplyM, flattenTel )
import Agda.TypeChecking.ProjectionLike ( reduceProjectionLike )
Expand All @@ -35,7 +41,7 @@ import Agda.Utils.Monad
import Agda.Utils.Size

import Agda2Hs.AgdaUtils
import Agda2Hs.Compile.Name ( compileQName )
import Agda2Hs.Compile.Name ( compileQName, importInstance )

import Agda2Hs.Compile.Type ( compileType, compileDom, DomOutput(..), compileTelSize )
import Agda2Hs.Compile.Types
Expand Down Expand Up @@ -615,3 +621,42 @@ compileLiteral (LitChar c) = return $ Hs.charE c
compileLiteral (LitString t) = return $ Hs.Lit () $ Hs.String () s s
where s = Text.unpack t
compileLiteral l = genericDocError =<< text "bad term:" <?> prettyTCM (Lit l)


checkInstance :: Term -> C ()
checkInstance u = do
reportSDoc "agda2hs.checkInstance" 12 $ text "checkInstance" <+> prettyTCM u
reduce u >>= \case
Var x es -> do
unlessM (isInstance <$> domOfBV x) illegalInstance
checkInstanceElims es
Def f es -> do
unlessM (isJust . defInstance <$> getConstInfo f) illegalInstance
importInstance f
checkInstanceElims es
-- We need to compile applications of `fromNat`, `fromNeg`, and
-- `fromString` where the constraint type is ⊤ or IsTrue .... Ideally
-- this constraint would be marked as erased but this would involve
-- changing Agda builtins.
Con c _ _
| prettyShow (conName c) == "Agda.Builtin.Unit.tt" ||
prettyShow (conName c) == "Haskell.Prim.IsTrue.itsTrue" ||
prettyShow (conName c) == "Haskell.Prim.IsFalse.itsFalse" -> return ()
_ -> illegalInstance

where
illegalInstance :: C ()
illegalInstance = do
reportSDoc "agda2hs.checkInstance" 15 $ text "illegal instance: " <+> pretty u
genericDocError =<< text "illegal instance: " <+> prettyTCM u

checkInstanceElims :: Elims -> C ()
checkInstanceElims = mapM_ checkInstanceElim

checkInstanceElim :: Elim -> C ()
checkInstanceElim (Apply v) =
when (isInstance v && usableQuantity v) $
checkInstance $ unArg v
checkInstanceElim IApply{} = illegalInstance
checkInstanceElim (Proj _ f) =
unlessM (isInstance . defArgInfo <$> getConstInfo f) illegalInstance
15 changes: 10 additions & 5 deletions src/Agda2Hs/Compile/Types.hs
Original file line number Diff line number Diff line change
Expand Up @@ -108,15 +108,20 @@ qualifiedAs = join

-- | Encoding of a Haskell module import statement.
data Import = Import
{ importModule :: Hs.ModuleName ()
, importQualified :: Qualifier
, importParent :: Maybe (Hs.Name ())
, importName :: Hs.Name ()
, importNamespace :: Hs.Namespace ()
{ _importModule :: Hs.ModuleName ()
, _importQualified :: Qualifier
, _importParent :: Maybe (Hs.Name ())
, _importName :: Hs.Name ()
, _importNamespace :: Hs.Namespace ()
-- ^^ if this is a type or something like that, we can add a namespace qualifier to the import spec
-- now it's only useful for differentiating type operators; so for others we always put Hs.NoNamespace () here
-- (we don't calculate it if it's not necessary)
}
| ImportInstances (Hs.ModuleName ())

importModule :: Import -> Hs.ModuleName ()
importModule (Import{ _importModule = mod }) = mod
importModule (ImportInstances mod) = mod

type Imports = [Import]

Expand Down
37 changes: 0 additions & 37 deletions src/Agda2Hs/Compile/Utils.hs
Original file line number Diff line number Diff line change
Expand Up @@ -253,43 +253,6 @@ isInlinedFunction q = do
(InlinePragma ==) <$> processPragma r
_ -> return False

checkInstance :: Term -> C ()
checkInstance u = do
reportSDoc "agda2hs.checkInstance" 12 $ text "checkInstance" <+> prettyTCM u
reduce u >>= \case
Var x es -> do
unlessM (isInstance <$> domOfBV x) illegalInstance
checkInstanceElims es
Def f es -> do
unlessM (isJust . defInstance <$> getConstInfo f) illegalInstance
checkInstanceElims es
-- We need to compile applications of `fromNat`, `fromNeg`, and
-- `fromString` where the constraint type is ⊤ or IsTrue .... Ideally
-- this constraint would be marked as erased but this would involve
-- changing Agda builtins.
Con c _ _
| prettyShow (conName c) == "Agda.Builtin.Unit.tt" ||
prettyShow (conName c) == "Haskell.Prim.IsTrue.itsTrue" ||
prettyShow (conName c) == "Haskell.Prim.IsFalse.itsFalse" -> return ()
_ -> illegalInstance

where
illegalInstance :: C ()
illegalInstance = do
reportSDoc "agda2hs.checkInstance" 15 $ text "illegal instance: " <+> pretty u
genericDocError =<< text "illegal instance: " <+> prettyTCM u

checkInstanceElims :: Elims -> C ()
checkInstanceElims = mapM_ checkInstanceElim

checkInstanceElim :: Elim -> C ()
checkInstanceElim (Apply v) =
when (isInstance v && usableQuantity v) $
checkInstance $ unArg v
checkInstanceElim IApply{} = illegalInstance
checkInstanceElim (Proj _ f) =
unlessM (isInstance . defArgInfo <$> getConstInfo f) illegalInstance

withNestedType :: C a -> C a
withNestedType = local $ \e -> e { isNestedInType = True }

Expand Down
2 changes: 1 addition & 1 deletion src/Agda2Hs/Render.hs
Original file line number Diff line number Diff line change
Expand Up @@ -104,7 +104,7 @@ prettyShowImportDecl (Hs.ImportDecl _ m qual src safe mbPkg mbName mbSpecs) =
++ parenList (map prettyShowSpec ispecs)

parenList :: [String] -> String
parenList [] = ""
parenList [] = "()"
parenList (x:xs) = '(' : (x ++ go xs)
where
go :: [String] -> String
Expand Down
2 changes: 2 additions & 0 deletions test/AllTests.agda
Original file line number Diff line number Diff line change
Expand Up @@ -92,6 +92,7 @@ import CustomTuples
import ProjectionLike
import FunCon
import Issue308
import Issue324

{-# FOREIGN AGDA2HS
import Issue14
Expand Down Expand Up @@ -181,4 +182,5 @@ import CustomTuples
import ProjectionLike
import FunCon
import Issue308
import Issue324
#-}
8 changes: 8 additions & 0 deletions test/Issue324.agda
Original file line number Diff line number Diff line change
@@ -0,0 +1,8 @@

open import Haskell.Prelude
open import Issue324instance

test : Bool
test = not == id

{-# COMPILE AGDA2HS test #-}
9 changes: 9 additions & 0 deletions test/Issue324instance.agda
Original file line number Diff line number Diff line change
@@ -0,0 +1,9 @@

open import Haskell.Prelude

instance
eqBoolFun : Eq (Bool Bool)
eqBoolFun ._==_ x y = x True == y True && x False == y False

{-# COMPILE AGDA2HS eqBoolFun #-}

1 change: 1 addition & 0 deletions test/golden/AllTests.hs
Original file line number Diff line number Diff line change
Expand Up @@ -87,4 +87,5 @@ import CustomTuples
import ProjectionLike
import FunCon
import Issue308
import Issue324

7 changes: 7 additions & 0 deletions test/golden/Issue324.hs
Original file line number Diff line number Diff line change
@@ -0,0 +1,7 @@
module Issue324 where

import Issue324instance ()

test :: Bool
test = not == id

5 changes: 5 additions & 0 deletions test/golden/Issue324instance.hs
Original file line number Diff line number Diff line change
@@ -0,0 +1,5 @@
module Issue324instance where

instance Eq (Bool -> Bool) where
x == y = x True == y True && x False == y False

1 change: 1 addition & 0 deletions test/golden/QualifiedImports.hs
Original file line number Diff line number Diff line change
@@ -1,6 +1,7 @@
module QualifiedImports where

import qualified Importee (Foo(MkFoo), foo)
import QualifiedImportee ()
import qualified QualifiedImportee as Qually (Foo, Fooable(defaultFoo, doTheFoo), foo, (!#))

-- ** simple qualification
Expand Down
1 change: 1 addition & 0 deletions test/golden/RequalifiedImports.hs
Original file line number Diff line number Diff line change
@@ -1,6 +1,7 @@
module RequalifiedImports where

import OtherImportee (OtherFoo(MkFoo))
import QualifiedImportee ()
import qualified QualifiedImportee as A (Foo, Fooable(defaultFoo, doTheFoo), foo, (!#))

-- ** conflicting imports are all replaced with the "smallest" qualifier
Expand Down

0 comments on commit 84c0c36

Please sign in to comment.