diff options
author | Helmut Grohne <helmut@subdivi.de> | 2019-03-31 21:14:50 +0200 |
---|---|---|
committer | Helmut Grohne <helmut@subdivi.de> | 2019-03-31 21:14:50 +0200 |
commit | 4d2b9ba79a5a35ad63ee941f0681697cf017dfd0 (patch) | |
tree | e2b8519ce63b03b470d970f8338048bdaac5597f /Precond.agda | |
parent | 25d4df9182c92ef26979566f06c7c9f17746f0fb (diff) | |
download | bidiragda-4d2b9ba79a5a35ad63ee941f0681697cf017dfd0.tar.gz |
replace FinMap.lemma-lookupM-empty with Data.Vec.Properties.lookup-replicate
Diffstat (limited to 'Precond.agda')
-rw-r--r-- | Precond.agda | 6 |
1 files changed, 3 insertions, 3 deletions
diff --git a/Precond.agda b/Precond.agda index 98ba7a6..8eec6da 100644 --- a/Precond.agda +++ b/Precond.agda @@ -15,7 +15,7 @@ import Data.Maybe.Categorical open Category.Monad.RawMonad {Level.zero} Data.Maybe.Categorical.monad using (_>>=_) open Category.Functor.RawFunctor {Level.zero} Data.Maybe.Categorical.functor using (_<$>_) open import Data.Vec using (Vec ; [] ; _∷_ ; map ; lookup ; toList) -open import Data.Vec.Properties using (map-cong ; map-∘ ; tabulate-∘) +open import Data.Vec.Properties using (map-cong ; map-∘ ; tabulate-∘ ; lookup-replicate) import Data.List.All open import Data.List.Any using (here ; there) open import Data.List.Membership.Setoid using (_∉_) @@ -28,7 +28,7 @@ open P.≡-Reasoning using (begin_ ; _≡⟨_⟩_ ; _∎) open import Relation.Nullary using (yes ; no) open import Structures using (IsFunctor ; module Shaped ; Shaped) -open import FinMap using (FinMapMaybe ; lookupM ; union ; fromFunc ; empty ; insert ; lemma-lookupM-empty ; delete-many ; lemma-tabulate-∘ ; delete ; lemma-lookupM-delete ; lemma-lookupM-fromFunc ; reshape ; lemma-reshape-id) +open import FinMap using (FinMapMaybe ; lookupM ; union ; fromFunc ; empty ; insert ; delete-many ; lemma-tabulate-∘ ; delete ; lemma-lookupM-delete ; lemma-lookupM-fromFunc ; reshape ; lemma-reshape-id) import CheckInsert open CheckInsert (P.decSetoid deq) using (checkInsert ; lemma-checkInsert-new ; lemma-lookupM-checkInsert-other) import BFF @@ -103,7 +103,7 @@ data All-different {A : Set} : List A → Set where different-∷ : {x : A} {xs : List A} → _∉_ (P.setoid A) x xs → All-different xs → All-different (x ∷ xs) lemma-∉-lookupM-assoc : {m n : ℕ} → (i : Fin n) → (is : Vec (Fin n) m) → (xs : Vec Carrier m) → {h : FinMapMaybe n Carrier} → assoc is xs ≡ just h → _∉_ (P.setoid (Fin n)) i (toList is) → lookupM i h ≡ nothing -lemma-∉-lookupM-assoc i [] [] P.refl i∉is = lemma-lookupM-empty i +lemma-∉-lookupM-assoc i [] [] P.refl i∉is = lookup-replicate i nothing lemma-∉-lookupM-assoc i (i' ∷ is') (x' ∷ xs') ph i∉is with assoc is' xs' | inspect (assoc is') xs' lemma-∉-lookupM-assoc i (i' ∷ is') (x' ∷ xs') () i∉is | nothing | [ ph' ] lemma-∉-lookupM-assoc i (i' ∷ is') (x' ∷ xs') {h} ph i∉is | just h' | [ ph' ] = begin |