Idris2/tests/idris2/record003/Record.idr
Edwin Brady a972778eab Add test script
They don't all pass yet, for minor reasons. Coming shortly...
Unfortunately the startup overhead for chez is really noticeable here!
2020-05-19 18:25:18 +01:00

34 lines
629 B
Idris

||| Test extended record syntax
data NotZero : Nat -> Type where
Is : {n : Nat} -> NotZero (S n)
record Positive (n : Nat) {auto 0 pos : NotZero n} where
constructor Check
a : Positive 1
a = Check
b : Positive 2
b = Check
{- Will not type-check
c : Positive 0
-}
data KindOfString = ASCII | UTF
UTForASCII : KindOfString -> Type
UTForASCII UTF = String
UTForASCII ASCII = List Char --- Just to demonstrate
record FlexiString { default UTF k : KindOfString} where
constructor MkFlexiString
val : UTForASCII k
c : FlexiString
c = MkFlexiString "hello"
d : FlexiString {k = ASCII}
d = MkFlexiString ['h', 'i']