Idris2/tests/idris2/linear002/Door.idr
Raoul Hidalgo Charman 3e85e23638
Namespace checks for resugaring (#1161)
This avoids resugaring to the wrong type when there are user defined
symbols which conflicts with builtins such as Pair.

Changed the test linear002 which was relying on this behaviour for a
user defined Unit.

Fixes #634.
2021-03-15 13:34:54 +00:00

43 lines
1.1 KiB
Idris

import Stuff
data Usage = Once | Many
data Unrestricted : Type -> Type where
Un : (x : a) -> Unrestricted a
data Use : Usage -> (Type -> Type) -> Type -> Type where
MkUse : m a -> Use l m a
contType : (Type -> Type) -> Usage -> Usage -> Type -> Type -> Type
contType m Once q a b = ((1 x : a) -> Use q m b)
contType m Many q a b = ((x : a) -> Use q m b)
(>>=) : (1 f : Use p m a) -> (1 k : contType m p q a b) -> Use q m b
pure : a -> Use u m a
One : (Type -> Type) -> Type -> Type
One = Use Once
Any : (Type -> Type) -> Type -> Type
Any = Use Many
data DoorState = Open | Closed
data Door : DoorState -> Type where
MkDoor : (isOpen : Bool) -> Door doorstate
newDoor : One m (Door Closed)
knock : (1 d : Door Closed) -> One m (Door Closed)
openDoor : (1 d : Door Closed) -> One m (Either (Door Closed) (Door Open))
closeDoor : (1 d : Door Open) -> One m (Door Closed)
deleteDoor : (1 d : Door Closed) -> Any m Unit
doorProg : Any m Unit
doorProg
= do d <- newDoor
Right d' <- openDoor d
| Left d => deleteDoor d
?foo