Skip to content

Commit d33702f

Browse files
committed
proper checks for unboxed datatypes
1 parent 46c3248 commit d33702f

File tree

3 files changed

+59
-19
lines changed

3 files changed

+59
-19
lines changed

src/Agda2Hs/Compile.hs

Lines changed: 4 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -20,10 +20,10 @@ import Agda.Utils.Monad ( whenM, anyM, when )
2020
import qualified Language.Haskell.Exts.Extension as Hs
2121

2222
import Agda2Hs.Compile.ClassInstance ( compileInstance )
23-
import Agda2Hs.Compile.Data ( compileData )
23+
import Agda2Hs.Compile.Data ( compileData, checkUnboxedDataPragma )
2424
import Agda2Hs.Compile.Function ( compileFun, checkTransparentPragma, checkInlinePragma )
2525
import Agda2Hs.Compile.Postulate ( compilePostulate )
26-
import Agda2Hs.Compile.Record ( compileRecord, checkUnboxPragma )
26+
import Agda2Hs.Compile.Record ( compileRecord, checkUnboxedRecordPragma )
2727
import Agda2Hs.Compile.Types
2828
import Agda2Hs.Compile.Utils ( setCurrentRangeQ, tellExtension, primModules, isClassName )
2929
import Agda2Hs.Pragma
@@ -88,8 +88,8 @@ compile opts tlm _ def =
8888
case (p , theDef def) of
8989
(NoPragma , _ ) -> return []
9090
(ExistingClassPragma , _ ) -> return []
91-
(UnboxPragma s , Record{} ) -> [] <$ checkUnboxPragma def
92-
(UnboxPragma s , Datatype{}) -> [] <$ checkUnboxPragma def
91+
(UnboxPragma s , Record{} ) -> [] <$ checkUnboxedRecordPragma def
92+
(UnboxPragma s , Datatype{}) -> [] <$ checkUnboxedDataPragma def
9393
(TransparentPragma , Function{}) -> [] <$ checkTransparentPragma def
9494
(InlinePragma , Function{}) -> [] <$ checkInlinePragma def
9595
(TuplePragma b , Record{} ) -> return []

src/Agda2Hs/Compile/Data.hs

Lines changed: 51 additions & 11 deletions
Original file line numberDiff line numberDiff line change
@@ -2,7 +2,7 @@ module Agda2Hs.Compile.Data where
22

33
import qualified Language.Haskell.Exts.Syntax as Hs
44

5-
import Control.Monad ( when )
5+
import Control.Monad ( unless, when )
66
import Agda.Compiler.Backend
77
import Agda.Syntax.Common
88
import Agda.Syntax.Internal
@@ -14,8 +14,9 @@ import Agda.TypeChecking.Substitute
1414
import Agda.TypeChecking.Telescope
1515

1616
import Agda.Utils.Impossible ( __IMPOSSIBLE__ )
17+
import Agda.Utils.Monad ( mapMaybeM )
1718

18-
import Agda2Hs.Compile.Type ( compileDomType, compileTeleBinds )
19+
import Agda2Hs.Compile.Type ( compileDomType, compileTeleBinds, compileDom, DomOutput(..) )
1920
import Agda2Hs.Compile.Types
2021
import Agda2Hs.Compile.Utils
2122
import Agda2Hs.HsUtils
@@ -46,15 +47,15 @@ compileData newtyp ds def = do
4647
when newtyp (checkNewtype d cs)
4748

4849
return [Hs.DataDecl () target Nothing hd cs ds]
49-
where
50-
allIndicesErased :: Type -> C ()
51-
allIndicesErased t = reduce (unEl t) >>= \case
52-
Pi dom t -> compileDomType (absName t) dom >>= \case
53-
DomDropped -> allIndicesErased (unAbs t)
54-
DomType{} -> genericDocError =<< text "Not supported: indexed datatypes"
55-
DomConstraint{} -> genericDocError =<< text "Not supported: constraints in types"
56-
DomForall{} -> genericDocError =<< text "Not supported: indexed datatypes"
57-
_ -> return ()
50+
51+
allIndicesErased :: Type -> C ()
52+
allIndicesErased t = reduce (unEl t) >>= \case
53+
Pi dom t -> compileDomType (absName t) dom >>= \case
54+
DomDropped -> allIndicesErased (unAbs t)
55+
DomType{} -> genericDocError =<< text "Not supported: indexed datatypes"
56+
DomConstraint{} -> genericDocError =<< text "Not supported: constraints in types"
57+
DomForall{} -> genericDocError =<< text "Not supported: indexed datatypes"
58+
_ -> return ()
5859

