@@ -1163,13 +1163,13 @@ module _ {P : Pred A p} (P? : Decidable P) where
1163
1163
filter-all : ∀ {xs} → All P xs → filter P? xs ≡ xs
1164
1164
filter-all {[]} [] = refl
1165
1165
filter-all {x ∷ xs} (px ∷ pxs) with P? x
1166
- ... | false because [ ¬px] = contradiction px (invert [ ¬px])
1167
- ... | true because _ = cong (x ∷_) (filter-all pxs)
1166
+ ... | no ¬px = contradiction px ¬px
1167
+ ... | true because _ = cong (x ∷_) (filter-all pxs)
1168
1168
1169
1169
filter-notAll : ∀ xs → Any (∁ P) xs → length (filter P? xs) < length xs
1170
1170
filter-notAll (x ∷ xs) (here ¬px) with P? x
1171
- ... | false because _ = s≤s (length-filter xs)
1172
- ... | true because [px] = contradiction (invert [px]) ¬px
1171
+ ... | false because _ = s≤s (length-filter xs)
1172
+ ... | yes px = contradiction px ¬px
1173
1173
filter-notAll (x ∷ xs) (there any) with ih ← filter-notAll xs any | does (P? x)
1174
1174
... | false = m≤n⇒m≤1+n ih
1175
1175
... | true = s≤s ih
@@ -1185,8 +1185,8 @@ module _ {P : Pred A p} (P? : Decidable P) where
1185
1185
filter-none : ∀ {xs} → All (∁ P) xs → filter P? xs ≡ []
1186
1186
filter-none {[]} [] = refl
1187
1187
filter-none {x ∷ xs} (¬px ∷ ¬pxs) with P? x
1188
- ... | false because _ = filter-none ¬pxs
1189
- ... | true because [px] = contradiction (invert [px]) ¬px
1188
+ ... | false because _ = filter-none ¬pxs
1189
+ ... | yes px = contradiction px ¬px
1190
1190
1191
1191
filter-complete : ∀ {xs} → length (filter P? xs) ≡ length xs →
1192
1192
filter P? xs ≡ xs
@@ -1197,13 +1197,13 @@ module _ {P : Pred A p} (P? : Decidable P) where
1197
1197
1198
1198
filter-accept : ∀ {x xs} → P x → filter P? (x ∷ xs) ≡ x ∷ (filter P? xs)
1199
1199
filter-accept {x} Px with P? x
1200
- ... | true because _ = refl
1201
- ... | false because [ ¬Px] = contradiction Px (invert [ ¬Px])
1200
+ ... | true because _ = refl
1201
+ ... | no ¬Px = contradiction Px ¬Px
1202
1202
1203
1203
filter-reject : ∀ {x xs} → ¬ P x → filter P? (x ∷ xs) ≡ filter P? xs
1204
1204
filter-reject {x} ¬Px with P? x
1205
- ... | true because [Px] = contradiction (invert [Px]) ¬Px
1206
- ... | false because _ = refl
1205
+ ... | yes Px = contradiction Px ¬Px
1206
+ ... | false because _ = refl
1207
1207
1208
1208
filter-idem : filter P? ∘ filter P? ≗ filter P?
1209
1209
filter-idem [] = refl
@@ -1241,13 +1241,13 @@ module _ {R : Rel A p} (R? : B.Decidable R) where
1241
1241
1242
1242
derun-reject : ∀ {x y} xs → R x y → derun R? (x ∷ y ∷ xs) ≡ derun R? (y ∷ xs)
1243
1243
derun-reject {x} {y} xs Rxy with R? x y
1244
- ... | true because _ = refl
1245
- ... | false because [ ¬Rxy] = contradiction Rxy (invert [ ¬Rxy])
1244
+ ... | yes _ = refl
1245
+ ... | no ¬Rxy = contradiction Rxy ¬Rxy
1246
1246
1247
1247
derun-accept : ∀ {x y} xs → ¬ R x y → derun R? (x ∷ y ∷ xs) ≡ x ∷ derun R? (y ∷ xs)
1248
1248
derun-accept {x} {y} xs ¬Rxy with R? x y
1249
- ... | true because [ Rxy] = contradiction (invert [ Rxy]) ¬Rxy
1250
- ... | false because _ = refl
1249
+ ... | yes Rxy = contradiction Rxy ¬Rxy
1250
+ ... | no _ = refl
1251
1251
1252
1252
------------------------------------------------------------------------
1253
1253
-- partition
@@ -1260,7 +1260,7 @@ module _ {P : Pred A p} (P? : Decidable P) where
1260
1260
... | true = cong (Product.map (x ∷_) id) ih
1261
1261
... | false = cong (Product.map id (x ∷_)) ih
1262
1262
1263
- length-partition : ∀ xs → (let ys , zs = partition P? xs) →
1263
+ length-partition : ∀ xs → (let ( ys , zs) = partition P? xs) →
1264
1264
length ys ≤ length xs × length zs ≤ length xs
1265
1265
length-partition [] = z≤n , z≤n
1266
1266
length-partition (x ∷ xs) with ih ← length-partition xs | does (P? x)
0 commit comments