3 open import Level using () renaming (zero to â„“â‚€)
4 open import Data.Nat using (â„• ; zero ; suc)
5 open import Data.Maybe using (Maybe ; just ; nothing ; maybe′)
6 open import Data.Maybe.Properties using (just-injective)
7 open import Data.Fin using (Fin ; zero ; suc)
8 open import Data.Fin.Properties using (_≟_)
9 open import Data.Vec using (Vec ; [] ; _∷_ ; _[_]≔_ ; replicate ; tabulate ; foldr ; zip ; toList) renaming (lookup to lookupVec ; map to mapV)
10 open import Data.Vec.Properties using (lookup∘update ; lookup∘update′ ; lookup-replicate ; tabulate-cong ; lookup∘tabulate)
11 open import Data.Product using (_×_ ; _,_)
12 open import Data.List.Relation.Unary.All as All using (All)
13 import Data.List.Relation.Unary.All.Properties as AllP
14 import Data.List.Relation.Unary.Any as Any
15 import Data.List.Membership.Setoid
16 open import Function using (id ; _∘_ ; flip ; const)
17 open import Function.Equality using (module Î )
18 open import Function.Surjection using (module Surjection)
19 open import Relation.Nullary using (yes ; no)
20 open import Relation.Nullary.Negation using (contradiction)
21 open import Relation.Binary using (Decidable)
22 open import Relation.Binary.PropositionalEquality as P using (_≡_ ; _≢_ ; _≗_ ; module ≡-Reasoning)
24 _∈_ : {A : Set} {n : ℕ} → A → Vec A n → Set
25 _∈_ {A} x xs = Data.List.Membership.Setoid._∈_ (P.setoid A) x (toList xs)
27 _∉_ : {A : Set} {n : ℕ} → A → Vec A n → Set
28 _∉_ {A} x xs = All (_≢_ x) (toList xs)
30 data Dec∈ {A : Set} {n : ℕ} (x : A) (xs : Vec A n) : Set where
31 yes-∈ : x ∈ xs → Dec∈ x xs
32 no-∉ : x ∉ xs → Dec∈ x xs
34 is-∈ : {A : Set} {n : ℕ} → Decidable (_≡_ {A = A}) → (x : A) → (xs : Vec A n) → Dec∈ x xs
35 is-∈ eq? x xs with Any.any (eq? x) (toList xs)
36 ... | yes x∈xs = yes-∈ x∈xs
37 ... | no x∉xs = no-∉ (Π._⟨$⟩_ (Surjection.to AllP.¬Any↠All¬) x∉xs)
39 FinMapMaybe : ℕ → Set → Set
40 FinMapMaybe n A = Vec (Maybe A) n
42 lookupM : {A : Set} {n : ℕ} → Fin n → FinMapMaybe n A → Maybe A
43 lookupM = flip lookupVec
45 insert : {A : Set} {n : ℕ} → Fin n → A → FinMapMaybe n A → FinMapMaybe n A
46 insert f a m = m [ f ]≔ (just a)
48 empty : {A : Set} {n : ℕ} → FinMapMaybe n A
49 empty = replicate nothing
51 fromAscList : {A : Set} {n m : ℕ} → Vec (Fin n × A) m → FinMapMaybe n A
52 fromAscList [] = empty
53 fromAscList ((f , a) ∷ xs) = insert f a (fromAscList xs)
55 fromFunc : {A : Set} {n : ℕ} → (Fin n → A) → FinMapMaybe n A
56 fromFunc = tabulate ∘ _∘_ Maybe.just
58 reshape : {n : ℕ} {A : Set} → FinMapMaybe n A → (l : ℕ) → FinMapMaybe l A
60 reshape [] (suc l) = nothing ∷ (reshape [] l)
61 reshape (x ∷ xs) (suc l) = x ∷ (reshape xs l)
63 union : {A : Set} {n : ℕ} → FinMapMaybe n A → FinMapMaybe n A → FinMapMaybe n A
64 union m1 m2 = tabulate (λ f → maybe′ just (lookupM f m2) (lookupM f m1))
66 restrict : {A : Set} {n m : ℕ} → (Fin n → A) → Vec (Fin n) m → FinMapMaybe n A
67 restrict f is = fromAscList (zip is (mapV f is))
69 delete : {A : Set} {n : ℕ} → Fin n → FinMapMaybe n A → FinMapMaybe n A
70 delete i m = m [ i ]≔ nothing
72 delete-many : {A : Set} {n m : ℕ} → Vec (Fin n) m → FinMapMaybe n A → FinMapMaybe n A
73 delete-many = flip (foldr (const _) delete)
75 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
76 lemma-insert-same [] () p
77 lemma-insert-same {suc n} (x ∷ xs) zero p = P.cong (flip _∷_ xs) p
78 lemma-insert-same (x ∷ xs) (suc i) p = P.cong (_∷_ x) (lemma-insert-same xs i p)
80 lemma-lookupM-restrict : {A : Set} {n m : ℕ} → (i : Fin n) → (f : Fin n → A) → (is : Vec (Fin n) m) → {a : A} → lookupM i (restrict f is) ≡ just a → f i ≡ a
81 lemma-lookupM-restrict i f [] p = contradiction (P.trans (P.sym p) (lookup-replicate i nothing)) (λ ())
82 lemma-lookupM-restrict i f (i' ∷ is) p with i ≟ i'
83 lemma-lookupM-restrict i f (.i ∷ is) {a} p | yes P.refl = just-injective (begin
85 ≡⟨ P.sym (lookup∘update i (restrict f is) (just (f i))) ⟩
86 lookupM i (insert i (f i) (restrict f is))
89 where open ≡-Reasoning
90 lemma-lookupM-restrict i f (i' ∷ is) {a} p | no i≢i' = lemma-lookupM-restrict i f is (begin
91 lookupM i (restrict f is)
92 ≡⟨ P.sym (lookup∘update′ i≢i' (restrict f is) (just (f i'))) ⟩
93 lookupM i (insert i' (f i') (restrict f is))
96 where open ≡-Reasoning
97 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)
98 lemma-lookupM-restrict-∈ i f [] ()
99 lemma-lookupM-restrict-∈ i f (j ∷ js) p with i ≟ j
100 lemma-lookupM-restrict-∈ i f (.i ∷ js) p | yes P.refl = lookup∘update i (restrict f js) (just (f i))
101 lemma-lookupM-restrict-∈ i f (j ∷ js) (Any.here i≡j) | no i≢j = contradiction i≡j i≢j
102 lemma-lookupM-restrict-∈ i f (j ∷ js) (Any.there p) | no i≢j =
103 P.trans (lookup∘update′ i≢j (restrict f js) (just (f j)))
104 (lemma-lookupM-restrict-∈ i f js p)
106 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) ≡ nothing
107 lemma-lookupM-restrict-∉ i f [] i∉[] = lookup-replicate i nothing
108 lemma-lookupM-restrict-∉ i f (j ∷ js) i∉jjs =
109 P.trans (lookup∘update′ (All.head i∉jjs) (restrict f js) (just (f j)))
110 (lemma-lookupM-restrict-∉ i f js (All.tail i∉jjs))
112 lemma-lookupM-delete-many : {n m : ℕ} {A : Set} (h : FinMapMaybe n A) → (i : Fin n) → (js : Vec (Fin n) m) → i ∉ js → lookupM i (delete-many js h) ≡ lookupM i h
113 lemma-lookupM-delete-many {n} h i [] i∉[] = P.refl
114 lemma-lookupM-delete-many {n} h i (j ∷ js) i∉jjs =
115 P.trans (lookup∘update′ (All.head i∉jjs) (delete-many js h) nothing)
116 (lemma-lookupM-delete-many h i js (All.tail i∉jjs))
118 lemma-reshape-id : {n : ℕ} {A : Set} → (m : FinMapMaybe n A) → reshape m n ≡ m
119 lemma-reshape-id [] = P.refl
120 lemma-reshape-id (x ∷ xs) = P.cong (_∷_ x) (lemma-reshape-id xs)
122 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
123 lemma-disjoint-union {n} f t = tabulate-cong inner
124 where open ≡-Reasoning
125 inner : (x : Fin n) → maybe′ just (lookupM x (delete-many t (fromFunc f))) (lookupM x (restrict f t)) ≡ just (f x)
126 inner x with is-∈ _≟_ x t
127 inner x | yes-∈ x∈t = P.cong (maybe′ just (lookupM x (delete-many t (fromFunc f)))) (lemma-lookupM-restrict-∈ x f t x∈t)
128 inner x | no-∉ x∉t = begin
129 maybe′ just (lookupM x (delete-many t (fromFunc f))) (lookupM x (restrict f t))
130 ≡⟨ P.cong₂ (maybe′ just) (lemma-lookupM-delete-many (fromFunc f) x t x∉t) (lemma-lookupM-restrict-∉ x f t x∉t) ⟩
131 maybe′ just (lookupM x (fromFunc f)) nothing
132 ≡⟨ P.cong (flip (maybe′ just) nothing) (lookup∘tabulate (just ∘ f) x) ⟩
135 lemma-exchange-maps : {n m : ℕ} → {A : Set} → (h h′ : FinMapMaybe n A) → {P : Fin n → Set} → (∀ j → P j → lookupM j h ≡ lookupM j h′) → {is : Vec (Fin n) m} → All P (toList is) → mapV (flip lookupM h) is ≡ mapV (flip lookupM h′) is
136 lemma-exchange-maps h h' h≈h′ {[]} All.[] = P.refl
137 lemma-exchange-maps h h' h≈h′ {i ∷ is} (pi All.∷ pis) = P.cong₂ _∷_ (h≈h′ i pi) (lemma-exchange-maps h h' h≈h′ pis)