Skip to content
Merged
Show file tree
Hide file tree
Changes from all commits
Commits
File filter

Filter by extension

Filter by extension

Conversations
Failed to load comments.
Loading
Jump to
Jump to file
Failed to load files.
Loading
Diff view
Diff view
8 changes: 8 additions & 0 deletions src/Agda2Hs/Compile/Function.hs
Original file line number Diff line number Diff line change
Expand Up @@ -332,6 +332,14 @@ compilePat ty (ConP ch i ps) = do
-- literal patterns
compilePat ty (LitP _ l) = compileLitPat l

-- "Inferred" dot patterns that the programmer has explicitly named are
-- compiled to variable patterns using that given name, wildcards otherwise
compilePat _ (DotP (PatternInfo (PatOVar n) _) _) = do
let n' = hsName $ prettyShow n
checkValidVarName n'
return $ Hs.PVar () n'
compilePat _ (DotP (PatternInfo PatODot _) _) =
return $ Hs.PWildCard ()

-- nothing else is supported
compilePat _ p = agda2hsErrorM $ "bad pattern:" <?> prettyTCM p
Expand Down
6 changes: 3 additions & 3 deletions src/Agda2Hs/Compile/Term.hs
Original file line number Diff line number Diff line change
Expand Up @@ -576,9 +576,7 @@ compileInlineFunctionApp f ty args = do
extractPatName _ = __IMPOSSIBLE__

extractName :: NamedArg DeBruijnPattern -> ArgName
extractName (unArg -> np)
| Just n <- nameOf np = rangedThing (woThing n)
| otherwise = extractPatName (namedThing np)
extractName p = fromMaybe (extractPatName $ namedArg p) (extractMaybeName p)

etaExpand :: NAPs -> Type -> [Term] -> C Term
etaExpand [] ty args = do
Expand All @@ -594,6 +592,8 @@ compileInlineFunctionApp f ty args = do
let ai = domInfo dom
Lam ai . mkAbs (extractName p) <$> etaExpand ps (absBody cod) (raise 1 args ++ [ var 0 ])

extractMaybeName :: NamedArg a -> Maybe ArgName
extractMaybeName p = rangedThing . woThing <$> nameOf (unArg p)

compileApp :: Hs.Exp () -> Type -> [Term] -> C (Hs.Exp ())
compileApp = compileApp' . eApp
Expand Down
6 changes: 6 additions & 0 deletions test/AllTests.agda
Original file line number Diff line number Diff line change
Expand Up @@ -103,6 +103,9 @@ import Issue408
import CompileTo
import RuntimeCast
import Issue306
import RelevantDotPattern1
import RelevantDotPattern2
import RelevantDotPattern3

