1 open import Level using () renaming (zero to ℓ₀)
2 open import Relation.Binary using (DecSetoid)
4 module Bidir (A : DecSetoid ℓ₀ ℓ₀) where
6 open import Data.Nat using (ℕ)
7 open import Data.Fin using (Fin)
10 import Category.Functor
11 open import Data.Maybe using (Maybe ; nothing ; just ; maybe′ ; drop-just) renaming (setoid to MaybeSetoid ; Eq to MaybeEq)
12 open Category.Monad.RawMonad {Level.zero} Data.Maybe.monad using (_>>=_)
13 open Category.Functor.RawFunctor {Level.zero} Data.Maybe.functor using (_<$>_)
14 open import Data.List using (List)
15 open import Data.List.All using (All)
16 open import Data.Vec using (Vec ; [] ; _∷_ ; toList ; map ; tabulate) renaming (lookup to lookupVec)
17 open import Data.Vec.Equality using () renaming (module Equality to VecEq)
18 open import Data.Vec.Properties using (tabulate-∘ ; lookup∘tabulate ; map-cong ; map-∘)
19 open import Data.Product using (∃ ; _×_ ; _,_ ; proj₁ ; proj₂)
20 open import Function using (id ; _∘_ ; flip)
21 open import Relation.Binary.Core using (refl ; _≡_)
22 open import Relation.Binary.Indexed using (_at_) renaming (Setoid to ISetoid)
23 open import Relation.Binary.PropositionalEquality using (cong ; sym ; inspect ; [_] ; trans ; cong₂ ; decSetoid ; module ≡-Reasoning) renaming (setoid to EqSetoid)
24 open import Relation.Binary using (Setoid ; module Setoid ; module DecSetoid)
25 import Relation.Binary.EqReasoning as EqR
28 open FreeTheorems.VecVec using (get-type ; free-theorem)
29 open import Generic using (mapMV ; mapMV-cong ; mapMV-purity ; sequenceV ; sequence-map ; VecISetoid)
34 open BFF.VecBFF A using (assoc ; enumerate ; denumerate ; bff)
35 open Setoid using () renaming (_≈_ to _∋_≈_)
36 open module A = DecSetoid A using (Carrier) renaming (_≟_ to deq)
38 module SetoidReasoning where
40 infixr 2 _≈⟨_⟩_ _≡⟨_⟩_
42 begin⟨_⟩_ : (X : Setoid ℓ₀ ℓ₀) → {x y : Setoid.Carrier X} → EqR._IsRelatedTo_ X x y → X ∋ x ≈ y
43 begin⟨_⟩_ X p = EqR.begin_ X p
44 _∎ : {X : Setoid ℓ₀ ℓ₀} → (x : Setoid.Carrier X) → EqR._IsRelatedTo_ X x x
46 _≈⟨_⟩_ : {X : Setoid ℓ₀ ℓ₀} → (x : Setoid.Carrier X) → {y z : Setoid.Carrier X} → X ∋ x ≈ y → EqR._IsRelatedTo_ X y z → EqR._IsRelatedTo_ X x z
47 _≈⟨_⟩_ {X} = EqR._≈⟨_⟩_ X
49 _≡⟨_⟩_ : {X : Setoid ℓ₀ ℓ₀} → (x : Setoid.Carrier X) → {y z : Setoid.Carrier X} → x ≡ y → EqR._IsRelatedTo_ X y z → EqR._IsRelatedTo_ X x z
50 _≡⟨_⟩_ {X} = EqR._≡⟨_⟩_ X
52 lemma-1 : {m n : ℕ} → (f : Fin n → Carrier) → (is : Vec (Fin n) m) → assoc is (map f is) ≡ just (restrict f (toList is))
54 lemma-1 f (i ∷ is′) = begin
55 (assoc is′ (map f is′) >>= checkInsert i (f i))
56 ≡⟨ cong (λ m → m >>= checkInsert i (f i)) (lemma-1 f is′) ⟩
57 checkInsert i (f i) (restrict f (toList is′))
58 ≡⟨ lemma-checkInsert-restrict f i (toList is′) ⟩
59 just (restrict f (toList (i ∷ is′))) ∎
60 where open ≡-Reasoning
62 lemma-lookupM-assoc : {m n : ℕ} → (i : Fin n) → (is : Vec (Fin n) m) → (x : Carrier) → (xs : Vec Carrier m) → (h : FinMapMaybe n Carrier) → assoc (i ∷ is) (x ∷ xs) ≡ just h → MaybeSetoid A.setoid ∋ lookupM i h ≈ just x
63 lemma-lookupM-assoc i is x xs h p with assoc is xs
64 lemma-lookupM-assoc i is x xs h () | nothing
65 lemma-lookupM-assoc i is x xs h p | just h' with checkInsert i x h' | insertionresult i x h'
66 lemma-lookupM-assoc i is x xs .h refl | just h | ._ | same x' x≈x' pl = begin
70 ≈⟨ MaybeEq.just (Setoid.sym A.setoid x≈x') ⟩
72 where open EqR (MaybeSetoid A.setoid)
73 lemma-lookupM-assoc i is x xs ._ refl | just h' | ._ | new _ = Setoid.reflexive (MaybeSetoid A.setoid) (lemma-lookupM-insert i x h')
74 lemma-lookupM-assoc i is x xs h () | just h' | ._ | wrong _ _ _
76 _in-domain-of_ : {n : ℕ} {A : Set} → (is : List (Fin n)) → (FinMapMaybe n A) → Set
77 _in-domain-of_ is h = All (λ i → ∃ λ x → lookupM i h ≡ just x) is
79 lemma-assoc-domain : {m n : ℕ} → (is : Vec (Fin n) m) → (xs : Vec Carrier m) → (h : FinMapMaybe n Carrier) → assoc is xs ≡ just h → (toList is) in-domain-of h
80 lemma-assoc-domain [] [] h ph = Data.List.All.[]
81 lemma-assoc-domain (i' ∷ is') (x' ∷ xs') h ph with assoc is' xs' | inspect (assoc is') xs'
82 lemma-assoc-domain (i' ∷ is') (x' ∷ xs') h () | nothing | [ ph' ]
83 lemma-assoc-domain (i' ∷ is') (x' ∷ xs') h ph | just h' | [ ph' ] with checkInsert i' x' h' | inspect (checkInsert i' x') h' | insertionresult i' x' h'
84 lemma-assoc-domain (i' ∷ is') (x' ∷ xs') .h refl | just h | [ ph' ] | ._ | _ | same x _ pl = All._∷_ (x , pl) (lemma-assoc-domain is' xs' h ph')
85 lemma-assoc-domain (i' ∷ is') (x' ∷ xs') ._ refl | just h' | [ ph' ] | ._ | [ cI≡ ] | new _ = All._∷_
86 (x' , lemma-lookupM-insert i' x' h')
88 (λ {i} p → proj₁ p , lemma-lookupM-checkInsert i i' (proj₁ p) x' h' (insert i' x' h') (proj₂ p) cI≡)
89 (lemma-assoc-domain is' xs' h' ph'))
90 lemma-assoc-domain (i' ∷ is') (x' ∷ xs') h () | just h' | [ ph' ] | ._ | _ | wrong _ _ _
92 lemma-map-lookupM-assoc : {m : ℕ} → (i : Fin m) → (x : Carrier) → (h : FinMapMaybe m Carrier) → (h' : FinMapMaybe m Carrier) → checkInsert i x h' ≡ just h → {n : ℕ} → (js : Vec (Fin m) n) → (toList js) in-domain-of h' → map (flip lookupM h) js ≡ map (flip lookupM h') js
93 lemma-map-lookupM-assoc i x h h' ph [] pj = refl
94 lemma-map-lookupM-assoc i x h h' ph (j ∷ js) (Data.List.All._∷_ (x' , pl) pj) = cong₂ _∷_
95 (trans (lemma-lookupM-checkInsert j i x' x h' h pl ph) (sym pl))
96 (lemma-map-lookupM-assoc i x h h' ph js pj)
98 lemma-2 : {m n : ℕ} → (is : Vec (Fin n) m) → (v : Vec Carrier m) → (h : FinMapMaybe n Carrier) → assoc is v ≡ just h → VecISetoid (MaybeSetoid A.setoid) at _ ∋ map (flip lookupM h) is ≈ map just v
99 lemma-2 [] [] h p = ISetoid.refl (VecISetoid (MaybeSetoid A.setoid))
100 lemma-2 (i ∷ is) (x ∷ xs) h p with assoc is xs | inspect (assoc is) xs
101 lemma-2 (i ∷ is) (x ∷ xs) h () | nothing | _
102 lemma-2 (i ∷ is) (x ∷ xs) h p | just h' | [ ir ] = begin
103 lookupM i h ∷ map (flip lookupM h) is
104 ≈⟨ lemma-lookupM-assoc i is x xs h (trans (cong (flip _>>=_ (checkInsert i x)) ir) p) VecEq.∷-cong ISetoid.refl (VecISetoid (MaybeSetoid A.setoid)) ⟩
105 just x ∷ map (flip lookupM h) is
106 ≡⟨ cong (_∷_ (just x)) (lemma-map-lookupM-assoc i x h h' p is (lemma-assoc-domain is xs h' ir)) ⟩
107 just x ∷ map (flip lookupM h') is
108 ≈⟨ Setoid.refl (MaybeSetoid A.setoid) VecEq.∷-cong (lemma-2 is xs h' ir) ⟩
109 just x ∷ map just xs ∎
110 where open EqR (VecISetoid (MaybeSetoid A.setoid) at _)
112 lemma-map-denumerate-enumerate : {m : ℕ} → (as : Vec Carrier m) → map (denumerate as) (enumerate as) ≡ as
113 lemma-map-denumerate-enumerate [] = refl
114 lemma-map-denumerate-enumerate (a ∷ as) = cong (_∷_ a) (begin
115 map (flip lookupVec (a ∷ as)) (tabulate Fin.suc)
116 ≡⟨ cong (map (flip lookupVec (a ∷ as))) (tabulate-∘ Fin.suc id) ⟩
117 map (flip lookupVec (a ∷ as)) (map Fin.suc (tabulate id))
119 map (flip lookupVec (a ∷ as)) (map Fin.suc (enumerate as))
120 ≡⟨ sym (map-∘ _ _ (enumerate as)) ⟩
121 map (flip lookupVec (a ∷ as) ∘ Fin.suc) (enumerate as)
123 map (denumerate as) (enumerate as)
124 ≡⟨ lemma-map-denumerate-enumerate as ⟩
126 where open ≡-Reasoning
128 theorem-1 : {getlen : ℕ → ℕ} → (get : get-type getlen) → {m : ℕ} → (s : Vec Carrier m) → bff get s (get s) ≡ just s
129 theorem-1 get s = begin
131 ≡⟨ cong (bff get s ∘ get) (sym (lemma-map-denumerate-enumerate s)) ⟩
132 bff get s (get (map (denumerate s) (enumerate s)))
133 ≡⟨ cong (bff get s) (free-theorem get (denumerate s) (enumerate s)) ⟩
134 bff get s (map (denumerate s) (get (enumerate s)))
136 (h′↦r ∘ h↦h′) (assoc (get (enumerate s)) (map (denumerate s) (get (enumerate s))))
137 ≡⟨ cong (h′↦r ∘ h↦h′) (lemma-1 (denumerate s) (get (enumerate s))) ⟩
138 (h′↦r ∘ h↦h′ ∘ just) (restrict (denumerate s) (toList (get (enumerate s))))
140 (h′↦r ∘ just) (union (restrict (denumerate s) (toList (get (enumerate s)))) (delete-many (get (enumerate s)) (fromFunc (denumerate s))))
141 ≡⟨ cong (h′↦r ∘ just) (lemma-disjoint-union (denumerate s) (get (enumerate s))) ⟩
142 (h′↦r ∘ just) (fromFunc (denumerate s))
144 mapMV (flip lookupVec (fromFunc (denumerate s))) (enumerate s)
145 ≡⟨ cong (flip mapMV (enumerate s) ∘ flip lookupVec) (lemma-fromFunc-tabulate (denumerate s)) ⟩
146 mapMV (flip lookupVec (tabulate (Maybe.just ∘ denumerate s))) (enumerate s)
147 ≡⟨ mapMV-cong (lookup∘tabulate (Maybe.just ∘ denumerate s)) (enumerate s) ⟩
148 mapMV (Maybe.just ∘ denumerate s) (enumerate s)
149 ≡⟨ mapMV-purity (denumerate s) (enumerate s) ⟩
150 just (map (denumerate s) (enumerate s))
151 ≡⟨ cong just (lemma-map-denumerate-enumerate s) ⟩
153 where open ≡-Reasoning
154 h↦h′ = _<$>_ (flip union (delete-many (get (enumerate s)) (fromFunc (denumerate s))))
155 h′↦r = flip _>>=_ (flip mapMV (enumerate s) ∘ flip lookupVec)
158 lemma-<$>-just : {A B : Set} {f : A → B} {b : B} (ma : Maybe A) → f <$> ma ≡ just b → ∃ λ a → ma ≡ just a
159 lemma-<$>-just (just x) f<$>ma≡just-b = x , refl
160 lemma-<$>-just nothing ()
162 lemma-union-not-used : {m n : ℕ} {A : Set} (h : FinMapMaybe n A) → (h' : FinMapMaybe n A) → (is : Vec (Fin n) m) → (toList is) in-domain-of h → map (flip lookupM (union h h')) is ≡ map (flip lookupM h) is
163 lemma-union-not-used h h' [] p = refl
164 lemma-union-not-used h h' (i ∷ is') (Data.List.All._∷_ (x , px) p') = cong₂ _∷_ (begin
165 lookupM i (union h h')
166 ≡⟨ lookup∘tabulate (λ j → maybe′ just (lookupM j h') (lookupM j h)) i ⟩
167 maybe′ just (lookupM i h') (lookupM i h)
168 ≡⟨ cong (maybe′ just (lookupM i h')) px ⟩
169 maybe′ just (lookupM i h') (just x)
172 (lemma-union-not-used h h' is' p')
173 where open ≡-Reasoning
175 lemma->>=-just : {A B : Set} (ma : Maybe A) {f : A → Maybe B} {b : B} → (ma >>= f) ≡ just b → ∃ λ a → ma ≡ just a
176 lemma->>=-just (just a) p = a , refl
177 lemma->>=-just nothing ()
179 lemma-just-sequence : {A : Set} {n : ℕ} → (v : Vec A n) → sequenceV (map just v) ≡ just v
180 lemma-just-sequence [] = refl
181 lemma-just-sequence (x ∷ xs) rewrite lemma-just-sequence xs = refl
183 lemma-mapM-successful : {A B : Set} {f : A → Maybe B} {n : ℕ} → (v : Vec A n) → {r : Vec B n} → mapMV f v ≡ just r → ∃ λ w → map f v ≡ map just w
184 lemma-mapM-successful [] p = [] , refl
185 lemma-mapM-successful {f = f} (x ∷ xs) p with f x | mapMV f xs | inspect (mapMV f) xs
186 lemma-mapM-successful (x ∷ xs) () | nothing | _ | _
187 lemma-mapM-successful (x ∷ xs) () | just y | nothing | _
188 lemma-mapM-successful (x ∷ xs) p | just y | just ys | [ p′ ] with lemma-mapM-successful xs p′
189 lemma-mapM-successful (x ∷ xs) p | just y | just ys | [ p′ ] | w , pw = y ∷ w , cong (_∷_ (just y)) pw
192 lemma-get-mapMV : {A B : Set} {f : A → Maybe B} {n : ℕ} {v : Vec A n} {r : Vec B n} → mapMV f v ≡ just r → {getlen : ℕ → ℕ} (get : get-type getlen) → get <$> mapMV f v ≡ mapMV f (get v)
193 lemma-get-mapMV {f = f} {v = v} p get = let w , pw = lemma-mapM-successful v p in begin
195 ≡⟨ cong (_<$>_ get) (sym (sequence-map f v)) ⟩
196 get <$> (sequenceV (map f v))
197 ≡⟨ cong (_<$>_ get ∘ sequenceV) pw ⟩
198 get <$> (sequenceV (map just w))
199 ≡⟨ cong (_<$>_ get) (lemma-just-sequence w) ⟩
201 ≡⟨ sym (lemma-just-sequence (get w)) ⟩
202 sequenceV (map just (get w))
203 ≡⟨ cong sequenceV (sym (free-theorem get just w)) ⟩
204 sequenceV (get (map just w))
205 ≡⟨ cong (sequenceV ∘ get) (sym pw) ⟩
206 sequenceV (get (map f v))
207 ≡⟨ cong sequenceV (free-theorem get f v) ⟩
208 sequenceV (map f (get v))
209 ≡⟨ sequence-map f (get v) ⟩
211 where open ≡-Reasoning
213 sequence-cong : {S : Setoid ℓ₀ ℓ₀} {n : ℕ} {m₁ m₂ : Setoid.Carrier (VecISetoid (MaybeSetoid S) at n)} → VecISetoid (MaybeSetoid S) at _ ∋ m₁ ≈ m₂ → MaybeSetoid (VecISetoid S at n) ∋ sequenceV m₁ ≈ sequenceV m₂
214 sequence-cong {S} VecEq.[]-cong = Setoid.refl (MaybeSetoid (VecISetoid S at _))
215 sequence-cong {S} {m₁ = just x ∷ xs} {m₂ = just y ∷ ys} (just x≈y VecEq.∷-cong xs≈ys) with sequenceV xs | sequenceV ys | sequence-cong xs≈ys
216 sequence-cong {S} {m₁ = just x ∷ xs} {m₂ = just y ∷ ys} (just x≈y VecEq.∷-cong xs≈ys) | just sxs | just sys | just p = MaybeEq.just (x≈y VecEq.∷-cong p)
217 sequence-cong {S} {m₁ = just x ∷ xs} {m₂ = just y ∷ ys} (just x≈y VecEq.∷-cong xs≈ys) | nothing | nothing | nothing = Setoid.refl (MaybeSetoid (VecISetoid S at _))
218 sequence-cong {S} (nothing VecEq.∷-cong xs≈ys) = Setoid.refl (MaybeSetoid (VecISetoid S at _))
220 theorem-2 : {getlen : ℕ → ℕ} (get : get-type getlen) → {m : ℕ} → (v : Vec Carrier (getlen m)) → (s u : Vec Carrier m) → bff get s v ≡ just u → VecISetoid A.setoid at _ ∋ get u ≈ v
221 theorem-2 get v s u p with (lemma->>=-just ((flip union (delete-many (get (enumerate s)) (fromFunc (denumerate s)))) <$> (assoc (get (enumerate s)) v)) p)
222 theorem-2 get v s u p | h′ , ph′ with (lemma-<$>-just (assoc (get (enumerate s)) v) ph′)
223 theorem-2 get v s u p | h′ , ph′ | h , ph = drop-just (begin⟨ MaybeSetoid (VecISetoid A.setoid at _) ⟩
225 ≡⟨ cong (_<$>_ get) (sym p) ⟩
226 get <$> (bff get s v)
227 ≡⟨ cong (_<$>_ get ∘ flip _>>=_ h′↦r ∘ _<$>_ h↦h′) ph ⟩
228 get <$> mapMV (flip lookupM (h↦h′ h)) s′
229 ≡⟨ lemma-get-mapMV (trans (cong (flip _>>=_ h′↦r ∘ _<$>_ h↦h′) (sym ph)) p) get ⟩
230 mapMV (flip lookupM (h↦h′ h)) (get s′)
231 ≡⟨ sym (sequence-map (flip lookupM (h↦h′ h)) (get s′)) ⟩
232 sequenceV (map (flip lookupM (h↦h′ h)) (get s′))
233 ≡⟨ cong sequenceV (lemma-union-not-used h g′ (get s′) (lemma-assoc-domain (get s′) v h ph)) ⟩
234 sequenceV (map (flip lookupM h) (get s′))
235 ≈⟨ sequence-cong (lemma-2 (get s′) v h ph) ⟩
236 sequenceV (map just v)
237 ≡⟨ lemma-just-sequence v ⟩
239 where open SetoidReasoning
241 g = fromFunc (denumerate s)
242 g′ = delete-many (get s′) g
244 h′↦r = flip mapMV s′ ∘ flip lookupM