diff options
author | Helmut Grohne <grohne@cs.uni-bonn.de> | 2014-01-27 13:46:47 +0100 |
---|---|---|
committer | Helmut Grohne <grohne@cs.uni-bonn.de> | 2014-01-27 13:46:47 +0100 |
commit | 09cf8a92e03a4cebffdf9bb58aaf1b01b55e73a3 (patch) | |
tree | d4087f8e162c52a706e07ec3aa9623c9d637984f /Generic.agda | |
parent | 00e60d0339d3e0044abfb06b8f77e9f7a3ffdcfb (diff) | |
parent | 71025b5f1d0a11b0cf373192210b293a77d45c04 (diff) | |
download | bidiragda-09cf8a92e03a4cebffdf9bb58aaf1b01b55e73a3.tar.gz |
Merge branch feature-delete
Most conflicts stem from varying imports or added functions and a
take-both approach merges them. A notable exception is theorem-2, where
a new result sequence-cong was required. Apart from that, theorem-2
could be taken almost verbatim from feature-delete albeit its type
coming from feature-decsetoid.
Conflicts:
Bidir.agda
FinMap.agda
Generic.agda
Precond.agda
Diffstat (limited to 'Generic.agda')
-rw-r--r-- | Generic.agda | 23 |
1 files changed, 11 insertions, 12 deletions
diff --git a/Generic.agda b/Generic.agda index f0606ac..f543256 100644 --- a/Generic.agda +++ b/Generic.agda @@ -8,7 +8,7 @@ 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 (_∘_) +open import Function using (_∘_ ; id) open import Level using () renaming (zero to ℓ₀) open import Relation.Binary using (Setoid ; module Setoid) open import Relation.Binary.Core using (_≡_ ; refl) @@ -18,23 +18,13 @@ open import Relation.Binary.PropositionalEquality using (_≗_ ; cong ; subst ; 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 - just-injective : {A : Set} → {x y : A} → Maybe.just x ≡ Maybe.just y → x ≡ y just-injective refl = refl length-replicate : {A : Set} {a : A} → (n : ℕ) → length (replicate n a) ≡ n -length-replicate zero = refl +length-replicate zero = refl length-replicate (suc n) = cong suc (length-replicate n) -map-just-injective : {A : Set} {n : ℕ} {xs ys : Vec A n} → - map Maybe.just xs ≡ map Maybe.just ys → xs ≡ ys -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)) @@ -57,6 +47,15 @@ maybeEq-to-≡ : {A : Set} {a b : Maybe A} → Setoid._≈_ (MaybeEq (PropEq A)) maybeEq-to-≡ (just refl) = refl maybeEq-to-≡ nothing = refl +sequenceV : {A : Set} {n : ℕ} → Vec (Maybe A) n → Maybe (Vec A n) +sequenceV = mapMV id + +sequence-map : {A B : Set} {n : ℕ} → (f : A → Maybe B) → sequenceV {n = n} ∘ map f ≗ mapMV f +sequence-map f []V = refl +sequence-map f (x ∷V xs) with f x +sequence-map f (x ∷V xs) | just y = cong (_<$>_ (_∷V_ y)) (sequence-map f xs) +sequence-map f (x ∷V xs) | nothing = 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 |