Skip to content

Commit 003d0df

Browse files
committed
[ #306 ] Error on module param dot pat
Delay applying getContextArgs until after we have checked for dot patterns Also stop throwing __IMPOSSIBLE__ on instances inside modules
1 parent 980c6f6 commit 003d0df

File tree

11 files changed

+129
-34
lines changed

11 files changed

+129
-34
lines changed

src/Agda2Hs/Compile/ClassInstance.hs

Lines changed: 17 additions & 15 deletions
Original file line numberDiff line numberDiff line change
@@ -46,7 +46,7 @@ disableCopatterns :: C a -> C a
4646
disableCopatterns = local $ \e -> e { copatternsEnabled = False }
4747

4848

49-
-- | Enable the approriate extensions for a given Haskell deriving strategy.
49+
-- | Enable the appropriate extensions for a given Haskell deriving strategy.
5050
enableStrategy :: Maybe (Hs.DerivStrategy ()) -> C ()
5151
enableStrategy Nothing = return ()
5252
enableStrategy (Just s) = do
@@ -66,14 +66,14 @@ compileInstance (ToDerivation strategy) def@Defn{..} =
6666
text "compiling instance" <+> prettyTCM defName <+> text "to standalone deriving"
6767
tellExtension Hs.StandaloneDeriving
6868
enableStrategy strategy
69-
ir <- compileInstRule [] (unEl defType)
69+
ir <- compileInstRule [] [] (unEl defType)
7070
return $ Hs.DerivDecl () strategy Nothing ir
7171

7272
compileInstance ToDefinition def@Defn{..} =
7373
enableCopatterns $ setCurrentRangeQ defName $ do
7474
reportSDoc "agda2hs.compile.instance" 13 $
7575
text "compiling instance" <+> prettyTCM defName <+> text "to instance definition"
76-
ir <- compileInstRule [] (unEl defType)
76+
ir <- compileInstRule [] [] (unEl defType)
7777
withFunctionLocals defName $ do
7878
reportSDoc "agda2hs.compile.instance" 20 $ vcat $
7979
text "compileInstance clauses: " :
@@ -83,11 +83,11 @@ compileInstance ToDefinition def@Defn{..} =
8383
tel <- lookupSection mod
8484
addContext tel $ do
8585
liftTCM $ setModuleCheckpoint mod
86-
pars <- getContextArgs
87-
ty <- defType `piApplyM` pars
88-
let clauses = funClauses `apply` pars
86+
-- We intentionally do not apply getContextArgs eagerly here so dotted
87+
-- patterns corresponding to module parameters are correctly detected
88+
-- (https://github.com/agda/agda2hs/issues/306)
8989
(ds, rs) <- concatUnzip
90-
<$> mapM (\c -> withClauseLocals mod c $ compileInstanceClause mod ty c) clauses
90+
<$> mapM (\c -> withClauseLocals mod c $ compileInstanceClause mod defType c) funClauses
9191
reportSDoc "agda2hs.compile.instance" 20 $ vcat $
9292
text "compileInstance compiled clauses: " :
9393
map (nest 2 . text . pp) ds
@@ -97,16 +97,16 @@ compileInstance ToDefinition def@Defn{..} =
9797
where Function{..} = theDef
9898

9999

100-
compileInstRule :: [Hs.Asst ()] -> Term -> C (Hs.InstRule ())
101-
compileInstRule cs ty = do
100+
compileInstRule :: [Hs.TyVarBind ()] -> [Hs.Asst ()] -> Term -> C (Hs.InstRule ())
101+
compileInstRule xs cs ty = do
102102
reportSDoc "agda2hs.compile.instance" 20 $ text "compileInstRule" <+> prettyTCM ty
103103
case unSpine1 ty of
104104
Def f es | Just args <- allApplyElims es -> do
105105
fty <- defType <$> getConstInfo f
106106
vs <- compileTypeArgs fty args
107107
f <- compileQName f
108108
return $
109-
Hs.IRule () Nothing (ctx cs) $ foldl (Hs.IHApp ()) (Hs.IHCon () f) (map pars vs)
109+
Hs.IRule () (Just xs) (ctx cs) $ foldl (Hs.IHApp ()) (Hs.IHCon () f) (map pars vs)
110110
where ctx [] = Nothing
111111
ctx cs = Just (Hs.CxTuple () cs)
112112
-- put parens around anything except a var or a constant
@@ -115,11 +115,12 @@ compileInstRule cs ty = do
115115
pars t@(Hs.TyCon () _) = t
116116
pars t = Hs.TyParen () t
117117
Pi a b -> compileDomType (absName b) a >>= \case
118-
DomDropped -> underAbstr a b (compileInstRule cs . unEl)
118+
DomDropped -> underAbstr a b (compileInstRule xs cs . unEl)
119119
DomConstraint hsA ->
120-
underAbstraction a b (compileInstRule (cs ++ [hsA]) . unEl)
120+
underAbstraction a b (compileInstRule xs (cs ++ [hsA]) . unEl)
121121
DomType _ t -> __IMPOSSIBLE__
122-
DomForall _ -> __IMPOSSIBLE__
122+
DomForall x ->
123+
underAbstraction a b (compileInstRule (xs ++ [x]) cs . unEl)
123124
_ -> __IMPOSSIBLE__
124125

125126

@@ -170,7 +171,7 @@ compileInstanceClause curModule ty c = ifNotM (keepClause c) (return ([], [])) $
170171
-- abuse compileClause:
171172
-- 1. drop any patterns before record projection to suppress the instance arg
172173
-- 2. use record proj. as function name
173-
-- 3. process remaing patterns as usual
174+
-- 3. process remaining patterns as usual
174175
compileInstanceClause' :: ModuleName -> Type -> NAPs -> Clause -> C ([Hs.InstDecl ()], [QName])
175176
compileInstanceClause' curModule ty [] c = __IMPOSSIBLE__
176177
compileInstanceClause' curModule ty (p:ps) c
@@ -289,7 +290,7 @@ compileInstanceClause' curModule ty (p:ps) c
289290
-- No minimal dictionary used, proceed with compiling as a regular clause.
290291
| otherwise -> do
291292
reportSDoc "agda2hs.compile.instance" 20 $ text "Compiling instance clause" <+> prettyTCM c'
292-
ms <- disableCopatterns $ compileClause curModule Nothing uf ty' c'
293+
ms <- disableCopatterns $ compileClause curModule [] Nothing uf ty' c'
293294
let cc = Hs.FunBind () (toList ms)
294295
reportSDoc "agda2hs.compile.instance" 20 $ vcat
295296
[ text "compileInstanceClause compiled clause"
@@ -302,6 +303,7 @@ compileInstanceClause' curModule ty (p:ps) c
302303
-- if there is something other than erased parameters
303304
compileInstanceClause' curModule ty (p:ps) c = do
304305
(a, b) <- mustBePi ty
306+
checkForced a (namedArg p)
305307
compileInstanceClause' curModule (absApp b (patternToTerm $ namedArg p)) ps c
306308

307309
fieldArgInfo :: QName -> C ArgInfo

src/Agda2Hs/Compile/Function.hs

Lines changed: 36 additions & 11 deletions
Original file line numberDiff line numberDiff line change
@@ -9,7 +9,7 @@ import Data.List
99
import Data.Maybe ( fromMaybe, isJust )
1010
import qualified Data.Text as Text
1111

12-
import Agda.Compiler.Backend
12+
import Agda.Compiler.Backend hiding (Args)
1313
import Agda.Compiler.Common
1414

1515
import Agda.Syntax.Common
@@ -152,9 +152,8 @@ compileFun' withSig def@Defn{..} = inTopContext $ withCurrentModule m $ do
152152
reportSDoc "agda2hs.compile.type" 8 $ "Compiled function type: " <+> text (Hs.prettyPrint ty)
153153
return [Hs.TypeSig () [x] ty]
154154

155-
let clauses = funClauses `apply` pars
156155
cs <- map (addPats paramPats) <$>
157-
mapMaybeM (compileClause m (Just defName) x typ) clauses
156+
mapMaybeM (compileClause m pars (Just defName) x defType) funClauses
158157

159158
when (null cs) $ agda2hsErrorM $
160159
text "Functions defined with absurd patterns exclusively are not supported."
@@ -185,13 +184,22 @@ compileModuleParams (ExtendTel a tel) = do
185184
return (Hs.TyFun () a, [Hs.PVar () n])
186185
((f .) *** (p++)) <$> underAbstraction a tel compileModuleParams
187186

188-
compileClause :: ModuleName -> Maybe QName -> Hs.Name () -> Type -> Clause -> C (Maybe (Hs.Match ()))
189-
compileClause curModule mproj x t c =
187+
compileClause :: ModuleName -> Args -> Maybe QName -> Hs.Name () -> Type -> Clause -> C (Maybe (Hs.Match ()))
188+
compileClause curModule pars mproj x t c =
190189
withClauseLocals curModule c $ do
191-
compileClause' curModule mproj x t c
192-
193-
compileClause' :: ModuleName -> Maybe QName -> Hs.Name () -> Type -> Clause -> C (Maybe (Hs.Match ()))
194-
compileClause' curModule projName x ty c@Clause{..} = do
190+
compileClause' curModule pars mproj x t c
191+
192+
compileClause' :: ModuleName -> Args -> Maybe QName -> Hs.Name () -> Type -> Clause -> C (Maybe (Hs.Match ()))
193+
compileClause' curModule pars projName x ty c = do
194+
-- Check whether any parameters (including module parameters) are forced
195+
-- See https://github.com/agda/agda2hs/issues/306
196+
annPats <- annPats ty (namedClausePats c)
197+
traverse (uncurry checkForced) annPats
198+
ty' <- piApplyM ty pars
199+
compileClause'' curModule projName x ty' (c `apply` pars)
200+
201+
compileClause'' :: ModuleName -> Maybe QName -> Hs.Name () -> Type -> Clause -> C (Maybe (Hs.Match ()))
202+
compileClause'' curModule projName x ty c@Clause{..} = do
195203
reportSDoc "agda2hs.compile" 7 $ "compiling clause: " <+> prettyTCM c
196204

197205
ifNotM (keepClause c) (pure Nothing) $ addContext (KeepNames clauseTel) $ do
@@ -251,6 +259,24 @@ keepClause c@Clause{..} = case (clauseBody, clauseType) of
251259
DOType -> __IMPOSSIBLE__
252260
DOTerm -> True
253261

262+
checkForced :: Dom Type -> DeBruijnPattern -> C ()
263+
checkForced a pat
264+
= when (usableDom a) $ when (isForcedPat pat)
265+
$ agda2hsError "not supported: forced (dot) patterns in non-erased positions"
266+
267+
annPats :: Type -> NAPs -> C [(Dom Type, DeBruijnPattern)]
268+
annPats ty [] = pure []
269+
annPats ty (p : ps) = do
270+
(a, b) <- recTy (namedArg p)
271+
ps' <- annPats b ps
272+
pure $ (a, namedArg p) : ps'
273+
where recTy (ProjP po pn) = do
274+
ty' <- fromMaybe __IMPOSSIBLE__ <$> getDefType pn ty
275+
(a, b) <- mustBePi ty'
276+
pure (a, absBody b)
277+
recTy pat = do
278+
(a, b) <- mustBePi ty
279+
pure (a, absApp b (patternToTerm pat))
254280

255281
-- TODO(flupe): projection-like definitions are missing the first (variable) patterns
256282
-- (that are however present in the type)
@@ -271,15 +297,14 @@ compilePats ty ((namedArg -> pat):ps) = do
271297
(a, b) <- mustBePi ty
272298
reportSDoc "agda2hs.compile.pattern" 10 $ text "Compiling pattern:" <+> prettyTCM pat
273299
let rest = compilePats (absApp b (patternToTerm pat)) ps
274-
when (usableDom a) checkForced
300+
checkForced a pat
275301
compileDom a >>= \case
276302
DOInstance -> rest
277303
DODropped -> rest
278304
DOType -> rest
279305
DOTerm -> do
280306
checkNoAsPatterns pat
281307
(:) <$> compilePat (unDom a) pat <*> rest
282-
where checkForced = when (isForcedPat pat) $ agda2hsError "not supported: forced (dot) patterns in non-erased positions"
283308

284309

285310
compilePat :: Type -> DeBruijnPattern -> C (Hs.Pat ())
Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -1,7 +1,7 @@
11
module Agda2Hs.Compile.Function where
22

33
import qualified Agda2Hs.Language.Haskell as Hs ( Match, Name )
4-
import Agda.Syntax.Internal ( Clause, ModuleName, QName, Type )
4+
import Agda.Syntax.Internal ( Clause, ModuleName, QName, Type, Args )
55
import Agda2Hs.Compile.Types ( C )
66

7-
compileClause' :: ModuleName -> Maybe QName -> Hs.Name () -> Type -> Clause -> C (Maybe (Hs.Match ()))
7+
compileClause' :: ModuleName -> Args -> Maybe QName -> Hs.Name () -> Type -> Clause -> C (Maybe (Hs.Match ()))

src/Agda2Hs/Compile/Term.hs

Lines changed: 2 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -80,11 +80,8 @@ lambdaCase q ty args = compileLocal $ setCurrentRangeQ q $ do
8080
npars <- size <$> lookupSection mname
8181

8282
let (pars, rest) = splitAt npars args
83-
cs = applys cls pars
8483

85-
ty' <- piApplyM ty pars
86-
87-
cs <- mapMaybeM (compileClause' mname (Just q) (hsName "(lambdaCase)") ty') cs
84+
cs <- mapMaybeM (compileClause' mname (map defaultArg pars) (Just q) (hsName "(lambdaCase)") ty) cls
8885

8986
case cs of
9087
-- If there is a single clause and all proper patterns got erased,
@@ -93,6 +90,7 @@ lambdaCase q ty args = compileLocal $ setCurrentRangeQ q $ do
9390
| null ps -> return rhs
9491
| all isVarPat ps -> return $ Hs.Lambda () ps rhs
9592
_ -> do
93+
ty' <- piApplyM ty pars
9694
lcase <- hsLCase =<< mapM clauseToAlt cs -- Pattern lambdas cannot have where blocks
9795
eApp lcase <$> compileArgs ty' rest
9896
-- undefined -- compileApp lcase (undefined, undefined, rest)

test/AllFailTests.agda

Lines changed: 4 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -41,3 +41,7 @@ import Fail.Issue125
4141
import Fail.Issue357a
4242
import Fail.Issue357b
4343
import Fail.DerivingParseFailure
44+
import Fail.ErasedPatternLambda
45+
import Fail.Issue306a
46+
import Fail.Issue306b
47+
import Fail.Issue306c

test/AllTests.agda

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -87,7 +87,6 @@ import Issue309
8787
import Issue317
8888
import Issue353
8989
import RankNTypes
90-
import ErasedPatternLambda
9190
import CustomTuples
9291
import ProjectionLike
9392
import FunCon
@@ -102,6 +101,7 @@ import Issue346
102101
import Issue408
103102
import CompileTo
104103
import RuntimeCast
104+
import Issue306
105105

106106
{-# FOREIGN AGDA2HS
107107
import Issue14
@@ -186,7 +186,6 @@ import Issue309
186186
import Issue317
187187
import Issue353
188188
import RankNTypes
189-
import ErasedPatternLambda
190189
import CustomTuples
191190
import ProjectionLike
192191
import FunCon
@@ -201,4 +200,5 @@ import Issue346
201200
import Issue408
202201
import CompileTo
203202
import RuntimeCast
203+
import Issue306
204204
#-}

test/ErasedPatternLambda.agda renamed to test/Fail/ErasedPatternLambda.agda

Lines changed: 2 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -1,3 +1,5 @@
1+
module Fail.ErasedPatternLambda where
2+
13
open import Haskell.Prelude
24

35
Scope = List Bool

test/Fail/Issue306a.agda

Lines changed: 11 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,11 @@
1+
open import Haskell.Prelude
2+
3+
module Fail.Issue306a where
4+
5+
cast : {a b : Set} @0 a ≡ b a b
6+
cast {a} {b} = cast'
7+
where
8+
cast' : @0 a ≡ b a b
9+
cast' refl x = x
10+
11+
{-# COMPILE AGDA2HS cast #-}

test/Fail/Issue306b.agda

Lines changed: 16 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,16 @@
1+
open import Haskell.Prelude
2+
3+
module Fail.Issue306b where
4+
5+
record Foo (a b : Set) : Set where
6+
no-eta-equality
7+
field
8+
coe : a b
9+
open Foo public
10+
11+
instance
12+
foo : {a b : Set} {{@0 _ : a ≡ b}} Foo a b
13+
foo {{refl}} .coe x = x
14+
15+
{-# COMPILE AGDA2HS Foo class #-}
16+
{-# COMPILE AGDA2HS foo #-}

test/Fail/Issue306c.agda

Lines changed: 17 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,17 @@
1+
open import Haskell.Prelude
2+
3+
module Fail.Issue306c where
4+
5+
record Foo (a b : Set) : Set where
6+
no-eta-equality
7+
field
8+
coe : a b
9+
open Foo public
10+
11+
module _ {a b : Set} where
12+
instance
13+
foo : {{@0 _ : a ≡ b}} Foo a b
14+
foo {{refl}} .coe x = x
15+
16+
{-# COMPILE AGDA2HS Foo class #-}
17+
{-# COMPILE AGDA2HS foo #-}

0 commit comments

Comments
 (0)