Merge pull request #1489 from buzden/some-uninhabiteds

[ base ] Some lacking implementations for `Uninhabited`
This commit is contained in:
Edwin Brady 2021-06-23 16:17:32 +01:00 committed by GitHub
commit c1057a19af
No known key found for this signature in database
GPG Key ID: 4AEE18F83AFDEB23
15 changed files with 148 additions and 3 deletions

View File

@ -204,6 +204,10 @@ Uninhabited (Data.Colist.InBounds k []) where
uninhabited InFirst impossible
uninhabited (InLater _) impossible
export
Uninhabited (Colist.InBounds k xs) => Uninhabited (Colist.InBounds (S k) (x::xs)) where
uninhabited (InLater y) = uninhabited y
||| Decide whether `k` is a valid index into Colist `xs`
public export
inBounds : (k : Nat) -> (xs : Colist a) -> Dec (InBounds k xs)

View File

@ -40,6 +40,10 @@ export
Uninhabited (FS k = FZ) where
uninhabited Refl impossible
export
Uninhabited (n = m) => Uninhabited (FS n = FS m) where
uninhabited Refl @{nm} = uninhabited Refl @{nm}
export
fsInjective : FS m = FS n -> m = n
fsInjective Refl = Refl

View File

@ -45,8 +45,12 @@ data InBounds : (k : Nat) -> (xs : List a) -> Type where
public export
Uninhabited (InBounds k []) where
uninhabited InFirst impossible
uninhabited (InLater _) impossible
uninhabited InFirst impossible
uninhabited (InLater _) impossible
export
Uninhabited (InBounds k xs) => Uninhabited (InBounds (S k) (x::xs)) where
uninhabited (InLater y) = uninhabited y
||| Decide whether `k` is a valid index into `xs`
public export
@ -678,6 +682,11 @@ export
Uninhabited (x :: xs = []) where
uninhabited Refl impossible
export
{0 xs : List a} -> Either (Uninhabited $ x === y) (Uninhabited $ xs === ys) => Uninhabited (x::xs = y::ys) where
uninhabited @{Left z} Refl = uninhabited @{z} Refl
uninhabited @{Right z} Refl = uninhabited @{z} Refl
||| (::) is injective
export
consInjective : forall x, xs, y, ys .

View File

@ -42,6 +42,11 @@ Uninhabited (Elem {a} x []) where
uninhabited Here impossible
uninhabited (There p) impossible
export
Uninhabited (x = z) => Uninhabited (Elem z xs) => Uninhabited (Elem z $ x::xs) where
uninhabited Here @{xz} = uninhabited Refl @{xz}
uninhabited (There y) = uninhabited y
||| An item not in the head and not in the tail is not in the list at all.
export
neitherHereNorThere : Not (x = y) -> Not (Elem x xs) -> Not (Elem x (y :: xs))

View File

@ -27,6 +27,10 @@ namespace Any
uninhabited (Here _) impossible
uninhabited (There _) impossible
export
{0 p : a -> Type} -> Uninhabited (p x) => Uninhabited (Any p xs) => Uninhabited (Any p $ x::xs) where
uninhabited (Here y) = uninhabited y
uninhabited (There y) = uninhabited y
||| Modify the property given a pointwise function
export
@ -67,6 +71,9 @@ namespace All
Nil : All p Nil
(::) : {0 xs : List a} -> p x -> All p xs -> All p (x :: xs)
Either (Uninhabited $ p x) (Uninhabited $ All p xs) => Uninhabited (All p $ x::xs) where
uninhabited @{Left _} (px::pxs) = uninhabited px
uninhabited @{Right _} (px::pxs) = uninhabited pxs
||| Modify the property given a pointwise function
export

View File

@ -201,3 +201,7 @@ Zippable List1 where
unzipWith3' (x :: xs) = let (b, c, d) = f x
(bs, cs, ds) = unzipWith3' xs in
(b :: bs, c :: cs, d :: ds)
export
Uninhabited a => Uninhabited (List1 a) where
uninhabited (hd ::: _) = uninhabited hd

View File

@ -10,6 +10,10 @@ export
Uninhabited (S n = Z) where
uninhabited Refl impossible
export
Uninhabited (a = b) => Uninhabited (S a = S b) where
uninhabited Refl @{ab} = uninhabited @{ab} Refl
public export
isZero : Nat -> Bool
isZero Z = True

View File

@ -15,10 +15,23 @@ data Elem : a -> Vect k a -> Type where
Here : Elem x (x::xs)
There : (later : Elem x xs) -> Elem x (y::xs)
export
Uninhabited (Here = There e) where
uninhabited Refl impossible
export
Uninhabited (There e = Here) where
uninhabited Refl impossible
export
Uninhabited (Elem x []) where
uninhabited Here impossible
export
Uninhabited (x = z) => Uninhabited (Elem z xs) => Uninhabited (Elem z $ x::xs) where
uninhabited Here @{xz} = uninhabited Refl @{xz}
uninhabited (There y) = uninhabited y
||| An item not in the head and not in the tail is not in the Vect at all
export
neitherHereNorThere : Not (x = y) -> Not (Elem x xs) -> Not (Elem x (y :: xs))

