summaryrefslogtreecommitdiff
path: root/BFF.agda
diff options
context:
space:
mode:
authorHelmut Grohne <helmut@subdivi.de>2012-06-05 15:35:09 +0200
committerHelmut Grohne <helmut@subdivi.de>2012-06-05 15:35:09 +0200
commitac8dee6708affea52aedc1ffcc5c83f8043ad91d (patch)
treef11c8161fe240233ecaee4f6c773f4075f7b564f /BFF.agda
parenta98ec27d280b41b86fad060aff60c3e9037fc669 (diff)
downloadbidiragda-ac8dee6708affea52aedc1ffcc5c83f8043ad91d.tar.gz
move bff and friends to submodule ListBFF
Diffstat (limited to 'BFF.agda')
-rw-r--r--BFF.agda37
1 files changed, 37 insertions, 0 deletions
diff --git a/BFF.agda b/BFF.agda
new file mode 100644
index 0000000..5c0e279
--- /dev/null
+++ b/BFF.agda
@@ -0,0 +1,37 @@
+module BFF where
+
+open import Data.Nat using (ℕ)
+open import Data.Fin using (Fin)
+open import Data.Maybe using (Maybe ; just ; nothing ; maybe′)
+open import Data.List using (List ; [] ; _∷_ ; map ; length)
+open import Data.Vec using (Vec ; toList ; fromList ; tabulate) renaming (lookup to lookupVec)
+open import Function using (id ; _∘_ ; flip)
+
+open import FinMap
+open import CheckInsert
+
+_>>=_ : {A B : Set} → Maybe A → (A → Maybe B) → Maybe B
+_>>=_ = flip (flip maybe′ nothing)
+
+fmap : {A B : Set} → (A → B) → Maybe A → Maybe B
+fmap f = maybe′ (λ a → just (f a)) nothing
+
+module ListBFF where
+
+ assoc : {A : Set} {n : ℕ} → EqInst A → List (Fin n) → List A → Maybe (FinMapMaybe n A)
+ assoc _ [] [] = just empty
+ assoc eq (i ∷ is) (b ∷ bs) = (assoc eq is bs) >>= (checkInsert eq i b)
+ assoc _ _ _ = nothing
+
+ enumerate : {A : Set} → (l : List A) → List (Fin (length l))
+ enumerate l = toList (tabulate id)
+
+ denumerate : {A : Set} (l : List A) → Fin (length l) → A
+ denumerate l = flip lookupVec (fromList l)
+
+ bff : ({A : Set} → List A → List A) → ({B : Set} → EqInst B → List B → List B → Maybe (List B))
+ bff get eq s v = let s′ = enumerate s
+ g = fromFunc (denumerate s)
+ h = assoc eq (get s′) v
+ h′ = fmap (flip union g) h
+ in fmap (flip map s′ ∘ flip lookup) h′