summaryrefslogtreecommitdiff
path: root/Generic.agda
diff options
context:
space:
mode:
authorHelmut Grohne <grohne@cs.uni-bonn.de>2013-12-16 18:10:17 +0100
committerHelmut Grohne <grohne@cs.uni-bonn.de>2013-12-16 18:10:17 +0100
commit2c37e0c2f32b4c6b5c121827a4abddf7fc1dd7e0 (patch)
tree4bdcaf9d73dbdd72c79226e6ffa40ae91051a619 /Generic.agda
parent2f999bfd6553cb31ebffe4c32d0a2a52dedaf4d3 (diff)
downloadbidiragda-2c37e0c2f32b4c6b5c121827a4abddf7fc1dd7e0.tar.gz
add a mapM variant on the Maybe monad on Vecs
Diffstat (limited to 'Generic.agda')
-rw-r--r--Generic.agda22
1 files changed, 21 insertions, 1 deletions
diff --git a/Generic.agda b/Generic.agda
index c7cbc45..11b9594 100644
--- a/Generic.agda
+++ b/Generic.agda
@@ -1,14 +1,20 @@
module Generic where
+import Category.Functor
+import Category.Monad
open import Data.List using (List ; length ; replicate) renaming ([] to []L ; _∷_ to _∷L_)
-open import Data.Maybe using (Maybe ; just)
+open import Data.Maybe using (Maybe ; just ; nothing)
open import Data.Nat using (ℕ ; zero ; suc)
open import Data.Product using (_×_ ; _,_)
open import Data.Vec using (Vec ; toList ; fromList ; map) renaming ([] to []V ; _∷_ to _∷V_)
open import Function using (_∘_)
+import Level
open import Relation.Binary.Core using (_≡_ ; refl)
open import Relation.Binary.PropositionalEquality using (_≗_ ; cong ; subst ; trans)
+open Category.Functor.RawFunctor {Level.zero} Data.Maybe.functor using (_<$>_)
+open Category.Monad.RawMonad {Level.zero} Data.Maybe.monad using (_>>=_)
+
∷-injective : {A : Set} {n : ℕ} {x y : A} {xs ys : Vec A n} →
(x ∷V xs) ≡ (y ∷V ys) → x ≡ y × xs ≡ ys
∷-injective refl = refl , refl
@@ -26,6 +32,20 @@ map-just-injective {xs = []V} {ys = []V} p = refl
map-just-injective {xs = x ∷V xs′} {ys = y ∷V ys′} p with ∷-injective p
map-just-injective {xs = x ∷V xs′} {ys = .x ∷V ys′} p | refl , p′ = cong (_∷V_ x) (map-just-injective p′)
+mapMV : {A B : Set} {n : ℕ} → (A → Maybe B) → Vec A n → Maybe (Vec B n)
+mapMV f []V = just []V
+mapMV f (x ∷V xs) = (f x) >>= (λ y → (_∷V_ y) <$> (mapMV f xs))
+
+mapMV-cong : {A B : Set} {f g : A → Maybe B} → f ≗ g → {n : ℕ} → mapMV {n = n} f ≗ mapMV g
+mapMV-cong f≗g []V = refl
+mapMV-cong {f = f} {g = g} f≗g (x ∷V xs) with f x | g x | f≗g x
+mapMV-cong f≗g (x ∷V xs) | just y | .(just y) | refl = cong (_<$>_ (_∷V_ y)) (mapMV-cong f≗g xs)
+mapMV-cong f≗g (x ∷V xs) | nothing | .nothing | refl = refl
+
+mapMV-purity : {A B : Set} {n : ℕ} → (f : A → B) → (v : Vec A n) → mapMV (just ∘ f) v ≡ just (map f v)
+mapMV-purity f []V = refl
+mapMV-purity f (x ∷V xs) rewrite mapMV-purity f xs = refl
+
subst-cong : {A : Set} → (T : A → Set) → {g : A → A} → {a b : A} → (f : {c : A} → T c → T (g c)) → (p : a ≡ b) →
f ∘ subst T p ≗ subst T (cong g p) ∘ f
subst-cong T f refl _ = refl