mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-11-28 11:05:17 +03:00
Vect reasoning library (#1439)
When working on Frex I needed a whole bunch of lemmata to do with Data.Vect. I hope it will be useful for others.
This commit is contained in:
parent
31a5175a02
commit
823230b77c
11
libs/contrib/Data/Vect/Extra.idr
Normal file
11
libs/contrib/Data/Vect/Extra.idr
Normal file
@ -0,0 +1,11 @@
|
|||||||
|
||| Additional functions about vectors
|
||||||
|
module Data.Vect.Extra
|
||||||
|
|
||||||
|
import Data.Vect
|
||||||
|
import Data.Fin
|
||||||
|
|
||||||
|
||| Version of `map` with access to the current position
|
||||||
|
public export
|
||||||
|
mapWithPos : (f : Fin n -> a -> b) -> Vect n a -> Vect n b
|
||||||
|
mapWithPos f [] = []
|
||||||
|
mapWithPos f (x :: xs) = f 0 x :: mapWithPos (f . FS) xs
|
7
libs/contrib/Data/Vect/Properties.idr
Normal file
7
libs/contrib/Data/Vect/Properties.idr
Normal file
@ -0,0 +1,7 @@
|
|||||||
|
||| Additional properties and lemmata to do with Vect
|
||||||
|
module Data.Vect.Properties
|
||||||
|
|
||||||
|
import public Data.Vect.Properties.Tabulate
|
||||||
|
import public Data.Vect.Properties.Index
|
||||||
|
import public Data.Vect.Properties.Foldr
|
||||||
|
import public Data.Vect.Properties.Map
|
82
libs/contrib/Data/Vect/Properties/Foldr.idr
Normal file
82
libs/contrib/Data/Vect/Properties/Foldr.idr
Normal file
@ -0,0 +1,82 @@
|
|||||||
|
|||
|
||||||
|
|||
|
||||||
|
||| foldr is the unique solution to the equation:
|
||||||
|
|||
|
||||||
|
||| h f e [] = e
|
||||||
|
||| h f e (x :: xs) = x `h` (foldr f e xs)
|
||||||
|
|||
|
||||||
|
||| (This fact is called 'the universal property of foldr'.)
|
||||||
|
|||
|
||||||
|
||| Since the prelude defines foldr tail-recursively, this fact isn't immediate
|
||||||
|
||| and we need some lemmata to prove it.
|
||||||
|
module Data.Vect.Properties.Foldr
|
||||||
|
|
||||||
|
import Data.Vect
|
||||||
|
import Data.Vect.Elem
|
||||||
|
import Data.Fin
|
||||||
|
|
||||||
|
import Syntax.PreorderReasoning
|
||||||
|
|
||||||
|
||| A function H : forall n. Vect n A -> B preserving the structure of vectors over A
|
||||||
|
public export
|
||||||
|
record VectHomomorphismProperty {0 A, B : Type} (F : A -> B -> B) (E : B) (H : forall n . Vect n A -> B) where
|
||||||
|
constructor ShowVectHomomorphismProperty
|
||||||
|
nil : H [] = E
|
||||||
|
cons : {0 n : Nat} -> (x : A) -> (xs : Vect n A) -> H (x :: xs) = x `F` (H xs)
|
||||||
|
|
||||||
|
||| There is an extensionally unique function preserving the vector structure
|
||||||
|
export
|
||||||
|
nilConsInitiality :
|
||||||
|
(f : a -> b -> b) -> (e : b)
|
||||||
|
-> (h1, h2 : forall n . Vect n a -> b)
|
||||||
|
-> (prf1 : VectHomomorphismProperty f e h1)
|
||||||
|
-> (prf2 : VectHomomorphismProperty f e h2)
|
||||||
|
-> (xs : Vect n a) -> h1 xs = h2 xs
|
||||||
|
nilConsInitiality f e h1 h2 prf1 prf2 [] = Calc $
|
||||||
|
|~ h1 []
|
||||||
|
~~ e ...(prf1.nil)
|
||||||
|
~~ h2 [] ...(sym prf2.nil)
|
||||||
|
|
||||||
|
nilConsInitiality f e h1 h2 prf1 prf2 (x :: xs) = Calc $
|
||||||
|
|~ h1 (x :: xs)
|
||||||
|
~~ (x `f` (h1 xs)) ...(prf1.cons _ _)
|
||||||
|
~~ (x `f` (h2 xs)) ...(cong (x `f`) $ nilConsInitiality f e h1 h2 prf1 prf2 xs)
|
||||||
|
~~ h2 (x :: xs) ...(sym $ prf2.cons _ _)
|
||||||
|
|
||||||
|
||| extensionality is a congruence with respect to Data.Vect.foldrImpl
|
||||||
|
foldrImplExtensional :
|
||||||
|
(f : a -> b -> b) -> (e : b)
|
||||||
|
-> (go1, go2 : b -> b)
|
||||||
|
-> ((y : b) -> go1 y = go2 y)
|
||||||
|
-> (xs : Vect n a)
|
||||||
|
-> foldrImpl f e go1 xs = foldrImpl f e go2 xs
|
||||||
|
foldrImplExtensional f e go1 go2 ext [] = ext e
|
||||||
|
foldrImplExtensional f e go1 go2 ext (x :: xs) =
|
||||||
|
foldrImplExtensional f e _ _
|
||||||
|
(\y => ext (f x y))
|
||||||
|
xs
|
||||||
|
|
||||||
|
||| foldrImpl f e x : (b -> -) -> - is natural
|
||||||
|
foldrImplNaturality : (f : a -> b -> b) -> (e : b) -> (xs : Vect n a) -> (go1, go2 : b -> b)
|
||||||
|
-> foldrImpl f e (go1 . go2) xs = go1 (foldrImpl f e go2 xs)
|
||||||
|
foldrImplNaturality f e [] go1 go2 = Refl
|
||||||
|
foldrImplNaturality f e (x :: xs) go1 go2 = foldrImplNaturality f e xs go1 (go2 . (f x))
|
||||||
|
|
||||||
|
||| Our tail-recursive foldr preserves the vector structure
|
||||||
|
export
|
||||||
|
foldrVectHomomorphism : VectHomomorphismProperty f e (foldr f e)
|
||||||
|
foldrVectHomomorphism = ShowVectHomomorphismProperty
|
||||||
|
{ nil = Refl
|
||||||
|
, cons = \x, xs => Calc $
|
||||||
|
|~ foldr f e (x :: xs)
|
||||||
|
~~ foldrImpl f e (id . (f x)) xs ...(Refl)
|
||||||
|
~~ foldrImpl f e ((f x) . id) xs ...(foldrImplExtensional f e _ _ (\y => Refl) xs)
|
||||||
|
~~ f x (foldrImpl f e id xs) ...(foldrImplNaturality f e xs (f x) _)
|
||||||
|
~~ f x (foldr f e xs) ...(Refl)
|
||||||
|
}
|
||||||
|
|
||||||
|
||| foldr is the unique function preserving the vector structure
|
||||||
|
export
|
||||||
|
foldrUniqueness : (h : forall n . Vect n a -> b) -> VectHomomorphismProperty f e h -> (xs : Vect n a) -> h xs = foldr f e xs
|
||||||
|
foldrUniqueness {f} h prf xs = irrelevantEq $
|
||||||
|
nilConsInitiality f e h (foldr f e) prf foldrVectHomomorphism xs
|
61
libs/contrib/Data/Vect/Properties/Index.idr
Normal file
61
libs/contrib/Data/Vect/Properties/Index.idr
Normal file
@ -0,0 +1,61 @@
|
|||||||
|
||| Properties of Data.Vect.index
|
||||||
|
module Data.Vect.Properties.Index
|
||||||
|
|
||||||
|
import Data.Vect.Properties.Tabulate
|
||||||
|
|
||||||
|
import Data.Vect
|
||||||
|
import Data.Vect.Elem
|
||||||
|
import Data.Fin
|
||||||
|
|
||||||
|
import Syntax.PreorderReasoning
|
||||||
|
|
||||||
|
||| Recall an element by its position, as we may not have the element
|
||||||
|
||| at runtime
|
||||||
|
public export
|
||||||
|
recallElem : {xs : Vect n a} -> x `Elem` xs -> a
|
||||||
|
recallElem {xs = x :: _ } Here = x
|
||||||
|
recallElem {xs = _ :: xs} (There later) = recallElem later
|
||||||
|
|
||||||
|
||| Recalling by a position of `x` does yield `x`
|
||||||
|
export
|
||||||
|
recallElemSpec : (pos : x `Elem` xs) -> recallElem pos = x
|
||||||
|
recallElemSpec Here = Refl
|
||||||
|
recallElemSpec (There later) = recallElemSpec later
|
||||||
|
|
||||||
|
||| `index i : Vect n a -> a` is a natural transformation
|
||||||
|
export
|
||||||
|
indexNaturality : (i : Fin n) -> (f : a -> b) -> (xs : Vect n a)
|
||||||
|
-> index i (map f xs) = f (index i xs)
|
||||||
|
indexNaturality FZ f (x :: xs) = Refl
|
||||||
|
indexNaturality (FS x) f (_ :: xs) = indexNaturality x f xs
|
||||||
|
|
||||||
|
||| Replication tabulates the constant function
|
||||||
|
export
|
||||||
|
indexReplicate : (i : Fin n) -> (x : a)
|
||||||
|
-> index i (replicate n x) = x
|
||||||
|
indexReplicate FZ x = Refl
|
||||||
|
indexReplicate (FS i) x = indexReplicate i x
|
||||||
|
|
||||||
|
||| `range` tabulates the identity function (by definition)
|
||||||
|
export
|
||||||
|
indexRange : (i : Fin n) -> index i (range {len = n}) === i
|
||||||
|
indexRange i = irrelevantEq $ indexTabulate id i
|
||||||
|
|
||||||
|
||| Inductive step auxiliary lemma for indexTranspose
|
||||||
|
indexZipWith_Cons : (i : Fin n) -> (xs : Vect n a) -> (xss : Vect n (Vect m a))
|
||||||
|
-> index i (zipWith (::) xs xss)
|
||||||
|
= (index i xs) :: (index i xss)
|
||||||
|
indexZipWith_Cons FZ (x :: _ ) (xs:: _ ) = Refl
|
||||||
|
indexZipWith_Cons (FS i) (_ :: xs) (_ :: xss) = indexZipWith_Cons i xs xss
|
||||||
|
|
||||||
|
||| The `i`-th vector in a transposed matrix is the vector of `i`-th components
|
||||||
|
export
|
||||||
|
indexTranspose : (xss : Vect m (Vect n a)) -> (i : Fin n)
|
||||||
|
-> index i (transpose xss) = map (index i) xss
|
||||||
|
indexTranspose [] i = indexReplicate i []
|
||||||
|
indexTranspose (xs :: xss) i = Calc $
|
||||||
|
|~ index i (transpose (xs :: xss))
|
||||||
|
~~ index i (zipWith (::) xs (transpose xss)) ...(Refl)
|
||||||
|
~~ index i xs :: index i (transpose xss) ...(indexZipWith_Cons i xs (transpose xss))
|
||||||
|
~~ index i xs :: map (index i) xss ...(cong (index i xs ::)
|
||||||
|
$ indexTranspose xss i)
|
71
libs/contrib/Data/Vect/Properties/Map.idr
Normal file
71
libs/contrib/Data/Vect/Properties/Map.idr
Normal file
@ -0,0 +1,71 @@
|
|||||||
|
||| Properties of Data.Vect.map
|
||||||
|
module Data.Vect.Properties.Map
|
||||||
|
|
||||||
|
import Data.Vect.Properties.Tabulate
|
||||||
|
import Data.Vect.Properties.Index
|
||||||
|
import Data.Vect.Properties.Foldr
|
||||||
|
|
||||||
|
import Data.Vect
|
||||||
|
import Data.Vect.Elem
|
||||||
|
import Data.Fin
|
||||||
|
import Data.Vect.Extra
|
||||||
|
|
||||||
|
import Syntax.PreorderReasoning
|
||||||
|
|
||||||
|
||| `map` functoriality: identity preservation
|
||||||
|
export
|
||||||
|
mapId : (xs : Vect n a) -> map Prelude.id xs = xs
|
||||||
|
mapId xs = vectorExtensionality _ _ \i => indexNaturality _ _ _
|
||||||
|
|
||||||
|
||| `mapWtihPos f` represents post-composition the tabulated function `f`
|
||||||
|
export
|
||||||
|
indexMapWithPos : (f : Fin n -> a -> b) -> (xs : Vect n a) -> (i : Fin n)
|
||||||
|
-> index i (mapWithPos f xs) = f i (index i xs)
|
||||||
|
indexMapWithPos f (x :: _ ) FZ = Refl
|
||||||
|
indexMapWithPos f (_ :: xs) (FS i) = indexMapWithPos _ _ _
|
||||||
|
|
||||||
|
||| `tabulate : (Fin n ->) -> Vect n` is a natural transformation
|
||||||
|
export
|
||||||
|
mapTabulate : (f : a -> b) -> (g : Fin n -> a)
|
||||||
|
-> tabulate (f . g) = map f (tabulate g)
|
||||||
|
mapTabulate f g = irrelevantEq $
|
||||||
|
vectorExtensionality _ _ \i => Calc $
|
||||||
|
|~ index i (tabulate (f . g))
|
||||||
|
~~ f (g i) ...(indexTabulate _ _)
|
||||||
|
~~ f (index i $ tabulate g) ...(cong f (sym $ indexTabulate _ _))
|
||||||
|
~~ index i (map f $ tabulate g) ...(sym $ indexNaturality _ _ _)
|
||||||
|
|
||||||
|
||| Tabulating with the constant function is replication
|
||||||
|
export
|
||||||
|
tabulateConstantly : (x : a) -> Fin.tabulate {len} (const x) === replicate len x
|
||||||
|
tabulateConstantly x = irrelevantEq $
|
||||||
|
vectorExtensionality _ _ \i => Calc $
|
||||||
|
|~ index i (Fin.tabulate (const x))
|
||||||
|
~~ x ...(indexTabulate _ _)
|
||||||
|
~~ index i (replicate _ x) ...(sym $ indexReplicate _ _)
|
||||||
|
|
||||||
|
||| It's enough that two functions agree on the elements of a vector for the maps to agree
|
||||||
|
export
|
||||||
|
mapRestrictedExtensional : (f, g : a -> b) -> (xs : Vect n a)
|
||||||
|
-> (prf : (i : Fin n) -> f (index i xs) = g (index i xs))
|
||||||
|
-> map f xs = map g xs
|
||||||
|
mapRestrictedExtensional f g xs prf = vectorExtensionality _ _ \i => Calc $
|
||||||
|
|~ index i (map f xs)
|
||||||
|
~~ f (index i xs) ...( indexNaturality _ _ _)
|
||||||
|
~~ g (index i xs) ...(prf _)
|
||||||
|
~~ index i (map g xs) ...(sym $ indexNaturality _ _ _)
|
||||||
|
|
||||||
|
||| function extensionality is a congruence wrt map
|
||||||
|
export
|
||||||
|
mapExtensional : (f, g : a -> b)
|
||||||
|
-> (prf : (x : a) -> f x = g x)
|
||||||
|
-> (xs : Vect n a)
|
||||||
|
-> map f xs = map g xs
|
||||||
|
mapExtensional f g prf xs = mapRestrictedExtensional f g xs (\i => prf (index i xs))
|
||||||
|
|
||||||
|
||| map-fusion property for vectors up to function extensionality
|
||||||
|
export
|
||||||
|
mapFusion : (f : b -> c) -> (g : a -> b) -> (xs : Vect n a)
|
||||||
|
-> map f (map g xs) = map (f . g) xs
|
||||||
|
mapFusion f g [] = Refl
|
||||||
|
mapFusion f g (x :: xs) = cong (f $ g x ::) $ mapFusion f g xs
|
37
libs/contrib/Data/Vect/Properties/Tabulate.idr
Normal file
37
libs/contrib/Data/Vect/Properties/Tabulate.idr
Normal file
@ -0,0 +1,37 @@
|
|||||||
|
||| Tabulation gives a bijection between functions `f : Fin n -> a`
|
||||||
|
||| (up to extensional equality) and vectors `tabulate f : Vect n a`.
|
||||||
|
module Data.Vect.Properties.Tabulate
|
||||||
|
|
||||||
|
import Data.Vect
|
||||||
|
import Data.Fin
|
||||||
|
|
||||||
|
||| Vectors are uniquely determined by their elements
|
||||||
|
export
|
||||||
|
vectorExtensionality
|
||||||
|
: (xs, ys : Vect n a) -> (ext : (i : Fin n) -> index i xs = index i ys)
|
||||||
|
-> xs = ys
|
||||||
|
vectorExtensionality [] [] ext = Refl
|
||||||
|
vectorExtensionality (x :: xs) (y :: ys) ext =
|
||||||
|
cong2 (::)
|
||||||
|
(ext FZ)
|
||||||
|
(vectorExtensionality xs ys (\i => ext (FS i)))
|
||||||
|
|
||||||
|
||| Extensionally equivalent functions tabulate to the same vector
|
||||||
|
export
|
||||||
|
tabulateExtensional
|
||||||
|
: {n : Nat} -> (f, g : Fin n -> a) -> (ext : (i : Fin n) -> f i = g i)
|
||||||
|
-> tabulate f = tabulate g
|
||||||
|
tabulateExtensional {n = 0 } f g ext = Refl
|
||||||
|
tabulateExtensional {n = S n} f g ext =
|
||||||
|
cong2 (::) (ext FZ) (tabulateExtensional (f . FS) (g . FS) (\ i => ext $ FS i))
|
||||||
|
|
||||||
|
||| Taking an index amounts to applying the tabulated function
|
||||||
|
export
|
||||||
|
indexTabulate : {n : Nat} -> (f : Fin n -> a) -> (i : Fin n) -> index i (tabulate f) = f i
|
||||||
|
indexTabulate f FZ = Refl
|
||||||
|
indexTabulate f (FS i) = indexTabulate (f . FS) i
|
||||||
|
|
||||||
|
||| The empty vector represents the unique function `Fin 0 -> a`
|
||||||
|
export
|
||||||
|
emptyInitial : (v : Vect 0 a) -> v = []
|
||||||
|
emptyInitial [] = Refl
|
@ -113,7 +113,12 @@ modules = Control.ANSI,
|
|||||||
Data.Validated,
|
Data.Validated,
|
||||||
|
|
||||||
Data.Vect.Binary,
|
Data.Vect.Binary,
|
||||||
|
Data.Vect.Properties,
|
||||||
|
Data.Vect.Properties.Tabulate,
|
||||||
|
Data.Vect.Properties.Index,
|
||||||
|
Data.Vect.Properties.Foldr,
|
||||||
|
Data.Vect.Properties.Map,
|
||||||
|
Data.Vect.Extra,
|
||||||
Data.Vect.Sort,
|
Data.Vect.Sort,
|
||||||
|
|
||||||
Data.Void,
|
Data.Void,
|
||||||
|
Loading…
Reference in New Issue
Block a user