mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-12-19 09:12:34 +03:00
18 lines
365 B
Idris
18 lines
365 B
Idris
%default total
|
|
|
|
data Desc : Type where
|
|
Prod : (d, e : Desc) -> Desc
|
|
Rec : Desc
|
|
|
|
Elem : (d : Desc) -> Type -> Type
|
|
Elem (Prod d e) x = (Elem d x, Elem e x)
|
|
Elem Rec x = x
|
|
|
|
failing "Mu is not total, not strictly positive"
|
|
|
|
data Mu : Desc -> Type where
|
|
MkMu : Elem d (Mu d) -> Mu d
|
|
|
|
data Mu : Desc -> Type where
|
|
MkMu : Elem d (assert_total (Mu d)) -> Mu d
|