Skip to content

Commit fd82f5c

Browse files
JacquesCarettegallais
authored andcommitted
A number of with are not really needed (#2363)
* a number of 'with' are not really needed. Many, many more were, so left as is. * whitespace * deal with a few outstanding comments. * actually re-introduce a 'with' as it is actually clearer. * with fits better here too. * tweak things a little from review by James. * Update src/Codata/Sized/Cowriter.agda layout for readability Co-authored-by: G. Allais <[email protected]> * Update src/Data/Fin/Base.agda layout for readability Co-authored-by: G. Allais <[email protected]> * Update src/Data/List/Fresh/Relation/Unary/All.agda layout for readability Co-authored-by: G. Allais <[email protected]> --------- Co-authored-by: G. Allais <[email protected]>
1 parent aad1d19 commit fd82f5c

File tree

28 files changed

+98
-126
lines changed

28 files changed

+98
-126
lines changed

src/Algebra/Construct/LexProduct.agda

Lines changed: 2 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -11,7 +11,7 @@ open import Algebra.Definitions
1111
open import Data.Bool.Base using (true; false)
1212
open import Data.Product.Base using (_×_; _,_)
1313
open import Data.Product.Relation.Binary.Pointwise.NonDependent using (Pointwise)
14-
open import Data.Sum.Base using (inj₁; inj₂)
14+
open import Data.Sum.Base using (inj₁; inj₂; map)
1515
open import Function.Base using (_∘_)
1616
open import Relation.Binary.Core using (Rel)
1717
open import Relation.Binary.Definitions using (Decidable)
@@ -107,6 +107,4 @@ sel ∙-sel ◦-sel (a , x) (b , y) with (a ∙ b) ≟₁ a | (a ∙ b) ≟₁ b
107107
... | no ab≉a | no ab≉b = contradiction₂ (∙-sel a b) ab≉a ab≉b
108108
... | yes ab≈a | no _ = inj₁ (ab≈a , ≈₂-refl)
109109
... | no _ | yes ab≈b = inj₂ (ab≈b , ≈₂-refl)
110-
... | yes ab≈a | yes ab≈b with ◦-sel x y
111-
... | inj₁ xy≈x = inj₁ (ab≈a , xy≈x)
112-
... | inj₂ xy≈y = inj₂ (ab≈b , xy≈y)
110+
... | yes ab≈a | yes ab≈b = map (ab≈a ,_) (ab≈b ,_) (◦-sel x y)

src/Algebra/Construct/LiftedChoice.agda

Lines changed: 6 additions & 7 deletions
Original file line numberDiff line numberDiff line change
@@ -11,13 +11,14 @@ open import Algebra
1111
module Algebra.Construct.LiftedChoice where
1212

1313
open import Algebra.Consequences.Base
14-
open import Relation.Binary.Core using (Rel; _⇒_; _Preserves_⟶_)
15-
open import Relation.Binary.Structures using (IsEquivalence)
16-
open import Relation.Nullary using (¬_; yes; no)
17-
open import Data.Sum.Base as Sum using (_⊎_; inj₁; inj₂; [_,_])
14+
open import Data.Sum.Base as Sum using (_⊎_; inj₁; inj₂; [_,_]; [_,_]′)
1815
open import Data.Product.Base using (_×_; _,_)
16+
open import Function.Base using (const; _$_)
1917
open import Level using (Level; _⊔_)
18+
open import Relation.Binary.Core using (Rel; _⇒_; _Preserves_⟶_)
19+
open import Relation.Nullary using (¬_; yes; no)
2020
open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_)
21+
open import Relation.Binary.Structures using (IsEquivalence)
2122
open import Relation.Unary using (Pred)
2223

2324
import Relation.Binary.Reasoning.Setoid as ≈-Reasoning
@@ -34,9 +35,7 @@ private
3435
module _ (_≈_ : Rel B ℓ) (_•_ : Op₂ B) where
3536

