Skip to content

A number of with are not really needed #2363

New issue

Have a question about this project? Sign up for a free GitHub account to open an issue and contact its maintainers and the community.

By clicking “Sign up for GitHub”, you agree to our terms of service and privacy statement. We’ll occasionally send you account related emails.

Already on GitHub? Sign in to your account

Merged
merged 10 commits into from
Apr 23, 2024
6 changes: 2 additions & 4 deletions src/Algebra/Construct/LexProduct.agda
Original file line number Diff line number Diff line change
Expand Up @@ -11,7 +11,7 @@ open import Algebra.Definitions
open import Data.Bool.Base using (true; false)
open import Data.Product.Base using (_×_; _,_)
open import Data.Product.Relation.Binary.Pointwise.NonDependent using (Pointwise)
open import Data.Sum.Base using (inj₁; inj₂)
open import Data.Sum.Base using (inj₁; inj₂; map)
open import Function.Base using (_∘_)
open import Relation.Binary.Core using (Rel)
open import Relation.Binary.Definitions using (Decidable)
Expand Down Expand Up @@ -107,6 +107,4 @@ sel ∙-sel ◦-sel (a , x) (b , y) with (a ∙ b) ≟₁ a | (a ∙ b) ≟₁ b
... | no ab≉a | no ab≉b = contradiction₂ (∙-sel a b) ab≉a ab≉b
... | yes ab≈a | no _ = inj₁ (ab≈a , ≈₂-refl)
... | no _ | yes ab≈b = inj₂ (ab≈b , ≈₂-refl)
... | yes ab≈a | yes ab≈b with ◦-sel x y
... | inj₁ xy≈x = inj₁ (ab≈a , xy≈x)
... | inj₂ xy≈y = inj₂ (ab≈b , xy≈y)
... | yes ab≈a | yes ab≈b = map (ab≈a ,_) (ab≈b ,_) (◦-sel x y)
13 changes: 6 additions & 7 deletions src/Algebra/Construct/LiftedChoice.agda
Original file line number Diff line number Diff line change
Expand Up @@ -11,13 +11,14 @@ open import Algebra
module Algebra.Construct.LiftedChoice where

open import Algebra.Consequences.Base
open import Relation.Binary.Core using (Rel; _⇒_; _Preserves_⟶_)
open import Relation.Binary.Structures using (IsEquivalence)
open import Relation.Nullary using (¬_; yes; no)
open import Data.Sum.Base as Sum using (_⊎_; inj₁; inj₂; [_,_])
open import Data.Sum.Base as Sum using (_⊎_; inj₁; inj₂; [_,_]; [_,_]′)
open import Data.Product.Base using (_×_; _,_)
open import Function.Base using (const; _$_)
open import Level using (Level; _⊔_)
open import Relation.Binary.Core using (Rel; _⇒_; _Preserves_⟶_)
open import Relation.Nullary using (¬_; yes; no)
open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_)
open import Relation.Binary.Structures using (IsEquivalence)
open import Relation.Unary using (Pred)

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

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

------------------------------------------------------------------------
-- Algebraic properties
Expand Down
4 changes: 2 additions & 2 deletions src/Algebra/Construct/NaturalChoice/MinOp.agda
Original file line number Diff line number Diff line change
Expand Up @@ -39,13 +39,13 @@ open import Relation.Binary.Reasoning.Preorder preorder

x⊓y≤x : ∀ x y → x ⊓ y ≤ x
x⊓y≤x x y with total x y
... | inj₁ x≤y = ≤-respˡ-≈ (Eq.sym (x≤y⇒x⊓y≈x x≤y)) refl
... | inj₁ x≤y = reflexive (x≤y⇒x⊓y≈x x≤y)
... | inj₂ y≤x = ≤-respˡ-≈ (Eq.sym (x≥y⇒x⊓y≈y y≤x)) y≤x

x⊓y≤y : ∀ x y → x ⊓ y ≤ y
x⊓y≤y x y with total x y
... | inj₁ x≤y = ≤-respˡ-≈ (Eq.sym (x≤y⇒x⊓y≈x x≤y)) x≤y
... | inj₂ y≤x = ≤-respˡ-≈ (Eq.sym (x≥y⇒x⊓y≈y y≤x)) refl
... | inj₂ y≤x = reflexive (x≥y⇒x⊓y≈y y≤x)

------------------------------------------------------------------------
-- Algebraic properties
Expand Down
7 changes: 4 additions & 3 deletions src/Codata/Sized/Cowriter.agda
Original file line number Diff line number Diff line change
Expand Up @@ -113,6 +113,7 @@ _>>=_ : ∀ {i} → Cowriter W A i → (A → Cowriter W X i) → Cowriter W X i
-- Construction.

