Idris2/tests/idris2/positivity004/Issue1771-2.idr

16 lines
252 B
Idris

%default total
data F : Type where
MkFix : ((0 g : Type -> Type) -> g === Not -> g F) -> F
yesF : Not F -> F
yesF nf = MkFix (\ g, Refl => nf)
notF : Not F
notF (MkFix f) =
let g = f Not Refl in
g (yesF g)
argh : Void
argh = notF (yesF notF)