3637
Lift : Selective _≈_ _•_ (A B) Op₂ A
37-
Lift ∙-sel f x y with ∙-sel (f x) (f y)
38-
... | inj₁ _ = x
39-
... | inj₂ _ = y
38+
Lift ∙-sel f x y = [ const x , const y ]′ $ ∙-sel (f x) (f y)
4039

4140
------------------------------------------------------------------------
4241
-- Algebraic properties

src/Algebra/Construct/NaturalChoice/MinOp.agda

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -39,13 +39,13 @@ open import Relation.Binary.Reasoning.Preorder preorder
3939

4040
x⊓y≤x : x y x ⊓ y ≤ x
4141
x⊓y≤x x y with total x y
42-
... | inj₁ x≤y = ≤-respˡ-≈ (Eq.sym (x≤y⇒x⊓y≈x x≤y)) refl
42+
... | inj₁ x≤y = reflexive (x≤y⇒x⊓y≈x x≤y)
4343
... | inj₂ y≤x = ≤-respˡ-≈ (Eq.sym (x≥y⇒x⊓y≈y y≤x)) y≤x
4444

4545
x⊓y≤y : x y x ⊓ y ≤ y
4646
x⊓y≤y x y with total x y
4747
... | inj₁ x≤y = ≤-respˡ-≈ (Eq.sym (x≤y⇒x⊓y≈x x≤y)) x≤y
48-
... | inj₂ y≤x = ≤-respˡ-≈ (Eq.sym (x≥y⇒x⊓y≈y y≤x)) refl
48+
... | inj₂ y≤x = reflexive (x≥y⇒x⊓y≈y y≤x)
4949

5050
------------------------------------------------------------------------
5151
-- Algebraic properties

src/Codata/Sized/Cowriter.agda

Lines changed: 4 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -113,6 +113,7 @@ _>>=_ : ∀ {i} → Cowriter W A i → (A → Cowriter W X i) → Cowriter W X i
113113
-- Construction.
114114

115115
unfold : {i} (X (W × X) ⊎ A) X Cowriter W A i
116-
unfold next seed with next seed
117-
... | inj₁ (w , seed′) = w ∷ λ where .force unfold next seed′
118-
... | inj₂ a = [ a ]
116+
unfold next seed =
117+
Sum.[ (λ where (w , seed′) w ∷ λ where .force unfold next seed′)
118+
, [_]
119+
] (next seed)

src/Data/Fin/Base.agda

Lines changed: 6 additions & 5 deletions
Original file line numberDiff line numberDiff line change
@@ -15,7 +15,7 @@ open import Data.Bool.Base using (Bool; T)
1515
open import Data.Nat.Base as ℕ using (ℕ; zero; suc)
1616
open import Data.Product.Base as Product using (_×_; _,_; proj₁; proj₂)
1717
open import Data.Sum.Base as Sum using (_⊎_; inj₁; inj₂; [_,_]′)
18-
open import Function.Base using (id; _∘_; _on_; flip)
18+
open import Function.Base using (id; _∘_; _on_; flip; _$_)
1919
open import Level using (0ℓ)
2020
open import Relation.Binary.Core
2121
open import Relation.Binary.PropositionalEquality.Core using (_≡_; _≢_; refl; cong)
@@ -134,7 +134,7 @@ strengthen (suc i) = suc (strengthen i)
134134
splitAt : m {n} Fin (m ℕ.+ n) Fin m ⊎ Fin n
135135
splitAt zero i = inj₂ i
136136
splitAt (suc m) zero = inj₁ zero
137-
splitAt (suc m) (suc i) = Sum.map suc id (splitAt m i)
137+
splitAt (suc m) (suc i) = Sum.map suc (splitAt m i)
138138

139139
-- inverse of above function
140140
join : m n Fin m ⊎ Fin n Fin (m ℕ.+ n)
@@ -144,9 +144,10 @@ join m n = [ _↑ˡ n , m ↑ʳ_ ]′
144144
-- This is dual to group from Data.Vec.
145145

