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_ ; _≡⟨_⟩_ ; _∎)
20 _>>=_ : {A B : Set} → Maybe A → (A → Maybe B) → Maybe B
21 _>>=_ = flip (flip maybe′ nothing)
23 fmap : {A B : Set} → (A → B) → Maybe A → Maybe B
24 fmap f = maybe′ (λ a → just (f a)) nothing
27 EqInst A = (x y : A) → Dec (x ≡ y)
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)
40 record checkInsertProof {A : Set} {n : â„•} (eq : EqInst A) (i : Fin n) (x : A) (m : FinMapMaybe n A) (P : Set) : Set where
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
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
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
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
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
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
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
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)
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'
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′))
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))
98 checkInsert eq i (f i) (restrict f is′)
99 ≡⟨ lemma-checkInsert-restrict eq f i is′ ⟩
100 just (restrict f (i ∷ is′)) ∎
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
108 ≡⟨ cong (lookupM i) (lemma-from-just (trans (sym p) (lemma-checkInsert-same eq i x h' lookupM≡justx))) ⟩
110 ≡⟨ lookupM≡justx ⟩
112 ; new = λ lookupM≡nothing → begin
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' ⟩
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'))
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)
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'
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
143 map just (x ∷ xs) ∎
145 enumerate : {A : Set} → (l : List A) → List (Fin (length l))
146 enumerate l = toList (tabulate id)
148 denumerate : {A : Set} (l : List A) → Fin (length l) → A
149 denumerate l = flip lookupVec (fromList l)
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′
159 free-theorem-list-list : {β γ : Set} → (get : {α : Set} → List α → List α) → (f : β → γ) → get ∘ map f ≗ map f ∘ get
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)
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)
176 map (denumerate as) (enumerate as)
177 ≡⟨ lemma-map-denumerate-enumerate as ⟩
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
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)))
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)))))
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)))
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) ⟩
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 = {!!}