Skip to content

Commit 66ee516

Browse files
committed
Allow dot patterns on non-dependent args
1 parent b26e677 commit 66ee516

File tree

10 files changed

+39
-14
lines changed

10 files changed

+39
-14
lines changed

src/Agda2Hs/Compile/ClassInstance.hs

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -303,7 +303,7 @@ compileInstanceClause' curModule ty (p:ps) c
303303
-- if there is something other than erased parameters
304304
compileInstanceClause' curModule ty (p:ps) c = do
305305
(a, b) <- mustBePi ty
306-
checkNonErasedForced a (namedArg p)
306+
checkIllegalForced a (namedArg p)
307307
compileInstanceClause' curModule (absApp b (patternToTerm $ namedArg p)) ps c
308308

309309
fieldArgInfo :: QName -> C ArgInfo

src/Agda2Hs/Compile/Function.hs

Lines changed: 7 additions & 6 deletions
Original file line numberDiff line numberDiff line change
@@ -37,7 +37,7 @@ import Agda.Utils.Size ( Sized(size) )
3737

3838
import Agda2Hs.AgdaUtils
3939
import Agda2Hs.Compile.Name ( compileQName )
40-
import Agda2Hs.Compile.Term ( compileTerm, usableDom )
40+
import Agda2Hs.Compile.Term ( compileTerm, usableDom, dependentDom )
4141
import Agda2Hs.Compile.Type ( compileType, compileDom, DomOutput(..), compileDomType )
4242
import Agda2Hs.Compile.TypeDefinition ( compileTypeDef )
4343
import Agda2Hs.Compile.Types
@@ -194,7 +194,7 @@ compileClause' curModule pars projName x ty c = do
194194
-- Check whether any parameters (including module parameters) are forced
195195
-- See https://github.com/agda/agda2hs/issues/306
196196
annPats <- annPats ty (namedClausePats c)
197-
traverse (uncurry checkNonErasedForced) annPats
197+
traverse (uncurry checkIllegalForced) annPats
198198
ty' <- piApplyM ty pars
199199
compileClause'' curModule projName x ty' (c `apply` pars)
200200

@@ -259,10 +259,11 @@ keepClause c@Clause{..} = case (clauseBody, clauseType) of
259259
DOType -> __IMPOSSIBLE__
260260
DOTerm -> True
261261

262-
checkNonErasedForced :: Dom Type -> DeBruijnPattern -> C ()
263-
checkNonErasedForced a pat
264-
= when (usableDom a && isForcedPat pat)
265-
$ agda2hsError "not supported: forced (dot) patterns in non-erased positions"
262+
checkIllegalForced :: Dom Type -> DeBruijnPattern -> C ()
263+
checkIllegalForced a pat = do
264+
dep <- dependentDom a
265+
when (dep && isForcedPat pat) $ agda2hsError
266+
"not supported: forced (dot) patterns in non-erased positions"
266267

267268
annPats :: Type -> NAPs -> C [(Dom Type, DeBruijnPattern)]
268269
annPats ty [] = pure []

src/Agda2Hs/Compile/Term.hs

Lines changed: 11 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -41,7 +41,7 @@ import Agda.Utils.Size
4141
import Agda2Hs.AgdaUtils
4242
import Agda2Hs.Compile.Name ( compileQName, importInstance )
4343

44-
import Agda2Hs.Compile.Type ( compileType, compileDom, DomOutput(..), compileTelSize )
44+
import Agda2Hs.Compile.Type ( compileType, compileDom, DomOutput(..), compileTelSize, compileDomType )
4545
import Agda2Hs.Compile.Types
4646
import Agda2Hs.Compile.Utils
4747
import Agda2Hs.Compile.Var ( compileDBVar )
@@ -499,6 +499,16 @@ usableDom :: Dom Type -> Bool
499499
usableDom dom | Prop _ <- getSort dom = False
500500
usableDom dom = usableModality dom
501501

502+
-- | Check whether a domain is dependent on the Haskell side (i.e. it is
503+
-- compiled to a forall)
504+
dependentDom :: Dom Type -> C Bool
505+
dependentDom dom = do
506+
d' <- compileDom dom
507+
pure $ case d' of
508+
DOType -> True
509+
DOTerm -> False
510+
DOInstance -> False
511+
DODropped -> False
502512

503513
compileLam :: Type -> ArgInfo -> Abs Term -> C (Hs.Exp ())
504514
compileLam ty argi abs = do

test/AllFailTests.agda

Lines changed: 0 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -41,7 +41,6 @@ import Fail.Issue125
4141
import Fail.Issue357a
4242
import Fail.Issue357b
4343
import Fail.DerivingParseFailure
44-
import Fail.ErasedPatternLambda
4544
import Fail.Issue306a
4645
import Fail.Issue306b
4746
import Fail.Issue306c

test/AllTests.agda

Lines changed: 2 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -87,6 +87,7 @@ import Issue309
8787
import Issue317
8888
import Issue353
8989
import RankNTypes
90+
import ErasedPatternLambda
9091
import CustomTuples
9192
import ProjectionLike
9293
import FunCon
@@ -186,6 +187,7 @@ import Issue309
186187
import Issue317
187188
import Issue353
188189
import RankNTypes
190+
import ErasedPatternLambda
189191
import CustomTuples
190192
import ProjectionLike
191193
import FunCon

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

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

53
Scope = List Bool

test/golden/AllTests.hs

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -82,6 +82,7 @@ import Issue309
8282
import Issue317
8383
import Issue353
8484
import RankNTypes
85+
import ErasedPatternLambda
8586
import CustomTuples
8687
import ProjectionLike
8788
import FunCon

test/golden/ErasedPatternLambda.err

Lines changed: 0 additions & 3 deletions
This file was deleted.

test/golden/ErasedPatternLambda.hs

Lines changed: 10 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,10 @@
1+
module ErasedPatternLambda where
2+
3+
data Telescope = ExtendTel Bool Telescope
4+
5+
caseTelBind :: Telescope -> (Bool -> Telescope -> d) -> d
6+
caseTelBind (ExtendTel a tel) f = f a tel
7+
8+
checkSubst :: Telescope -> Bool
9+
checkSubst t = caseTelBind t (\ ty rest -> True)
10+

test/golden/StreamFusion.hs

Lines changed: 7 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,7 @@
1+
module Cubical.StreamFusion where
2+
3+
data Stream a = (:>){shead :: a, stail :: Stream a}
4+
5+
smap :: (a -> b) -> Stream a -> Stream b
6+
smap f (x :> xs) = f x :> smap f xs
7+

0 commit comments

Comments
 (0)