unfold : ∀ {i} → (X → (W × X) ⊎ A) → X → Cowriter W A i
unfold next seed with next seed
... | inj₁ (w , seed′) = w ∷ λ where .force → unfold next seed′
... | inj₂ a = [ a ]
unfold next seed =
Sum.[ (λ where (w , seed′) → w ∷ λ where .force → unfold next seed′)
, [_]
] (next seed)
11 changes: 6 additions & 5 deletions src/Data/Fin/Base.agda
Original file line number Diff line number Diff line change
Expand Up @@ -15,7 +15,7 @@ open import Data.Bool.Base using (Bool; T)
open import Data.Nat.Base as ℕ using (ℕ; zero; suc)
open import Data.Product.Base as Product using (_×_; _,_; proj₁; proj₂)
open import Data.Sum.Base as Sum using (_⊎_; inj₁; inj₂; [_,_]′)
open import Function.Base using (id; _∘_; _on_; flip)
open import Function.Base using (id; _∘_; _on_; flip; _$_)
open import Level using (0ℓ)
open import Relation.Binary.Core
open import Relation.Binary.PropositionalEquality.Core using (_≡_; _≢_; refl; cong)
Expand Down Expand Up @@ -134,7 +134,7 @@ strengthen (suc i) = suc (strengthen i)
splitAt : ∀ m {n} → Fin (m ℕ.+ n) → Fin m ⊎ Fin n
splitAt zero i = inj₂ i
splitAt (suc m) zero = inj₁ zero
splitAt (suc m) (suc i) = Sum.map suc id (splitAt m i)
splitAt (suc m) (suc i) = Sum.map suc (splitAt m i)

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

quotRem : ∀ n → Fin (m ℕ.* n) → Fin n × Fin m
quotRem {suc m} n i with splitAt n i
... | inj₁ j = j , zero
... | inj₂ j = Product.map₂ suc (quotRem {m} n j)
quotRem {suc m} n i =
[ (_, zero)
, Product.map₂ suc ∘ quotRem {m} n
]′ $ splitAt n i

-- a variant of quotRem the type of whose result matches the order of multiplication
remQuot : ∀ n → Fin (m ℕ.* n) → Fin m × Fin n
Expand Down
19 changes: 9 additions & 10 deletions src/Data/List/Fresh.agda
Original file line number Diff line number Diff line change
Expand Up @@ -14,7 +14,7 @@
module Data.List.Fresh where

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