146146
quotRem : n Fin (m ℕ.* n) Fin n × Fin m
147-
quotRem {suc m} n i with splitAt n i
148-
... | inj₁ j = j , zero
149-
... | inj₂ j = Product.map₂ suc (quotRem {m} n j)
147+
quotRem {suc m} n i =
148+
[ (_, zero)
149+
, Product.map₂ suc ∘ quotRem {m} n
150+
]′ $ splitAt n i
150151

151152
-- a variant of quotRem the type of whose result matches the order of multiplication
152153
remQuot : n Fin (m ℕ.* n) Fin m × Fin n

src/Data/List/Fresh.agda

Lines changed: 9 additions & 10 deletions
Original file line numberDiff line numberDiff line change
@@ -14,7 +14,7 @@
1414
module Data.List.Fresh where
1515

1616
open import Level using (Level; _⊔_)
17-
open import Data.Bool.Base using (true; false)
17+
open import Data.Bool.Base using (true; false; if_then_else_)
1818
open import Data.Unit.Polymorphic.Base using (⊤)
1919
open import Data.Product.Base using (∃; _×_; _,_; -,_; proj₁; proj₂)
2020
open import Data.List.Relation.Unary.All using (All; []; _∷_)
@@ -161,29 +161,28 @@ module _ {P : Pred A p} (P? : U.Decidable P) where
161161
takeWhile-# : {R : Rel A r} a (as : List# A R) a # as a # takeWhile as
162162

