summaryrefslogtreecommitdiff
diff options
context:
space:
mode:
authorHelmut Grohne <grohne@cs.uni-bonn.de>2014-01-16 11:32:12 +0100
committerHelmut Grohne <grohne@cs.uni-bonn.de>2014-01-16 11:32:12 +0100
commit48a000f3dc05a9117a9b72e250569c204a4d1371 (patch)
tree538e2b674eec4b9424683f0c0192ca2e8b31dcdb
parent2c37e0c2f32b4c6b5c121827a4abddf7fc1dd7e0 (diff)
downloadbidiragda-48a000f3dc05a9117a9b72e250569c204a4d1371.tar.gz
generalize lemma-insert-same to arbitrary Setoids
The general idea is to enable bff to use arbitrary DecSetoids. * assoc needs to learn about DecSetoid * checkInsert needs to learn about DecSetoid * InsertionResult needs to learn about Setoid * Parameters to InsertionResult.same become weaker * lemma-checkInsert-restrict should work with weaker same * lemma-insert-same needs to learn about Setoid
-rw-r--r--CheckInsert.agda15
-rw-r--r--FinMap.agda24
-rw-r--r--Generic.agda28
3 files changed, 55 insertions, 12 deletions
diff --git a/CheckInsert.agda b/CheckInsert.agda
index 6926587..d82c40b 100644
--- a/CheckInsert.agda
+++ b/CheckInsert.agda
@@ -5,15 +5,19 @@ module CheckInsert (Carrier : Set) (deq : Decidable {A = Carrier} _≡_) 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 MaybeEq)
open import Data.List using (List ; [] ; _∷_)
+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 using (Setoid)
open import Relation.Binary.Core using (refl ; _≢_)
-open import Relation.Binary.PropositionalEquality using (cong ; sym ; inspect ; [_] ; trans)
+open import Relation.Binary.PropositionalEquality using (cong ; sym ; inspect ; [_] ; trans) renaming (setoid to PropEq)
open Relation.Binary.PropositionalEquality.≡-Reasoning using (begin_ ; _≡⟨_⟩_ ; _∎)
open import FinMap
+open import Generic using (maybeEq-from-≡ ; vecIsSetoid)
checkInsert : {n : ℕ} → Fin n → Carrier → FinMapMaybe n Carrier → Maybe (FinMapMaybe n Carrier)
checkInsert i b m with lookupM i m
@@ -50,9 +54,14 @@ 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
lemma-checkInsert-wrong i x m x' d refl | .(just x') | no ¬q = refl
+vecSetoidToProp : {A : Set} {n : ℕ} {x y : Setoid.Carrier (vecIsSetoid (MaybeEq (PropEq A)) n)} → Setoid._≈_ (vecIsSetoid (MaybeEq (PropEq A)) n) x y → x ≡ y
+vecSetoidToProp VecEq.[]-cong = refl
+vecSetoidToProp (just refl VecEq.∷-cong p) = cong (_∷V_ _) (vecSetoidToProp p)
+vecSetoidToProp (nothing VecEq.∷-cong p) = cong (_∷V_ _) (vecSetoidToProp p)
+
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 p = cong just (vecSetoidToProp (lemma-insert-same _ i (f i) (maybeEq-from-≡ 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
diff --git a/FinMap.agda b/FinMap.agda
index 46dbd1c..a85f119 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,18 @@ 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)
+module Private {S : Setoid ℓ₀ ℓ₀} where
+ private
+ open Setoid S
+ reflMaybe = Setoid.refl (MaybeEq S)
+ _≈Maybe_ = Setoid._≈_ (MaybeEq S)
+
+ lemma-insert-same : {n : ℕ} → (m : FinMapMaybe n Carrier) → (f : Fin n) → (a : Carrier) → lookupM f m ≈Maybe just a → Setoid._≈_ (vecIsSetoid (MaybeEq S) n) m (insert f a m)
+ lemma-insert-same [] () a p
+ lemma-insert-same {suc n} (x ∷ xs) zero a p = p ∷-cong Setoid.refl (vecIsSetoid (MaybeEq S) n)
+ lemma-insert-same (x ∷ xs) (suc i) a p = reflMaybe ∷-cong lemma-insert-same xs i a p
+
+open Private public
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..33cd1ac 100644
--- a/Generic.agda
+++ b/Generic.agda
@@ -3,14 +3,16 @@ 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.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 +48,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 +75,15 @@ 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
+
+vecIsSetoid : Setoid ℓ₀ ℓ₀ → ℕ → Setoid ℓ₀ ℓ₀
+vecIsSetoid S n = record
+ { Carrier = Vec S.Carrier n
+ ; _≈_ = λ x y → VecEq._≈_ S {n} x {n} y
+ ; isEquivalence = record
+ { refl = VecEq.refl S _
+ ; sym = VecEq.sym S
+ ; trans = VecEq.trans S }
+ }
+ where
+ module S = Setoid S