summaryrefslogtreecommitdiff
path: root/FinMap.agda
diff options
context:
space:
mode:
authorHelmut Grohne <grohne@cs.uni-bonn.de>2013-12-17 08:49:24 +0100
committerHelmut Grohne <grohne@cs.uni-bonn.de>2013-12-17 08:49:24 +0100
commit88ac52ceabaf2bf4a3f27293a8e04a3daaf8b907 (patch)
tree434e67b2e485bee51268cedef1b417424466f958 /FinMap.agda
parent066861f9cdde4ded6c5442508bef1a27576c68d7 (diff)
downloadbidiragda-88ac52ceabaf2bf4a3f27293a8e04a3daaf8b907.tar.gz
refactor to get rid of FinMap without Maybe entirely
The union was the only user of this type and now it uses only partial mappings. So drop remaining uses of FinMap and make everything work with FinMapMaybe instead.
Diffstat (limited to 'FinMap.agda')
-rw-r--r--FinMap.agda49
1 files changed, 20 insertions, 29 deletions
diff --git a/FinMap.agda b/FinMap.agda
index 8cde5a6..c125c47 100644
--- a/FinMap.agda
+++ b/FinMap.agda
@@ -33,17 +33,11 @@ fromAscList : {A : Set} {n : ℕ} → List (Fin n × A) → FinMapMaybe n A
fromAscList [] = empty
fromAscList ((f , a) ∷ xs) = insert f a (fromAscList xs)
-FinMap : ℕ → Set → Set
-FinMap n A = Vec A n
-
-lookup : {A : Set} {n : ℕ} → Fin n → FinMap n A → A
-lookup = lookupVec
-
-fromFunc : {A : Set} {n : ℕ} → (Fin n → A) → FinMap n A
-fromFunc = tabulate
+fromFunc : {A : Set} {n : ℕ} → (Fin n → A) → FinMapMaybe n A
+fromFunc = mapV just ∘ tabulate
union : {A : Set} {n : ℕ} → FinMapMaybe n A → FinMapMaybe n A → FinMapMaybe n A
-union m1 m2 = fromFunc (λ f → maybe′ just (lookupM f m2) (lookupM f m1))
+union m1 m2 = tabulate (λ f → maybe′ just (lookupM f m2) (lookupM f m1))
restrict : {A : Set} {n : ℕ} → (Fin n → A) → List (Fin n) → FinMapMaybe n A
restrict f is = fromAscList (zip is (map f is))
@@ -54,9 +48,6 @@ delete i m = m [ i ]≔ nothing
delete-many : {A : Set} {n m : ℕ} → Vec (Fin n) m → FinMapMaybe n A → FinMapMaybe n A
delete-many = flip (foldr (const _) delete)
-partialize : {A : Set} {n : ℕ} → FinMap n A → FinMapMaybe n A
-partialize = mapV just
-
lemma-just≢nothing : {A Whatever : Set} {a : A} {ma : Maybe A} → ma ≡ just a → ma ≡ nothing → Whatever
lemma-just≢nothing refl ()
@@ -99,9 +90,9 @@ lemma-tabulate-∘ : {n : ℕ} {A : Set} → {f g : Fin n → A} → f ≗ g →
lemma-tabulate-∘ {zero} {_} {f} {g} f≗g = refl
lemma-tabulate-∘ {suc n} {_} {f} {g} f≗g = cong₂ _∷_ (f≗g zero) (lemma-tabulate-∘ (f≗g ∘ suc))
-lemma-partialize-fromFunc : {n : ℕ} {A : Set} → (f : Fin n → A) → partialize (fromFunc f) ≡ fromFunc (just ∘ f)
-lemma-partialize-fromFunc {zero} f = refl
-lemma-partialize-fromFunc {suc _} f = cong (_∷_ (just (f zero))) (lemma-partialize-fromFunc (f ∘ suc))
+lemma-fromFunc-tabulate : {n : ℕ} {A : Set} → (f : Fin n → A) → fromFunc f ≡ tabulate (just ∘ f)
+lemma-fromFunc-tabulate {zero} f = refl
+lemma-fromFunc-tabulate {suc _} f = cong (_∷_ (just (f zero))) (lemma-fromFunc-tabulate (f ∘ suc))
lemma-lookupM-delete : {n : ℕ} {A : Set} {i j : Fin n} → (f : FinMapMaybe n A) → i ≢ j → lookupM i (delete j f) ≡ lookupM i f
lemma-lookupM-delete {i = zero} {j = zero} (_ ∷ _) p with p refl
@@ -110,24 +101,24 @@ lemma-lookupM-delete {i = zero} {j = suc j} (_ ∷ _) p = refl
lemma-lookupM-delete {i = suc i} {j = zero} (x ∷ xs) p = refl
lemma-lookupM-delete {i = suc i} {j = suc j} (x ∷ xs) p = lemma-lookupM-delete xs (p ∘ cong suc)
-lemma-disjoint-union : {n m : ℕ} {A : Set} → (f : Fin n → A) → (t : Vec (Fin n) m) → union (restrict f (toList t)) (delete-many t (partialize (fromFunc f))) ≡ partialize (fromFunc f)
-lemma-disjoint-union {n} {m} f t = trans (lemma-tabulate-∘ (lemma-inner t)) (sym (lemma-partialize-fromFunc f))
- where lemma-inner : {m : ℕ} → (t : Vec (Fin n) m) → (x : Fin n) → maybe′ just (lookupM x (delete-many t (partialize (fromFunc f)))) (lookupM x (restrict f (toList t))) ≡ just (f x)
+lemma-disjoint-union : {n m : ℕ} {A : Set} → (f : Fin n → A) → (t : Vec (Fin n) m) → union (restrict f (toList t)) (delete-many t (fromFunc f)) ≡ fromFunc f
+lemma-disjoint-union {n} {m} f t = trans (lemma-tabulate-∘ (lemma-inner t)) (sym (lemma-fromFunc-tabulate f))
+ where lemma-inner : {m : ℕ} → (t : Vec (Fin n) m) → (x : Fin n) → maybe′ just (lookupM x (delete-many t (fromFunc f))) (lookupM x (restrict f (toList t))) ≡ just (f x)
lemma-inner [] x = begin
- maybe′ just (lookupM x (partialize (fromFunc f))) (lookupM x empty)
- ≡⟨ cong (maybe′ just (lookupM x (partialize (fromFunc f)))) (lemma-lookupM-empty x) ⟩
- lookupM x (partialize (fromFunc f))
- ≡⟨ cong (lookupM x) (lemma-partialize-fromFunc f) ⟩
- lookupM x (fromFunc (just ∘ f))
+ maybe′ just (lookupM x (fromFunc f)) (lookupM x empty)
+ ≡⟨ cong (maybe′ just (lookupM x (fromFunc f))) (lemma-lookupM-empty x) ⟩
+ lookupM x (fromFunc f)
+ ≡⟨ cong (lookupM x) (lemma-fromFunc-tabulate f) ⟩
+ lookupM x (tabulate (just ∘ f))
≡⟨ lookup∘tabulate (just ∘ f) x ⟩
just (f x) ∎
lemma-inner (t ∷ ts) x with x ≟ t
- lemma-inner (.x ∷ ts) x | yes refl = cong (maybe′ just (lookupM x (delete-many (x ∷ ts) (partialize (fromFunc f))))) (lemma-lookupM-insert x (f x) (restrict f (toList ts)))
+ lemma-inner (.x ∷ ts) x | yes refl = cong (maybe′ just (lookupM x (delete-many (x ∷ ts) (fromFunc f)))) (lemma-lookupM-insert x (f x) (restrict f (toList ts)))
lemma-inner (t ∷ ts) x | no ¬p = begin
- maybe′ just (lookupM x (delete-many (t ∷ ts) (partialize (fromFunc f)))) (lookupM x (restrict f (toList (t ∷ ts))))
- ≡⟨ cong (maybe′ just (lookupM x (delete-many (t ∷ ts) (partialize (fromFunc f))))) (sym (lemma-lookupM-insert-other x t (f t) (restrict f (toList ts)) ¬p)) ⟩
- maybe′ just (lookupM x (delete-many (t ∷ ts) (partialize (fromFunc f)))) (lookupM x (restrict f (toList ts)))
- ≡⟨ cong (flip (maybe′ just) (lookupM x (restrict f (toList ts)))) (lemma-lookupM-delete (delete-many ts (partialize (fromFunc f))) ¬p) ⟩
- maybe′ just (lookupM x (delete-many ts (partialize (fromFunc f)))) (lookupM x (restrict f (toList ts)))
+ maybe′ just (lookupM x (delete-many (t ∷ ts) (fromFunc f))) (lookupM x (restrict f (toList (t ∷ ts))))
+ ≡⟨ cong (maybe′ just (lookupM x (delete-many (t ∷ ts) (fromFunc f)))) (sym (lemma-lookupM-insert-other x t (f t) (restrict f (toList ts)) ¬p)) ⟩
+ maybe′ just (lookupM x (delete-many (t ∷ ts) (fromFunc f))) (lookupM x (restrict f (toList ts)))
+ ≡⟨ cong (flip (maybe′ just) (lookupM x (restrict f (toList ts)))) (lemma-lookupM-delete (delete-many ts (fromFunc f)) ¬p) ⟩
+ maybe′ just (lookupM x (delete-many ts (fromFunc f))) (lookupM x (restrict f (toList ts)))
≡⟨ lemma-inner ts x ⟩
just (f x) ∎