5960
compileConstructor :: [Arg Term] -> QName -> C (Hs.QualConDecl ())
6061
compileConstructor params c = do
@@ -77,3 +78,42 @@ compileConstructorArgs (ExtendTel a tel) = compileDomType (absName tel) a >>= \c
7778
DomConstraint hsA -> genericDocError =<< text "Not supported: constructors with class constraints"
7879
DomDropped -> underAbstraction a tel compileConstructorArgs
7980
DomForall{} -> __IMPOSSIBLE__
81+
82+
83+
checkUnboxedDataPragma :: Definition -> C ()
84+
checkUnboxedDataPragma def = do
85+
let Datatype{..} = theDef def
86+
87+
-- unboxed datatypes shouldn't be recursive
88+
unless (all null dataMutual) $ genericDocError
89+
=<< text "Unboxed datatype" <+> prettyTCM (defName def)
90+
<+> text "cannot be recursive"
91+
92+
TelV tel t <- telViewUpTo dataPars (defType def)
93+
let params :: [Arg Term] = teleArgs tel
94+
95+
allIndicesErased t
96+
97+
case dataCons of
98+
[con] -> do
99+
info <- getConstInfo con
100+
let Constructor {..} = theDef info
101+
ty <- defType info `piApplyM` params
102+
TelV conTel _ <- telView ty
103+
args <- nonErasedArgs conTel
104+
unless (length args == 1) $ genericDocError
105+
=<< text "Unboxed datatype" <+> prettyTCM (defName def)
106+
<+> text "should have a single constructor with exactly one non-erased argument."
107+
108+
_ -> genericDocError =<< text "Unboxed datatype" <+> prettyTCM (defName def)
109+
<+> text "must have exactly one constructor."
110+
111+
where
112+
nonErasedArgs :: Telescope -> C [String]
113+
nonErasedArgs EmptyTel = return []
114+
nonErasedArgs (ExtendTel a tel) = compileDom a >>= \case
115+
DODropped -> underAbstraction a tel nonErasedArgs
116+
DOType -> genericDocError =<< text "Type argument in unboxed datatype not supported"
117+
DOInstance -> genericDocError =<< text "Instance argument in unboxed datatype not supported"
118+
DOTerm -> (absName tel:) <$> underAbstraction a tel nonErasedArgs
119+

src/Agda2Hs/Compile/Record.hs

Lines changed: 4 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -29,6 +29,7 @@ import Agda2Hs.Compile.Function ( compileFun )
2929
import Agda2Hs.Compile.Type ( compileDomType, compileTeleBinds, compileDom, DomOutput(..) )
3030
import Agda2Hs.Compile.Types
3131
import Agda2Hs.Compile.Utils
32+
import Agda2Hs.Compile.Data ( allIndicesErased )
3233
import Agda2Hs.HsUtils
3334

3435
-- | Primitive fields and default implementations
@@ -170,8 +171,9 @@ compileRecord target def = do
170171
let conDecl = Hs.QualConDecl () Nothing Nothing $ Hs.RecDecl () cName fieldDecls
171172
return $ Hs.DataDecl () don Nothing hd [conDecl] ds
172173

173-
checkUnboxPragma :: Definition -> C ()
174-
checkUnboxPragma def | Record{..} <- theDef def = do
174+
checkUnboxedRecordPragma :: Definition -> C ()
175+
checkUnboxedRecordPragma def = do
176+
let Record{..} = theDef def
175177
-- recRecursive can be used again after agda 2.6.4.2 is released
176178
-- see agda/agda#7042
177179
unless (all null recMutual) $ genericDocError
@@ -196,5 +198,3 @@ checkUnboxPragma def | Record{..} <- theDef def = do
196198
DOInstance -> genericDocError =<< text "Instance field in unboxed record not supported"
197199
DOTerm -> (absName tel:) <$> underAbstraction a tel nonErasedFields
198200

199-
checkUnboxPragma def | Datatype{..} <- theDef def = do
200-
pure () -- TOOD(flupe):

0 commit comments

Comments
 (0)