summaryrefslogtreecommitdiff
path: root/Generic.agda
diff options
context:
space:
mode:
authorHelmut Grohne <grohne@cs.uni-bonn.de>2014-01-28 10:53:45 +0100
committerHelmut Grohne <grohne@cs.uni-bonn.de>2014-01-28 10:53:45 +0100
commitb0eb7ada208d33eb82ec27cb7d40b9fa59646c92 (patch)
tree00da19e818e122a279c46818f025156903c8a8ac /Generic.agda
parent2a4ee38b0506121d52e259ae88a7b470cdac2206 (diff)
downloadbidiragda-b0eb7ada208d33eb82ec27cb7d40b9fa59646c92.tar.gz
improve readability using _∋_≈_ instead of Setoid._≈_
Diffstat (limited to 'Generic.agda')
-rw-r--r--Generic.agda5
1 files changed, 3 insertions, 2 deletions
diff --git a/Generic.agda b/Generic.agda
index d757c95..81292ff 100644
--- a/Generic.agda
+++ b/Generic.agda
@@ -15,6 +15,7 @@ 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 PropEq)
+open Setoid using () renaming (_≈_ to _∋_≈_)
open Category.Functor.RawFunctor {Level.zero} Data.Maybe.functor using (_<$>_)
open Category.Monad.RawMonad {Level.zero} Data.Maybe.monad using (_>>=_)
@@ -39,11 +40,11 @@ mapMV-purity : {A B : Set} {n : ℕ} → (f : A → B) → (v : Vec A n) → map
mapMV-purity f []V = refl
mapMV-purity f (x ∷V xs) rewrite mapMV-purity f xs = refl
-maybeEq-from-≡ : {A : Set} {a b : Maybe A} → Setoid._≈_ (PropEq (Maybe A)) a b → Setoid._≈_ (MaybeEq (PropEq A)) a b
+maybeEq-from-≡ : {A : Set} {a b : Maybe A} → a ≡ b → MaybeEq (PropEq A) ∋ a ≈ b
maybeEq-from-≡ {a = just x} {b = .(just x)} refl = just refl
maybeEq-from-≡ {a = nothing} {b = .nothing} refl = nothing
-maybeEq-to-≡ : {A : Set} {a b : Maybe A} → Setoid._≈_ (MaybeEq (PropEq A)) a b → Setoid._≈_ (PropEq (Maybe A)) a b
+maybeEq-to-≡ : {A : Set} {a b : Maybe A} → MaybeEq (PropEq A) ∋ a ≈ b → a ≡ b
maybeEq-to-≡ (just refl) = refl
maybeEq-to-≡ nothing = refl