163163
takeWhile [] = []
164-
takeWhile (cons a as ps) with does (P? a)
165-
... | true = cons a (takeWhile as) (takeWhile-# a as ps)
166-
... | false = []
164+
takeWhile (cons a as ps) =
165+
if does (P? a) then cons a (takeWhile as) (takeWhile-# a as ps) else []
167166

167+
-- this 'with' is needed to cause reduction in the type of 'takeWhile (a ∷# as)'
168168
takeWhile-# a [] _ = _
169169
takeWhile-# a (x ∷# xs) (p , ps) with does (P? x)
170170
... | true = p , takeWhile-# a xs ps
171171
... | false = _
172172

173173
dropWhile : {R : Rel A r} List# A R List# A R
174174
dropWhile [] = []
175-
dropWhile aas@(a ∷# as) with does (P? a)
176-
... | true = dropWhile as
177-
... | false = aas
175+
dropWhile aas@(a ∷# as) = if does (P? a) then dropWhile as else aas
178176

179177
filter : {R : Rel A r} List# A R List# A R
180178
filter-# : {R : Rel A r} a (as : List# A R) a # as a # filter as
181179

182180
filter [] = []
183-
filter (cons a as ps) with does (P? a)
184-
... | true = cons a (filter as) (filter-# a as ps)
185-
... | false = filter as
181+
filter (cons a as ps) =
182+
let l = filter as in
183+
if does (P? a) then cons a l (filter-# a as ps) else l
186184

185+
-- this 'with' is needed to cause reduction in the type of 'filter-# a (x ∷# xs)'
187186
filter-# a [] _ = _
188187
filter-# a (x ∷# xs) (p , ps) with does (P? x)
189188
... | true = p , filter-# a xs ps

src/Data/List/Fresh/Relation/Unary/All.agda

Lines changed: 6 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -10,7 +10,8 @@ module Data.List.Fresh.Relation.Unary.All where
1010

1111
open import Level using (Level; _⊔_; Lift)
1212
open import Data.Product.Base using (_×_; _,_; proj₁; uncurry)
13-
open import Data.Sum.Base as Sum using (inj₁; inj₂)
13+
open import Data.Sum.Base as Sum using (inj₁; inj₂; [_,_]′)
14+
open import Function.Base using (_∘_; _$_)
1415
open import Relation.Nullary.Decidable as Dec using (Dec; yes; no; _×-dec_)
1516
open import Relation.Unary as U
1617
open import Relation.Binary.Core using (Rel)
@@ -74,6 +75,7 @@ module _ {R : Rel A r} {P : Pred A p} {Q : Pred A q} where
7475

7576
decide : Π[ P ∪ Q ] Π[ All {R = R} P ∪ Any Q ]
7677
decide p∪q [] = inj₁ []
77-
decide p∪q (x ∷# xs) with p∪q x
78-
... | inj₂ qx = inj₂ (here qx)
79-
... | inj₁ px = Sum.map (px ∷_) there (decide p∪q xs)
78+
decide p∪q (x ∷# xs) =
79+
[ (λ px Sum.map (px ∷_) there (decide p∪q xs))
80+
, inj₂ ∘ here
81+
]′ $ p∪q x

src/Data/List/Kleene/Base.agda

Lines changed: 3 additions & 5 deletions
Original file line numberDiff line numberDiff line change
@@ -11,7 +11,7 @@ module Data.List.Kleene.Base where
1111
open import Data.Product.Base as Product
1212
using (_×_; _,_; map₂; map₁; proj₁; proj₂)
1313
open import Data.Nat.Base as ℕ using (ℕ; suc; zero)
14-
open import Data.Maybe.Base as Maybe using (Maybe; just; nothing)
14+
open import Data.Maybe.Base as Maybe using (Maybe; just; nothing; maybe′)
1515
open import Data.Sum.Base as Sum using (_⊎_; inj₁; inj₂)
1616
open import Level using (Level)
1717

@@ -137,9 +137,7 @@ module _ (f : A → Maybe B) where
137137
mapMaybe+ : A + B *
138138
mapMaybe* : A * B *
139139

140-
mapMaybe+ (x & xs) with f x
141-
... | just y = ∹ y & mapMaybe* xs
142-
... | nothing = mapMaybe* xs
140+
mapMaybe+ (x & xs) = maybe′ (λ y z ∹ y & z) id (f x) $ mapMaybe* xs
143141

144142
mapMaybe* [] = []
145143
mapMaybe* (∹ xs) = mapMaybe+ xs
@@ -268,7 +266,7 @@ module _ (f : A → Maybe B → B) where
268266
foldrMaybe* [] = nothing
269267
foldrMaybe* (∹ xs) = just (foldrMaybe+ xs)
270268

271-
foldrMaybe+ xs = f (head xs) (foldrMaybe* (tail xs))
269+
foldrMaybe+ (x & xs) = f x (foldrMaybe* xs)
272270

273271
------------------------------------------------------------------------
274272
-- Indexing

src/Data/List/NonEmpty/Base.agda

Lines changed: 5 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -203,9 +203,12 @@ snocView (x ∷ .(xs List.∷ʳ y)) | xs List.∷ʳ′ y = (x ∷ xs) ∷ʳ′ y
203203

204204
-- The last element in the list.
205205

206+
private
207+
last′ : {l} SnocView {A = A} l A
208+
last′ (_ ∷ʳ′ y) = y
209+
206210
last : List⁺ A A
207-
last xs with snocView xs
208-
last .(ys ∷ʳ y) | ys ∷ʳ′ y = y
211+
last = last′ ∘ snocView
209212

210213
-- Groups all contiguous elements for which the predicate returns the
211214
-- same result into lists. The left sums are the ones for which the

src/Data/List/Relation/Binary/Infix/Heterogeneous.agda

Lines changed: 2 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -49,10 +49,8 @@ map R⇒S (here pref) = here (Prefix.map R⇒S pref)
4949
map R⇒S (there inf) = there (map R⇒S inf)
5050

5151
toView : {as bs} Infix R as bs View R as bs
52-
toView (here p) with Prefix.toView p
53-
...| inf Prefix.++ suff = MkView [] inf suff
54-
toView (there p) with toView p
55-
... | MkView pref inf suff = MkView (_ ∷ pref) inf suff
52+
toView (here p) with inf Prefix.++ suff Prefix.toView p = MkView [] inf suff
53+
toView (there p) with MkView pref inf suff toView p = MkView (_ ∷ pref) inf suff
5654

5755
fromView : {as bs} View R as bs Infix R as bs
5856
fromView (MkView [] inf suff) = here (Prefix.fromView (inf Prefix.++ suff))

0 commit comments

Comments
 (0)