mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-12-16 07:34:45 +03:00
d0c0698c45
It's fine to allow positive occurences in (strictly positive) parameters but we do need to check that these occurences are strictly positive!
20 lines
293 B
Idris
20 lines
293 B
Idris
%default total
|
|
|
|
data Wrap : Type -> Type where
|
|
MkWrap : a -> Wrap a
|
|
|
|
unwrap : Wrap a -> a
|
|
unwrap (MkWrap v) = v
|
|
|
|
data F : Type where
|
|
MkF : Wrap (Not F) -> F
|
|
|
|
yesF : Not F -> F
|
|
yesF = MkF . MkWrap
|
|
|
|
notF : Not F
|
|
notF (MkF f) = unwrap f (yesF $ unwrap f)
|
|
|
|
argh : Void
|
|
argh = notF (yesF notF)
|