3 open import Data.Nat using (ℕ)
4 open import Data.Fin using (Fin)
5 open import Data.Fin.Props using (_≟_)
6 open import Data.Maybe using (Maybe ; nothing ; just ; maybe′)
7 open import Data.List using (List ; [] ; _∷_ ; map ; length)
8 open import Data.List.Properties using (map-cong ; ∷-injective) renaming (map-compose to map-∘)
9 open import Data.List.Any using (Any ; any ; here ; there)
10 open import Data.List.All using (All)
11 open Data.List.Any.Membership-≡ using (_∈_ ; _∉_)
12 open import Data.Vec using (toList ; fromList ; tabulate) renaming (lookup to lookupVec ; _∷_ to _∷V_)
13 open import Data.Vec.Properties using (tabulate-∘ ; lookup∘tabulate)
14 open import Data.Product using (∃ ; _,_ ; proj₁ ; proj₂)
15 open import Data.Empty using (⊥-elim)
16 open import Function using (id ; _∘_ ; flip)
17 open import Relation.Nullary using (Dec ; yes ; no ; _)
18 open import Relation.Nullary.Negation using (contradiction)
19 open import Relation.Binary.Core using (_≡_ ; refl)
20 open import Relation.Binary.PropositionalEquality using (cong ; sym ; inspect ; Reveal_is_ ; _≗_ ; trans)
21 open Relation.Binary.PropositionalEquality.≡-Reasoning using (begin_ ; _≡⟨_⟩_ ; _∎)
25 _>>=_ : {A B : Set} → Maybe A → (A → Maybe B) → Maybe B
26 _>>=_ = flip (flip maybe′ nothing)
28 fmap : {A B : Set} → (A → B) → Maybe A → Maybe B
29 fmap f = maybe′ (λ a → just (f a)) nothing
32 EqInst A = (x y : A) → Dec (x ≡ y)
34 checkInsert : {A : Set} {n : ℕ} → EqInst A → Fin n → A → FinMapMaybe n A → Maybe (FinMapMaybe n A)
35 checkInsert eq i b m with lookupM i m
36 checkInsert eq i b m | just c with eq b c
37 checkInsert eq i b m | just .b | yes refl = just m
38 checkInsert eq i b m | just c | no p = nothing
39 checkInsert eq i b m | nothing = just (insert i b m)
40 assoc : {A : Set} {n : ℕ} → EqInst A → List (Fin n) → List A → Maybe (FinMapMaybe n A)
41 assoc _ [] [] = just empty
42 assoc eq (i ∷ is) (b ∷ bs) = (assoc eq is bs) >>= (checkInsert eq i b)
45 record checkInsertProof {A : Set} {n : ℕ} (eq : EqInst A) (i : Fin n) (x : A) (m : FinMapMaybe n A) (P : Set) : Set where
47 same : lookupM i m ≡ just x → P
48 new : lookupM i m ≡ nothing → P
49 wrong : (x' : A) → ¬(x ≡ x') → lookupM i m ≡ just x' → P
51 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
52 apply-checkInsertProof eq i x m rp with lookupM i m | inspect (lookupM i) m
53 apply-checkInsertProof eq i x m rp | just x' | il with eq x x'
54 apply-checkInsertProof eq i x m rp | just .x | Reveal_is_.[_] il | yes refl = checkInsertProof.same rp il
55 apply-checkInsertProof eq i x m rp | just x' | Reveal_is_.[_] il | no x≢x' = checkInsertProof.wrong rp x' x≢x' il
56 apply-checkInsertProof eq i x m rp | nothing | Reveal_is_.[_] il = checkInsertProof.new rp il
58 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
59 lemma-checkInsert-same eq i x m p with lookupM i m
60 lemma-checkInsert-same eq i x m refl | .(just x) with eq x x
61 lemma-checkInsert-same eq i x m refl | .(just x) | yes refl = refl
62 lemma-checkInsert-same eq i x m refl | .(just x) | no x≢x = contradiction refl x≢x
64 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)
65 lemma-checkInsert-new eq i x m p with lookupM i m
66 lemma-checkInsert-new eq i x m refl | .nothing = refl
68 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
69 lemma-checkInsert-wrong eq i x m x' d p with lookupM i m
70 lemma-checkInsert-wrong eq i x m x' d refl | .(just x') with eq x x'
71 lemma-checkInsert-wrong eq i x m x' d refl | .(just x') | yes q = contradiction q d
72 lemma-checkInsert-wrong eq i x m x' d refl | .(just x') | no q = refl
74 record checkInsertEqualProof {A : Set} {n : ℕ} (eq : EqInst A) (i : Fin n) (x : A) (m : FinMapMaybe n A) (e : Maybe (FinMapMaybe n A)) : Set where
76 same : lookupM i m ≡ just x → just m ≡ e
77 new : lookupM i m ≡ nothing → just (insert i x m) ≡ e
78 wrong : (x' : A) → ¬(x ≡ x') → lookupM i m ≡ just x' → nothing ≡ e
80 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)
81 lift-checkInsertProof {_} {_} {eq} {i} {x} {m} o = record
82 { same = λ p → trans (lemma-checkInsert-same eq i x m p) (checkInsertEqualProof.same o p)
83 ; new = λ p → trans (lemma-checkInsert-new eq i x m p) (checkInsertEqualProof.new o p)
84 ; wrong = λ x' q p → trans (lemma-checkInsert-wrong eq i x m x' q p) (checkInsertEqualProof.wrong o x' q p)
87 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))
88 lemma-checkInsert-restrict {τ} eq f i is = apply-checkInsertProof eq i (f i) (restrict f is) (lift-checkInsertProof record
89 { same = λ lookupM≡justx → cong just (lemma-insert-same (restrict f is) i (f i) lookupM≡justx)
90 ; new = λ lookupM≡nothing → refl
91 ; wrong = λ x' x≢x' lookupM≡justx' → contradiction (lemma-lookupM-restrict i f is x' lookupM≡justx') x≢x'
94 lemma-1 : {τ : Set} {n : ℕ} → (eq : EqInst τ) → (f : Fin n → τ) → (is : List (Fin n)) → assoc eq is (map f is) ≡ just (restrict f is)
95 lemma-1 eq f [] = refl
96 lemma-1 eq f (i ∷ is′) = begin
97 assoc eq (i ∷ is′) (map f (i ∷ is′))
99 assoc eq is′ (map f is′) >>= checkInsert eq i (f i)
100 ≡⟨ cong (λ m → m >>= checkInsert eq i (f i)) (lemma-1 eq f is′) ⟩
101 just (restrict f is′) >>= (checkInsert eq i (f i))
103 checkInsert eq i (f i) (restrict f is′)
104 ≡⟨ lemma-checkInsert-restrict eq f i is′ ⟩
105 just (restrict f (i ∷ is′)) ∎
107 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
108 lemma-lookupM-assoc eq i is x xs h p with assoc eq is xs
109 lemma-lookupM-assoc eq i is x xs h () | nothing
110 lemma-lookupM-assoc eq i is x xs h p | just h' = apply-checkInsertProof eq i x h' record
111 { same = λ lookupM≡justx → begin
113 ≡⟨ cong (lookupM i) (lemma-from-just (trans (sym p) (lemma-checkInsert-same eq i x h' lookupM≡justx))) ⟩
117 ; new = λ lookupM≡nothing → begin
119 ≡⟨ cong (lookupM i) (lemma-from-just (trans (sym p) (lemma-checkInsert-new eq i x h' lookupM≡nothing))) ⟩
120 lookupM i (insert i x h')
121 ≡⟨ lemma-lookupM-insert i x h' ⟩
123 ; wrong = λ x' x≢x' lookupM≡justx' → lemma-just≢nothing (trans (sym p) (lemma-checkInsert-wrong eq i x h' x' x≢x' lookupM≡justx'))
126 lemma-lookupM-checkInsert : {A : Set} {n : ℕ} → (eq : EqInst A) → (i j : Fin n) → (x y : A) → (h h' : FinMapMaybe n A) → lookupM i h ≡ just x → checkInsert eq j y h ≡ just h' → lookupM i h' ≡ just x
127 lemma-lookupM-checkInsert eq i j x y h h' pl ph' with lookupM j h | inspect (lookupM j) h
128 lemma-lookupM-checkInsert eq i j x y h .(insert j y h) pl refl | nothing | pl' with i ≟ j
129 lemma-lookupM-checkInsert eq i .i x y h .(insert i y h) pl refl | nothing | Reveal_is_.[_] pl' | yes refl with begin just x ≡⟨ sym pl ⟩ lookupM i h ≡⟨ pl' ⟩ nothing ∎
131 lemma-lookupM-checkInsert eq i j x y h .(insert j y h) pl refl | nothing | pl' | no p = begin
132 lookupM i (insert j y h)
133 ≡⟨ sym (lemma-lookupM-insert-other i j y h ¬p) ⟩
137 lemma-lookupM-checkInsert eq i j x y h h' pl ph' | just z | pl' with eq y z
138 lemma-lookupM-checkInsert eq i j x y h h' pl ph' | just .y | pl' | yes refl = begin
140 ≡⟨ cong (lookupM i) (lemma-from-just (sym ph')) ⟩
144 lemma-lookupM-checkInsert eq i j x y h h' pl () | just z | pl' | no p
146 lemma-∉-lookupM-assoc : {A : Set} {n : ℕ} → (eq : EqInst A) → (i : Fin n) → (is : List (Fin n)) → (xs : List A) → (h : FinMapMaybe n A) → assoc eq is xs ≡ just h → (i ∉ is) → lookupM i h ≡ nothing
147 lemma-∉-lookupM-assoc eq i [] [] h ph i∉is = begin
149 ≡⟨ cong (lookupM i) (sym (lemma-from-just ph)) ⟩
151 ≡⟨ lemma-lookupM-empty i ⟩
153 lemma-∉-lookupM-assoc eq i [] (x' ∷ xs') h () i∉is
154 lemma-∉-lookupM-assoc eq i (i' ∷ is') [] h () i∉is
155 lemma-∉-lookupM-assoc eq i (i' ∷ is') (x' ∷ xs') h ph i∉is with i ≟ i'
156 lemma-∉-lookupM-assoc eq i (i' ∷ is') (x' ∷ xs') h ph i∉is | yes p = contradiction (here p) i∉is
157 lemma-∉-lookupM-assoc eq i (i' ∷ is') (x' ∷ xs') h ph i∉is | no ¬p with assoc eq is' xs' | inspect (assoc eq is') xs'
158 lemma-∉-lookupM-assoc eq i (i' ∷ is') (x' ∷ xs') h () i∉is | no ¬p | nothing | Reveal_is_.[_] ph'
159 lemma-∉-lookupM-assoc eq i (i' ∷ is') (x' ∷ xs') h ph i∉is | no ¬p | just h' | Reveal_is_.[_] ph' = apply-checkInsertProof eq i' x' h' record {
160 same = λ lookupM-i'-h'≡just-x' → begin
162 ≡⟨ cong (lookupM i) (lemma-from-just (trans (sym ph) (lemma-checkInsert-same eq i' x' h' lookupM-i'-h'≡just-x'))) ⟩
164 ≡⟨ lemma-∉-lookupM-assoc eq i is' xs' h' ph' (i∉is ∘ there) ⟩
166 ; new = λ lookupM-i'-h'≡nothing → begin
168 ≡⟨ cong (lookupM i) (lemma-from-just (trans (sym ph) (lemma-checkInsert-new eq i' x' h' lookupM-i'-h'≡nothing))) ⟩
169 lookupM i (insert i' x' h')
170 ≡⟨ sym (lemma-lookupM-insert-other i i' x' h' ¬p) ⟩
172 ≡⟨ lemma-∉-lookupM-assoc eq i is' xs' h' ph' (i∉is ∘ there) ⟩
174 ; wrong = λ x'' x'≢x'' lookupM-i'-h'≡just-x'' → lemma-just≢nothing (trans (sym ph) (lemma-checkInsert-wrong eq i' x' h' x'' x'≢x'' lookupM-i'-h'≡just-x''))
177 _in-domain-of_ : {n : ℕ} {A : Set} → (is : List (Fin n)) → (FinMapMaybe n A) → Set
178 _in-domain-of_ is h = All (λ i → ∃ λ x → lookupM i h ≡ just x) is
180 lemma-assoc-domain : {n : ℕ} {A : Set} → (eq : EqInst A) → (is : List (Fin n)) → (xs : List A) → (h : FinMapMaybe n A) → assoc eq is xs ≡ just h → is in-domain-of h
181 lemma-assoc-domain eq [] [] h ph = Data.List.All.[]
182 lemma-assoc-domain eq [] (x' ∷ xs') h ()
183 lemma-assoc-domain eq (i' ∷ is') [] h ()
184 lemma-assoc-domain eq (i' ∷ is') (x' ∷ xs') h ph with assoc eq is' xs' | inspect (assoc eq is') xs'
185 lemma-assoc-domain eq (i' ∷ is') (x' ∷ xs') h () | nothing | ph'
186 lemma-assoc-domain eq (i' ∷ is') (x' ∷ xs') h ph | just h' | Reveal_is_.[_] ph' = apply-checkInsertProof eq i' x' h' record {
187 same = λ lookupM-i'-h'≡just-x' → Data.List.All._∷_
188 (x' , (trans (cong (lookupM i') (lemma-from-just (trans (sym ph) (lemma-checkInsert-same eq i' x' h' lookupM-i'-h'≡just-x')))) lookupM-i'-h'≡just-x'))
189 (lemma-assoc-domain eq is' xs' h (trans ph' (trans (sym (lemma-checkInsert-same eq i' x' h' lookupM-i'-h'≡just-x')) ph)))
190 ; new = λ lookupM-i'-h'≡nothing → Data.List.All._∷_
191 (x' , (trans (cong (lookupM i') (lemma-from-just (trans (sym ph) (lemma-checkInsert-new eq i' x' h' lookupM-i'-h'≡nothing)))) (lemma-lookupM-insert i' x' h')))
193 (λ {i} p → proj₁ p , lemma-lookupM-checkInsert eq i i' (proj₁ p) x' h' h (proj₂ p) ph)
194 (lemma-assoc-domain eq is' xs' h' ph'))
195 ; wrong = λ x'' x'≢x'' lookupM-i'-h'≡just-x'' → lemma-just≢nothing (trans (sym ph) (lemma-checkInsert-wrong eq i' x' h' x'' x'≢x'' lookupM-i'-h'≡just-x''))
198 lemma-map-lookupM-insert : {n : ℕ} {A : Set} → (eq : EqInst A) → (i : Fin n) → (is : List (Fin n)) → (x : A) → (h : FinMapMaybe n A) → ¬(i ∈ is) → is in-domain-of h → map (flip lookupM (insert i x h)) is ≡ map (flip lookupM h) is
199 lemma-map-lookupM-insert eq i [] x h i∉is ph = refl
200 lemma-map-lookupM-insert eq i (i' ∷ is') x h i∉is ph = begin
201 lookupM i' (insert i x h) ∷ map (flip lookupM (insert i x h)) is'
202 ≡⟨ cong (flip _∷_ (map (flip lookupM (insert i x h)) is')) (sym (lemma-lookupM-insert-other i' i x h (i∉is ∘ here ∘ sym))) ⟩
203 lookupM i' h ∷ map (flip lookupM (insert i x h)) is'
204 ≡⟨ cong (_∷_ (lookupM i' h)) (lemma-map-lookupM-insert eq i is' x h (i∉is ∘ there) (Data.List.All.tail ph)) ⟩
205 lookupM i' h ∷ map (flip lookupM h) is' ∎
207 lemma-map-lookupM-assoc : {n : ℕ} {A : Set} → (eq : EqInst A) → (i : Fin n) → (is : List (Fin n)) → (x : A) → (xs : List A) → (h : FinMapMaybe n A) → (h' : FinMapMaybe n A) → assoc eq is xs ≡ just h' → checkInsert eq i x h' ≡ just h → map (flip lookupM h) is ≡ map (flip lookupM h') is
208 lemma-map-lookupM-assoc eq i [] x [] h h' ph' ph = refl
209 lemma-map-lookupM-assoc eq i [] x (x' ∷ xs') h h' () ph
210 lemma-map-lookupM-assoc eq i (i' ∷ is') x [] h h' () ph
211 lemma-map-lookupM-assoc eq i (i' ∷ is') x (x' ∷ xs') h h' ph' ph with any (_≟_ i) (i' ∷ is')
212 lemma-map-lookupM-assoc eq i (i' ∷ is') x (x' ∷ xs') h h' ph' ph | yes p with Data.List.All.lookup (lemma-assoc-domain eq (i' ∷ is') (x' ∷ xs') h' ph') p
213 lemma-map-lookupM-assoc eq i (i' ∷ is') x (x' ∷ xs') h h' ph' ph | yes p | (x'' , p') with lookupM i h'
214 lemma-map-lookupM-assoc eq i (i' ∷ is') x (x' ∷ xs') h h' ph' ph | yes p | (x'' , refl) | .(just x'') with eq x x''
215 lemma-map-lookupM-assoc eq i (i' ∷ is') x (x' ∷ xs') h .h ph' refl | yes p | (.x , refl) | .(just x) | yes refl = refl
216 lemma-map-lookupM-assoc eq i (i' ∷ is') x (x' ∷ xs') h h' ph' () | yes p | (x'' , refl) | .(just x'') | no ¬p
217 lemma-map-lookupM-assoc eq i (i' ∷ is') x (x' ∷ xs') h h' ph' ph | no ¬p with lookupM i h' | lemma-∉-lookupM-assoc eq i (i' ∷ is') (x' ∷ xs') h' ph' ¬p
218 lemma-map-lookupM-assoc eq i (i' ∷ is') x (x' ∷ xs') h h' ph' ph | no ¬p | .nothing | refl = begin
219 map (flip lookupM h) (i' ∷ is')
220 ≡⟨ map-cong (λ i'' → cong (lookupM i'') (lemma-from-just (sym ph))) (i' ∷ is') ⟩
221 map (flip lookupM (insert i x h')) (i' ∷ is')
222 ≡⟨ lemma-map-lookupM-insert eq i (i' ∷ is') x h' ¬p (lemma-assoc-domain eq (i' ∷ is') (x' ∷ xs') h' ph') ⟩
223 map (flip lookupM h') (i' ∷ is') ∎
225 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
226 lemma-2 eq [] [] h p = refl
227 lemma-2 eq [] (x ∷ xs) h ()
228 lemma-2 eq (x ∷ xs) [] h ()
229 lemma-2 eq (i ∷ is) (x ∷ xs) h p with assoc eq is xs | inspect (assoc eq is) xs
230 lemma-2 eq (i ∷ is) (x ∷ xs) h () | nothing | _
231 lemma-2 eq (i ∷ is) (x ∷ xs) h p | just h' | Reveal_is_.[_] ir = begin
232 map (flip lookupM h) (i ∷ is)
234 lookupM i h ∷ map (flip lookupM h) is
235 ≡⟨ cong (flip _∷_ (map (flip lookup h) is)) (lemma-lookupM-assoc eq i is x xs h (begin
236 assoc eq (i ∷ is) (x ∷ xs)
237 ≡⟨ cong (flip _>>=_ (checkInsert eq i x)) ir ⟩
238 checkInsert eq i x h'
241 just x ∷ map (flip lookupM h) is
242 ≡⟨ cong (_∷_ (just x)) (lemma-map-lookupM-assoc eq i is x xs h h' ir p) ⟩
243 just x ∷ map (flip lookupM h') is
244 ≡⟨ cong (_∷_ (just x)) (lemma-2 eq is xs h' ir) ⟩
249 enumerate : {A : Set} → (l : List A) → List (Fin (length l))
250 enumerate l = toList (tabulate id)
252 denumerate : {A : Set} (l : List A) → Fin (length l) → A
253 denumerate l = flip lookupVec (fromList l)
255 bff : ({A : Set} → List A → List A) → ({B : Set} → EqInst B → List B → List B → Maybe (List B))
256 bff get eq s v = let s′ = enumerate s
257 g = fromFunc (denumerate s)
258 h = assoc eq (get s′) v
259 h′ = fmap (flip union g) h
260 in fmap (flip map s′ ∘ flip lookup) h′
263 free-theorem-list-list : {β γ : Set} → (get : {α : Set} → List α → List α) → (f : β → γ) → get ∘ map f ≗ map f ∘ get
265 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)
266 toList-map-commutes f Data.Vec.[] = refl
267 toList-map-commutes f (x ∷V xs) = cong (_∷_ (f x)) (toList-map-commutes f xs)
269 lemma-map-denumerate-enumerate : {A : Set} → (as : List A) → map (denumerate as) (enumerate as) ≡ as
270 lemma-map-denumerate-enumerate [] = refl
271 lemma-map-denumerate-enumerate (a ∷ as) = cong (_∷_ a) (begin
272 map (flip lookupVec (a ∷V (fromList as))) (toList (tabulate Fin.suc))
273 ≡⟨ cong (map (flip lookupVec (a ∷V (fromList as))) ∘ toList) (tabulate-∘ Fin.suc id) ⟩
274 map (flip lookupVec (a ∷V (fromList as))) (toList (Data.Vec.map Fin.suc (tabulate id)))
275 ≡⟨ cong (map (flip lookupVec (a ∷V fromList as))) (toList-map-commutes Data.Fin.suc (tabulate id)) ⟩
276 map (flip lookupVec (a ∷V fromList as)) (map Fin.suc (enumerate as))
277 ≡⟨ sym (map-∘ (enumerate as)) ⟩
278 map (flip lookupVec (a ∷V (fromList as)) ∘ Fin.suc) (enumerate as)
280 map (denumerate as) (enumerate as)
281 ≡⟨ lemma-map-denumerate-enumerate as ⟩
284 theorem-1 : (get : {α : Set} → List α → List α) → {τ : Set} → (eq : EqInst τ) → (s : List τ) → bff get eq s (get s) ≡ just s
285 theorem-1 get eq s = begin
287 ≡⟨ cong (bff get eq s ∘ get) (sym (lemma-map-denumerate-enumerate s)) ⟩
288 bff get eq s (get (map (denumerate s) (enumerate s)))
289 ≡⟨ cong (bff get eq s) (free-theorem-list-list get (denumerate s) (enumerate s)) ⟩
290 bff get eq s (map (denumerate s) (get (enumerate s)))
292 fmap (flip map (enumerate s) ∘ flip lookup) (fmap (flip union (fromFunc (denumerate s))) (assoc eq (get (enumerate s)) (map (denumerate s) (get (enumerate s)))))
293 ≡⟨ cong (fmap (flip map (enumerate s) ∘ flip lookup) ∘ fmap (flip union (fromFunc (denumerate s)))) (lemma-1 eq (denumerate s) (get (enumerate s))) ⟩
294 fmap (flip map (enumerate s) ∘ flip lookup) (fmap (flip union (fromFunc (flip lookupVec (fromList s)))) (just (restrict (denumerate s) (get (enumerate s)))))
296 just ((flip map (enumerate s) ∘ flip lookup) (union (restrict (denumerate s) (get (enumerate s))) (fromFunc (denumerate s))))
297 ≡⟨ cong just (cong (flip map (enumerate s) ∘ flip lookup) (lemma-union-restrict (denumerate s) (get (enumerate s)))) ⟩
298 just ((flip map (enumerate s) ∘ flip lookup) (fromFunc (denumerate s)))
300 just (map (flip lookup (fromFunc (denumerate s))) (enumerate s))
301 ≡⟨ cong just (map-cong (lookup∘tabulate (denumerate s)) (enumerate s)) ⟩
302 just (map (denumerate s) (enumerate s))
303 ≡⟨ cong just (lemma-map-denumerate-enumerate s) ⟩
306 lemma-fmap-just : {A B : Set} → {f : A → B} {b : B} → (ma : Maybe A) → fmap f ma ≡ just b → ∃ λ a → ma ≡ just a
307 lemma-fmap-just (just x) fmap-f-ma≡just-b = x , refl
308 lemma-fmap-just nothing ()
310 lemma-from-map-just : {A : Set} → {xs ys : List A} → map Maybe.just xs ≡ map Maybe.just ys → xs ≡ ys
311 lemma-from-map-just {xs = []} {ys = []} p = refl
312 lemma-from-map-just {xs = []} {ys = y ∷ ys'} ()
313 lemma-from-map-just {xs = x ∷ xs'} {ys = []} ()
314 lemma-from-map-just {xs = x ∷ xs'} {ys = y ∷ ys'} p with ∷-injective p
315 lemma-from-map-just {xs = x ∷ xs'} {ys = .x ∷ ys'} p | refl , p' = cong (_∷_ x) (lemma-from-map-just p')
317 lemma-union-not-used : {n : ℕ} {A : Set} (h : FinMapMaybe n A) → (h' : FinMap n A) → (is : List (Fin n)) → is in-domain-of h → map just (map (flip lookup (union h h')) is) ≡ map (flip lookupM h) is
318 lemma-union-not-used h h' [] p = refl
319 lemma-union-not-used h h' (i ∷ is') p with Data.List.All.head p
320 lemma-union-not-used h h' (i ∷ is') p | x , lookupM-i-h≡just-x = begin
321 just (lookup i (union h h')) ∷ map just (map (flip lookup (union h h')) is')
322 ≡⟨ cong (flip _∷_ (map just (map (flip lookup (union h h')) is'))) (begin
323 just (lookup i (union h h'))
324 ≡⟨ cong just (lookup∘tabulate (λ j → maybe′ id (lookup j h') (lookupM j h)) i) ⟩
325 just (maybe′ id (lookup i h') (lookupM i h))
326 ≡⟨ cong just (cong (maybe′ id (lookup i h')) lookupM-i-h≡just-x) ⟩
327 just (maybe′ id (lookup i h') (just x))
330 ≡⟨ sym lookupM-i-h≡just-x ⟩
332 lookupM i h ∷ map just (map (flip lookup (union h h')) is')
333 ≡⟨ cong (_∷_ (lookupM i h)) (lemma-union-not-used h h' is' (Data.List.All.tail p)) ⟩
334 lookupM i h ∷ map (flip lookupM h) is' ∎
336 theorem-2 : (get : {α : Set} → List α → List α) → {τ : Set} → (eq : EqInst τ) → (v s u : List τ) → bff get eq s v ≡ just u → get u ≡ v
337 theorem-2 get eq v s u p with lemma-fmap-just (assoc eq (get (enumerate s)) v) (proj₂ (lemma-fmap-just (fmap (flip union (fromFunc (denumerate s))) (assoc eq (get (enumerate s)) v)) p))
338 theorem-2 get eq v s u p | h , ph = begin
340 ≡⟨ lemma-from-just (begin
344 ≡⟨ cong (fmap get) (sym p) ⟩
345 fmap get (bff get eq s v)
346 ≡⟨ cong (fmap get ∘ fmap (flip map (enumerate s) ∘ flip lookup) ∘ fmap (flip union (fromFunc (denumerate s)))) ph ⟩
347 fmap get (fmap (flip map (enumerate s) ∘ flip lookup) (fmap (flip union (fromFunc (denumerate s))) (just h)))
349 just (get (map (flip lookup (union h (fromFunc (denumerate s)))) (enumerate s)))
351 get (map (flip lookup (union h (fromFunc (denumerate s)))) (enumerate s))
352 ≡⟨ free-theorem-list-list get (flip lookup (union h (fromFunc (denumerate s)))) (enumerate s) ⟩
353 map (flip lookup (union h (fromFunc (denumerate s)))) (get (enumerate s))
354 ≡⟨ lemma-from-map-just (begin
355 map just (map (flip lookup (union h (fromFunc (denumerate s)))) (get (enumerate s)))
356 ≡⟨ lemma-union-not-used h (fromFunc (denumerate s)) (get (enumerate s)) (lemma-assoc-domain eq (get (enumerate s)) v h ph) ⟩
357 map (flip lookupM h) (get (enumerate s))
358 ≡⟨ lemma-2 eq (get (enumerate s)) v h ph ⟩