@@ -1156,13 +1156,13 @@ module _ {P : Pred A p} (P? : Decidable P) where
1156
1156
filter-all : ∀ {xs} → All P xs → filter P? xs ≡ xs
1157
1157
filter-all {[]} [] = refl
1158
1158
filter-all {x ∷ xs} (px ∷ pxs) with P? x
1159
- ... | no ¬px = contradiction px ¬px
1160
- ... | true because _ = cong (x ∷_) (filter-all pxs)
1159
+ ... | false because [ ¬px] = contradiction px (invert [ ¬px])
1160
+ ... | true because _ = cong (x ∷_) (filter-all pxs)
1161
1161
1162
1162
filter-notAll : ∀ xs → Any (∁ P) xs → length (filter P? xs) < length xs
1163
1163
filter-notAll (x ∷ xs) (here ¬px) with P? x
1164
- ... | false because _ = s≤s (length-filter xs)
1165
- ... | yes px = contradiction px ¬px
1164
+ ... | false because _ = s≤s (length-filter xs)
1165
+ ... | true because [px] = contradiction (invert [px]) ¬px
1166
1166
filter-notAll (x ∷ xs) (there any) with ih ← filter-notAll xs any | does (P? x)
1167
1167
... | false = m≤n⇒m≤1+n ih
1168
1168
... | true = s≤s ih
@@ -1178,8 +1178,8 @@ module _ {P : Pred A p} (P? : Decidable P) where
1178
1178
filter-none : ∀ {xs} → All (∁ P) xs → filter P? xs ≡ []
1179
1179
filter-none {[]} [] = refl
1180
1180
filter-none {x ∷ xs} (¬px ∷ ¬pxs) with P? x
1181
- ... | false because _ = filter-none ¬pxs
1182
- ... | yes px = contradiction px ¬px
1181
+ ... | false because _ = filter-none ¬pxs
1182
+ ... | true because [px] = contradiction (invert [px]) ¬px
1183
1183
1184
1184
filter-complete : ∀ {xs} → length (filter P? xs) ≡ length xs →
1185
1185
filter P? xs ≡ xs
@@ -1190,13 +1190,13 @@ module _ {P : Pred A p} (P? : Decidable P) where
1190
1190
1191
1191
filter-accept : ∀ {x xs} → P x → filter P? (x ∷ xs) ≡ x ∷ (filter P? xs)
1192
1192
filter-accept {x} Px with P? x
1193
- ... | true because _ = refl
1194
- ... | no ¬Px = contradiction Px ¬Px
1193
+ ... | true because _ = refl
1194
+ ... | false because [ ¬Px] = contradiction Px (invert [ ¬Px])
1195
1195
1196
1196
filter-reject : ∀ {x xs} → ¬ P x → filter P? (x ∷ xs) ≡ filter P? xs
1197
1197
filter-reject {x} ¬Px with P? x
1198
- ... | yes Px = contradiction Px ¬Px
1199
- ... | false because _ = refl
1198
+ ... | true because [Px] = contradiction (invert [Px]) ¬Px
1199
+ ... | false because _ = refl
1200
1200
1201
1201
filter-idem : filter P? ∘ filter P? ≗ filter P?
1202
1202
filter-idem [] = refl
@@ -1234,13 +1234,13 @@ module _ {R : Rel A p} (R? : B.Decidable R) where
1234
1234
1235
1235
derun-reject : ∀ {x y} xs → R x y → derun R? (x ∷ y ∷ xs) ≡ derun R? (y ∷ xs)
1236
1236
derun-reject {x} {y} xs Rxy with R? x y
1237
- ... | yes _ = refl
1238
- ... | no ¬Rxy = contradiction Rxy ¬Rxy
1237
+ ... | true because _ = refl
1238
+ ... | false because [ ¬Rxy] = contradiction Rxy (invert [ ¬Rxy])
1239
1239
1240
1240
derun-accept : ∀ {x y} xs → ¬ R x y → derun R? (x ∷ y ∷ xs) ≡ x ∷ derun R? (y ∷ xs)
1241
1241
derun-accept {x} {y} xs ¬Rxy with R? x y
1242
- ... | yes Rxy = contradiction Rxy ¬Rxy
1243
- ... | no _ = refl
1242
+ ... | true because [ Rxy] = contradiction (invert [ Rxy]) ¬Rxy
1243
+ ... | false because _ = refl
1244
1244
1245
1245
------------------------------------------------------------------------
1246
1246
-- partition
@@ -1253,7 +1253,7 @@ module _ {P : Pred A p} (P? : Decidable P) where
1253
1253
... | true = cong (Product.map (x ∷_) id) ih
1254
1254
... | false = cong (Product.map id (x ∷_)) ih
1255
1255
1256
- length-partition : ∀ xs → (let ( ys , zs) = partition P? xs) →
1256
+ length-partition : ∀ xs → (let ys , zs = partition P? xs) →
1257
1257
length ys ≤ length xs × length zs ≤ length xs
1258
1258
length-partition [] = z≤n , z≤n
1259
1259
length-partition (x ∷ xs) with ih ← length-partition xs | does (P? x)
0 commit comments