Skip to content

Commit 10a696b

Browse files
Improve Data.List.Base.mapMaybe (#2359; deprecate use of with #2123) (#2361)
* `with`-free definitions plus tests * `CHANGELOG` * use `foldr` on @JacquesCarette 's solution * tidied up unsolved metas * factrored out comparison as removable module `MapMaybeTest` * tidied up; removed `mapMaybeTest` * tidied up; removed v2.1 deprecation section * tidy up long line * Update src/Data/List/Base.agda Co-authored-by: G. Allais <[email protected]> * @gallais 's comments * Update src/Data/List/Base.agda Oops! Co-authored-by: G. Allais <[email protected]> --------- Co-authored-by: G. Allais <[email protected]>
1 parent 69de98b commit 10a696b

File tree

3 files changed

+49
-47
lines changed

3 files changed

+49
-47
lines changed

CHANGELOG.md

+1
Original file line numberDiff line numberDiff line change
@@ -273,6 +273,7 @@ Additions to existing modules
273273

274274
* In `Data.List.Properties`:
275275
```agda
276+
length-catMaybes : length (catMaybes xs) ≤ length xs
276277
applyUpTo-∷ʳ : applyUpTo f n ∷ʳ f n ≡ applyUpTo f (suc n)
277278
applyDownFrom-∷ʳ : applyDownFrom (f ∘ suc) n ∷ʳ f 0 ≡ applyDownFrom f (suc n)
278279
upTo-∷ʳ : upTo n ∷ʳ n ≡ upTo (suc n)

src/Data/List/Base.agda

+11-14
Original file line numberDiff line numberDiff line change
@@ -49,15 +49,6 @@ map : (A → B) → List A → List B
4949
map f [] = []
5050
map f (x ∷ xs) = f x ∷ map f xs
5151

52-
mapMaybe : (A Maybe B) List A List B
53-
mapMaybe p [] = []
54-
mapMaybe p (x ∷ xs) with p x
55-
... | just y = y ∷ mapMaybe p xs
56-
... | nothing = mapMaybe p xs
57-
58-
catMaybes : List (Maybe A) List A
59-
catMaybes = mapMaybe id
60-
6152
infixr 5 _++_
6253

6354
_++_ : List A List A List A
@@ -123,11 +114,11 @@ partitionSums : List (A ⊎ B) → List A × List B
123114
partitionSums = partitionSumsWith id
124115

125116
merge : {R : Rel A ℓ} B.Decidable R List A List A List A
126-
merge R? [] ys = ys
127-
merge R? xs [] = xs
128-
merge R? (x ∷ xs) (y ∷ ys) = if does (R? x y)
129-
then x ∷ merge R? xs (y ∷ ys)
130-
else y ∷ merge R? (x ∷ xs) ys
117+
merge R? [] ys = ys
118+
merge R? xs [] = xs
119+
merge R? x∷xs@(x ∷ xs) y∷ys@(y ∷ ys) = if does (R? x y)
120+
then x ∷ merge R? xs y∷ys
121+
else y ∷ merge R? x∷xs ys
131122

132123
------------------------------------------------------------------------
133124
-- Operations for reducing lists
@@ -149,6 +140,12 @@ concatMap f = concat ∘ map f
149140
ap : List (A B) List A List B
150141
ap fs as = concatMap (flip map as) fs
151142

143+
catMaybes : List (Maybe A) List A
144+
catMaybes = foldr (maybe′ _∷_ id) []
145+
146+
mapMaybe : (A Maybe B) List A List B
147+
mapMaybe p = catMaybes ∘ map p
148+
152149
null : List A Bool
153150
null [] = true
154151
null (x ∷ xs) = false

src/Data/List/Properties.agda

+37-33
Original file line numberDiff line numberDiff line change
@@ -21,7 +21,7 @@ open import Data.List.Base as List
2121
open import Data.List.Membership.Propositional using (_∈_)
2222
open import Data.List.Relation.Unary.All using (All; []; _∷_)
2323
open import Data.List.Relation.Unary.Any using (Any; here; there)
24-
open import Data.Maybe.Base as Maybe using (Maybe; just; nothing)
24+
open import Data.Maybe.Base as Maybe using (Maybe; just; nothing; maybe)
2525
open import Data.Nat.Base
2626
open import Data.Nat.Divisibility using (_∣_; divides; ∣n⇒∣m*n)
2727
open import Data.Nat.Properties
@@ -56,23 +56,23 @@ private
5656
C : Set c
5757
D : Set d
5858
E : Set e
59+
x y z w : A
60+
xs ys zs ws : List A
5961

6062
------------------------------------------------------------------------
6163
-- _∷_
6264

63-
module _ {x y : A} {xs ys : List A} where
65+
∷-injective : x ∷ xs ≡ y List.∷ ys x ≡ y × xs ≡ ys
66+
∷-injective refl = refl , refl
6467

65-
∷-injective : x ∷ xs ≡ y List.∷ ys x ≡ y × xs ≡ ys
66-
∷-injective refl = (refl , refl)
68+
∷-injectiveˡ : x ∷ xs ≡ y List.∷ ys x ≡ y
69+
∷-injectiveˡ refl = refl
6770

68-
∷-injectiveˡ : x ∷ xs ≡ y List.∷ ys xy
69-
∷-injectiveˡ refl = refl
71+
∷-injectiveʳ : x ∷ xs ≡ y List.∷ ys xsys
72+
∷-injectiveʳ refl = refl
7073

71-
∷-injectiveʳ : x ∷ xs ≡ y List.∷ ys xs ≡ ys
72-
∷-injectiveʳ refl = refl
73-
74-
∷-dec : Dec (x ≡ y) Dec (xs ≡ ys) Dec (x List.∷ xs ≡ y ∷ ys)
75-
∷-dec x≟y xs≟ys = Decidable.map′ (uncurry (cong₂ _∷_)) ∷-injective (x≟y ×-dec xs≟ys)
74+
∷-dec : Dec (x ≡ y) Dec (xs ≡ ys) Dec (x ∷ xs ≡ y List.∷ ys)
75+
∷-dec x≟y xs≟ys = Decidable.map′ (uncurry (cong₂ _∷_)) ∷-injective (x≟y ×-dec xs≟ys)
7676

7777
≡-dec : DecidableEquality A DecidableEquality (List A)
7878
≡-dec _≟_ [] [] = yes refl
@@ -122,28 +122,34 @@ map-injective finj {x ∷ xs} {y ∷ ys} eq =
122122
------------------------------------------------------------------------
123123
-- mapMaybe
124124

125+
length-catMaybes : xs length (catMaybes {A = A} xs) ≤ length xs
126+
length-catMaybes [] = ≤-refl
127+
length-catMaybes (just x ∷ xs) = s≤s (length-catMaybes xs)
128+
length-catMaybes (nothing ∷ xs) = m≤n⇒m≤1+n (length-catMaybes xs)
129+
125130
mapMaybe-just : (xs : List A) mapMaybe just xs ≡ xs
126131
mapMaybe-just [] = refl
127132
mapMaybe-just (x ∷ xs) = cong (x ∷_) (mapMaybe-just xs)
128133

129134
mapMaybe-nothing : (xs : List A)
130-
mapMaybe {B = A} (λ _ nothing) xs ≡ []
135+
mapMaybe {B = B} (λ _ nothing) xs ≡ []
131136
mapMaybe-nothing [] = refl
132137
mapMaybe-nothing (x ∷ xs) = mapMaybe-nothing xs
133138

134139
module _ (f : A Maybe B) where
135140

136141
mapMaybe-concatMap : mapMaybe f ≗ concatMap (fromMaybe ∘ f)
137-
mapMaybe-concatMap [] = refl
142+
mapMaybe-concatMap [] = refl
138143
mapMaybe-concatMap (x ∷ xs) with ih mapMaybe-concatMap xs | f x
139144
... | just y = cong (y ∷_) ih
140145
... | nothing = ih
141146

142147
length-mapMaybe : xs length (mapMaybe f xs) ≤ length xs
143-
length-mapMaybe [] = z≤n
144-
length-mapMaybe (x ∷ xs) with ih length-mapMaybe xs | f x
145-
... | just y = s≤s ih
146-
... | nothing = m≤n⇒m≤1+n ih
148+
length-mapMaybe xs = ≤-begin
149+
length (mapMaybe f xs) ≤⟨ length-catMaybes (map f xs) ⟩
150+
length (map f xs) ≤⟨ ≤-reflexive (length-map f xs) ⟩
151+
length xs ≤-∎
152+
where open ≤-Reasoning renaming (begin_ to ≤-begin_; _∎ to _≤-∎)
147153

148154
------------------------------------------------------------------------
149155
-- _++_
@@ -175,14 +181,14 @@ module _ {A : Set a} where
175181
++-identityʳ-unique : (xs : List A) {ys} xs ≡ xs ++ ys ys ≡ []
176182
++-identityʳ-unique [] refl = refl
177183
++-identityʳ-unique (x ∷ xs) eq =
178-
++-identityʳ-unique xs (proj₂ (∷-injective eq))
184+
++-identityʳ-unique xs (∷-injectiveʳ eq)
179185

180186
++-identityˡ-unique : {xs} (ys : List A) xs ≡ ys ++ xs ys ≡ []
181187
++-identityˡ-unique [] _ = refl
182188
++-identityˡ-unique {xs = x ∷ xs} (y ∷ ys) eq
183189
with ++-identityˡ-unique (ys ++ [ x ]) (begin
184-
xs ≡⟨ proj₂ (∷-injective eq)
185-
ys ++ x ∷ xs ≡⟨ sym (++-assoc ys [ x ] xs) ⟩
190+
xs ≡⟨ ∷-injectiveʳ eq ⟩
191+
ys ++ x ∷ xs ≡⟨ ++-assoc ys [ x ] xs
186192
(ys ++ [ x ]) ++ xs ∎)
187193
++-identityˡ-unique {xs = x ∷ xs} (y ∷ [] ) eq | ()
188194
++-identityˡ-unique {xs = x ∷ xs} (y ∷ _ ∷ _) eq | ()
@@ -1215,22 +1221,20 @@ reverse-downFrom = reverse-applyDownFrom id
12151221
------------------------------------------------------------------------
12161222
-- _∷ʳ_
12171223

1218-
module _ {x y : A} where
1219-
1220-
∷ʳ-injective : xs ys xs ∷ʳ x ≡ ys ∷ʳ y xs ≡ ys × x ≡ y
1221-
∷ʳ-injective [] [] refl = (refl , refl)
1222-
∷ʳ-injective (x ∷ xs) (y ∷ ys) eq with refl , eq′ ∷-injective eq
1223-
= Product.map (cong (x ∷_)) id (∷ʳ-injective xs ys eq′)
1224-
∷ʳ-injective [] (_ ∷ _ ∷ _) ()
1225-
∷ʳ-injective (_ ∷ _ ∷ _) [] ()
1224+
∷ʳ-injective : xs ys xs ∷ʳ x ≡ ys ∷ʳ y xs ≡ ys × x ≡ y
1225+
∷ʳ-injective [] [] refl = refl , refl
1226+
∷ʳ-injective (x ∷ xs) (y ∷ ys) eq with refl , eq′ ∷-injective eq
1227+
= Product.map (cong (x ∷_)) id (∷ʳ-injective xs ys eq′)
1228+
∷ʳ-injective [] (_ ∷ _ ∷ _) ()
1229+
∷ʳ-injective (_ ∷ _ ∷ _) [] ()
12261230

1227-
∷ʳ-injectiveˡ : (xs ys : List A) xs ∷ʳ x ≡ ys ∷ʳ y xs ≡ ys
1228-
∷ʳ-injectiveˡ xs ys eq = proj₁ (∷ʳ-injective xs ys eq)
1231+
∷ʳ-injectiveˡ : xs ys xs ∷ʳ x ≡ ys ∷ʳ y xs ≡ ys
1232+
∷ʳ-injectiveˡ xs ys eq = proj₁ (∷ʳ-injective xs ys eq)
12291233

1230-
∷ʳ-injectiveʳ : (xs ys : List A) xs ∷ʳ x ≡ ys ∷ʳ y x ≡ y
1231-
∷ʳ-injectiveʳ xs ys eq = proj₂ (∷ʳ-injective xs ys eq)
1234+
∷ʳ-injectiveʳ : xs ys xs ∷ʳ x ≡ ys ∷ʳ y x ≡ y
1235+
∷ʳ-injectiveʳ xs ys eq = proj₂ (∷ʳ-injective xs ys eq)
12321236

1233-
∷ʳ-++ : (xs : List A) (a : A) (ys : List A) xs ∷ʳ a ++ ys ≡ xs ++ a ∷ ys
1237+
∷ʳ-++ : xs (a : A) ys xs ∷ʳ a ++ ys ≡ xs ++ a ∷ ys
12341238
∷ʳ-++ xs a ys = ++-assoc xs [ a ] ys
12351239

12361240

0 commit comments

Comments
 (0)