summaryrefslogtreecommitdiff
path: root/BFF.agda
diff options
context:
space:
mode:
authorHelmut Grohne <grohne@cs.uni-bonn.de>2014-02-17 11:04:51 +0100
committerHelmut Grohne <grohne@cs.uni-bonn.de>2014-02-17 11:17:01 +0100
commit1561b1288b93c1353972d08e91f02101b2ccebfc (patch)
treed0c1e18521af040efffaf9019a8287a369ed1000 /BFF.agda
parent248dc87e7c282a56bcc13fc28701a572288bc3ec (diff)
parentce4bbcc0c06b088a10881fcd66da5422571e7995 (diff)
downloadbidiragda-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 'BFF.agda')
-rw-r--r--BFF.agda28
1 files changed, 19 insertions, 9 deletions
diff --git a/BFF.agda b/BFF.agda
index 56ba359..6943f0d 100644
--- a/BFF.agda
+++ b/BFF.agda
@@ -11,15 +11,15 @@ open Category.Functor.RawFunctor {Level.zero} Data.Maybe.functor using (_<$>_)
open import Data.List using (List ; [] ; _∷_ ; map ; length)
open import Data.Vec using (Vec ; toList ; fromList ; tabulate ; allFin) renaming (lookup to lookupV ; map to mapV ; [] to []V ; _∷_ to _∷V_)
open import Function using (id ; _∘_ ; flip)
-open import Relation.Binary using (DecSetoid ; module DecSetoid)
+open import Relation.Binary using (Setoid ; DecSetoid ; module DecSetoid)
open import FinMap
-open import Generic using (mapMV)
+open import Generic using (mapMV ; ≡-to-Π)
import CheckInsert
-import GetTypes
+open import GetTypes using (VecVec-to-PartialVecVec)
-module VecBFF (A : DecSetoid ℓ₀ ℓ₀) where
- open GetTypes.VecVec public using (Get)
+module PartialVecBFF (A : DecSetoid ℓ₀ ℓ₀) where
+ open GetTypes.PartialVecVec public using (Get)
open module A = DecSetoid A using (Carrier) renaming (_≟_ to deq)
open CheckInsert A
@@ -36,12 +36,22 @@ module VecBFF (A : DecSetoid ℓ₀ ℓ₀) where
denumerate : {n : ℕ} → Vec Carrier n → Fin n → Carrier
denumerate = flip lookupV
- bff : (G : Get) → {n : ℕ} → (m : ℕ) → Vec Carrier n → Vec Carrier (Get.getlen G m) → Maybe (Vec Carrier m)
- bff G m s v = let s′ = enumerate s
+ bff : (G : Get) → {i : Get.|I| G} → (j : Get.|I| G) → Vec Carrier (Get.|gl₁| G i) → Vec Carrier (Get.|gl₂| G j) → Maybe (Vec Carrier (Get.|gl₁| G j))
+ bff G i s v = let s′ = enumerate s
t′ = Get.get G s′
g = fromFunc (denumerate s)
g′ = delete-many t′ g
- t = enumeratel m
+ t = enumeratel (Get.|gl₁| G i)
h = assoc (Get.get G t) v
- h′ = (flip union (reshape g′ m)) <$> h
+ h′ = (flip union (reshape g′ (Get.|gl₁| G i))) <$> h
in h′ >>= flip mapMV t ∘ flip lookupM
+
+module VecBFF (A : DecSetoid ℓ₀ ℓ₀) where
+ open GetTypes.VecVec public using (Get)
+ open module A = DecSetoid A using (Carrier) renaming (_≟_ to deq)
+ open CheckInsert A
+
+ open PartialVecBFF A public using (assoc ; enumerate ; denumerate)
+
+ bff : (G : Get) → {n : ℕ} → (m : ℕ) → Vec Carrier n → Vec Carrier (Get.getlen G m) → Maybe (Vec Carrier m)
+ bff G = PartialVecBFF.bff A (VecVec-to-PartialVecVec G)