Skip to content

Commit

Permalink
Compile Erase to ()
Browse files Browse the repository at this point in the history
  • Loading branch information
jespercockx committed Dec 20, 2023
1 parent 3bc7d8e commit 84862e0
Show file tree
Hide file tree
Showing 8 changed files with 57 additions and 1 deletion.
1 change: 0 additions & 1 deletion lib/Haskell/Extra/Erase.agda
Original file line number Diff line number Diff line change
Expand Up @@ -13,7 +13,6 @@ module Haskell.Extra.Erase where
constructor Erased
field @0 get : a
open Erase public
{-# COMPILE AGDA2HS Erase #-}

infixr 4 ⟨_⟩_
record Σ0 (@0 a : Set) (b : @0 a Set) : Set where
Expand Down
4 changes: 4 additions & 0 deletions src/Agda2Hs/Compile/Term.hs
Original file line number Diff line number Diff line change
Expand Up @@ -55,6 +55,7 @@ isSpecialCon :: QName -> Maybe (ConHead -> ConInfo -> Elims -> C (Hs.Exp ()))
isSpecialCon = prettyShow >>> \case
"Haskell.Prim.Tuple._,_" -> Just tupleTerm
"Haskell.Prim.Tuple._×_×_._,_,_" -> Just tupleTerm
"Haskell.Extra.Erase.Erased" -> Just (\_ _ _ -> erasedTerm)
_ -> Nothing

tupleTerm :: ConHead -> ConInfo -> Elims -> C (Hs.Exp ())
Expand Down Expand Up @@ -158,6 +159,9 @@ sequ q (e:es) = do
vs -> return $ hsVar "_>>_" `eApp` vs
sequ q [] = return $ hsVar "_>>_"

erasedTerm :: C (Hs.Exp ())
erasedTerm = return $ Hs.Tuple () Hs.Boxed []

caseOf :: QName -> Elims -> C (Hs.Exp ())
caseOf _ es = compileElims es >>= \case
-- applied to pattern lambda
Expand Down
4 changes: 4 additions & 0 deletions src/Agda2Hs/Compile/Type.hs
Original file line number Diff line number Diff line change
Expand Up @@ -42,6 +42,7 @@ isSpecialType :: QName -> Maybe (QName -> Elims -> C (Hs.Type ()))
isSpecialType = prettyShow >>> \case
"Haskell.Prim.Tuple._×_" -> Just tupleType
"Haskell.Prim.Tuple._×_×_" -> Just tupleType
"Haskell.Extra.Erase.Erase" -> Just erasedType
_ -> Nothing

tupleType :: QName -> Elims -> C (Hs.Type ())
Expand All @@ -50,6 +51,9 @@ tupleType q es = do
ts <- mapM (compileType . unArg) as
return $ Hs.TyTuple () Hs.Boxed ts

erasedType :: QName -> Elims -> C (Hs.Type ())
erasedType _ _ = return $ Hs.TyTuple () Hs.Boxed []

-- | Add a class constraint to a Haskell type.
constrainType
:: Hs.Asst () -- ^ The class assertion.
Expand Down
2 changes: 2 additions & 0 deletions test/AllTests.agda
Original file line number Diff line number Diff line change
Expand Up @@ -63,6 +63,7 @@ import Issue210
import ModuleParameters
import ModuleParametersImports
import Coerce
import EraseType

{-# FOREIGN AGDA2HS
import Issue14
Expand Down Expand Up @@ -126,4 +127,5 @@ import Issue210
import ModuleParameters
import ModuleParametersImports
import Coerce
import EraseType
#-}
19 changes: 19 additions & 0 deletions test/EraseType.agda
Original file line number Diff line number Diff line change
@@ -0,0 +1,19 @@
module EraseType where

open import Haskell.Prelude
open import Haskell.Extra.Erase

testErase : Erase Int
testErase = Erased 42

{-# COMPILE AGDA2hs testErase #-}

testRezz : Rezz Int (get testErase)
testRezz = rezz 42

{-# COMPILE AGDA2HS testRezz #-}

testRezzErase : Rezz (Erase Int) testErase
testRezzErase = rezzErase

{-# COMPILE AGDA2HS testRezzErase #-}
1 change: 1 addition & 0 deletions test/golden/AllTests.hs
Original file line number Diff line number Diff line change
Expand Up @@ -61,4 +61,5 @@ import Issue210
import ModuleParameters
import ModuleParametersImports
import Coerce
import EraseType

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

import Haskell.Extra.Erase (rezzErase)

testRezz :: Int
testRezz = 42

testRezzErase :: ()
testRezzErase = rezzErase

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

rezzCong :: (a -> b) -> a -> b
rezzCong f x = f x

rezzCong2 :: (a -> b -> c) -> a -> b -> c
rezzCong2 f x y = f x y

rezzHead :: [a] -> a
rezzHead (x : xs) = x

rezzTail :: [a] -> [a]
rezzTail (x : xs) = xs

rezzErase :: ()
rezzErase = ()

0 comments on commit 84862e0

Please sign in to comment.