View File

@ -19,6 +19,11 @@ implementation Uninhabited (Any p Nil) where
uninhabited (Here _) impossible
uninhabited (There _) impossible
export
implementation {0 p : a -> Type} -> Uninhabited (p x) => Uninhabited (Any p xs) => Uninhabited (Any p $ x::xs) where
uninhabited (Here y) = uninhabited y
uninhabited (There y) = uninhabited y
||| Eliminator for `Any`
public export
anyElim : {0 xs : Vect n a} -> {0 p : a -> Type} -> (Any p xs -> b) -> (p x -> b) -> Any p (x :: xs) -> b
@ -80,3 +85,8 @@ all d (x::xs) with (d x)
case all d xs of
Yes prf' => Yes (prf :: prf')
No prf' => No (notAllThere prf')
export
Either (Uninhabited $ p x) (Uninhabited $ All p xs) => Uninhabited (All p $ x::xs) where
uninhabited @{Left _} (px::pxs) = uninhabited px
uninhabited @{Right _} (px::pxs) = uninhabited pxs

View File

@ -257,6 +257,16 @@ data Either : (a : Type) -> (b : Type) -> Type where
export Uninhabited (Left p === Right q) where uninhabited eq impossible
export Uninhabited (Right p === Left q) where uninhabited eq impossible
export
Either (Uninhabited a) (Uninhabited b) => Uninhabited (a, b) where
uninhabited (x, _) @{Left _} = uninhabited x
uninhabited (_, y) @{Right _} = uninhabited y
export
Uninhabited a => Uninhabited b => Uninhabited (Either a b) where
uninhabited (Left l) = uninhabited l
uninhabited (Right r) = uninhabited r
||| Simply-typed eliminator for Either.
||| @ f the action to take on Left
||| @ g the action to take on Right

View File

@ -90,7 +90,7 @@ idrisTestsInterface = MkTestPool "Interface" [] Nothing
"interface013", "interface014", "interface015", "interface016",
"interface017", "interface018", "interface019", "interface020",
"interface021", "interface022", "interface023", "interface024",
"interface025"]
"interface025", "interface026"]
idrisTestsLinear : TestPool
idrisTestsLinear = MkTestPool "Quantities" [] Nothing

View File

@ -0,0 +1,49 @@
module UninhabitedRec
import Data.Nat
import Data.List.Elem
ff : Uninhabited (a, b) => Int
ff = 4
callFGood : Int
callFGood = ff {b = (Left 4 = Right 4)} {a = 5 = 5}
------------------
data Lookup : a -> List (a, b) -> Type where
Here : (y : b) -> Lookup x $ (x, y)::xys
There : (0 _ : Uninhabited $ x === z) => Lookup z xys -> Lookup z $ (x, y)::xys
fff : (xs : List (Nat, String)) -> (n : Nat) -> (0 _ : Lookup n xs) => String
xxs : List (Nat, String)
xxs = [(1, "one"), (2, "two"), (4, "four")]
lkup1Good : String
lkup1Good = fff xxs 1
lkup2Good : String
lkup2Good = fff xxs 2
lkup3Bad : String
lkup3Bad = fff xxs 3
------------------
data Uniq : Type -> Type
toList : Uniq a -> List a
data Uniq : Type -> Type where
Nil : Uniq a
(::) : (x : a) -> (xs : Uniq a) -> Uninhabited (Elem x $ toList xs) => Uniq a
toList [] = []
toList (x::xs) = x :: toList xs
uniqGood : Uniq Nat
uniqGood = [1, 2, 3]
uniqBad : Uniq Nat
uniqBad = [1, 2, 1]

View File

@ -0,0 +1,22 @@
1/1: Building UninhabitedRec (UninhabitedRec.idr)
Error: While processing right hand side of lkup3Bad. Can't find an implementation for Lookup 3 [(1, "one"), (2, "two"), (4, "four")].
UninhabitedRec:30:12--30:21
26 | lkup2Good : String
27 | lkup2Good = fff xxs 2
28 |
29 | lkup3Bad : String
30 | lkup3Bad = fff xxs 3
^^^^^^^^^
Error: While processing right hand side of uniqBad. Can't find an implementation for Uninhabited (Elem 1 (toList [2, 1])).
UninhabitedRec:49:11--49:12
45 | uniqGood : Uniq Nat
46 | uniqGood = [1, 2, 3]
47 |
48 | uniqBad : Uniq Nat
49 | uniqBad = [1, 2, 1]
^
UninhabitedRec> Bye for now!

View File

@ -0,0 +1 @@
:q

View File

@ -0,0 +1,3 @@
$1 --no-banner --no-color --console-width 0 UninhabitedRec.idr < input
rm -rf build