@@ -24,13 +24,14 @@ open import Data.List.Membership.Propositional using (_∈_)
2424open import Data.List.Relation.Unary.All using (All; []; _∷_)
2525open import Data.List.Relation.Unary.Any using (Any; here; there)
2626open import Data.Maybe.Base as Maybe using (Maybe; just; nothing; maybe)
27+ open import Data.Maybe.Relation.Unary.Any using (just) renaming (Any to MAny)
2728open import Data.Nat.Base
2829open import Data.Nat.Divisibility using (_∣_; divides; ∣n⇒∣m*n)
2930open import Data.Nat.Properties
3031open import Data.Product.Base as Product
3132 using (_×_; _,_; uncurry; uncurry′; proj₁; proj₂; <_,_>)
3233import Data.Product.Relation.Unary.All as Product using (All)
33- open import Data.Sum.Base using (_⊎_; inj₁; inj₂)
34+ open import Data.Sum using (_⊎_; inj₁; inj₂; isInj₁; isInj ₂)
3435open import Data.These.Base as These using (These; this; that; these)
3536open import Data.Fin.Properties using (toℕ-cast)
3637open import Function.Base using (id; _∘_; _∘′_; _∋_; _-⟨_∣; ∣_⟩-_; _$_; const; flip)
@@ -48,19 +49,17 @@ open import Relation.Unary using (Pred; Decidable; ∁)
4849open import Relation.Unary.Properties using (∁?)
4950import Data.Nat.GeneralisedArithmetic as ℕ
5051
51-
5252open ≡-Reasoning
5353
54- private
55- variable
56- a b c d e p : Level
57- A : Set a
58- B : Set b
59- C : Set c
60- D : Set d
61- E : Set e
62- x y z w : A
63- xs ys zs ws : List A
54+ private variable
55+ a b c d e p ℓ : Level
56+ A : Set a
57+ B : Set b
58+ C : Set c
59+ D : Set d
60+ E : Set e
61+ x y z w : A
62+ xs ys zs ws : List A
6463
6564------------------------------------------------------------------------
6665-- _∷_
@@ -83,71 +82,6 @@ private
8382≡-dec _≟_ [] (y ∷ ys) = no λ ()
8483≡-dec _≟_ (x ∷ xs) (y ∷ ys) = ∷-dec (x ≟ y) (≡-dec _≟_ xs ys)
8584
86- ------------------------------------------------------------------------
87- -- map
88-
89- map-id : map id ≗ id {A = List A}
90- map-id [] = refl
91- map-id (x ∷ xs) = cong (x ∷_) (map-id xs)
92-
93- map-id-local : ∀ {f : A → A} {xs} → All (λ x → f x ≡ x) xs → map f xs ≡ xs
94- map-id-local [] = refl
95- map-id-local (fx≡x ∷ pxs) = cong₂ _∷_ fx≡x (map-id-local pxs)
96-
97- map-++ : ∀ (f : A → B) xs ys →
98- map f (xs ++ ys) ≡ map f xs ++ map f ys
99- map-++ f [] ys = refl
100- map-++ f (x ∷ xs) ys = cong (f x ∷_) (map-++ f xs ys)
101-
102- map-cong : ∀ {f g : A → B} → f ≗ g → map f ≗ map g
103- map-cong f≗g [] = refl
104- map-cong f≗g (x ∷ xs) = cong₂ _∷_ (f≗g x) (map-cong f≗g xs)
105-
106- map-cong-local : ∀ {f g : A → B} {xs} →
107- All (λ x → f x ≡ g x) xs → map f xs ≡ map g xs
108- map-cong-local [] = refl
109- map-cong-local (fx≡gx ∷ fxs≡gxs) = cong₂ _∷_ fx≡gx (map-cong-local fxs≡gxs)
110-
111- length-map : ∀ (f : A → B) xs → length (map f xs) ≡ length xs
112- length-map f [] = refl
113- length-map f (x ∷ xs) = cong suc (length-map f xs)
114-
115- map-∘ : {g : B → C} {f : A → B} → map (g ∘ f) ≗ map g ∘ map f
116- map-∘ [] = refl
117- map-∘ (x ∷ xs) = cong (_ ∷_) (map-∘ xs)
118-
119- map-injective : ∀ {f : A → B} → Injective _≡_ _≡_ f → Injective _≡_ _≡_ (map f)
120- map-injective finj {[]} {[]} eq = refl
121- map-injective finj {x ∷ xs} {y ∷ ys} eq =
122- let fx≡fy , fxs≡fys = ∷-injective eq in
123- cong₂ _∷_ (finj fx≡fy) (map-injective finj fxs≡fys)
124-
125- ------------------------------------------------------------------------
126- -- catMaybes
127-
128- catMaybes-concatMap : catMaybes {A = A} ≗ concatMap fromMaybe
129- catMaybes-concatMap [] = refl
130- catMaybes-concatMap (just x ∷ xs) = cong (x ∷_) (catMaybes-concatMap xs)
131- catMaybes-concatMap (nothing ∷ xs) = catMaybes-concatMap xs
132-
133- length-catMaybes : ∀ xs → length (catMaybes {A = A} xs) ≤ length xs
134- length-catMaybes [] = ≤-refl
135- length-catMaybes (just x ∷ xs) = s≤s (length-catMaybes xs)
136- length-catMaybes (nothing ∷ xs) = m≤n⇒m≤1+n (length-catMaybes xs)
137-
138- catMaybes-++ : (xs ys : List (Maybe A)) →
139- catMaybes (xs ++ ys) ≡ catMaybes xs ++ catMaybes ys
140- catMaybes-++ [] ys = refl
141- catMaybes-++ (just x ∷ xs) ys = cong (x ∷_) (catMaybes-++ xs ys)
142- catMaybes-++ (nothing ∷ xs) ys = catMaybes-++ xs ys
143-
144- module _ (f : A → B) where
145-
146- map-catMaybes : map f ∘ catMaybes ≗ catMaybes ∘ map (Maybe.map f)
147- map-catMaybes [] = refl
148- map-catMaybes (just x ∷ xs) = cong (f x ∷_) (map-catMaybes xs)
149- map-catMaybes (nothing ∷ xs) = map-catMaybes xs
150-
15185------------------------------------------------------------------------
15286-- _++_
15387
@@ -263,6 +197,46 @@ module _ (A : Set a) where
263197 ; ε-homo = refl
264198 }
265199
200+
201+ ------------------------------------------------------------------------
202+ -- map
203+
204+ map-id : map id ≗ id {A = List A}
205+ map-id [] = refl
206+ map-id (x ∷ xs) = cong (x ∷_) (map-id xs)
207+
208+ map-id-local : ∀ {f : A → A} {xs} → All (λ x → f x ≡ x) xs → map f xs ≡ xs
209+ map-id-local [] = refl
210+ map-id-local (fx≡x ∷ pxs) = cong₂ _∷_ fx≡x (map-id-local pxs)
211+
212+ map-++ : ∀ (f : A → B) xs ys →
213+ map f (xs ++ ys) ≡ map f xs ++ map f ys
214+ map-++ f [] ys = refl
215+ map-++ f (x ∷ xs) ys = cong (f x ∷_) (map-++ f xs ys)
216+
217+ map-cong : ∀ {f g : A → B} → f ≗ g → map f ≗ map g
218+ map-cong f≗g [] = refl
219+ map-cong f≗g (x ∷ xs) = cong₂ _∷_ (f≗g x) (map-cong f≗g xs)
220+
221+ map-cong-local : ∀ {f g : A → B} {xs} →
222+ All (λ x → f x ≡ g x) xs → map f xs ≡ map g xs
223+ map-cong-local [] = refl
224+ map-cong-local (fx≡gx ∷ fxs≡gxs) = cong₂ _∷_ fx≡gx (map-cong-local fxs≡gxs)
225+
226+ length-map : ∀ (f : A → B) xs → length (map f xs) ≡ length xs
227+ length-map f [] = refl
228+ length-map f (x ∷ xs) = cong suc (length-map f xs)
229+
230+ map-∘ : {g : B → C} {f : A → B} → map (g ∘ f) ≗ map g ∘ map f
231+ map-∘ [] = refl
232+ map-∘ (x ∷ xs) = cong (_ ∷_) (map-∘ xs)
233+
234+ map-injective : ∀ {f : A → B} → Injective _≡_ _≡_ f → Injective _≡_ _≡_ (map f)
235+ map-injective finj {[]} {[]} eq = refl
236+ map-injective finj {x ∷ xs} {y ∷ ys} eq =
237+ let fx≡fy , fxs≡fys = ∷-injective eq in
238+ cong₂ _∷_ (finj fx≡fy) (map-injective finj fxs≡fys)
239+
266240------------------------------------------------------------------------
267241-- cartesianProductWith
268242
@@ -740,6 +714,39 @@ map-concatMap f g xs = begin
740714 concatMap (map f ∘′ g) xs
741715 ∎
742716
717+ ------------------------------------------------------------------------
718+ -- catMaybes
719+
720+ catMaybes-concatMap : catMaybes {A = A} ≗ concatMap fromMaybe
721+ catMaybes-concatMap [] = refl
722+ catMaybes-concatMap (just x ∷ xs) = cong (x ∷_) (catMaybes-concatMap xs)
723+ catMaybes-concatMap (nothing ∷ xs) = catMaybes-concatMap xs
724+
725+ length-catMaybes : ∀ xs → length (catMaybes {A = A} xs) ≤ length xs
726+ length-catMaybes [] = ≤-refl
727+ length-catMaybes (just x ∷ xs) = s≤s (length-catMaybes xs)
728+ length-catMaybes (nothing ∷ xs) = m≤n⇒m≤1+n (length-catMaybes xs)
729+
730+ catMaybes-++ : (xs ys : List (Maybe A)) →
731+ catMaybes (xs ++ ys) ≡ catMaybes xs ++ catMaybes ys
732+ catMaybes-++ [] ys = refl
733+ catMaybes-++ (just x ∷ xs) ys = cong (x ∷_) (catMaybes-++ xs ys)
734+ catMaybes-++ (nothing ∷ xs) ys = catMaybes-++ xs ys
735+
736+ module _ (f : A → B) where
737+
738+ map-catMaybes : map f ∘ catMaybes ≗ catMaybes ∘ map (Maybe.map f)
739+ map-catMaybes [] = refl
740+ map-catMaybes (just x ∷ xs) = cong (f x ∷_) (map-catMaybes xs)
741+ map-catMaybes (nothing ∷ xs) = map-catMaybes xs
742+
743+ Any-catMaybes⁺ : ∀ {P : Pred A ℓ} {xs : List (Maybe A)}
744+ → Any (MAny P) xs → Any P (catMaybes xs)
745+ Any-catMaybes⁺ {xs = nothing ∷ xs} (there x∈) = Any-catMaybes⁺ x∈
746+ Any-catMaybes⁺ {xs = just x ∷ xs} = λ where
747+ (here (just px)) → here px
748+ (there x∈) → there $ Any-catMaybes⁺ x∈
749+
743750------------------------------------------------------------------------
744751-- mapMaybe
745752
@@ -792,6 +799,26 @@ module _ (g : B → C) (f : A → Maybe B) where
792799 mapMaybe (Maybe.map g) (map f xs) ≡⟨ mapMaybe-map _ f xs ⟩
793800 mapMaybe (Maybe.map g ∘ f) xs ∎
794801
802+ mapMaybeIsInj₁∘mapInj₁ : (xs : List A) → mapMaybe (isInj₁ {B = B}) (map inj₁ xs) ≡ xs
803+ mapMaybeIsInj₁∘mapInj₁ = λ where
804+ [] → refl
805+ (x ∷ xs) → cong (x ∷_) (mapMaybeIsInj₁∘mapInj₁ xs)
806+
807+ mapMaybeIsInj₁∘mapInj₂ : (xs : List B) → mapMaybe (isInj₁ {A = A}) (map inj₂ xs) ≡ []
808+ mapMaybeIsInj₁∘mapInj₂ = λ where
809+ [] → refl
810+ (x ∷ xs) → mapMaybeIsInj₁∘mapInj₂ xs
811+
812+ mapMaybeIsInj₂∘mapInj₂ : (xs : List B) → mapMaybe (isInj₂ {A = A}) (map inj₂ xs) ≡ xs
813+ mapMaybeIsInj₂∘mapInj₂ = λ where
814+ [] → refl
815+ (x ∷ xs) → cong (x ∷_) (mapMaybeIsInj₂∘mapInj₂ xs)
816+
817+ mapMaybeIsInj₂∘mapInj₁ : (xs : List A) → mapMaybe (isInj₂ {B = B}) (map inj₁ xs) ≡ []
818+ mapMaybeIsInj₂∘mapInj₁ = λ where
819+ [] → refl
820+ (x ∷ xs) → mapMaybeIsInj₂∘mapInj₁ xs
821+
795822------------------------------------------------------------------------
796823-- sum
797824
0 commit comments