@@ -21,7 +21,7 @@ open import Data.List.Base as List
21
21
open import Data.List.Membership.Propositional using (_∈_)
22
22
open import Data.List.Relation.Unary.All using (All; []; _∷_)
23
23
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 )
25
25
open import Data.Nat.Base
26
26
open import Data.Nat.Divisibility using (_∣_; divides; ∣n⇒∣m*n)
27
27
open import Data.Nat.Properties
@@ -56,23 +56,23 @@ private
56
56
C : Set c
57
57
D : Set d
58
58
E : Set e
59
+ x y z w : A
60
+ xs ys zs ws : List A
59
61
60
62
------------------------------------------------------------------------
61
63
-- _∷_
62
64
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
64
67
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
67
70
68
- ∷-injectiveˡ : x ∷ xs ≡ y List.∷ ys → x ≡ y
69
- ∷-injectiveˡ refl = refl
71
+ ∷-injectiveʳ : x ∷ xs ≡ y List.∷ ys → xs ≡ ys
72
+ ∷-injectiveʳ refl = refl
70
73
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)
76
76
77
77
≡-dec : DecidableEquality A → DecidableEquality (List A)
78
78
≡-dec _≟_ [] [] = yes refl
@@ -122,28 +122,34 @@ map-injective finj {x ∷ xs} {y ∷ ys} eq =
122
122
------------------------------------------------------------------------
123
123
-- mapMaybe
124
124
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
+
125
130
mapMaybe-just : (xs : List A) → mapMaybe just xs ≡ xs
126
131
mapMaybe-just [] = refl
127
132
mapMaybe-just (x ∷ xs) = cong (x ∷_) (mapMaybe-just xs)
128
133
129
134
mapMaybe-nothing : (xs : List A) →
130
- mapMaybe {B = A } (λ _ → nothing) xs ≡ []
135
+ mapMaybe {B = B } (λ _ → nothing) xs ≡ []
131
136
mapMaybe-nothing [] = refl
132
137
mapMaybe-nothing (x ∷ xs) = mapMaybe-nothing xs
133
138
134
139
module _ (f : A → Maybe B) where
135
140
136
141
mapMaybe-concatMap : mapMaybe f ≗ concatMap (fromMaybe ∘ f)
137
- mapMaybe-concatMap [] = refl
142
+ mapMaybe-concatMap [] = refl
138
143
mapMaybe-concatMap (x ∷ xs) with ih ← mapMaybe-concatMap xs | f x
139
144
... | just y = cong (y ∷_) ih
140
145
... | nothing = ih
141
146
142
147
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 _≤-∎)
147
153
148
154
------------------------------------------------------------------------
149
155
-- _++_
@@ -175,14 +181,14 @@ module _ {A : Set a} where
175
181
++-identityʳ-unique : ∀ (xs : List A) {ys} → xs ≡ xs ++ ys → ys ≡ []
176
182
++-identityʳ-unique [] refl = refl
177
183
++-identityʳ-unique (x ∷ xs) eq =
178
- ++-identityʳ-unique xs (proj₂ (∷-injective eq) )
184
+ ++-identityʳ-unique xs (∷-injectiveʳ eq)
179
185
180
186
++-identityˡ-unique : ∀ {xs} (ys : List A) → xs ≡ ys ++ xs → ys ≡ []
181
187
++-identityˡ-unique [] _ = refl
182
188
++-identityˡ-unique {xs = x ∷ xs} (y ∷ ys) eq
183
189
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 ⟨
186
192
(ys ++ [ x ]) ++ xs ∎)
187
193
++-identityˡ-unique {xs = x ∷ xs} (y ∷ [] ) eq | ()
188
194
++-identityˡ-unique {xs = x ∷ xs} (y ∷ _ ∷ _) eq | ()
@@ -1215,22 +1221,20 @@ reverse-downFrom = reverse-applyDownFrom id
1215
1221
------------------------------------------------------------------------
1216
1222
-- _∷ʳ_
1217
1223
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 (_ ∷ _ ∷ _) [] ()
1226
1230
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)
1229
1233
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)
1232
1236
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
1234
1238
∷ʳ-++ xs a ys = ++-assoc xs [ a ] ys
1235
1239
1236
1240
0 commit comments