Skip to content

Commit

Permalink
further work on compiling unboxed datatypes
Browse files Browse the repository at this point in the history
  • Loading branch information
flupe committed Aug 14, 2024
1 parent 16f3398 commit 7f873ce
Show file tree
Hide file tree
Showing 7 changed files with 70 additions and 14 deletions.
4 changes: 2 additions & 2 deletions lib/Haskell/Extra/Singleton.agda
Original file line number Diff line number Diff line change
Expand Up @@ -8,9 +8,9 @@ data Singleton (a : Set) : @0 a → Set where
{-# COMPILE AGDA2HS Singleton unboxed #-}

pureSingl : {a : Set} (x : a) Singleton a x
pureSingl = MkSingl
pureSingl x = MkSingl x

{-# COMPILE AGDA2HS pureSingl inline #-}
{-# COMPILE AGDA2HS pureSingl transparent #-}

fmapSingl
: {a b : Set} (f : a b) {@0 x : a}
Expand Down
1 change: 1 addition & 0 deletions src/Agda2Hs/Compile.hs
Original file line number Diff line number Diff line change
Expand Up @@ -89,6 +89,7 @@ compile opts tlm _ def =
(NoPragma , _ ) -> return []
(ExistingClassPragma , _ ) -> return []
(UnboxPragma s , Record{} ) -> [] <$ checkUnboxPragma def
(UnboxPragma s , Datatype{}) -> [] <$ checkUnboxPragma def
(TransparentPragma , Function{}) -> [] <$ checkTransparentPragma def
(InlinePragma , Function{}) -> [] <$ checkInlinePragma def
(TuplePragma b , Record{} ) -> return []
Expand Down
11 changes: 8 additions & 3 deletions src/Agda2Hs/Compile/Function.hs
Original file line number Diff line number Diff line change
Expand Up @@ -399,7 +399,7 @@ checkInlinePragma def@Defn{defName = f} = do
[c] ->
unlessM (allowedPats (namedClausePats c)) $ genericDocError =<<
"Cannot make function" <+> prettyTCM (defName def) <+> "inlinable." <+>
"Inline functions can only use variable patterns or transparent record constructor patterns."
"Inline functions can only use variable patterns or unboxed constructor patterns."
_ ->
genericDocError =<<
"Cannot make function" <+> prettyTCM f <+> "inlinable." <+>
Expand All @@ -408,11 +408,16 @@ checkInlinePragma def@Defn{defName = f} = do
where allowedPat :: DeBruijnPattern -> C Bool
allowedPat VarP{} = pure True
-- only allow matching on (unboxed) record constructors
allowedPat (ConP ch ci cargs) =
allowedPat (ConP ch ci cargs) = do
reportSDoc "agda2hs.compile.inline" 18 $ "Checking unboxing of constructor: "<+> prettyTCM ch
isUnboxConstructor (conName ch) >>= \case
Just _ -> allowedPats cargs
Nothing -> pure False
allowedPat _ = pure False
-- NOTE(flupe): dot patterns should really only be allowed for *erased* arguments, I think.
allowedPat (DotP _ _) = pure True
allowedPat p = do
reportSDoc "agda2hs.compile.inline" 18 $ "Unsupported pattern" <+> prettyTCM p
pure False

allowedPats :: NAPs -> C Bool
allowedPats pats = allM pats (allowedPat . dget . dget)
7 changes: 4 additions & 3 deletions src/Agda2Hs/Compile/Record.hs
Original file line number Diff line number Diff line change
Expand Up @@ -171,9 +171,7 @@ compileRecord target def = do
return $ Hs.DataDecl () don Nothing hd [conDecl] ds

checkUnboxPragma :: Definition -> C ()
checkUnboxPragma def = do
let Record{..} = theDef def

checkUnboxPragma def | Record{..} <- theDef def = do
-- recRecursive can be used again after agda 2.6.4.2 is released
-- see agda/agda#7042
unless (all null recMutual) $ genericDocError
Expand All @@ -197,3 +195,6 @@ checkUnboxPragma def = do
DOType -> genericDocError =<< text "Type field in unboxed record not supported"
DOInstance -> genericDocError =<< text "Instance field in unboxed record not supported"
DOTerm -> (absName tel:) <$> underAbstraction a tel nonErasedFields

checkUnboxPragma def | Datatype{..} <- theDef def = do
pure () -- TOOD(flupe):
29 changes: 25 additions & 4 deletions src/Agda2Hs/Compile/Type.hs
Original file line number Diff line number Diff line change
Expand Up @@ -9,6 +9,7 @@ import Control.Monad.Trans ( lift )
import Control.Monad.Reader ( asks )
import Data.List ( find )
import Data.Maybe ( mapMaybe, isJust )
import Data.Foldable ( toList )
import qualified Data.Set as Set ( singleton )

import qualified Language.Haskell.Exts.Syntax as Hs
Expand All @@ -21,6 +22,7 @@ import Agda.Syntax.Common
import Agda.Syntax.Internal
import Agda.Syntax.Common.Pretty ( prettyShow )

import Agda.TypeChecking.Datatypes
import Agda.TypeChecking.Pretty
import Agda.TypeChecking.Reduce ( reduce, unfoldDefinitionStep )
import Agda.TypeChecking.Substitute
Expand Down Expand Up @@ -76,7 +78,7 @@ delayType [] = genericDocError =<< text "Cannot compile unapplied Delay type"
compileTypeWithStrictness :: Term -> C (Strictness, Hs.Type ())
compileTypeWithStrictness t = do
s <- case t of
Def f es -> fromMaybe Lazy <$> isUnboxRecord f
Def f es -> fromMaybe Lazy <$> isUnboxType f
_ -> return Lazy
ty <- compileType t
pure (s, ty)
Expand Down Expand Up @@ -106,7 +108,7 @@ compileType t = do
<+> text "in Haskell type"
| Just semantics <- isSpecialType f -> setCurrentRange f $ semantics es
| Just args <- allApplyElims es ->
ifJustM (isUnboxRecord f) (\_ -> compileUnboxType f args) $
ifJustM (isUnboxType f) (\_ -> compileUnboxType f args) $
ifJustM (isTupleRecord f) (\b -> compileTupleType f b args) $
ifM (isTransparentFunction f) (compileTransparentType (defType def) args) $
ifM (isInlinedFunction f) (compileInlineType f es) $ do
Expand Down Expand Up @@ -174,11 +176,31 @@ compileTelSize (ExtendTel a tel) = compileDom a >>= \case
compileUnboxType :: QName -> Args -> C (Hs.Type ())
compileUnboxType r pars = do
def <- getConstInfo r
let tel = recTel (theDef def) `apply` pars

tel <- case theDef def of

Record{recTel} -> pure (recTel `apply` pars)

Datatype{dataCons = [con]} -> do
-- NOTE(flupe): con is the *only* constructor of the unboxed datatype
Constructor { conSrcCon } <- theDef <$> getConstInfo con
-- we retrieve its type
Just (_, ty) <- getConType conSrcCon (El (DummyS "some level") (apply (Def r []) pars))
theTel <$> telView ty

compileTel tel >>= \case
[t] -> return t
_ -> __IMPOSSIBLE__

where
compileTel :: Telescope -> C [Hs.Type ()]
compileTel EmptyTel = return []
compileTel (ExtendTel a tel) = compileDom a >>= \case
DODropped -> underAbstraction a tel compileTel
DOInstance -> __IMPOSSIBLE__
DOType -> __IMPOSSIBLE__
DOTerm -> (:) <$> compileType (unEl $ unDom a) <*> underAbstraction a tel compileTel

compileTupleType :: QName -> Hs.Boxed -> Args -> C (Hs.Type ())
compileTupleType r b pars = do
tellUnboxedTuples b
Expand All @@ -192,7 +214,6 @@ compileTransparentType ty args = compileTypeArgs ty args >>= \case
(v:vs) -> return $ v `tApp` vs
[] -> __IMPOSSIBLE__


compileInlineType :: QName -> Elims -> C (Hs.Type ())
compileInlineType f args = do
Function { funClauses = cs } <- theDef <$> getConstInfo f
Expand Down
19 changes: 17 additions & 2 deletions src/Agda2Hs/Compile/Utils.hs
Original file line number Diff line number Diff line change
Expand Up @@ -192,9 +192,24 @@ isUnboxRecord q = do
_ -> Nothing
_ -> return Nothing

isUnboxType :: QName -> C (Maybe Strictness)
isUnboxType q = do
getConstInfo q >>= \case
Defn{defName = r} ->
processPragma r <&> \case
UnboxPragma s -> Just s
_ -> Nothing
-- _ -> return Nothing

isUnboxConstructor :: QName -> C (Maybe Strictness)
isUnboxConstructor q =
caseMaybeM (isRecordConstructor q) (return Nothing) $ isUnboxRecord . fst
isUnboxConstructor q = do
def <- getConstInfo q
case theDef def of
Constructor {conData = r} -> isUnboxType r
_ -> pure Nothing

-- caseMaybeM (isRecordConstructor) isUnboxType q
-- caseMaybeM (isRecordConstructor q) (return Nothing) $ isUnboxRecord . fst

isUnboxProjection :: QName -> C (Maybe Strictness)
isUnboxProjection q =
Expand Down
13 changes: 13 additions & 0 deletions test/Singleton.agda
Original file line number Diff line number Diff line change
@@ -0,0 +1,13 @@
open import Haskell.Prelude
open import Haskell.Extra.Singleton

test1 : (@0 x : ⊤) Singleton ⊤ x
test1 x = MkSingl tt

{-# COMPILE AGDA2HS test1 #-}

-- doesn't work because inline functions only reduce because of inline record transformation
-- test2 : {a b : Set} (f : a → b) {@0 x : a} → Singleton a x → Singleton b (f x)
-- test2 f sx = fmapSingl f sx
--
-- {-# COMPILE AGDA2HS test2 #-}

0 comments on commit 7f873ce

Please sign in to comment.