mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-11-24 15:07:37 +03:00
[ base ] implement decEq
for SnocList
(#2630)
Co-authored-by: Guillaume Allais <guillaume.allais@ens-lyon.org>
This commit is contained in:
parent
d1ae9d9fd5
commit
c0153e72cd
@ -36,7 +36,7 @@ public export
|
||||
||| This is more strict than injectivity on each of arguments.
|
||||
||| For instance, list appending is injective on both arguments but is not biinjective.
|
||||
public export
|
||||
interface Biinjective (f : a -> b -> c) | f where
|
||||
interface Biinjective (0 f : a -> b -> c) | f where
|
||||
constructor MkBiinjective
|
||||
biinjective : {x, y : a} -> {v, w : b} -> f x v = f y w -> (x = y, v = w)
|
||||
|
||||
|
@ -21,6 +21,14 @@ public export
|
||||
asList : SnocList type -> List type
|
||||
asList = (reverse . cast)
|
||||
|
||||
export
|
||||
Uninhabited (Lin = x :< xs) where
|
||||
uninhabited Refl impossible
|
||||
|
||||
export
|
||||
Uninhabited (x :< xs = Lin) where
|
||||
uninhabited Refl impossible
|
||||
|
||||
||| True iff input is Lin
|
||||
public export
|
||||
isLin : SnocList a -> Bool
|
||||
@ -87,6 +95,11 @@ Alternative SnocList where
|
||||
empty = Lin
|
||||
xs <|> ys = xs ++ ys
|
||||
|
||||
-- Why can't we just use an implementation here?!
|
||||
export %hint
|
||||
SnocBiinjective : Biinjective (:<)
|
||||
SnocBiinjective = MkBiinjective $ \case Refl => (Refl, Refl)
|
||||
|
||||
||| Find the rightmost element of the snoc-list that satisfies the predicate.
|
||||
public export
|
||||
find : (a -> Bool) -> SnocList a -> Maybe a
|
||||
|
@ -7,6 +7,7 @@ import Data.Nat
|
||||
import Data.List
|
||||
import Data.List1
|
||||
import Data.List1.Properties
|
||||
import Data.SnocList
|
||||
import Data.These
|
||||
|
||||
import public Decidable.Equality.Core as Decidable.Equality
|
||||
@ -111,6 +112,17 @@ public export
|
||||
DecEq a => DecEq (List1 a) where
|
||||
decEq (x ::: xs) (y ::: ys) = decEqCong2 (decEq x y) (decEq xs ys)
|
||||
|
||||
--------------------------------------------------------------------------------
|
||||
-- SnocList
|
||||
--------------------------------------------------------------------------------
|
||||
|
||||
public export
|
||||
DecEq a => DecEq (SnocList a) where
|
||||
decEq Lin Lin = Yes Refl
|
||||
decEq (xs :< x) Lin = No absurd
|
||||
decEq Lin (xs :< x) = No absurd
|
||||
decEq (xs :< x) (ys :< y) = decEqCong2 (decEq xs ys) (decEq x y)
|
||||
|
||||
-- TODO: Other prelude data types
|
||||
|
||||
-- For the primitives, we have to cheat because we don't have access to their
|
||||
|
Loading…
Reference in New Issue
Block a user