takeWhile [] = []
takeWhile (cons a as ps) with does (P? a)
... | true = cons a (takeWhile as) (takeWhile-# a as ps)
... | false = []
takeWhile (cons a as ps) =
if does (P? a) then cons a (takeWhile as) (takeWhile-# a as ps) else []

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

dropWhile : {R : Rel A r} → List# A R → List# A R
dropWhile [] = []
dropWhile aas@(a ∷# as) with does (P? a)
... | true = dropWhile as
... | false = aas
dropWhile aas@(a ∷# as) = if does (P? a) then dropWhile as else aas

filter : {R : Rel A r} → List# A R → List# A R
filter-# : ∀ {R : Rel A r} a (as : List# A R) → a # as → a # filter as

filter [] = []
filter (cons a as ps) with does (P? a)
... | true = cons a (filter as) (filter-# a as ps)
... | false = filter as
filter (cons a as ps) =
let l = filter as in
if does (P? a) then cons a l (filter-# a as ps) else l

-- this 'with' is needed to cause reduction in the type of 'filter-# a (x ∷# xs)'
filter-# a [] _ = _
filter-# a (x ∷# xs) (p , ps) with does (P? x)
... | true = p , filter-# a xs ps
Expand Down
10 changes: 6 additions & 4 deletions src/Data/List/Fresh/Relation/Unary/All.agda
Original file line number Diff line number Diff line change
Expand Up @@ -10,7 +10,8 @@ module Data.List.Fresh.Relation.Unary.All where

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

decide : Π[ P ∪ Q ] → Π[ All {R = R} P ∪ Any Q ]
decide p∪q [] = inj₁ []
decide p∪q (x ∷# xs) with p∪q x
... | inj₂ qx = inj₂ (here qx)
... | inj₁ px = Sum.map (px ∷_) there (decide p∪q xs)
decide p∪q (x ∷# xs) =
[ (λ px → Sum.map (px ∷_) there (decide p∪q xs))
, inj₂ ∘ here
]′ $ p∪q x
8 changes: 3 additions & 5 deletions src/Data/List/Kleene/Base.agda
Original file line number Diff line number Diff line change
Expand Up @@ -11,7 +11,7 @@ module Data.List.Kleene.Base where
open import Data.Product.Base as Product
using (_×_; _,_; map₂; map₁; proj₁; proj₂)
open import Data.Nat.Base as ℕ using (ℕ; suc; zero)
open import Data.Maybe.Base as Maybe using (Maybe; just; nothing)
open import Data.Maybe.Base as Maybe using (Maybe; just; nothing; maybe′)
open import Data.Sum.Base as Sum using (_⊎_; inj₁; inj₂)
open import Level using (Level)

Expand Down Expand Up @@ -137,9 +137,7 @@ module _ (f : A → Maybe B) where
mapMaybe+ : A + → B *
mapMaybe* : A * → B *

mapMaybe+ (x & xs) with f x
... | just y = ∹ y & mapMaybe* xs
... | nothing = mapMaybe* xs
mapMaybe+ (x & xs) = maybe′ (λ y z → ∹ y & z) id (f x) $ mapMaybe* xs

mapMaybe* [] = []
mapMaybe* (∹ xs) = mapMaybe+ xs
Expand Down Expand Up @@ -268,7 +266,7 @@ module _ (f : A → Maybe B → B) where
foldrMaybe* [] = nothing
foldrMaybe* (∹ xs) = just (foldrMaybe+ xs)

foldrMaybe+ xs = f (head xs) (foldrMaybe* (tail xs))
foldrMaybe+ (x & xs) = f x (foldrMaybe* xs)

------------------------------------------------------------------------
-- Indexing
Expand Down
7 changes: 5 additions & 2 deletions src/Data/List/NonEmpty/Base.agda
Original file line number Diff line number Diff line change
Expand Up @@ -203,9 +203,12 @@ snocView (x ∷ .(xs List.∷ʳ y)) | xs List.∷ʳ′ y = (x ∷ xs) ∷ʳ′ y

-- The last element in the list.

private
last′ : ∀ {l} → SnocView {A = A} l → A
last′ (_ ∷ʳ′ y) = y

last : List⁺ A → A
last xs with snocView xs
last .(ys ∷ʳ y) | ys ∷ʳ′ y = y
last = last′ ∘ snocView

-- Groups all contiguous elements for which the predicate returns the
-- same result into lists. The left sums are the ones for which the
Expand Down
6 changes: 2 additions & 4 deletions src/Data/List/Relation/Binary/Infix/Heterogeneous.agda
Original file line number Diff line number Diff line change
Expand Up @@ -49,10 +49,8 @@ map R⇒S (here pref) = here (Prefix.map R⇒S pref)
map R⇒S (there inf) = there (map R⇒S inf)

toView : ∀ {as bs} → Infix R as bs → View R as bs
toView (here p) with Prefix.toView p
...| inf Prefix.++ suff = MkView [] inf suff
toView (there p) with toView p
... | MkView pref inf suff = MkView (_ ∷ pref) inf suff
toView (here p) with inf Prefix.++ suff ← Prefix.toView p = MkView [] inf suff
toView (there p) with MkView pref inf suff ← toView p = MkView (_ ∷ pref) inf suff

fromView : ∀ {as bs} → View R as bs → Infix R as bs
fromView (MkView [] inf suff) = here (Prefix.fromView (inf Prefix.++ suff))
Expand Down
3 changes: 1 addition & 2 deletions src/Data/List/Relation/Binary/Prefix/Heterogeneous.agda
Original file line number Diff line number Diff line change
Expand Up @@ -53,8 +53,7 @@ module _ {a b r} {A : Set a} {B : Set b} {R : REL A B r} where

toView : ∀ {as bs} → Prefix R as bs → PrefixView R as bs
toView [] = [] ++ _
toView (r ∷ rs) with toView rs
... | rs′ ++ ds = (r ∷ rs′) ++ ds
toView (r ∷ rs) with rs′ ++ ds ← toView rs = (r ∷ rs′) ++ ds

fromView : ∀ {as bs} → PrefixView R as bs → Prefix R as bs
fromView ([] ++ ds) = []
Expand Down
3 changes: 1 addition & 2 deletions src/Data/List/Relation/Binary/Suffix/Heterogeneous.agda
Original file line number Diff line number Diff line change
Expand Up @@ -47,8 +47,7 @@ module _ {a b r} {A : Set a} {B : Set b} {R : REL A B r} where

toView : ∀ {as bs} → Suffix R as bs → SuffixView R as bs
toView (here rs) = [] ++ rs
toView (there {c} suf) with toView suf
... | cs ++ rs = (c ∷ cs) ++ rs
toView (there {c} suf) with cs ++ rs ← toView suf = (c ∷ cs) ++ rs

fromView : ∀ {as bs} → SuffixView R as bs → Suffix R as bs
fromView ([] ++ rs) = here rs
Expand Down
Original file line number Diff line number Diff line change
Expand Up @@ -40,8 +40,7 @@ module _ {a b r} {A : Set a} {B : Set b} {R : REL A B r} where

fromPrefix : ∀ {as bs} → Prefix R as bs →
Suffix R (reverse as) (reverse bs)
fromPrefix {as} {bs} p with Prefix.toView p
... | Prefix._++_ {cs} rs ds =
fromPrefix {as} {bs} p with Prefix._++_ {cs} rs ds ← Prefix.toView p =
subst (Suffix R (reverse as))
(sym (List.reverse-++ cs ds))
(Suffix.fromView (reverse ds Suffix.++ Pw.reverse⁺ rs))
Expand All @@ -56,8 +55,7 @@ module _ {a b r} {A : Set a} {B : Set b} {R : REL A B r} where

toPrefix-rev : ∀ {as bs} → Suffix R as bs →
Prefix R (reverse as) (reverse bs)
toPrefix-rev {as} {bs} s with Suffix.toView s
... | Suffix._++_ cs {ds} rs =
toPrefix-rev {as} {bs} s with Suffix._++_ cs {ds} rs ← Suffix.toView s =
subst (Prefix R (reverse as))
(sym (List.reverse-++ cs ds))
(Prefix.fromView (Pw.reverse⁺ rs Prefix.++ reverse cs))
Expand Down
Original file line number Diff line number Diff line change
Expand Up @@ -40,8 +40,8 @@ module _ (S : Setoid a ℓ₁) (T : Setoid b ℓ₂) (surj : Surjection S T) whe
open Surjection surj

map⁺ : ∀ {xs} → IsEnumeration S xs → IsEnumeration T (map to xs)
map⁺ _∈xs y with strictlySurjective y
... | (x , fx≈y) = ∈-resp-≈ T fx≈y (∈-map⁺ S T cong (x ∈xs))
map⁺ _∈xs y with (x , fx≈y) ← strictlySurjective y =
∈-resp-≈ T fx≈y (∈-map⁺ S T cong (x ∈xs))

------------------------------------------------------------------------
-- _++_
Expand Down
12 changes: 4 additions & 8 deletions src/Data/Nat/Binary/Subtraction.agda
Original file line number Diff line number Diff line change
Expand Up @@ -12,8 +12,8 @@ open import Algebra.Core using (Op₂)
open import Algebra.Bundles using (Magma)
open import Algebra.Consequences.Propositional using (comm∧distrˡ⇒distrʳ)
open import Algebra.Morphism.Consequences using (homomorphic₂-inv)
open import Data.Bool.Base using (true; false)
open import Data.Nat.Base as ℕ using (ℕ)
open import Data.Bool.Base using (true; false; if_then_else_)
open import Data.Nat as ℕ using (ℕ)
open import Data.Nat.Binary.Base using (ℕᵇ; 0ᵇ; 2[1+_]; 1+[2_]; double;
pred; toℕ; fromℕ; even<odd; odd<even; _≥_; _>_; _≤_; _<_; _+_; zero; suc; 1ᵇ;
_*_)
Expand Down Expand Up @@ -53,12 +53,8 @@ zero ∸ _ = 0ᵇ
x ∸ zero = x
2[1+ x ] ∸ 2[1+ y ] = double (x ∸ y)
1+[2 x ] ∸ 1+[2 y ] = double (x ∸ y)
2[1+ x ] ∸ 1+[2 y ] with does (x <? y)
... | true = 0ᵇ
... | false = 1+[2 (x ∸ y) ]
1+[2 x ] ∸ 2[1+ y ] with does (x ≤? y)
... | true = 0ᵇ
... | false = pred (double (x ∸ y))
2[1+ x ] ∸ 1+[2 y ] = if does (x <? y) then 0ᵇ else 1+[2 (x ∸ y) ]
1+[2 x ] ∸ 2[1+ y ] = if does (x ≤? y) then 0ᵇ else pred (double (x ∸ y))

------------------------------------------------------------------------
-- Properties of _∸_ and _≡_
Expand Down
8 changes: 2 additions & 6 deletions src/Data/Rational/Base.agda
Original file line number Diff line number Diff line change
Expand Up @@ -269,15 +269,11 @@ ceiling p@record{} = ℤ.- floor (- p)

-- Truncate (round towards 0)
truncate : ℚ → ℤ
truncate p with p ≤ᵇ 0ℚ
... | true = ceiling p
... | false = floor p
truncate p = if p ≤ᵇ 0ℚ then ceiling p else floor p

-- Round (to nearest integer)
round : ℚ → ℤ
round p with p ≤ᵇ 0ℚ
... | true = ceiling (p - ½)
... | false = floor (p + ½)
round p = if p ≤ᵇ 0ℚ then ceiling (p - ½) else floor (p + ½)

-- Fractional part (remainder after floor)
fracPart : ℚ → ℚ
Expand Down
8 changes: 2 additions & 6 deletions src/Data/Rational/Unnormalised/Base.agda
Original file line number Diff line number Diff line change
Expand Up @@ -276,15 +276,11 @@ ceiling p@record{} = ℤ.- floor (- p)

-- Truncate (round towards 0)
truncate : ℚᵘ → ℤ
truncate p with p ≤ᵇ 0ℚᵘ
... | true = ceiling p
... | false = floor p
truncate p = if p ≤ᵇ 0ℚᵘ then ceiling p else floor p

-- Round (to nearest integer)
round : ℚᵘ → ℤ
round p with p ≤ᵇ 0ℚᵘ
... | true = ceiling (p - ½)
... | false = floor (p + ½)
round p = if p ≤ᵇ 0ℚᵘ then ceiling (p - ½) else floor (p + ½)

-- Fractional part (remainder after floor)
fracPart : ℚᵘ → ℚᵘ
Expand Down
5 changes: 2 additions & 3 deletions src/Data/Star/Environment.agda
Original file line number Diff line number Diff line change
Expand Up @@ -13,7 +13,7 @@ open import Data.Star.List using (List)
open import Data.Star.Decoration using (All)
open import Data.Star.Pointer as Pointer using (Any; this; that; result)
open import Data.Unit.Polymorphic.Base using (⊤)
open import Function.Base using (const)
open import Function.Base using (const; case_of_)
open import Relation.Binary.PropositionalEquality.Core using (_≡_; refl)
open import Relation.Binary.Construct.Closure.ReflexiveTransitive
using (_▻_)
Expand Down Expand Up @@ -44,5 +44,4 @@ Env T Γ = All T Γ
-- A safe lookup function for environments.

lookup : ∀ {Γ σ} {T : Ty → Set} → Env T Γ → Γ ∋ σ → T σ
lookup ρ i with Pointer.lookup ρ i
... | result refl x = x
lookup ρ i with result refl x ← Pointer.lookup ρ i = x
6 changes: 3 additions & 3 deletions src/Data/Star/Pointer.agda
Original file line number Diff line number Diff line change
Expand Up @@ -11,7 +11,7 @@ module Data.Star.Pointer {ℓ} {I : Set ℓ} where
open import Data.Maybe.Base using (Maybe; nothing; just)
open import Data.Star.Decoration
open import Data.Unit.Base
open import Function.Base using (const)
open import Function.Base using (const; case_of_)
open import Level
open import Relation.Binary.Core using (Rel)
open import Relation.Binary.Definitions using (NonEmpty; nonEmpty)
Expand Down Expand Up @@ -92,5 +92,5 @@ module _ {T : Rel I r} {P : EdgePred p T} {Q : EdgePred q T} where

last : ∀ {i j} {xs : Star T i j} →
Any P Q xs → NonEmptyEdgePred T Q
last ps with lookup {r = p} (decorate (const (lift tt)) _) ps
... | result q _ = nonEmptyEdgePred q
last ps with result q _ ← lookup {r = p} (decorate (const (lift tt)) _) ps =
nonEmptyEdgePred q
5 changes: 2 additions & 3 deletions src/Data/Star/Vec.agda
Original file line number Diff line number Diff line change
Expand Up @@ -16,7 +16,7 @@ open import Data.Star.List using (List)
open import Level using (Level)
open import Relation.Binary.Construct.Closure.ReflexiveTransitive
using (ε; _◅_; gmap)
open import Function.Base using (const)
open import Function.Base using (const; case_of_)
open import Data.Unit.Base using (tt)

private
Expand Down Expand Up @@ -58,8 +58,7 @@ _++_ = _◅◅◅_
-- Safe lookup.

lookup : ∀ {n} → Vec A n → Fin n → A
lookup xs i with Pointer.lookup xs i
... | result _ x = x
lookup xs i with result _ x ← Pointer.lookup xs i = x

------------------------------------------------------------------------
-- Conversions
Expand Down
Loading