summaryrefslogtreecommitdiff
diff options
context:
space:
mode:
authorHelmut Grohne <grohne@cs.uni-bonn.de>2014-01-27 10:50:15 +0100
committerHelmut Grohne <grohne@cs.uni-bonn.de>2014-01-27 10:50:15 +0100
commit00e60d0339d3e0044abfb06b8f77e9f7a3ffdcfb (patch)
tree1a634eae43fb1656bda3f5a3bb856b744131c73d
parent2c37e0c2f32b4c6b5c121827a4abddf7fc1dd7e0 (diff)
parentaf1ea86b6e817a85d4d160833fc5d4bb89e2df7b (diff)
downloadbidiragda-00e60d0339d3e0044abfb06b8f77e9f7a3ffdcfb.tar.gz
Merge branch feature-decsetoid
-rw-r--r--BFF.agda9
-rw-r--r--Bidir.agda87
-rw-r--r--CheckInsert.agda50
-rw-r--r--FinMap.agda16
-rw-r--r--Generic.agda31
-rw-r--r--Precond.agda6
6 files changed, 134 insertions, 65 deletions
diff --git a/BFF.agda b/BFF.agda
index 0cdb231..88d9244 100644
--- a/BFF.agda
+++ b/BFF.agda
@@ -2,7 +2,7 @@ module BFF where
open import Data.Nat using (ℕ)
open import Data.Fin using (Fin)
-import Level
+open import Level using () renaming (zero to ℓ₀)
import Category.Monad
import Category.Functor
open import Data.Maybe using (Maybe ; just ; nothing ; maybe′)
@@ -11,15 +11,16 @@ 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.Core using (Decidable ; _≡_)
+open import Relation.Binary using (DecSetoid ; module DecSetoid)
open import FinMap
import CheckInsert
import FreeTheorems
-module VecBFF (Carrier : Set) (deq : Decidable {A = Carrier} _≡_) where
+module VecBFF (A : DecSetoid ℓ₀ ℓ₀) where
open FreeTheorems.VecVec public using (get-type)
- open CheckInsert Carrier deq
+ open module A = DecSetoid A using (Carrier) renaming (_≟_ to deq)
+ open CheckInsert A
assoc : {n m : ℕ} → Vec (Fin n) m → Vec Carrier m → Maybe (FinMapMaybe n Carrier)
assoc []V []V = just empty
diff --git a/Bidir.agda b/Bidir.agda
index 9cc0ca6..03daf9d 100644
--- a/Bidir.agda
+++ b/Bidir.agda
@@ -1,33 +1,51 @@
-open import Relation.Binary.Core using (Decidable ; _≡_)
+open import Level using () renaming (zero to ℓ₀)
+open import Relation.Binary using (DecSetoid)
-module Bidir (Carrier : Set) (deq : Decidable {A = Carrier} _≡_) where
+module Bidir (A : DecSetoid ℓ₀ ℓ₀) where
open import Data.Nat using (ℕ)
open import Data.Fin using (Fin)
import Level
import Category.Monad
import Category.Functor
-open import Data.Maybe using (Maybe ; nothing ; just ; maybe′)
+open import Data.Maybe using (Maybe ; nothing ; just ; maybe′) renaming (setoid to MaybeSetoid ; Eq to MaybeEq)
open Category.Monad.RawMonad {Level.zero} Data.Maybe.monad using (_>>=_)
open Category.Functor.RawFunctor {Level.zero} Data.Maybe.functor using (_<$>_)
open import Data.List using (List)
open import Data.List.All using (All)
open import Data.Vec using (Vec ; [] ; _∷_ ; toList ; map ; tabulate) renaming (lookup to lookupVec)
+open import Data.Vec.Equality using () renaming (module Equality to VecEq)
open import Data.Vec.Properties using (tabulate-∘ ; lookup∘tabulate ; map-cong ; map-∘)
open import Data.Product using (∃ ; _×_ ; _,_ ; proj₁ ; proj₂)
open import Function using (id ; _∘_ ; flip)
-open import Relation.Binary.Core using (refl)
-open import Relation.Binary.PropositionalEquality using (cong ; sym ; inspect ; [_] ; trans ; cong₂)
-open Relation.Binary.PropositionalEquality.≡-Reasoning using (begin_ ; _≡⟨_⟩_ ; _∎)
+open import Relation.Binary.Core using (refl ; _≡_)
+open import Relation.Binary.PropositionalEquality using (cong ; sym ; inspect ; [_] ; trans ; cong₂ ; decSetoid ; module ≡-Reasoning) renaming (setoid to EqSetoid)
+open import Relation.Binary using (Setoid ; module Setoid ; module DecSetoid)
+import Relation.Binary.EqReasoning as EqR
import FreeTheorems
open FreeTheorems.VecVec using (get-type ; free-theorem)
-open import Generic using (just-injective ; map-just-injective)
+open import Generic using (just-injective ; map-just-injective ; vecIsSetoid)
open import FinMap
import CheckInsert
-open CheckInsert Carrier deq
+open CheckInsert A
import BFF
-open BFF.VecBFF Carrier deq using (assoc ; enumerate ; denumerate ; bff)
+open BFF.VecBFF A using (assoc ; enumerate ; denumerate ; bff)
+open module A = DecSetoid A using (Carrier) renaming (_≟_ to deq)
+
+module SetoidReasoning where
+ infix 1 begin⟨_⟩_
+ infixr 2 _≈⟨_⟩_ _≡⟨_⟩_
+ infix 2 _∎
+ begin⟨_⟩_ : (X : Setoid ℓ₀ ℓ₀) → {x y : Setoid.Carrier X} → EqR._IsRelatedTo_ X x y → Setoid._≈_ X x y
+ begin⟨_⟩_ X p = EqR.begin_ X p
+ _∎ : {X : Setoid ℓ₀ ℓ₀} → (x : Setoid.Carrier X) → EqR._IsRelatedTo_ X x x
+ _∎ {X} = EqR._∎ X
+ _≈⟨_⟩_ : {X : Setoid ℓ₀ ℓ₀} → (x : Setoid.Carrier X) → {y z : Setoid.Carrier X} → Setoid._≈_ X x y → EqR._IsRelatedTo_ X y z → EqR._IsRelatedTo_ X x z
+ _≈⟨_⟩_ {X} = EqR._≈⟨_⟩_ X
+
+ _≡⟨_⟩_ : {X : Setoid ℓ₀ ℓ₀} → (x : Setoid.Carrier X) → {y z : Setoid.Carrier X} → x ≡ y → EqR._IsRelatedTo_ X y z → EqR._IsRelatedTo_ X x z
+ _≡⟨_⟩_ {X} = EqR._≡⟨_⟩_ X
lemma-1 : {m n : ℕ} → (f : Fin n → Carrier) → (is : Vec (Fin n) m) → assoc is (map f is) ≡ just (restrict f (toList is))
lemma-1 f [] = refl
@@ -37,13 +55,20 @@ lemma-1 f (i ∷ is′) = begin
checkInsert i (f i) (restrict f (toList is′))
≡⟨ lemma-checkInsert-restrict f i (toList is′) ⟩
just (restrict f (toList (i ∷ is′))) ∎
+ where open ≡-Reasoning
-lemma-lookupM-assoc : {m n : ℕ} → (i : Fin n) → (is : Vec (Fin n) m) → (x : Carrier) → (xs : Vec Carrier m) → (h : FinMapMaybe n Carrier) → assoc (i ∷ is) (x ∷ xs) ≡ just h → lookupM i h ≡ just x
+lemma-lookupM-assoc : {m n : ℕ} → (i : Fin n) → (is : Vec (Fin n) m) → (x : Carrier) → (xs : Vec Carrier m) → (h : FinMapMaybe n Carrier) → assoc (i ∷ is) (x ∷ xs) ≡ just h → Setoid._≈_ (MaybeSetoid A.setoid) (lookupM i h) (just x)
lemma-lookupM-assoc i is x xs h p with assoc is xs
lemma-lookupM-assoc i is x xs h () | nothing
lemma-lookupM-assoc i is x xs h p | just h' with checkInsert i x h' | insertionresult i x h'
-lemma-lookupM-assoc i is x xs .h refl | just h | ._ | same pl = pl
-lemma-lookupM-assoc i is x xs ._ refl | just h' | ._ | new _ = lemma-lookupM-insert i x h'
+lemma-lookupM-assoc i is x xs .h refl | just h | ._ | same x' x≈x' pl = begin
+ lookupM i h
+ ≡⟨ pl ⟩
+ just x'
+ ≈⟨ MaybeEq.just (Setoid.sym A.setoid x≈x') ⟩
+ just x ∎
+ where open EqR (MaybeSetoid A.setoid)
+lemma-lookupM-assoc i is x xs ._ refl | just h' | ._ | new _ = Setoid.reflexive (MaybeSetoid A.setoid) (lemma-lookupM-insert i x h')
lemma-lookupM-assoc i is x xs h () | just h' | ._ | wrong _ _ _
_in-domain-of_ : {n : ℕ} {A : Set} → (is : List (Fin n)) → (FinMapMaybe n A) → Set
@@ -54,7 +79,7 @@ lemma-assoc-domain [] [] h ph = Data.List.All.[]
lemma-assoc-domain (i' ∷ is') (x' ∷ xs') h ph with assoc is' xs' | inspect (assoc is') xs'
lemma-assoc-domain (i' ∷ is') (x' ∷ xs') h () | nothing | [ ph' ]
lemma-assoc-domain (i' ∷ is') (x' ∷ xs') h ph | just h' | [ ph' ] with checkInsert i' x' h' | inspect (checkInsert i' x') h' | insertionresult i' x' h'
-lemma-assoc-domain (i' ∷ is') (x' ∷ xs') .h refl | just h | [ ph' ] | ._ | _ | same pl = All._∷_ (x' , pl) (lemma-assoc-domain is' xs' h ph')
+lemma-assoc-domain (i' ∷ is') (x' ∷ xs') .h refl | just h | [ ph' ] | ._ | _ | same x _ pl = All._∷_ (x , pl) (lemma-assoc-domain is' xs' h ph')
lemma-assoc-domain (i' ∷ is') (x' ∷ xs') ._ refl | just h' | [ ph' ] | ._ | [ cI≡ ] | new _ = All._∷_
(x' , lemma-lookupM-insert i' x' h')
(Data.List.All.map
@@ -68,23 +93,19 @@ lemma-map-lookupM-assoc i x h h' ph (j ∷ js) (Data.List.All._∷_ (x' , pl) pj
(trans (lemma-lookupM-checkInsert j i x' x h' h pl ph) (sym pl))
(lemma-map-lookupM-assoc i x h h' ph js pj)
-lemma-2 : {m n : ℕ} → (is : Vec (Fin n) m) → (v : Vec Carrier m) → (h : FinMapMaybe n Carrier) → assoc is v ≡ just h → map (flip lookupM h) is ≡ map just v
-lemma-2 [] [] h p = refl
+lemma-2 : {m n : ℕ} → (is : Vec (Fin n) m) → (v : Vec Carrier m) → (h : FinMapMaybe n Carrier) → assoc is v ≡ just h → Setoid._≈_ (vecIsSetoid (MaybeSetoid A.setoid) m) (map (flip lookupM h) is) (map just v)
+lemma-2 [] [] h p = Setoid.refl (vecIsSetoid (MaybeSetoid A.setoid) _)
lemma-2 (i ∷ is) (x ∷ xs) h p with assoc is xs | inspect (assoc is) xs
lemma-2 (i ∷ is) (x ∷ xs) h () | nothing | _
lemma-2 (i ∷ is) (x ∷ xs) h p | just h' | [ ir ] = begin
lookupM i h ∷ map (flip lookupM h) is
- ≡⟨ cong (flip _∷_ (map (flip lookupM h) is)) (lemma-lookupM-assoc i is x xs h (begin
- assoc (i ∷ is) (x ∷ xs)
- ≡⟨ cong (flip _>>=_ (checkInsert i x)) ir ⟩
- checkInsert i x h'
- ≡⟨ p ⟩
- just h ∎) ) ⟩
+ ≈⟨ lemma-lookupM-assoc i is x xs h (trans (cong (flip _>>=_ (checkInsert i x)) ir) p) VecEq.∷-cong Setoid.refl (vecIsSetoid (MaybeSetoid A.setoid) _) ⟩
just x ∷ map (flip lookupM h) is
≡⟨ cong (_∷_ (just x)) (lemma-map-lookupM-assoc i x h h' p is (lemma-assoc-domain is xs h' ir)) ⟩
just x ∷ map (flip lookupM h') is
- ≡⟨ cong (_∷_ (just x)) (lemma-2 is xs h' ir) ⟩
+ ≈⟨ Setoid.refl (MaybeSetoid A.setoid) VecEq.∷-cong (lemma-2 is xs h' ir) ⟩
just x ∷ map just xs ∎
+ where open EqR (vecIsSetoid (MaybeSetoid A.setoid) _)
lemma-map-denumerate-enumerate : {m : ℕ} → (as : Vec Carrier m) → map (denumerate as) (enumerate as) ≡ as
lemma-map-denumerate-enumerate [] = refl
@@ -100,6 +121,7 @@ lemma-map-denumerate-enumerate (a ∷ as) = cong (_∷_ a) (begin
map (denumerate as) (enumerate as)
≡⟨ lemma-map-denumerate-enumerate as ⟩
as ∎)
+ where open ≡-Reasoning
theorem-1 : {getlen : ℕ → ℕ} → (get : get-type getlen) → {m : ℕ} → (s : Vec Carrier m) → bff get s (get s) ≡ just s
theorem-1 get s = begin
@@ -122,7 +144,8 @@ theorem-1 get s = begin
just (map (denumerate s) (enumerate s))
≡⟨ cong just (lemma-map-denumerate-enumerate s) ⟩
just s ∎
- where h↦h′ = _<$>_ (flip union (fromFunc (denumerate s)))
+ where open ≡-Reasoning
+ h↦h′ = _<$>_ (flip union (fromFunc (denumerate s)))
h′↦r = _<$>_ (flip map (enumerate s) ∘ flip lookupVec)
lemma-<$>-just : {A B : Set} {f : A → B} {b : B} {ma : Maybe A} → f <$> ma ≡ just b → ∃ λ a → ma ≡ just a
@@ -140,12 +163,17 @@ lemma-union-not-used h h' (i ∷ is') (Data.List.All._∷_ (x , px) p') = cong
≡⟨ sym px ⟩
lookupM i h ∎)
(lemma-union-not-used h h' is' p')
+ where open ≡-Reasoning
+
+map-just-≈-injective : {n : ℕ} {x y : Vec Carrier n} → Setoid._≈_ (vecIsSetoid (MaybeSetoid A.setoid) n) (map just x) (map just y) → Setoid._≈_ (vecIsSetoid A.setoid n) x y
+map-just-≈-injective {x = []} {y = []} VecEq.[]-cong = VecEq.[]-cong
+map-just-≈-injective {x = _ ∷ _} {y = _ ∷ _} (just x≈y VecEq.∷-cong ps) = x≈y VecEq.∷-cong map-just-≈-injective ps
-theorem-2 : {getlen : ℕ → ℕ} (get : get-type getlen) → {m : ℕ} → (v : Vec Carrier (getlen m)) → (s u : Vec Carrier m) → bff get s v ≡ just u → get u ≡ v
+theorem-2 : {getlen : ℕ → ℕ} (get : get-type getlen) → {m : ℕ} → (v : Vec Carrier (getlen m)) → (s u : Vec Carrier m) → bff get s v ≡ just u → Setoid._≈_ (vecIsSetoid A.setoid (getlen m)) (get u) v
theorem-2 get v s u p with lemma-<$>-just (proj₂ (lemma-<$>-just p))
-theorem-2 get v s u p | h , ph = begin
+theorem-2 get v s u p | h , ph = begin⟨ vecIsSetoid A.setoid _ ⟩
get u
- ≡⟨ just-injective (begin
+ ≡⟨ just-injective (begin⟨ EqSetoid _ ⟩
get <$> (just u)
≡⟨ cong (_<$>_ get) (sym p) ⟩
get <$> (bff get s v)
@@ -154,15 +182,16 @@ theorem-2 get v s u p | h , ph = begin
get (map (flip lookup (h↦h′ h)) s′)
≡⟨ free-theorem get (flip lookup (h↦h′ h)) s′ ⟩
map (flip lookup (h↦h′ h)) (get s′)
- ≡⟨ map-just-injective (begin
+ ≈⟨ map-just-≈-injective (begin⟨ vecIsSetoid (MaybeSetoid A.setoid) _ ⟩
map just (map (flip lookup (union h g)) (get s′))
≡⟨ lemma-union-not-used h g (get s′) (lemma-assoc-domain (get s′) v h ph) ⟩
map (flip lookupM h) (get s′)
- ≡⟨ lemma-2 (get s′) v h ph ⟩
+ ≈⟨ lemma-2 (get s′) v h ph ⟩
map just v
∎) ⟩
v ∎
- where s′ = enumerate s
+ where open SetoidReasoning
+ s′ = enumerate s
g = fromFunc (denumerate s)
h↦h′ = flip union g
h′↦r = flip map s′ ∘ flip lookupVec
diff --git a/CheckInsert.agda b/CheckInsert.agda
index 6926587..47af215 100644
--- a/CheckInsert.agda
+++ b/CheckInsert.agda
@@ -1,50 +1,58 @@
-open import Relation.Binary.Core using (Decidable ; _≡_)
+open import Level using () renaming (zero to ℓ₀)
+open import Relation.Binary using (DecSetoid)
-module CheckInsert (Carrier : Set) (deq : Decidable {A = Carrier} _≡_) where
+module CheckInsert (A : DecSetoid ℓ₀ ℓ₀) where
open import Data.Nat using (ℕ)
open import Data.Fin using (Fin)
open import Data.Fin.Props using (_≟_)
-open import Data.Maybe using (Maybe ; nothing ; just)
+open import Data.Maybe using (Maybe ; nothing ; just) renaming (setoid to MaybeSetoid ; Eq to MaybeEq)
open import Data.List using (List ; [] ; _∷_)
-open import Relation.Nullary using (Dec ; yes ; no)
+open import Data.Vec using () renaming (_∷_ to _∷V_)
+open import Data.Vec.Equality using () renaming (module Equality to VecEq)
+open import Relation.Nullary using (Dec ; yes ; no ; ¬_)
open import Relation.Nullary.Negation using (contradiction)
-open import Relation.Binary.Core using (refl ; _≢_)
+open import Relation.Binary using (Setoid ; IsPreorder ; module DecSetoid)
+open import Relation.Binary.Core using (refl ; _≡_ ; _≢_)
+import Relation.Binary.EqReasoning as EqR
open import Relation.Binary.PropositionalEquality using (cong ; sym ; inspect ; [_] ; trans)
-open Relation.Binary.PropositionalEquality.≡-Reasoning using (begin_ ; _≡⟨_⟩_ ; _∎)
open import FinMap
+open import Generic using (vecIsSetoid)
+
+private
+ open module A = DecSetoid A using (Carrier ; _≈_) renaming (_≟_ to deq)
checkInsert : {n : ℕ} → Fin n → Carrier → FinMapMaybe n Carrier → Maybe (FinMapMaybe n Carrier)
checkInsert i b m with lookupM i m
... | nothing = just (insert i b m)
... | just c with deq b c
-... | yes b≡c = just m
-... | no b≢c = nothing
+... | yes b≈c = just m
+... | no b≉c = nothing
data InsertionResult {n : ℕ} (i : Fin n) (x : Carrier) (h : FinMapMaybe n Carrier) : Maybe (FinMapMaybe n Carrier) → Set where
- same : lookupM i h ≡ just x → InsertionResult i x h (just h)
+ same : (x' : Carrier) → x ≈ x' → lookupM i h ≡ just x' → InsertionResult i x h (just h)
new : lookupM i h ≡ nothing → InsertionResult i x h (just (insert i x h))
- wrong : (x' : Carrier) → x ≢ x' → lookupM i h ≡ just x' → InsertionResult i x h nothing
+ wrong : (x' : Carrier) → ¬ (x ≈ x') → lookupM i h ≡ just x' → InsertionResult i x h nothing
insertionresult : {n : ℕ} → (i : Fin n) → (x : Carrier) → (h : FinMapMaybe n Carrier) → InsertionResult i x h (checkInsert i x h)
insertionresult i x h with lookupM i h | inspect (lookupM i) h
insertionresult i x h | just x' | _ with deq x x'
-insertionresult i x h | just .x | [ il ] | yes refl = same il
-insertionresult i x h | just x' | [ il ] | no x≢x' = wrong x' x≢x' il
+insertionresult i x h | just x' | [ il ] | yes x≈x' = same x' x≈x' il
+insertionresult i x h | just x' | [ il ] | no x≉x' = wrong x' x≉x' il
insertionresult i x h | nothing | [ il ] = new il
lemma-checkInsert-same : {n : ℕ} → (i : Fin n) → (x : Carrier) → (m : FinMapMaybe n Carrier) → lookupM i m ≡ just x → checkInsert i x m ≡ just m
lemma-checkInsert-same i x m p with lookupM i m
lemma-checkInsert-same i x m refl | .(just x) with deq x x
-lemma-checkInsert-same i x m refl | .(just x) | yes refl = refl
-lemma-checkInsert-same i x m refl | .(just x) | no x≢x = contradiction refl x≢x
+lemma-checkInsert-same i x m refl | .(just x) | yes x≈x = refl
+lemma-checkInsert-same i x m refl | .(just x) | no x≉x = contradiction A.refl x≉x
lemma-checkInsert-new : {n : ℕ} → (i : Fin n) → (x : Carrier) → (m : FinMapMaybe n Carrier) → lookupM i m ≡ nothing → checkInsert i x m ≡ just (insert i x m)
lemma-checkInsert-new i x m p with lookupM i m
lemma-checkInsert-new i x m refl | .nothing = refl
-lemma-checkInsert-wrong : {n : ℕ} → (i : Fin n) → (x : Carrier) → (m : FinMapMaybe n Carrier) → (x' : Carrier) → x ≢ x' → lookupM i m ≡ just x' → checkInsert i x m ≡ nothing
+lemma-checkInsert-wrong : {n : ℕ} → (i : Fin n) → (x : Carrier) → (m : FinMapMaybe n Carrier) → (x' : Carrier) → ¬ (x ≈ x') → lookupM i m ≡ just x' → checkInsert i x m ≡ nothing
lemma-checkInsert-wrong i x m x' d p with lookupM i m
lemma-checkInsert-wrong i x m x' d refl | .(just x') with deq x x'
lemma-checkInsert-wrong i x m x' d refl | .(just x') | yes q = contradiction q d
@@ -52,13 +60,13 @@ lemma-checkInsert-wrong i x m x' d refl | .(just x') | no ¬q = refl
lemma-checkInsert-restrict : {n : ℕ} → (f : Fin n → Carrier) → (i : Fin n) → (is : List (Fin n)) → checkInsert i (f i) (restrict f is) ≡ just (restrict f (i ∷ is))
lemma-checkInsert-restrict f i is with checkInsert i (f i) (restrict f is) | insertionresult i (f i) (restrict f is)
-lemma-checkInsert-restrict f i is | ._ | same p = cong just (lemma-insert-same _ i (f i) p)
+lemma-checkInsert-restrict f i is | ._ | same x fi≈x p = cong just (lemma-insert-same _ i (f i) (trans p (cong just (sym (lemma-lookupM-restrict i f is x p)))))
lemma-checkInsert-restrict f i is | ._ | new _ = refl
-lemma-checkInsert-restrict f i is | ._ | wrong x fi≢x p = contradiction (lemma-lookupM-restrict i f is x p) fi≢x
+lemma-checkInsert-restrict f i is | ._ | wrong x fi≉x p = contradiction (IsPreorder.reflexive (Setoid.isPreorder A.setoid) (lemma-lookupM-restrict i f is x p)) fi≉x
lemma-lookupM-checkInsert : {n : ℕ} → (i j : Fin n) → (x y : Carrier) → (h h' : FinMapMaybe n Carrier) → lookupM i h ≡ just x → checkInsert j y h ≡ just h' → lookupM i h' ≡ just x
lemma-lookupM-checkInsert i j x y h h' pl ph' with checkInsert j y h | insertionresult j y h
-lemma-lookupM-checkInsert i j x y h .h pl refl | ._ | same _ = pl
+lemma-lookupM-checkInsert i j x y h .h pl refl | ._ | same _ _ _ = pl
lemma-lookupM-checkInsert i j x y h h' pl ph' | ._ | new _ with i ≟ j
lemma-lookupM-checkInsert i .i x y h h' pl ph' | ._ | new pl' | yes refl = lemma-just≢nothing pl pl'
lemma-lookupM-checkInsert i j x y h .(insert j y h) pl refl | ._ | new _ | no i≢j = begin
@@ -67,11 +75,13 @@ lemma-lookupM-checkInsert i j x y h .(insert j y h) pl refl | ._ | new _ | no i
lookupM i h
≡⟨ pl ⟩
just x ∎
+ where open Relation.Binary.PropositionalEquality.≡-Reasoning
+
lemma-lookupM-checkInsert i j x y h h' pl () | ._ | wrong _ _ _
lemma-lookupM-checkInsert-other : {n : ℕ} → (i j : Fin n) → i ≢ j → (x : Carrier) → (h h' : FinMapMaybe n Carrier) → checkInsert j x h ≡ just h' → lookupM i h ≡ lookupM i h'
lemma-lookupM-checkInsert-other i j i≢j x h h' ph' with lookupM j h
lemma-lookupM-checkInsert-other i j i≢j x h h' ph' | just y with deq x y
-lemma-lookupM-checkInsert-other i j i≢j x h .h refl | just .x | yes refl = refl
-lemma-lookupM-checkInsert-other i j i≢j x h h' () | just y | no x≢y
+lemma-lookupM-checkInsert-other i j i≢j x h .h refl | just y | yes x≈y = refl
+lemma-lookupM-checkInsert-other i j i≢j x h h' () | just y | no x≉y
lemma-lookupM-checkInsert-other i j i≢j x h .(insert j x h) refl | nothing = lemma-lookupM-insert-other i j x h i≢j
diff --git a/FinMap.agda b/FinMap.agda
index 46dbd1c..1fc2d16 100644
--- a/FinMap.agda
+++ b/FinMap.agda
@@ -1,21 +1,25 @@
module FinMap where
+open import Level using () renaming (zero to ℓ₀)
open import Data.Nat using (ℕ ; zero ; suc)
-open import Data.Maybe using (Maybe ; just ; nothing ; maybe′)
+open import Data.Maybe using (Maybe ; just ; nothing ; maybe′) renaming (setoid to MaybeEq)
open import Data.Fin using (Fin ; zero ; suc)
open import Data.Fin.Props using (_≟_)
open import Data.Vec using (Vec ; [] ; _∷_ ; _[_]≔_ ; replicate ; tabulate ; foldr) renaming (lookup to lookupVec ; map to mapV)
+open import Data.Vec.Equality using ()
+open Data.Vec.Equality.Equality using (_∷-cong_)
open import Data.Vec.Properties using (lookup∘tabulate)
open import Data.List using (List ; [] ; _∷_ ; map ; zip)
open import Data.Product using (_×_ ; _,_)
open import Function using (id ; _∘_ ; flip ; const)
open import Relation.Nullary using (yes ; no)
open import Relation.Nullary.Negation using (contradiction)
+open import Relation.Binary using (Setoid ; module Setoid)
open import Relation.Binary.Core using (_≡_ ; refl ; _≢_)
open import Relation.Binary.PropositionalEquality using (cong ; sym ; _≗_ ; trans ; cong₂)
open Relation.Binary.PropositionalEquality.≡-Reasoning using (begin_ ; _≡⟨_⟩_ ; _∎)
-open import Generic using (just-injective)
+open import Generic using (just-injective ; vecIsSetoid)
FinMapMaybe : ℕ → Set → Set
FinMapMaybe n A = Vec (Maybe A) n
@@ -60,10 +64,10 @@ partialize = mapV just
lemma-just≢nothing : {A Whatever : Set} {a : A} {ma : Maybe A} → ma ≡ just a → ma ≡ nothing → Whatever
lemma-just≢nothing refl ()
-lemma-insert-same : {τ : Set} {n : ℕ} → (m : FinMapMaybe n τ) → (f : Fin n) → (a : τ) → lookupM f m ≡ just a → m ≡ insert f a m
-lemma-insert-same [] () a p
-lemma-insert-same (.(just a) ∷ xs) zero a refl = refl
-lemma-insert-same (x ∷ xs) (suc i) a p = cong (_∷_ x) (lemma-insert-same xs i a p)
+lemma-insert-same : {n : ℕ} {A : Set} → (m : FinMapMaybe n A) → (f : Fin n) → (a : A) → lookupM f m ≡ just a → m ≡ insert f a m
+lemma-insert-same [] () a p
+lemma-insert-same {suc n} (x ∷ xs) zero a p = cong (flip _∷_ xs) p
+lemma-insert-same (x ∷ xs) (suc i) a p = cong (_∷_ x) (lemma-insert-same xs i a p)
lemma-lookupM-empty : {A : Set} {n : ℕ} → (i : Fin n) → lookupM {A} i empty ≡ nothing
lemma-lookupM-empty zero = refl
diff --git a/Generic.agda b/Generic.agda
index 11b9594..f0606ac 100644
--- a/Generic.agda
+++ b/Generic.agda
@@ -3,14 +3,17 @@ module Generic where
import Category.Functor
import Category.Monad
open import Data.List using (List ; length ; replicate) renaming ([] to []L ; _∷_ to _∷L_)
-open import Data.Maybe using (Maybe ; just ; nothing)
+open import Data.Maybe using (Maybe ; just ; nothing) renaming (setoid to MaybeEq)
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 (_∘_)
-import Level
+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.PropositionalEquality using (_≗_ ; cong ; subst ; trans)
+open import Relation.Binary.Indexed using (_at_) renaming (Setoid to ISetoid)
+open import Relation.Binary.PropositionalEquality using (_≗_ ; cong ; subst ; trans) renaming (setoid to PropEq)
open Category.Functor.RawFunctor {Level.zero} Data.Maybe.functor using (_<$>_)
open Category.Monad.RawMonad {Level.zero} Data.Maybe.monad using (_>>=_)
@@ -46,6 +49,14 @@ 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) rewrite mapMV-purity f xs = refl
+maybeEq-from-≡ : {A : Set} {a b : Maybe A} → Setoid._≈_ (PropEq (Maybe A)) a b → Setoid._≈_ (MaybeEq (PropEq 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} → Setoid._≈_ (MaybeEq (PropEq A)) a b → Setoid._≈_ (PropEq (Maybe A)) a b
+maybeEq-to-≡ (just refl) = refl
+maybeEq-to-≡ 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
@@ -65,3 +76,17 @@ toList-fromList (x ∷L xs) = cong (_∷L_ x) (toList-fromList xs)
toList-subst : {A : Set} → {n m : ℕ} (v : Vec A n) → (p : n ≡ m) →
toList (subst (Vec A) p v) ≡ toList v
toList-subst v refl = refl
+
+vecIsISetoid : Setoid ℓ₀ ℓ₀ → ISetoid ℕ ℓ₀ ℓ₀
+vecIsISetoid S = record
+ { Carrier = Vec (Setoid.Carrier S)
+ ; _≈_ = λ x → S VecEq.≈ x
+ ; isEquivalence = record
+ { refl = VecEq.refl S _
+ ; sym = VecEq.sym S
+ ; trans = VecEq.trans S }
+ }
+
+
+vecIsSetoid : Setoid ℓ₀ ℓ₀ → ℕ → Setoid ℓ₀ ℓ₀
+vecIsSetoid S n = (vecIsISetoid S) at n
diff --git a/Precond.agda b/Precond.agda
index e4699dc..9bae83b 100644
--- a/Precond.agda
+++ b/Precond.agda
@@ -17,16 +17,16 @@ open Data.List.Any.Membership-≡ using (_∉_)
open import Data.Maybe using (just)
open import Data.Product using (∃ ; _,_)
open import Function using (flip ; _∘_)
-open import Relation.Binary.PropositionalEquality using (refl ; cong ; inspect ; [_] ; sym)
+open import Relation.Binary.PropositionalEquality using (refl ; cong ; inspect ; [_] ; sym ; decSetoid)
open Relation.Binary.PropositionalEquality.≡-Reasoning using (begin_ ; _≡⟨_⟩_ ; _∎)
open import FinMap using (FinMap ; FinMapMaybe ; lookupM ; union ; fromFunc ; empty ; insert ; lemma-lookupM-empty)
import CheckInsert
-open CheckInsert Carrier deq using (checkInsert ; lemma-checkInsert-new ; lemma-lookupM-checkInsert-other)
+open CheckInsert (decSetoid deq) using (checkInsert ; lemma-checkInsert-new ; lemma-lookupM-checkInsert-other)
import BFF
import Bidir
-open BFF.VecBFF Carrier deq using (get-type ; assoc ; enumerate ; denumerate ; bff)
+open BFF.VecBFF (decSetoid deq) using (get-type ; assoc ; enumerate ; denumerate ; bff)
assoc-enough : {getlen : ℕ → ℕ} (get : get-type getlen) → {m : ℕ} → (s : Vec Carrier m) → (v : Vec Carrier (getlen m)) → ∃ (λ h → assoc (get (enumerate s)) v ≡ just h) → ∃ λ u → bff get s v ≡ just u
assoc-enough get s v (h , p) = u , cong (_<$>_ (flip map s′ ∘ flip lookup) ∘ (_<$>_ (flip union g))) p