Skip to content

Commit debfec1

Browse files
propagate use of shortcut to where it increases readability
1 parent 856fbf8 commit debfec1

File tree

7 files changed

+24
-22
lines changed

7 files changed

+24
-22
lines changed

src/Data/Product/Function/Dependent/Setoid.agda

+5-3
Original file line numberDiff line numberDiff line change
@@ -14,7 +14,9 @@ module Data.Product.Function.Dependent.Setoid where
1414
open import Data.Product.Base using (map; _,_; proj₁; proj₂)
1515
open import Data.Product.Relation.Binary.Pointwise.Dependent as Σ
1616
open import Level using (Level)
17-
open import Function
17+
open import Function.Base using (_$_; _∘_)
18+
open import Function.Bundles
19+
open import Function.Definitions
1820
open import Function.Consequences.Setoid
1921
open import Function.Properties.Injection using (mkInjection)
2022
open import Function.Properties.Surjection using (mkSurjection; ↠⇒⇔)
@@ -67,8 +69,8 @@ module _ where
6769

6870
function :
6971
(f : I ⟶ J)
70-
( {i} Func (A atₛ i) (B atₛ (to f i)))
71-
Func (I ×ₛ A) (J ×ₛ B)
72+
( {i} (A atₛ i) ⟶ₛ (B atₛ (to f i)))
73+
I ×ₛ A ⟶ₛ J ×ₛ B
7274
function {I = I} {J = J} {A = A} {B = B} I⟶J A⟶B = record
7375
{ to = to′
7476
; cong = cong′

src/Data/Product/Function/NonDependent/Setoid.agda

+8-8
Original file line numberDiff line numberDiff line change
@@ -10,10 +10,10 @@
1010
module Data.Product.Function.NonDependent.Setoid where
1111

1212
open import Data.Product.Base as Prod
13-
open import Data.Product.Relation.Binary.Pointwise.NonDependent
13+
open import Data.Product.Relation.Binary.Pointwise.NonDependent using (_×ₛ_)
1414
open import Level using (Level)
15-
open import Relation.Binary
16-
open import Function
15+
open import Relation.Binary using (Setoid)
16+
open import Function.Bundles -- much of it is used
1717

1818
private
1919
variable
@@ -24,25 +24,25 @@ private
2424
------------------------------------------------------------------------
2525
-- Combinators for equality preserving functions
2626

27-
proj₁ₛ : Func (A ×ₛ B) A
27+
proj₁ₛ : A ×ₛ B ⟶ₛ A
2828
proj₁ₛ = record { to = proj₁ ; cong = proj₁ }
2929

30-
proj₂ₛ : Func (A ×ₛ B) B
30+
proj₂ₛ : A ×ₛ B ⟶ₛ B
3131
proj₂ₛ = record { to = proj₂ ; cong = proj₂ }
3232

33-
<_,_>ₛ : Func A B Func A C Func A (B ×ₛ C)
33+
<_,_>ₛ : A ⟶ₛ B A ⟶ₛ C A ⟶ₛ B ×ₛ C
3434
< f , g >ₛ = record
3535
{ to = < to f , to g >
3636
; cong = < cong f , cong g >
3737
} where open Func
3838

39-
swapₛ : Func (A ×ₛ B) (B ×ₛ A)
39+
swapₛ : A ×ₛ B ⟶ₛ B ×ₛ A
4040
swapₛ = < proj₂ₛ , proj₁ₛ >ₛ
4141

4242
------------------------------------------------------------------------
4343
-- Function bundles
4444

45-
_×-function_ : Func A B Func C D Func (A ×ₛ C) (B ×ₛ D)
45+
_×-function_ : A ⟶ₛ B C ⟶ₛ D A ×ₛ C ⟶ₛ B ×ₛ D
4646
f ×-function g = record
4747
{ to = Prod.map (to f) (to g)
4848
; cong = Prod.map (cong f) (cong g)

src/Data/Sum/Function/Setoid.agda

+6-6
Original file line numberDiff line numberDiff line change
@@ -11,8 +11,8 @@ module Data.Sum.Function.Setoid where
1111
open import Data.Product.Base as Prod using (_,_)
1212
open import Data.Sum.Base as Sum
1313
open import Data.Sum.Relation.Binary.Pointwise as Pointwise
14-
open import Relation.Binary
15-
open import Function.Base
14+
open import Relation.Binary using (Setoid; Rel)
15+
open import Function.Base using (_$_; _∘_)
1616
open import Function.Bundles
1717
open import Function.Definitions
1818
open import Level
@@ -28,21 +28,21 @@ private
2828
------------------------------------------------------------------------
2929
-- Combinators for equality preserving functions
3030

31-
inj₁ₛ : Func S (S ⊎ₛ T)
31+
inj₁ₛ : S ⟶ₛ (S ⊎ₛ T)
3232
inj₁ₛ = record { to = inj₁ ; cong = inj₁ }
3333

34-
inj₂ₛ : Func T (S ⊎ₛ T)
34+
inj₂ₛ : T ⟶ₛ (S ⊎ₛ T)
3535
inj₂ₛ = record { to = inj₂ ; cong = inj₂ }
3636

37-
[_,_]ₛ : Func S U Func T U Func (S ⊎ₛ T) U
37+
[_,_]ₛ : S ⟶ₛ U T ⟶ₛ U S ⊎ₛ T ⟶ₛ U
3838
[ f , g ]ₛ = record
3939
{ to = [ to f , to g ]
4040
; cong = λ where
4141
(inj₁ x∼₁y) cong f x∼₁y
4242
(inj₂ x∼₂y) cong g x∼₂y
4343
} where open Func
4444

45-
swapₛ : Func (S ⊎ₛ T) (T ⊎ₛ S)
45+
swapₛ : (S ⊎ₛ T) ⟶ₛ (T ⊎ₛ S)
4646
swapₛ = [ inj₂ₛ , inj₁ₛ ]ₛ
4747

4848
------------------------------------------------------------------------

src/Function/Properties/Equivalence.agda

+1-1
Original file line numberDiff line numberDiff line change
@@ -29,7 +29,7 @@ private
2929
------------------------------------------------------------------------
3030
-- Constructors
3131

32-
mkEquivalence : Func S T Func T S Equivalence S T
32+
mkEquivalence : S ⟶ₛ T T ⟶ₛ S Equivalence S T
3333
mkEquivalence f g = record
3434
{ to = to f
3535
; from = to g

src/Function/Properties/Injection.agda

+1-1
Original file line numberDiff line numberDiff line change
@@ -29,7 +29,7 @@ private
2929
------------------------------------------------------------------------
3030
-- Constructors
3131

32-
mkInjection : (f : Func S T) (open Func f)
32+
mkInjection : (f : S ⟶ₛ T) (open Func f)
3333
Injective Eq₁._≈_ Eq₂._≈_ to
3434
Injection S T
3535
mkInjection f injective = record

src/Function/Properties/Inverse.agda

+2-2
Original file line numberDiff line numberDiff line change
@@ -68,11 +68,11 @@ isEquivalence = record
6868
------------------------------------------------------------------------
6969
-- Conversion functions
7070

71-
toFunction : Inverse S T Func S T
71+
toFunction : Inverse S T S ⟶ₛ T
7272
toFunction I = record { to = to ; cong = to-cong }
7373
where open Inverse I
7474

75-
fromFunction : Inverse S T Func T S
75+
fromFunction : Inverse S T T ⟶ₛ S
7676
fromFunction I = record { to = from ; cong = from-cong }
7777
where open Inverse I
7878

src/Function/Properties/Surjection.agda

+1-1
Original file line numberDiff line numberDiff line change
@@ -29,7 +29,7 @@ private
2929
------------------------------------------------------------------------
3030
-- Constructors
3131

32-
mkSurjection : (f : Func S T) (open Func f)
32+
mkSurjection : (f : S ⟶ₛ T) (open Func f)
3333
Surjective Eq₁._≈_ Eq₂._≈_ to
3434
Surjection S T
3535
mkSurjection f surjective = record

0 commit comments

Comments
 (0)