summaryrefslogtreecommitdiff
path: root/Generic.agda
diff options
context:
space:
mode:
Diffstat (limited to 'Generic.agda')
-rw-r--r--Generic.agda18
1 files changed, 8 insertions, 10 deletions
diff --git a/Generic.agda b/Generic.agda
index a734ec2..c458483 100644
--- a/Generic.agda
+++ b/Generic.agda
@@ -8,13 +8,13 @@ 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 Data.Vec.Equality using () renaming (module Equality to VecEq)
-open import Function using (_∘_ ; id)
+open import Function using (_∘_ ; id ; flip)
open import Function.Equality using (_⟶_)
open import Level using () renaming (zero to ℓ₀)
open import Relation.Binary using (Setoid ; module Setoid)
open import Relation.Binary.Core using (_≡_ ; refl)
open import Relation.Binary.Indexed using (_at_) renaming (Setoid to ISetoid)
-open import Relation.Binary.PropositionalEquality using (_≗_ ; cong ; subst ; trans) renaming (setoid to EqSetoid)
+open import Relation.Binary.PropositionalEquality using (_≗_ ; cong ; subst ; trans ; cong₂) renaming (setoid to EqSetoid)
open Setoid using () renaming (_≈_ to _∋_≈_)
open Category.Functor.RawFunctor {Level.zero} Data.Maybe.functor using (_<$>_)
@@ -35,14 +35,12 @@ 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-cong f≗g []V = refl
+mapMV-cong f≗g (x ∷V xs) = cong₂ _>>=_ (f≗g x) (cong (flip (_<$>_ ∘ _∷V_)) (mapMV-cong f≗g xs))
-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
+mapMV-purity : {A B : Set} {n : ℕ} → (f : A → B) → (v : Vec A n) → mapMV (Maybe.just ∘ f) v ≡ just (map f v)
+mapMV-purity f []V = refl
+mapMV-purity f (x ∷V xs) = cong (_<$>_ (_∷V_ (f x))) (mapMV-purity f xs)
maybeEq-from-≡ : {A : Set} {a b : Maybe A} → a ≡ b → MaybeEq (EqSetoid A) ∋ a ≈ b
maybeEq-from-≡ {a = just x} {b = .(just x)} refl = just refl
@@ -84,7 +82,7 @@ toList-subst v refl = refl
VecISetoid : Setoid ℓ₀ ℓ₀ → ISetoid ℕ ℓ₀ ℓ₀
VecISetoid S = record
{ Carrier = Vec (Setoid.Carrier S)
- ; _≈_ = λ x → S VecEq.≈ x
+ ; _≈_ = λ x → VecEq._≈_ S x
; isEquivalence = record
{ refl = VecEq.refl S _
; sym = VecEq.sym S