diff options
author | Helmut Grohne <grohne@cs.uni-bonn.de> | 2014-02-17 11:04:51 +0100 |
---|---|---|
committer | Helmut Grohne <grohne@cs.uni-bonn.de> | 2014-02-17 11:17:01 +0100 |
commit | 1561b1288b93c1353972d08e91f02101b2ccebfc (patch) | |
tree | d0c1e18521af040efffaf9019a8287a369ed1000 /Generic.agda | |
parent | 248dc87e7c282a56bcc13fc28701a572288bc3ec (diff) | |
parent | ce4bbcc0c06b088a10881fcd66da5422571e7995 (diff) | |
download | bidiragda-1561b1288b93c1353972d08e91f02101b2ccebfc.tar.gz |
Merge branch feature-partial-getlen into master
It allows defining get functions that are only defined for some vector
lengths. It retains backwards compatibility with the previous state via
a VecVec compatibility module. The biggest benefit is that now standard
library functions such as tail, take and drop can be passed to bff.
Conflicts: heavy
BFF.agda (imports, bff type clash)
Bidir.agda (imports, heavy bff type clash in theorem-{1,2} and
lemma-get-mapMV)
Generic.agda (imports)
Precond.agda (imports, bff type clash in assoc-enough)
Diffstat (limited to 'Generic.agda')
-rw-r--r-- | Generic.agda | 10 |
1 files changed, 7 insertions, 3 deletions
diff --git a/Generic.agda b/Generic.agda index b29f47d..c458483 100644 --- a/Generic.agda +++ b/Generic.agda @@ -9,16 +9,20 @@ 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 ; 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 ; cong₂) renaming (setoid to PropEq) +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 (_<$>_) open Category.Monad.RawMonad {Level.zero} Data.Maybe.monad using (_>>=_) +≡-to-Π : {A B : Set} → (A → B) → EqSetoid A ⟶ EqSetoid B +≡-to-Π f = record { _⟨$⟩_ = f; cong = cong f } + just-injective : {A : Set} → {x y : A} → Maybe.just x ≡ Maybe.just y → x ≡ y just-injective refl = refl @@ -38,11 +42,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) = cong (_<$>_ (_∷V_ (f x))) (mapMV-purity f xs) -maybeEq-from-≡ : {A : Set} {a b : Maybe A} → a ≡ b → MaybeEq (PropEq A) ∋ a ≈ b +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 maybeEq-from-≡ {a = nothing} {b = .nothing} refl = nothing -maybeEq-to-≡ : {A : Set} {a b : Maybe A} → MaybeEq (PropEq A) ∋ a ≈ b → a ≡ b +maybeEq-to-≡ : {A : Set} {a b : Maybe A} → MaybeEq (EqSetoid A) ∋ a ≈ b → a ≡ b maybeEq-to-≡ (just refl) = refl maybeEq-to-≡ nothing = refl |