reorganize equality imports
[~helmut/bidiragda.git] / FinMap.agda
1 module FinMap where
2
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.Fin using (Fin ; zero ; suc)
7 open import Data.Fin.Properties using (_≟_)
8 open import Data.Vec using (Vec ; [] ; _∷_ ; _[_]≔_ ; replicate ; tabulate ; foldr ; zip ; toList) renaming (lookup to lookupVec ; map to mapV)
9 open import Data.Vec.Equality using ()
10 open import Data.Vec.Properties using (lookup∘update ; lookup∘update′)
11 open import Data.Product using (_×_ ; _,_)
12 open import Data.List.All as All using (All)
13 import Data.List.All.Properties as AllP
14 import Data.List.Any as Any
15 import Data.List.Any.Membership.Propositional
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.Core using (Decidable)
22 open import Relation.Binary.PropositionalEquality as P using (_≡_ ; _≢_ ; _≗_)
23 open P.≡-Reasoning using (begin_ ; _≡⟨_⟩_ ; _∎)
24
25 open import Generic using (just-injective)
26
27 _∈_ : {A : Set} {n : â„•} â†’ A â†’ Vec A n â†’ Set
28 _∈_ {A} x xs = x Data.List.Any.Membership.Propositional.∈ (toList xs)
29
30 _∉_ : {A : Set} {n : â„•} â†’ A â†’ Vec A n â†’ Set
31 _∉_ {A} x xs = All (_≢_ x) (toList xs)
32
33 data Dec∈ {A : Set} {n : â„•} (x : A) (xs : Vec A n) : Set where
34   yes-∈ : x âˆˆ xs â†’ Dec∈ x xs
35   no-∉ : x âˆ‰ xs â†’ Dec∈ x xs
36
37 is-∈ : {A : Set} {n : â„•} â†’ Decidable (_≡_ {A = A}) â†’ (x : A) â†’ (xs : Vec A n) â†’ Dec∈ x xs
38 is-∈ eq? x xs with Any.any (eq? x) (toList xs)
39 ...     | yes x∈xs = yes-∈ x∈xs
40 ...     | no  x∉xs = no-∉ (Π._⟨$⟩_ (Surjection.to AllP.¬Any↠All¬) x∉xs)
41
42 FinMapMaybe : â„• â†’ Set â†’ Set
43 FinMapMaybe n A = Vec (Maybe A) n
44
45 lookupM : {A : Set} {n : â„•} â†’ Fin n â†’ FinMapMaybe n A â†’ Maybe A
46 lookupM = lookupVec
47
48 insert : {A : Set} {n : â„•} â†’ Fin n â†’ A â†’ FinMapMaybe n A â†’ FinMapMaybe n A
49 insert f a m = m [ f ]≔ (just a)
50
51 empty : {A : Set} {n : â„•} â†’ FinMapMaybe n A
52 empty = replicate nothing
53
54 fromAscList : {A : Set} {n m : â„•} â†’ Vec (Fin n Ã— A) m â†’ FinMapMaybe n A
55 fromAscList []             = empty
56 fromAscList ((f , a) âˆ· xs) = insert f a (fromAscList xs)
57
58 fromFunc : {A : Set} {n : â„•} â†’ (Fin n â†’ A) â†’ FinMapMaybe n A
59 fromFunc = tabulate âˆ˜ _∘_ Maybe.just
60
61 reshape : {n : â„•} {A : Set} â†’ FinMapMaybe n A â†’ (l : â„•) â†’ FinMapMaybe l A
62 reshape m        zero    = []
63 reshape []       (suc l) = nothing âˆ· (reshape [] l)
64 reshape (x âˆ· xs) (suc l) = x âˆ· (reshape xs l)
65
66 union : {A : Set} {n : â„•} â†’ FinMapMaybe n A â†’ FinMapMaybe n A â†’ FinMapMaybe n A
67 union m1 m2 = tabulate (λ f â†’ maybe′ just (lookupM f m2) (lookupM f m1))
68
69 restrict : {A : Set} {n m : â„•} â†’ (Fin n â†’ A) â†’ Vec (Fin n) m â†’ FinMapMaybe n A
70 restrict f is = fromAscList (zip is (mapV f is))
71
72 delete : {A : Set} {n : â„•} â†’ Fin n â†’ FinMapMaybe n A â†’ FinMapMaybe n A
73 delete i m = m [ i ]≔ nothing
74
75 delete-many : {A : Set} {n m : â„•} â†’ Vec (Fin n) m â†’ FinMapMaybe n A â†’ FinMapMaybe n A
76 delete-many = flip (foldr (const _) delete)
77
78 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
79 lemma-insert-same         []       ()      p
80 lemma-insert-same {suc n} (x âˆ· xs) zero    p = P.cong (flip _∷_ xs) p
81 lemma-insert-same         (x âˆ· xs) (suc i) p = P.cong (_∷_ x) (lemma-insert-same xs i p)
82
83 lemma-lookupM-empty : {A : Set} {n : â„•} â†’ (i : Fin n) â†’ lookupM {A} i empty â‰¡ nothing
84 lemma-lookupM-empty zero    = P.refl
85 lemma-lookupM-empty (suc i) = lemma-lookupM-empty i
86
87 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
88 lemma-lookupM-restrict i f []            p = contradiction (P.trans (P.sym p) (lemma-lookupM-empty i)) (λ ())
89 lemma-lookupM-restrict i f (i' âˆ· is)     p with i â‰Ÿ i'
90 lemma-lookupM-restrict i f (.i âˆ· is) {a} p | yes P.refl = just-injective (begin
91    just (f i)
92      â‰¡âŸ¨ P.sym (lookup∘update i (restrict f is) (just (f i))) âŸ©
93    lookupM i (insert i (f i) (restrict f is))
94      â‰¡âŸ¨ p âŸ©
95    just a âˆŽ)
96 lemma-lookupM-restrict i f (i' âˆ· is) {a} p | no i≢i' = lemma-lookupM-restrict i f is (begin
97   lookupM i (restrict f is)
98     â‰¡âŸ¨ P.sym (lookup∘update′ i≢i' (restrict f is) (just (f i'))) âŸ©
99   lookupM i (insert i' (f i') (restrict f is))
100     â‰¡âŸ¨ p âŸ©
101   just a âˆŽ)
102 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)
103 lemma-lookupM-restrict-∈ i f [] ()
104 lemma-lookupM-restrict-∈ i f (j âˆ· js)  p             with i â‰Ÿ j
105 lemma-lookupM-restrict-∈ i f (.i âˆ· js) p             | yes P.refl = lookup∘update i (restrict f js) (just (f i))
106 lemma-lookupM-restrict-∈ i f (j âˆ· js) (Any.here i≡j) | no i≢j = contradiction i≡j i≢j
107 lemma-lookupM-restrict-∈ i f (j âˆ· js) (Any.there p)  | no i≢j =
108   P.trans (lookup∘update′ i≢j (restrict f js) (just (f j)))
109           (lemma-lookupM-restrict-∈ i f js p)
110
111 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
112 lemma-lookupM-restrict-∉ i f []       i∉[]  = lemma-lookupM-empty i
113 lemma-lookupM-restrict-∉ i f (j âˆ· js) i∉jjs =
114   P.trans (lookup∘update′ (All.head i∉jjs) (restrict f js) (just (f j)))
115           (lemma-lookupM-restrict-∉ i f js (All.tail i∉jjs))
116
117 lemma-tabulate-∘ : {n : â„•} {A : Set} â†’ {f g : Fin n â†’ A} â†’ f â‰— g â†’ tabulate f â‰¡ tabulate g
118 lemma-tabulate-∘ {zero}  {_} {f} {g} f≗g = P.refl
119 lemma-tabulate-∘ {suc n} {_} {f} {g} f≗g = P.congâ‚‚ _∷_ (f≗g zero) (lemma-tabulate-∘ (f≗g âˆ˜ suc))
120
121 lemma-lookupM-fromFunc : {n : â„•} {A : Set} â†’ (f : Fin n â†’ A) â†’ flip lookupM (fromFunc f) â‰— Maybe.just âˆ˜ f
122 lemma-lookupM-fromFunc f zero = P.refl
123 lemma-lookupM-fromFunc f (suc i) = lemma-lookupM-fromFunc (f âˆ˜ suc) i
124
125 lemma-lookupM-delete : {n : â„•} {A : Set} {i j : Fin n} â†’ (f : FinMapMaybe n A) â†’ i â‰¢ j â†’ lookupM i (delete j f) â‰¡ lookupM i f
126 lemma-lookupM-delete {i = zero}  {j = zero}  (_ âˆ· _)  p = contradiction P.refl p
127 lemma-lookupM-delete {i = zero}  {j = suc j} (_ âˆ· _)  p = P.refl
128 lemma-lookupM-delete {i = suc i} {j = zero}  (x âˆ· xs) p = P.refl
129 lemma-lookupM-delete {i = suc i} {j = suc j} (x âˆ· xs) p = lemma-lookupM-delete xs (p âˆ˜ P.cong suc)
130
131 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
132 lemma-lookupM-delete-many {n} h i []       i∉[]  = P.refl
133 lemma-lookupM-delete-many {n} h i (j âˆ· js) i∉jjs =
134   P.trans (lemma-lookupM-delete (delete-many js h) (All.head i∉jjs))
135           (lemma-lookupM-delete-many h i js (All.tail i∉jjs))
136
137 lemma-reshape-id : {n : â„•} {A : Set} â†’ (m : FinMapMaybe n A) â†’ reshape m n â‰¡ m
138 lemma-reshape-id []       = P.refl
139 lemma-reshape-id (x âˆ· xs) = P.cong (_∷_ x) (lemma-reshape-id xs)
140
141 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
142 lemma-disjoint-union {n} f t = lemma-tabulate-∘ inner
143   where inner : (x : Fin n) â†’ maybe′ just (lookupM x (delete-many t (fromFunc f))) (lookupM x (restrict f t)) â‰¡ just (f x)
144         inner x with is-∈ _≟_ x t
145         inner x | yes-∈ x∈t = P.cong (maybe′ just (lookupM x (delete-many t (fromFunc f)))) (lemma-lookupM-restrict-∈ x f t x∈t)
146         inner x | no-∉ x∉t = begin
147           maybe′ just (lookupM x (delete-many t (fromFunc f))) (lookupM x (restrict f t))
148             â‰¡âŸ¨ P.congâ‚‚ (maybe′ just) (lemma-lookupM-delete-many (fromFunc f) x t x∉t) (lemma-lookupM-restrict-∉ x f t x∉t) âŸ©
149           maybe′ just (lookupM x (fromFunc f)) nothing
150             â‰¡âŸ¨ lemma-lookupM-fromFunc f x âŸ©
151           just (f x) âˆŽ
152
153 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
154 lemma-exchange-maps h≈h′ {[]}     All.[]         = P.refl
155 lemma-exchange-maps h≈h′ {i âˆ· is} (pi All.∷ pis) = P.congâ‚‚ _∷_ (h≈h′ i pi) (lemma-exchange-maps h≈h′ pis)