[ linear ] These seem useful (#2316)

This commit is contained in:
G. Allais 2022-02-11 09:28:15 +00:00 committed by GitHub
parent bec4a0a88e
commit 33d99adb53
No known key found for this signature in database
GPG Key ID: 4AEE18F83AFDEB23
5 changed files with 64 additions and 15 deletions

View File

@ -17,6 +17,22 @@ Consumable Void where consume v impossible
export
Consumable () where consume u = u
||| Unions can be consumed by pattern matching
export
Consumable Bool where
consume True = ()
consume False = ()
export
Consumable (!* a) where
consume (MkBang v) = ()
||| We can cheat to make built-in types consumable
export
Consumable Int where
consume = believe_me (\ 0 i : Int => ())
-- But crucially we don't have `Consumable World` or `Consumable Handle`.
infixr 5 `seq`
@ -27,29 +43,27 @@ export
seq : Consumable a => a -@ b -@ b
a `seq` b = let 1 () = consume a in b
||| Unions can be consumed by pattern matching
export
Consumable Bool where
consume True = ()
consume False = ()
||| We can cheat to make built-in types consumable
export
Consumable Int where
consume = believe_me (\ 0 i : Int => ())
public export
interface Duplicable a where
duplicate : (1 v : a) -> 2 `Copies` v
export
Duplicable Void where
duplicate Void impossible
duplicate v impossible
export
Duplicable () where
duplicate () = [(), ()]
export
Duplicable Bool where
duplicate True = [True, True]
duplicate False = [False, False]
export
Duplicable (!* a) where
duplicate (MkBang v) = [MkBang v, MkBang v]
||| Comonoid is the dual of Monoid, it can consume a value linearly and duplicate a value linearly
||| `comult` returns a pair instead of 2 copies, because we do not guarantee that the two values
||| are identical, unlike with `duplicate`. For example if we build a comonoid out of a group, with
@ -64,7 +78,7 @@ interface Comonoid a where
||| If a value is consumable and duplicable we can make an instance of Comonoid for it
export
Consumable a => Duplicable a => Comonoid a where
[AutoComonoid] Consumable a => Duplicable a => Comonoid a where
counit = consume
comult x = pair (duplicate x)
@ -72,4 +86,3 @@ export
Comonoid (!* a) where
counit (MkBang _) = ()
comult (MkBang v) = MkBang v # MkBang v

View File

@ -1,5 +1,7 @@
module Data.Linear.LVect
import Data.Fin
import Data.Linear.Bifunctor
import Data.Linear.Interface
import Data.Linear.Notation
@ -15,6 +17,11 @@ data LVect : Nat -> Type -> Type where
%name LVect xs, ys, zs, ws
export
lookup : Fin (S n) -@ LVect (S n) a -@ LPair a (LVect n a)
lookup FZ (v :: vs) = (v # vs)
lookup (FS k) (v :: vs@(_ :: _)) = bimap id (v ::) (lookup k vs)
export
(<$>) : (f : a -@ b) -> LVect n a -@ LVect n b
f <$> [] = []
@ -50,6 +57,20 @@ export
[] ++ ys = ys
(x :: xs) ++ ys = x :: (xs ++ ys)
export
lfoldr : (0 p : Nat -> Type) -> (forall n. a -@ p n -@ p (S n)) -> p Z -@ LVect n a -@ p n
lfoldr p c n [] = n
lfoldr p c n (x :: xs) = c x (lfoldr p c n xs)
export
lfoldl : (0 p : Nat -> Type) -> (forall n. a -@ p n -@ p (S n)) -> p Z -@ LVect n a -@ p n
lfoldl p c n [] = n
lfoldl p c n (x :: xs) = lfoldl (p . S) c (c x n) xs
export
reverse : LVect m a -@ LVect m a
reverse = lfoldl (\ m => LVect m a) (::) []
export
Consumable a => Consumable (LVect n a) where
consume [] = ()

View File

@ -0,0 +1,10 @@
module Data.Linear.List.LQuantifiers
import Data.Linear.Notation
%default total
public export
data LAll : (p : a -> Type) -> List a -> Type where
Nil : LAll p []
(::) : p x -@ LAll p xs -@ LAll p (x :: xs)

View File

@ -8,6 +8,11 @@ public export
(-@) : Type -> Type -> Type
a -@ b = (1 _ : a) -> b
||| Linear identity function
public export
id : a -@ a
id x = x
||| Linear function composition
public export
(.) : (b -@ c) -@ (a -@ b) -@ (a -@ c)

View File

@ -9,8 +9,8 @@ modules = Data.Linear,
Data.Linear.Interface,
Data.Linear.LEither,
Data.Linear.LList,
Data.Linear.List.LQuantifiers,
Data.Linear.LMaybe,
Data.Linear.LNat,
Data.Linear.LVect,
Data.Linear.Notation