Idris2/tests/idris2/positivity004/Issue1771-1.idr
Guillaume ALLAIS 230f42b697 [ re #1771 ] Do not use Erased to go under binders
In the `MkFix : f (Fix f) -> Fix f` example, using `Erased` for `f`
makes the type reduce to `[__] (Fix [__]) -> Fix [__]` and because
`[__] e1 ... en` computes to `[__]`, we end up with `[__] -> Fix [__]`
which does not reference `Fix` anymore.
2021-07-23 13:30:24 +01:00

17 lines
224 B
Idris

%default total
data Fix : (Type -> Type) -> Type where
MkFix : f (Fix f) -> Fix f
F : Type
F = Fix Not
yesF : Not F -> F
yesF nf = MkFix nf
notF : Not F
notF (MkFix f) = f (yesF f)
argh : Void
argh = notF (yesF notF)