Idris2/tests/idris2/positivity004/Issue1771-3.idr
Guillaume ALLAIS d0c0698c45 [ re #1771 ] Check parameters for positive uses
It's fine to allow positive occurences in (strictly positive)
parameters but we do need to check that these occurences are
strictly positive!
2021-07-23 13:30:24 +01:00

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)