summaryrefslogtreecommitdiff
path: root/FinMap.agda
diff options
context:
space:
mode:
authorHelmut Grohne <helmut@subdivi.de>2020-08-01 09:05:13 +0200
committerHelmut Grohne <helmut@subdivi.de>2020-08-01 09:05:13 +0200
commit85865ec3c7c3e3a458dc233d4c28e4db97191f3d (patch)
tree0a14606aecf7bc228f0ddd9d9af78bf37650188f /FinMap.agda
parent6ce567bf63a61bce4ccf71e3ec402d94d1da2fb1 (diff)
downloadbidiragda-85865ec3c7c3e3a458dc233d4c28e4db97191f3d.tar.gz
individually open ≡-Reasoning
_≡⟨_⟩_ will be turned into a syntax, so it cannot be imported in future. Avoid doing so now by opening it where needed.
Diffstat (limited to 'FinMap.agda')
-rw-r--r--FinMap.agda28
1 files changed, 15 insertions, 13 deletions
diff --git a/FinMap.agda b/FinMap.agda
index 051014c..1241252 100644
--- a/FinMap.agda
+++ b/FinMap.agda
@@ -19,8 +19,7 @@ open import Function.Surjection using (module Surjection)
open import Relation.Nullary using (yes ; no)
open import Relation.Nullary.Negation using (contradiction)
open import Relation.Binary.Core using (Decidable)
-open import Relation.Binary.PropositionalEquality as P using (_≡_ ; _≢_ ; _≗_)
-open P.≡-Reasoning using (begin_ ; _≡⟨_⟩_ ; _∎)
+open import Relation.Binary.PropositionalEquality as P using (_≡_ ; _≢_ ; _≗_ ; module ≡-Reasoning)
_∈_ : {A : Set} {n : ℕ} → A → Vec A n → Set
_∈_ {A} x xs = Data.List.Membership.Setoid._∈_ (P.setoid A) x (toList xs)
@@ -82,17 +81,19 @@ lemma-lookupM-restrict : {A : Set} {n m : â„•} → (i : Fin n) → (f : Fin n â†
lemma-lookupM-restrict i f [] p = contradiction (P.trans (P.sym p) (lookup-replicate i nothing)) (λ ())
lemma-lookupM-restrict i f (i' ∷ is) p with i ≟ i'
lemma-lookupM-restrict i f (.i ∷ is) {a} p | yes P.refl = just-injective (begin
- just (f i)
- ≡⟨ P.sym (lookup∘update i (restrict f is) (just (f i))) ⟩
- lookupM i (insert i (f i) (restrict f is))
- ≡⟨ p ⟩
- just a ∎)
+ just (f i)
+ ≡⟨ P.sym (lookup∘update i (restrict f is) (just (f i))) ⟩
+ lookupM i (insert i (f i) (restrict f is))
+ ≡⟨ p ⟩
+ just a ∎)
+ where open ≡-Reasoning
lemma-lookupM-restrict i f (i' ∷ is) {a} p | no i≢i' = lemma-lookupM-restrict i f is (begin
- lookupM i (restrict f is)
- ≡⟨ P.sym (lookup∘update′ i≢i' (restrict f is) (just (f i'))) ⟩
- lookupM i (insert i' (f i') (restrict f is))
- ≡⟨ p ⟩
- just a ∎)
+ lookupM i (restrict f is)
+ ≡⟨ P.sym (lookup∘update′ i≢i' (restrict f is) (just (f i'))) ⟩
+ lookupM i (insert i' (f i') (restrict f is))
+ ≡⟨ p ⟩
+ just a ∎)
+ where open ≡-Reasoning
lemma-lookupM-restrict-∈ : {A : Set} {n m : ℕ} → (i : Fin n) → (f : Fin n → A) → (js : Vec (Fin n) m) → i ∈ js → lookupM i (restrict f js) ≡ just (f i)
lemma-lookupM-restrict-∈ i f [] ()
lemma-lookupM-restrict-∈ i f (j ∷ js) p with i ≟ j
@@ -120,7 +121,8 @@ lemma-reshape-id (x ∷ xs) = P.cong (_∷_ x) (lemma-reshape-id xs)
lemma-disjoint-union : {n m : ℕ} {A : Set} → (f : Fin n → A) → (t : Vec (Fin n) m) → union (restrict f t) (delete-many t (fromFunc f)) ≡ fromFunc f
lemma-disjoint-union {n} f t = tabulate-cong inner
- where inner : (x : Fin n) → maybe′ just (lookupM x (delete-many t (fromFunc f))) (lookupM x (restrict f t)) ≡ just (f x)
+ where open ≡-Reasoning
+ inner : (x : Fin n) → maybe′ just (lookupM x (delete-many t (fromFunc f))) (lookupM x (restrict f t)) ≡ just (f x)
inner x with is-∈ _≟_ x t
inner x | yes-∈ x∈t = P.cong (maybe′ just (lookupM x (delete-many t (fromFunc f)))) (lemma-lookupM-restrict-∈ x f t x∈t)
inner x | no-∉ x∉t = begin