Skip to content

Commit 7cae101

Browse files
committed
Add assert function to Haskell.Control.Exception
1 parent 1bf9d3a commit 7cae101

File tree

6 files changed

+33
-0
lines changed

6 files changed

+33
-0
lines changed

lib/Haskell/Control/Exception.agda

Lines changed: 11 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,11 @@
1+
module Haskell.Control.Exception where
2+
3+
open import Haskell.Prim
4+
5+
open import Haskell.Extra.Dec
6+
open import Haskell.Extra.Refinement
7+
8+
assert : (@0 b : Set ℓ) {{Dec b}} (@0 {{b}} a) a
9+
assert _ {{True ⟨ p ⟩}} x = x {{p}}
10+
assert _ {{False ⟨ _ ⟩}} x = oops
11+
where postulate oops : _

src/Agda2Hs/Compile/Name.hs

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -67,6 +67,7 @@ defaultSpecialRules :: SpecialRules
6767
defaultSpecialRules = Map.fromList
6868
[ "Agda.Builtin.Nat.Nat" `to` "Natural" `importing` Just "Numeric.Natural"
6969
, "Haskell.Control.Monad.guard" `to` "guard" `importing` Just "Control.Monad"
70+
, "Haskell.Control.Exception.assert" `to` "assert" `importing` Just "Control.Exception"
7071
, "Haskell.Prelude.coerce" `to` "unsafeCoerce" `importing` Just "Unsafe.Coerce"
7172
, "Agda.Builtin.Int.Int" `to` "Integer" `importing` Nothing
7273
, "Agda.Builtin.Word.Word64" `to` "Word" `importing` Nothing

test/AllTests.agda

Lines changed: 2 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -93,6 +93,7 @@ import ProjectionLike
9393
import FunCon
9494
import Issue308
9595
import Issue324
96+
import Assert
9697

9798
{-# FOREIGN AGDA2HS
9899
import Issue14
@@ -183,4 +184,5 @@ import ProjectionLike
183184
import FunCon
184185
import Issue308
185186
import Issue324
187+
import Assert
186188
#-}

test/Assert.agda

Lines changed: 10 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,10 @@
1+
2+
open import Haskell.Prelude
3+
open import Haskell.Control.Exception
4+
open import Haskell.Law.Ord
5+
open import Haskell.Extra.Dec
6+
7+
subtractChecked : Nat Nat Nat
8+
subtractChecked x y = assert (IsFalse (x < y)) (x - y)
9+
10+
{-# COMPILE AGDA2HS subtractChecked #-}

test/golden/AllTests.hs

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -88,4 +88,5 @@ import ProjectionLike
8888
import FunCon
8989
import Issue308
9090
import Issue324
91+
import Assert
9192

test/golden/Assert.hs

Lines changed: 8 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,8 @@
1+
module Assert where
2+
3+
import Control.Exception (assert)
4+
import Numeric.Natural (Natural)
5+
6+
subtractChecked :: Natural -> Natural -> Natural
7+
subtractChecked x y = assert (not (x < y)) (x - y)
8+

0 commit comments

Comments
 (0)