move lemma-just!=nothing to FinMap and use it there
[~helmut/bidiragda.git] / Bidir.agda
1 module Bidir where
2
3 open import Data.Nat using (â„•)
4 open import Data.Fin using (Fin)
5 open import Data.Maybe using (Maybe ; nothing ; just ; maybe′)
6 open import Data.List using (List ; [] ; _∷_ ; map ; length)
7 open import Data.List.Properties using (map-cong) renaming (map-compose to map-∘)
8 open import Data.Vec using (toList ; fromList ; tabulate) renaming (lookup to lookupVec ; _∷_ to _∷V_)
9 open import Data.Vec.Properties using (tabulate-∘ ; lookup∘tabulate)
10 open import Data.Empty using (⊥-elim)
11 open import Function using (id ; _∘_ ; flip)
12 open import Relation.Nullary using (Dec ; yes ; no ; ¬_)
13 open import Relation.Nullary.Negation using (contradiction)
14 open import Relation.Binary.Core using (_≡_ ; refl)
15 open import Relation.Binary.PropositionalEquality using (cong ; sym ; inspect ; Reveal_is_ ; _≗_ ; trans)
16 open Relation.Binary.PropositionalEquality.≡-Reasoning using (begin_ ; _≡⟨_⟩_ ; _∎)
17
18 open import FinMap
19
20 _>>=_ : {A B : Set} â†’ Maybe A â†’ (A â†’ Maybe B) â†’ Maybe B
21 _>>=_ = flip (flip maybe′ nothing)
22
23 fmap : {A B : Set} â†’ (A â†’ B) â†’ Maybe A â†’ Maybe B
24 fmap f = maybe′ (λ a â†’ just (f a)) nothing
25
26 EqInst : Set â†’ Set
27 EqInst A = (x y : A) â†’ Dec (x â‰¡ y)
28
29 checkInsert : {A : Set} {n : â„•} â†’ EqInst A â†’ Fin n â†’ A â†’ FinMapMaybe n A â†’ Maybe (FinMapMaybe n A)
30 checkInsert eq i b m with lookupM i m
31 checkInsert eq i b m | just c with eq b c
32 checkInsert eq i b m | just .b | yes refl = just m
33 checkInsert eq i b m | just c  | no ¬p    = nothing
34 checkInsert eq i b m | nothing = just (insert i b m)
35 assoc : {A : Set} {n : â„•} â†’ EqInst A â†’ List (Fin n) â†’ List A â†’ Maybe (FinMapMaybe n A)
36 assoc _  []       []       = just empty
37 assoc eq (i âˆ· is) (b âˆ· bs) = (assoc eq is bs) >>= (checkInsert eq i b)
38 assoc _  _        _        = nothing
39
40 record checkInsertProof {A : Set} {n : â„•} (eq : EqInst A) (i : Fin n) (x : A) (m : FinMapMaybe n A) (P : Set) : Set where
41   field
42      same : lookupM i m â‰¡ just x â†’ P
43      new : lookupM i m â‰¡ nothing â†’ P
44      wrong : (x' : A) â†’ Â¬(x â‰¡ x') â†’ lookupM i m â‰¡ just x'  â†’ P
45
46 apply-checkInsertProof : {A P : Set} {n : â„•} â†’ (eq : EqInst A) â†’ (i : Fin n) â†’ (x : A) â†’ (m : FinMapMaybe n A) â†’ checkInsertProof eq i x m P â†’ P
47 apply-checkInsertProof eq i x m rp with lookupM i m | inspect (lookupM i) m
48 apply-checkInsertProof eq i x m rp | just x' | il with eq x x'
49 apply-checkInsertProof eq i x m rp | just .x | Reveal_is_.[_] il | yes refl = checkInsertProof.same rp il
50 apply-checkInsertProof eq i x m rp | just x' | Reveal_is_.[_] il | no x≢x' = checkInsertProof.wrong rp x' x≢x' il
51 apply-checkInsertProof eq i x m rp | nothing | Reveal_is_.[_] il = checkInsertProof.new rp il
52
53 lemma-checkInsert-same : {A : Set} {n : â„•} â†’ (eq : EqInst A) â†’ (i : Fin n) â†’ (x : A) â†’ (m : FinMapMaybe n A) â†’ lookupM i m â‰¡ just x â†’ checkInsert eq i x m â‰¡ just m
54 lemma-checkInsert-same eq i x m p with lookupM i m
55 lemma-checkInsert-same eq i x m refl | .(just x) with eq x x
56 lemma-checkInsert-same eq i x m refl | .(just x) | yes refl = refl
57 lemma-checkInsert-same eq i x m refl | .(just x) | no x≢x = contradiction refl x≢x
58
59 lemma-checkInsert-new : {A : Set} {n : â„•} â†’ (eq : EqInst A) â†’ (i : Fin n) â†’ (x : A) â†’ (m : FinMapMaybe n A) â†’ lookupM i m â‰¡ nothing â†’ checkInsert eq i x m â‰¡ just (insert i x m)
60 lemma-checkInsert-new eq i x m p with lookupM i m
61 lemma-checkInsert-new eq i x m refl | .nothing = refl
62
63 lemma-checkInsert-wrong : {A : Set} {n : â„•} â†’ (eq : EqInst A) â†’ (i : Fin n) â†’ (x : A) â†’ (m : FinMapMaybe n A) â†’ (x' : A) â†’ Â¬(x â‰¡ x') â†’ lookupM i m â‰¡ just x' â†’ checkInsert eq i x m â‰¡ nothing
64 lemma-checkInsert-wrong eq i x m x' d p with lookupM i m
65 lemma-checkInsert-wrong eq i x m x' d refl | .(just x') with eq x x'
66 lemma-checkInsert-wrong eq i x m x' d refl | .(just x') | yes q = contradiction q d
67 lemma-checkInsert-wrong eq i x m x' d refl | .(just x') | no ¬q = refl
68
69 record checkInsertEqualProof {A : Set} {n : â„•} (eq : EqInst A) (i : Fin n) (x : A) (m : FinMapMaybe n A) (e : Maybe (FinMapMaybe n A)) : Set where
70   field
71      same : lookupM i m â‰¡ just x â†’ just m â‰¡ e
72      new : lookupM i m â‰¡ nothing â†’ just (insert i x m) â‰¡ e
73      wrong : (x' : A) â†’ Â¬(x â‰¡ x') â†’ lookupM i m â‰¡ just x' â†’ nothing â‰¡ e
74
75 lift-checkInsertProof : {A : Set} {n : â„•} {eq : EqInst A} {i : Fin n} {x : A} {m : FinMapMaybe n A} {e : Maybe (FinMapMaybe n A)} â†’ checkInsertEqualProof eq i x m e â†’ checkInsertProof eq i x m (checkInsert eq i x m â‰¡ e)
76 lift-checkInsertProof {_} {_} {eq} {i} {x} {m} o = record
77   { same  = Î» p â†’ trans (lemma-checkInsert-same eq i x m p) (checkInsertEqualProof.same o p)
78   ; new   = Î» p â†’ trans (lemma-checkInsert-new eq i x m p) (checkInsertEqualProof.new o p)
79   ; wrong = Î» x' q p â†’ trans (lemma-checkInsert-wrong eq i x m x' q p) (checkInsertEqualProof.wrong o x' q p)
80   }
81
82 lemma-checkInsert-restrict : {Ï„ : Set} {n : â„•} â†’ (eq : EqInst Ï„) â†’ (f : Fin n â†’ Ï„) â†’ (i : Fin n) â†’ (is : List (Fin n)) â†’ checkInsert eq i (f i) (restrict f is) â‰¡ just (restrict f (i âˆ· is))
83 lemma-checkInsert-restrict {Ï„} eq f i is = apply-checkInsertProof eq i (f i) (restrict f is) (lift-checkInsertProof record
84   { same  = Î» lookupM≡justx â†’ cong just (lemma-insert-same (restrict f is) i (f i) lookupM≡justx)
85   ; new   = Î» lookupM≡nothing â†’ refl
86   ; wrong = Î» x' x≢x' lookupM≡justx' â†’ contradiction (lemma-lookupM-restrict i f is x' lookupM≡justx') x≢x'
87   })
88
89 lemma-1 : {Ï„ : Set} {n : â„•} â†’ (eq : EqInst Ï„) â†’ (f : Fin n â†’ Ï„) â†’ (is : List (Fin n)) â†’ assoc eq is (map f is) â‰¡ just (restrict f is)
90 lemma-1 eq f []        = refl
91 lemma-1 eq f (i âˆ· is′) = begin
92   assoc eq (i âˆ· is′) (map f (i âˆ· is′))
93     â‰¡âŸ¨ refl âŸ©
94   assoc eq is′ (map f is′) >>= checkInsert eq i (f i)
95     â‰¡âŸ¨ cong (λ m â†’ m >>= checkInsert eq i (f i)) (lemma-1 eq f is′) âŸ©
96   just (restrict f is′) >>= (checkInsert eq i (f i))
97     â‰¡âŸ¨ refl âŸ©
98   checkInsert eq i (f i) (restrict f is′)
99     â‰¡âŸ¨ lemma-checkInsert-restrict eq f i is′ âŸ©
100   just (restrict f (i âˆ· is′)) âˆŽ
101
102 lemma-lookupM-assoc : {A : Set} {n : â„•} â†’ (eq : EqInst A) â†’ (i : Fin n) â†’ (is : List (Fin n)) â†’ (x : A) â†’ (xs : List A) â†’ (h : FinMapMaybe n A) â†’ assoc eq (i âˆ· is) (x âˆ· xs) â‰¡ just h â†’ lookupM i h â‰¡ just x
103 lemma-lookupM-assoc eq i is x xs h    p with assoc eq is xs
104 lemma-lookupM-assoc eq i is x xs h    () | nothing
105 lemma-lookupM-assoc eq i is x xs h    p | just h' = apply-checkInsertProof eq i x h' record
106   { same  = Î» lookupM≡justx â†’ begin
107       lookupM i h
108         â‰¡âŸ¨ cong (lookupM i) (lemma-from-just (trans (sym p) (lemma-checkInsert-same eq i x h' lookupM≡justx))) âŸ©
109       lookupM i h'
110         â‰¡âŸ¨ lookupM≡justx âŸ©
111       just x âˆŽ
112   ; new   = Î» lookupM≡nothing â†’ begin
113       lookupM i h
114         â‰¡âŸ¨ cong (lookupM i) (lemma-from-just (trans (sym p) (lemma-checkInsert-new eq i x h' lookupM≡nothing))) âŸ©
115       lookupM i (insert i x h')
116         â‰¡âŸ¨ lemma-lookupM-insert i x h' âŸ©
117       just x âˆŽ
118   ; wrong = Î» x' x≢x' lookupM≡justx' â†’ lemma-just≢nothing (trans (sym p) (lemma-checkInsert-wrong eq i x h' x' x≢x' lookupM≡justx'))
119   }
120
121 lemma-2 : {Ï„ : Set} {n : â„•} â†’ (eq : EqInst Ï„) â†’ (is : List (Fin n)) â†’ (v : List Ï„) â†’ (h : FinMapMaybe n Ï„) â†’ assoc eq is v â‰¡ just h â†’ map (flip lookupM h) is â‰¡ map just v
122 lemma-2 eq []       []       h p = refl
123 lemma-2 eq []       (x âˆ· xs) h ()
124 lemma-2 eq (x âˆ· xs) []       h ()
125 lemma-2 eq (i âˆ· is) (x âˆ· xs) h p with assoc eq is xs | inspect (assoc eq is) xs
126 lemma-2 eq (i âˆ· is) (x âˆ· xs) h () | nothing | _
127 lemma-2 eq (i âˆ· is) (x âˆ· xs) h p | just h' | Reveal_is_.[_] ir = begin
128   map (flip lookupM h) (i âˆ· is)
129     â‰¡âŸ¨ refl âŸ©
130   lookupM i h âˆ· map (flip lookupM h) is
131     â‰¡âŸ¨ cong (flip _∷_ (map (flip lookup h) is)) (lemma-lookupM-assoc eq i is x xs h (begin
132       assoc eq (i âˆ· is) (x âˆ· xs)
133         â‰¡âŸ¨ cong (flip _>>=_ (checkInsert eq i x)) ir âŸ©
134       checkInsert eq i x h'
135         â‰¡âŸ¨ p âŸ©
136       just h âˆŽ) ) âŸ©
137   just x âˆ· map (flip lookupM h) is
138     â‰¡âŸ¨  cong (_∷_ (just x)) {!!} âŸ©
139   just x âˆ· map (flip lookupM h') is
140     â‰¡âŸ¨ cong (_∷_ (just x)) (lemma-2 eq is xs h' ir) âŸ©
141   just x âˆ· map just xs
142     â‰¡âŸ¨ refl âŸ©
143   map just (x âˆ· xs) âˆŽ
144
145 enumerate : {A : Set} â†’ (l : List A) â†’ List (Fin (length l))
146 enumerate l = toList (tabulate id)
147
148 denumerate : {A : Set} (l : List A) â†’ Fin (length l) â†’ A
149 denumerate l = flip lookupVec (fromList l)
150
151 bff : ({A : Set} â†’ List A â†’ List A) â†’ ({B : Set} â†’ EqInst B â†’ List B â†’ List B â†’ Maybe (List B))
152 bff get eq s v = let s′ = enumerate s
153                      g  = fromFunc (denumerate s)
154                      h  = assoc eq (get s′) v
155                      h′ = fmap (flip union g) h
156                  in fmap (flip map s′ âˆ˜ flip lookup) h′
157
158 postulate
159   free-theorem-list-list : {β Î³ : Set} â†’ (get : {α : Set} â†’ List Î± â†’ List Î±) â†’ (f : Î² â†’ Î³) â†’ get âˆ˜ map f â‰— map f âˆ˜ get
160
161 toList-map-commutes : {A B : Set} {n : â„•} â†’ (f : A â†’ B) â†’ (v : Data.Vec.Vec A n) â†’ (toList (Data.Vec.map f v)) â‰¡ map f (toList v)
162 toList-map-commutes f Data.Vec.[] = refl
163 toList-map-commutes f (x âˆ·V xs) = cong (_∷_ (f x)) (toList-map-commutes f xs)
164
165 lemma-map-denumerate-enumerate : {A : Set} â†’ (as : List A) â†’ map (denumerate as) (enumerate as) â‰¡ as
166 lemma-map-denumerate-enumerate [] = refl
167 lemma-map-denumerate-enumerate (a âˆ· as) = cong (_∷_ a) (begin
168   map (flip lookupVec (a âˆ·V (fromList as))) (toList (tabulate Fin.suc))
169     â‰¡âŸ¨ cong (map (flip lookupVec (a âˆ·V (fromList as))) âˆ˜ toList) (tabulate-∘ Fin.suc id) âŸ©
170   map (flip lookupVec (a âˆ·V (fromList as))) (toList (Data.Vec.map Fin.suc (tabulate id)))
171     â‰¡âŸ¨ cong (map (flip lookupVec (a âˆ·V fromList as))) (toList-map-commutes Data.Fin.suc (tabulate id)) âŸ©
172   map (flip lookupVec (a âˆ·V fromList as)) (map Fin.suc (enumerate as))
173     â‰¡âŸ¨ sym (map-∘ (enumerate as)) âŸ©
174   map (flip lookupVec (a âˆ·V (fromList as)) âˆ˜ Fin.suc) (enumerate as)
175     â‰¡âŸ¨ refl âŸ©
176   map (denumerate as) (enumerate as)
177     â‰¡âŸ¨ lemma-map-denumerate-enumerate as âŸ©
178   as âˆŽ)
179
180 theorem-1 : (get : {α : Set} â†’ List Î± â†’ List Î±) â†’ {Ï„ : Set} â†’ (eq : EqInst Ï„) â†’ (s : List Ï„) â†’ bff get eq s (get s) â‰¡ just s
181 theorem-1 get eq s = begin
182   bff get eq s (get s)
183     â‰¡âŸ¨ cong (bff get eq s âˆ˜ get) (sym (lemma-map-denumerate-enumerate s)) âŸ©
184   bff get eq s (get (map (denumerate s) (enumerate s)))
185     â‰¡âŸ¨ cong (bff get eq s) (free-theorem-list-list get (denumerate s) (enumerate s)) âŸ©
186   bff get eq s (map (denumerate s) (get (enumerate s)))
187     â‰¡âŸ¨ refl âŸ©
188   fmap (flip map (enumerate s) âˆ˜ flip lookup) (fmap (flip union (fromFunc (denumerate s))) (assoc eq (get (enumerate s)) (map (denumerate s) (get (enumerate s)))))
189     â‰¡âŸ¨ cong (fmap (flip map (enumerate s) âˆ˜ flip lookup) âˆ˜ fmap (flip union (fromFunc (denumerate s)))) (lemma-1 eq (denumerate s) (get (enumerate s))) âŸ©
190   fmap (flip map (enumerate s) âˆ˜ flip lookup) (fmap (flip union (fromFunc (flip lookupVec (fromList s)))) (just (restrict (denumerate s) (get (enumerate s)))))
191     â‰¡âŸ¨ refl âŸ©
192   just ((flip map (enumerate s) âˆ˜ flip lookup) (union (restrict (denumerate s) (get (enumerate s))) (fromFunc (denumerate s))))
193     â‰¡âŸ¨ cong just (cong (flip map (enumerate s) âˆ˜ flip lookup) (lemma-union-restrict (denumerate s) (get (enumerate s)))) âŸ©
194   just ((flip map (enumerate s) âˆ˜ flip lookup) (fromFunc (denumerate s)))
195     â‰¡âŸ¨ refl âŸ©
196   just (map (flip lookup (fromFunc (denumerate s))) (enumerate s))
197     â‰¡âŸ¨ cong just (map-cong (lookup∘tabulate (denumerate s)) (enumerate s)) âŸ©
198   just (map (denumerate s) (enumerate s))
199     â‰¡âŸ¨ cong just (lemma-map-denumerate-enumerate s) âŸ©
200   just s âˆŽ
201
202 theorem-2 : (get : {α : Set} â†’ List Î± â†’ List Î±) â†’ {Ï„ : Set} â†’ (eq : EqInst Ï„) â†’ (v s u : List Ï„) â†’ bff get eq s v â‰¡ just u â†’ get u â‰¡ v
203 theorem-2 get eq v s u p = {!!}