{-# FOREIGN AGDA2HS
import Issue14
Expand Down Expand Up @@ -203,4 +206,7 @@ import Issue408
import CompileTo
import RuntimeCast
import Issue306
import RelevantDotPattern1
import RelevantDotPattern2
import RelevantDotPattern3
#-}
23 changes: 23 additions & 0 deletions test/RelevantDotPattern1.agda
Original file line number Diff line number Diff line change
@@ -0,0 +1,23 @@
-- Derived from "ErasedPatternLambda"
open import Haskell.Prelude

Scope = List Bool

data Telescope (@0 α : Scope) : @0 Scope → Type where
ExtendTel : ∀ {@0 x β} → Bool → Telescope (x ∷ α) β → Telescope α (x ∷ β)
{-# COMPILE AGDA2HS Telescope #-}

caseTelBind : ∀ {@0 x α β} (tel : Telescope α (x ∷ β))
→ ((a : Bool) (rest : Telescope (x ∷ α) β) → @0 tel ≡ ExtendTel a rest → d)
→ d
caseTelBind (ExtendTel a tel) f = f a tel refl

{-# COMPILE AGDA2HS caseTelBind #-}

checkSubst' : ∀ {@0 x α β} (t : Telescope α (x ∷ β)) (ty : Bool) (rest : Telescope (x ∷ α) β) → @0 t ≡ ExtendTel ty rest → Bool
checkSubst' t ty rest refl = True
{-# COMPILE AGDA2HS checkSubst' #-}

checkSubst : ∀ {@0 x α β} (t : Telescope α (x ∷ β)) → Bool
checkSubst t = caseTelBind t (checkSubst' t)
{-# COMPILE AGDA2HS checkSubst #-}
24 changes: 24 additions & 0 deletions test/RelevantDotPattern2.agda
Original file line number Diff line number Diff line change
@@ -0,0 +1,24 @@
-- Derived from "ErasedPatternLambda"
open import Haskell.Prelude

Scope = List Bool

data Telescope (@0 α : Scope) : @0 Scope → Type where
ExtendTel : ∀ {@0 x β} → Bool → Telescope (x ∷ α) β → Telescope α (x ∷ β)
{-# COMPILE AGDA2HS Telescope #-}

caseTelBind : ∀ {@0 x α β} (tel : Telescope α (x ∷ β))
→ ((a : Bool) (rest : Telescope (x ∷ α) β) → @0 tel ≡ ExtendTel a rest → d)
→ d
caseTelBind (ExtendTel a tel) f = f a tel refl

{-# COMPILE AGDA2HS caseTelBind #-}

module _ {@0 x α β} (t : Telescope α (x ∷ β)) where
checkSubst' : ∀ (ty : Bool) (rest : Telescope (x ∷ α) β) → @0 t ≡ ExtendTel ty rest → Bool
checkSubst' ty rest refl = True
{-# COMPILE AGDA2HS checkSubst' #-}

checkSubst : Bool
checkSubst = caseTelBind t checkSubst'
{-# COMPILE AGDA2HS checkSubst #-}
23 changes: 23 additions & 0 deletions test/RelevantDotPattern3.agda
Original file line number Diff line number Diff line change
@@ -0,0 +1,23 @@
-- Derived from "ErasedPatternLambda"
open import Haskell.Prelude

Scope = List Bool

data Telescope (@0 α : Scope) : @0 Scope → Type where
ExtendTel : ∀ {@0 x β} → Bool → Telescope (x ∷ α) β → Telescope α (x ∷ β)
{-# COMPILE AGDA2HS Telescope #-}

caseTelBind : ∀ {@0 x α β} (tel : Telescope α (x ∷ β))
→ ((a : Bool) (rest : Telescope (x ∷ α) β) → @0 tel ≡ ExtendTel a rest → d)
→ d
caseTelBind (ExtendTel a tel) f = f a tel refl

{-# COMPILE AGDA2HS caseTelBind #-}

checkSubst' : ∀ {@0 x α β} (t : Telescope α (x ∷ β)) (ty : Bool) (rest : Telescope (x ∷ α) β) → @0 t ≡ ExtendTel ty rest → Bool
checkSubst' (ExtendTel ty' rest') ty rest refl = True
{-# COMPILE AGDA2HS checkSubst' #-}

checkSubst : ∀ {@0 x α β} (t : Telescope α (x ∷ β)) → Bool
checkSubst t = caseTelBind t (checkSubst' t)
{-# COMPILE AGDA2HS checkSubst #-}
3 changes: 3 additions & 0 deletions test/golden/AllTests.hs
Original file line number Diff line number Diff line change
Expand Up @@ -98,4 +98,7 @@ import Issue408
import CompileTo
import RuntimeCast
import Issue306
import RelevantDotPattern1
import RelevantDotPattern2
import RelevantDotPattern3

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

data Telescope = ExtendTel Bool Telescope

caseTelBind :: Telescope -> (Bool -> Telescope -> d) -> d
caseTelBind (ExtendTel a tel) f = f a tel

checkSubst' :: Telescope -> Bool -> Telescope -> Bool
checkSubst' t ty rest = True

checkSubst :: Telescope -> Bool
checkSubst t = caseTelBind t (checkSubst' t)

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

data Telescope = ExtendTel Bool Telescope

caseTelBind :: Telescope -> (Bool -> Telescope -> d) -> d
caseTelBind (ExtendTel a tel) f = f a tel

checkSubst' :: Telescope -> Bool -> Telescope -> Bool
checkSubst' t ty rest = True

checkSubst :: Telescope -> Bool
checkSubst t = caseTelBind t (checkSubst' t)

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

data Telescope = ExtendTel Bool Telescope

caseTelBind :: Telescope -> (Bool -> Telescope -> d) -> d
caseTelBind (ExtendTel a tel) f = f a tel

checkSubst' :: Telescope -> Bool -> Telescope -> Bool
checkSubst' (ExtendTel ty' rest') ty rest = True

checkSubst :: Telescope -> Bool
checkSubst t = caseTelBind t (checkSubst' t)

Loading