Idris2/tests/idris2/total003/Total.idr
Edwin Brady 498421a236 All functions now need to be covering by default
This has caught a couple of things in the Idris 2 code base itself. Some
tests needed partial annotations too.
2020-05-24 19:58:20 +01:00

28 lines
844 B
Idris

%default partial
onlyOne : Int -> Int
onlyOne 1 = 2
covered : Int -> Int
covered 1 = 2
covered 2 = 3
covered _ = 94
data IntNat : Integer -> Nat -> Type where
IsZero : IntNat 0 Z
IsSuc : IntNat 1 (S k)
-- Only identified as covering if 'x' has multiplicity 0 because then it
-- doesn't show up in the case tree!
matchInt : (0 x : Integer) -> (n : Nat) -> IntNat x n -> String
matchInt 0 Z IsZero = "Zero"
matchInt 1 (S k) IsSuc = "Non Zero"
-- Should be identified as covering but isn't yet since the checker requires
-- a catch all case. This does at least test that the declared 'impossible'
-- case really is impossible; we can update it when the checker improves!
matchInt' : (x : Integer) -> (n : Nat) -> IntNat x n -> String
matchInt' 0 Z IsZero = "Zero"
matchInt' 1 (S k) IsSuc = "Non Zero"
matchInt' 0 (S k) x impossible