give the type of different gets a name
[~helmut/bidiragda.git] / FinMap.agda
1 module FinMap where
2
3 open import Data.Nat using (â„• ; zero ; suc)
4 open import Data.Maybe using (Maybe ; just ; nothing ; maybe′)
5 open import Data.Fin using (Fin ; zero ; suc)
6 open import Data.Fin.Props using (_≟_)
7 open import Data.Vec using (Vec ; [] ; _∷_ ; _[_]≔_ ; replicate ; tabulate) renaming (lookup to lookupVec)
8 open import Data.Vec.Properties using (lookup∘tabulate)
9 open import Data.List using (List ; [] ; _∷_ ; map ; zip)
10 open import Data.Product using (_×_ ; _,_)
11 open import Function using (id ; _∘_ ; flip)
12 open import Relation.Nullary using (¬_ ; yes ; no)
13 open import Relation.Nullary.Negation using (contradiction ; contraposition)
14 open import Relation.Binary.Core using (_≡_ ; refl)
15 open import Relation.Binary.PropositionalEquality using (cong ; sym ; _≗_ ; trans)
16 open Relation.Binary.PropositionalEquality.≡-Reasoning using (begin_ ; _≡⟨_⟩_ ; _∎)
17
18 FinMapMaybe : â„• â†’ Set â†’ Set
19 FinMapMaybe n A = Vec (Maybe A) n
20
21 lookupM : {A : Set} {n : â„•} â†’ Fin n â†’ FinMapMaybe n A â†’ Maybe A
22 lookupM = lookupVec
23
24 insert : {A : Set} {n : â„•} â†’ Fin n â†’ A â†’ FinMapMaybe n A â†’ FinMapMaybe n A
25 insert f a m = m [ f ]≔ (just a)
26
27 empty : {A : Set} {n : â„•} â†’ FinMapMaybe n A
28 empty = replicate nothing
29
30 fromAscList : {A : Set} {n : â„•} â†’ List (Fin n Ã— A) â†’ FinMapMaybe n A
31 fromAscList []             = empty
32 fromAscList ((f , a) âˆ· xs) = insert f a (fromAscList xs)
33
34 FinMap : â„• â†’ Set â†’ Set
35 FinMap n A = Vec A n
36
37 lookup : {A : Set} {n : â„•} â†’ Fin n â†’ FinMap n A â†’ A
38 lookup = lookupVec
39
40 fromFunc : {A : Set} {n : â„•} â†’ (Fin n â†’ A) â†’ FinMap n A
41 fromFunc = tabulate
42
43 union : {A : Set} {n : â„•} â†’ FinMapMaybe n A â†’ FinMap n  A â†’ FinMap n A
44 union m1 m2 = fromFunc (λ f â†’ maybe′ id (lookup f m2) (lookupM f m1))
45
46 restrict : {A : Set} {n : â„•} â†’ (Fin n â†’ A) â†’ List (Fin n) â†’ FinMapMaybe n A
47 restrict f is = fromAscList (zip is (map f is))
48
49 lemma-just≢nothing : {A Whatever : Set} {a : A} â†’ _≡_ {_} {Maybe A} (just a) nothing â†’ Whatever
50 lemma-just≢nothing ()
51
52 lemma-insert-same : {Ï„ : Set} {n : â„•} â†’ (m : FinMapMaybe n Ï„) â†’ (f : Fin n) â†’ (a : Ï„) â†’ lookupM f m â‰¡ just a â†’ m â‰¡ insert f a m
53 lemma-insert-same []               ()      a p
54 lemma-insert-same (.(just a) âˆ· xs) zero    a refl = refl
55 lemma-insert-same (x âˆ· xs)         (suc i) a p    = cong (_∷_ x) (lemma-insert-same xs i a p)
56
57 lemma-lookupM-empty : {A : Set} {n : â„•} â†’ (i : Fin n) â†’ lookupM {A} i empty â‰¡ nothing
58 lemma-lookupM-empty zero    = refl
59 lemma-lookupM-empty (suc i) = lemma-lookupM-empty i
60
61 lemma-lookupM-insert : {A : Set} {n : â„•} â†’ (i : Fin n) â†’ (a : A) â†’ (m : FinMapMaybe n A) â†’ lookupM i (insert i a m) â‰¡ just a
62 lemma-lookupM-insert zero    _ (_ âˆ· _)  = refl
63 lemma-lookupM-insert (suc i) a (_ âˆ· xs) = lemma-lookupM-insert i a xs
64
65 lemma-lookupM-insert-other : {A : Set} {n : â„•} â†’ (i j : Fin n) â†’ (a : A) â†’ (m : FinMapMaybe n A) â†’ Â¬(i â‰¡ j) â†’ lookupM i m â‰¡ lookupM i (insert j a m)
66 lemma-lookupM-insert-other zero    zero    a m        p = contradiction refl p
67 lemma-lookupM-insert-other zero    (suc j) a (x âˆ· xs) p = refl
68 lemma-lookupM-insert-other (suc i) zero    a (x âˆ· xs) p = refl
69 lemma-lookupM-insert-other (suc i) (suc j) a (x âˆ· xs) p = lemma-lookupM-insert-other i j a xs (contraposition (cong suc) p)
70
71 lemma-from-just : {A : Set} â†’ {x y : A} â†’ _≡_ {_} {Maybe A} (just x) (just y) â†’ x â‰¡ y
72 lemma-from-just refl = refl
73
74 lemma-lookupM-restrict : {A : Set} {n : â„•} â†’ (i : Fin n) â†’ (f : Fin n â†’ A) â†’ (is : List (Fin n)) â†’ (a : A) â†’ lookupM i (restrict f is) â‰¡ just a â†’ f i â‰¡ a
75 lemma-lookupM-restrict i f [] a p = lemma-just≢nothing (trans (sym p) (lemma-lookupM-empty i))
76 lemma-lookupM-restrict i f (i' âˆ· is) a p with i â‰Ÿ i'
77 lemma-lookupM-restrict i f (.i âˆ· is) a p | yes refl = lemma-from-just (begin
78    just (f i)
79      â‰¡âŸ¨ sym (lemma-lookupM-insert i (f i) (restrict f is)) âŸ©
80    lookupM i (insert i (f i) (restrict f is))
81      â‰¡âŸ¨ refl âŸ©
82    lookupM i (restrict f (i âˆ· is))
83      â‰¡âŸ¨ p âŸ©
84    just a âˆŽ)
85 lemma-lookupM-restrict i f (i' âˆ· is) a p | no Â¬p2 = lemma-lookupM-restrict i f is a (begin
86   lookupM i (restrict f is)
87     â‰¡âŸ¨ lemma-lookupM-insert-other i i' (f i') (restrict f is) Â¬p2 âŸ©
88   lookupM i (insert i' (f i') (restrict f is))
89     â‰¡âŸ¨ refl âŸ©
90   lookupM i (restrict f (i' âˆ· is))
91     â‰¡âŸ¨ p âŸ©
92   just a âˆŽ)
93
94 lemma-tabulate-∘ : {n : â„•} {A : Set} â†’ {f g : Fin n â†’ A} â†’ f â‰— g â†’ tabulate f â‰¡ tabulate g
95 lemma-tabulate-∘ {zero}  {_} {f} {g} f≗g = refl
96 lemma-tabulate-∘ {suc n} {_} {f} {g} f≗g = begin
97   f zero âˆ· tabulate (f âˆ˜ suc)
98     â‰¡âŸ¨ cong (flip Vec._∷_ (tabulate (f âˆ˜ suc))) (f≗g zero) âŸ©
99   g zero âˆ· tabulate (f âˆ˜ suc)
100     â‰¡âŸ¨ cong (Vec._∷_ (g zero)) (lemma-tabulate-∘ (f≗g âˆ˜ suc)) âŸ©
101   g zero âˆ· tabulate (g âˆ˜ suc) âˆŽ
102
103 lemma-union-restrict : {n : â„•} {A : Set} â†’ (f : Fin n â†’ A) â†’ (is : List (Fin n)) â†’ union (restrict f is) (fromFunc f) â‰¡ fromFunc f
104 lemma-union-restrict f is = begin
105   union (restrict f is) (fromFunc f)
106     â‰¡âŸ¨ refl âŸ©
107   tabulate (λ j â†’ maybe′ id (lookup j (fromFunc f)) (lookupM j (restrict f is)))
108     â‰¡âŸ¨ lemma-tabulate-∘ (lemma-inner f is) âŸ©
109   tabulate f âˆŽ
110     where lemma-inner : {n : â„•} {A : Set} (f : Fin n â†’ A) â†’ (is : List (Fin n)) â†’ (j : Fin n) â†’ maybe′ id (lookup j (fromFunc f)) (lookupM j (restrict f is)) â‰¡ f j
111           lemma-inner f []       j = begin
112             maybe′ id (lookup j (fromFunc f)) (lookupM j empty)
113               â‰¡âŸ¨ cong (maybe′ id (lookup j (fromFunc f))) (lemma-lookupM-empty j) âŸ©
114             maybe′ id (lookup j (fromFunc f)) nothing
115               â‰¡âŸ¨ refl âŸ©
116             lookup j (fromFunc f)
117               â‰¡âŸ¨ lookup∘tabulate f j âŸ©
118             f j âˆŽ
119           lemma-inner f (i âˆ· is)  j with j â‰Ÿ i
120           lemma-inner f (.j âˆ· is) j | yes refl = cong (maybe′ id (lookup j (fromFunc f))) (lemma-lookupM-insert j (f j) (restrict f is))
121           lemma-inner f (i âˆ· is)  j | no j≢i = begin
122             maybe′ id (lookup j (fromFunc f)) (lookupM j (insert i (f i) (restrict f is)))
123               â‰¡âŸ¨ cong (maybe′ id (lookup j (fromFunc f))) (sym (lemma-lookupM-insert-other j i (f i) (restrict f is) j≢i)) âŸ©
124             maybe′ id (lookup j (fromFunc f)) (lookupM j (restrict f is))
125               â‰¡âŸ¨ lemma-inner f is j âŸ©
126             f j âˆŽ