Skip to content

Commit e8dcbfd

Browse files
MatthewDaggittandreasabel
authored andcommitted
Add consistent deprecation warnings to old function hierarchy (#2163)
1 parent b391a3d commit e8dcbfd

13 files changed

+340
-31
lines changed

CHANGELOG.md

+1
Original file line numberDiff line numberDiff line change
@@ -1993,6 +1993,7 @@ New modules
19931993
Function.Properties.Bijection
19941994
Function.Properties.RightInverse
19951995
Function.Properties.Surjection
1996+
Function.Construct.Constant
19961997
```
19971998
19981999
* In order to improve modularity, the contents of `Relation.Binary.Lattice` has been

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

+4-5
Original file line numberDiff line numberDiff line change
@@ -16,11 +16,10 @@ open import Data.Product.Relation.Binary.Pointwise.Dependent as Σ
1616
open import Level using (Level)
1717
open import Function
1818
open import Function.Consequences
19-
open import Function.Properties.Injection
20-
open import Function.Properties.Surjection
21-
open import Function.Properties.Equivalence
22-
open import Function.Properties.RightInverse
23-
import Function.Properties.Inverse as InverseProperties
19+
open import Function.Properties.Injection using (mkInjection)
20+
open import Function.Properties.Surjection using (mkSurjection; ↠⇒⇔)
21+
open import Function.Properties.Equivalence using (mkEquivalence; ⇔⇒⟶; ⇔⇒⟵)
22+
open import Function.Properties.RightInverse using (mkRightInverse)
2423
open import Relation.Binary.Core using (_=[_]⇒_)
2524
open import Relation.Binary.Bundles as B
2625
open import Relation.Binary.Indexed.Heterogeneous

src/Function/Bijection.agda

+27
Original file line numberDiff line numberDiff line change
@@ -38,6 +38,10 @@ record Bijective {f₁ f₂ t₁ t₂}
3838

3939
left-inverse-of : from LeftInverseOf to
4040
left-inverse-of x = injective (right-inverse-of (to ⟨$⟩ x))
41+
{-# WARNING_ON_USAGE Bijective
42+
"Warning: Bijective was deprecated in v2.0.
43+
Please use Function.(Structures.)IsBijection instead."
44+
#-}
4145

4246
------------------------------------------------------------------------
4347
-- The set of all bijections between two setoids.
@@ -74,6 +78,10 @@ record Bijection {f₁ f₂ t₁ t₂}
7478
}
7579

7680
open LeftInverse left-inverse public using (to-from)
81+
{-# WARNING_ON_USAGE Bijection
82+
"Warning: Bijection was deprecated in v2.0.
83+
Please use Function.(Bundles.)Bijection instead."
84+
#-}
7785

7886
------------------------------------------------------------------------
7987
-- The set of all bijections between two sets (i.e. bijections with
@@ -83,6 +91,10 @@ infix 3 _⤖_
8391

8492
_⤖_ : {f t} Set f Set t Set _
8593
From ⤖ To = Bijection (P.setoid From) (P.setoid To)
94+
{-# WARNING_ON_USAGE _⤖_
95+
"Warning: _⤖_ was deprecated in v2.0.
96+
Please use Function.(Bundles.)mk⤖ instead."
97+
#-}
8698

8799
bijection : {f t} {From : Set f} {To : Set t}
88100
(to : From To) (from : To From)
@@ -99,6 +111,11 @@ bijection to from inj invʳ = record
99111
}
100112
}
101113
}
114+
{-# WARNING_ON_USAGE bijection
115+
"Warning: bijection was deprecated in v2.0.
116+
Please use either Function.Properties.Bijection.trans or
117+
Function.Construct.Composition.bijection instead."
118+
#-}
102119

103120
------------------------------------------------------------------------
104121
-- Identity and composition. (Note that these proofs are superfluous,
@@ -112,6 +129,11 @@ id {S = S} = record
112129
; surjective = Surjection.surjective (Surj.id {S = S})
113130
}
114131
}
132+
{-# WARNING_ON_USAGE id
133+
"Warning: id was deprecated in v2.0.
134+
Please use either Function.Properties.Bijection.refl or
135+
Function.Construct.Identity.bijection instead."
136+
#-}
115137

116138
infixr 9 _∘_
117139

@@ -125,3 +147,8 @@ f ∘ g = record
125147
; surjective = Surjection.surjective (Surj._∘_ (surjection f) (surjection g))
126148
}
127149
} where open Bijection
150+
{-# WARNING_ON_USAGE _∘_
151+
"Warning: _∘_ was deprecated in v2.0.
152+
Please use either Function.Properties.Bijection.trans or
153+
Function.Construct.Composition.bijection instead."
154+
#-}

src/Function/Construct/Constant.agda

+64
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,64 @@
1+
------------------------------------------------------------------------
2+
-- The Agda standard library
3+
--
4+
-- The constant function
5+
------------------------------------------------------------------------
6+
7+
{-# OPTIONS --cubical-compatible --safe #-}
8+
9+
module Function.Construct.Constant where
10+
11+
open import Function.Base using (const)
12+
open import Function.Bundles
13+
import Function.Definitions as Definitions
14+
import Function.Structures as Structures
15+
open import Level using (Level)
16+
open import Relation.Binary.Core using (Rel)
17+
open import Relation.Binary.Bundles using (Setoid)
18+
open import Relation.Binary.Structures as B hiding (IsEquivalence)
19+
20+
private
21+
variable
22+
a b ℓ₁ ℓ₂ : Level
23+
A B : Set a
24+
25+
------------------------------------------------------------------------
26+
-- Properties
27+
28+
module _ (_≈₁_ : Rel A ℓ₁) (_≈₂_ : Rel B ℓ₂) where
29+
30+
open Definitions
31+
32+
congruent : {b} b ≈₂ b Congruent _≈₁_ _≈₂_ (const b)
33+
congruent refl _ = refl
34+
35+
------------------------------------------------------------------------
36+
-- Structures
37+
38+
module _
39+
{≈₁ : Rel A ℓ₁} {≈₂ : Rel B ℓ₂}
40+
(isEq₁ : B.IsEquivalence ≈₁)
41+
(isEq₂ : B.IsEquivalence ≈₂) where
42+
43+
open Structures ≈₁ ≈₂
44+
open B.IsEquivalence
45+
46+
isCongruent : b IsCongruent (const b)
47+
isCongruent b = record
48+
{ cong = congruent ≈₁ ≈₂ (refl isEq₂)
49+
; isEquivalence₁ = isEq₁
50+
; isEquivalence₂ = isEq₂
51+
}
52+
53+
------------------------------------------------------------------------
54+
-- Setoid bundles
55+
56+
module _ (S : Setoid a ℓ₂) (T : Setoid b ℓ₂) where
57+
58+
open Setoid
59+
60+
function : Carrier T Func S T
61+
function b = record
62+
{ to = const b
63+
; cong = congruent (_≈_ S) (_≈_ T) (refl T)
64+
}

src/Function/Equality.agda

+21
Original file line numberDiff line numberDiff line change
@@ -5,6 +5,7 @@
55
------------------------------------------------------------------------
66

77
{-# OPTIONS --cubical-compatible --safe #-}
8+
{-# OPTIONS --warn=noUserWarning #-}
89

910
module Function.Equality where
1011

@@ -34,19 +35,31 @@ record Π {f₁ f₂ t₁ t₂}
3435
field
3536
_⟨$⟩_ : (x : Setoid.Carrier From) IndexedSetoid.Carrier To x
3637
cong : Setoid._≈_ From =[ _⟨$⟩_ ]⇒ IndexedSetoid._≈_ To
38+
{-# WARNING_ON_USAGE Π
39+
"Warning: Π was deprecated in v2.0.
40+
Please use Function.Dependent.Bundles.Func instead."
41+
#-}
3742

3843
open Π public
3944

4045
infixr 0 _⟶_
4146

4247
_⟶_ : {f₁ f₂ t₁ t₂} Setoid f₁ f₂ Setoid t₁ t₂ Set _
4348
From ⟶ To = Π From (Trivial.indexedSetoid To)
49+
{-# WARNING_ON_USAGE _⟶_
50+
"Warning: _⟶_ was deprecated in v2.0.
51+
Please use Function.(Bundles.)Func instead."
52+
#-}
4453

4554
------------------------------------------------------------------------
4655
-- Identity and composition.
4756

4857
id : {a₁ a₂} {A : Setoid a₁ a₂} A ⟶ A
4958
id = record { _⟨$⟩_ = Fun.id; cong = Fun.id }
59+
{-# WARNING_ON_USAGE id
60+
"Warning: id was deprecated in v2.0.
61+
Please use Function.Construct.Identity.function instead."
62+
#-}
5063

5164
infixr 9 _∘_
5265

@@ -58,6 +71,10 @@ f ∘ g = record
5871
{ _⟨$⟩_ = Fun._∘_ (_⟨$⟩_ f) (_⟨$⟩_ g)
5972
; cong = Fun._∘_ (cong f) (cong g)
6073
}
74+
{-# WARNING_ON_USAGE _∘_
75+
"Warning: _∘_ was deprecated in v2.0.
76+
Please use Function.Construct.Composition.function instead."
77+
#-}
6178

6279
-- Constant equality-preserving function.
6380

@@ -68,6 +85,10 @@ const {B = B} b = record
6885
{ _⟨$⟩_ = Fun.const b
6986
; cong = Fun.const (Setoid.refl B)
7087
}
88+
{-# WARNING_ON_USAGE const
89+
"Warning: const was deprecated in v2.0.
90+
Please use Function.Construct.Constant.function instead."
91+
#-}
7192

7293
------------------------------------------------------------------------
7394
-- Function setoids

src/Function/Equivalence.agda

+35
Original file line numberDiff line numberDiff line change
@@ -31,6 +31,10 @@ record Equivalence {f₁ f₂ t₁ t₂}
3131
field
3232
to : From ⟶ To
3333
from : To ⟶ From
34+
{-# WARNING_ON_USAGE Equivalence
35+
"Warning: Equivalence was deprecated in v2.0.
36+
Please use Function.(Bundles.)Equivalence instead."
37+
#-}
3438

3539
------------------------------------------------------------------------
3640
-- The set of all equivalences between two sets (i.e. equivalences
@@ -40,13 +44,21 @@ infix 3 _⇔_
4044

4145
_⇔_ : {f t} Set f Set t Set _
4246
From ⇔ To = Equivalence (P.setoid From) (P.setoid To)
47+
{-# WARNING_ON_USAGE _⇔_
48+
"Warning: _⇔_ was deprecated in v2.0.
49+
Please use Function.(Bundles.)_⇔_ instead."
50+
#-}
4351

4452
equivalence : {f t} {From : Set f} {To : Set t}
4553
(From To) (To From) From ⇔ To
4654
equivalence to from = record
4755
{ to = →-to-⟶ to
4856
; from = →-to-⟶ from
4957
}
58+
{-# WARNING_ON_USAGE equivalence
59+
"Warning: equivalence was deprecated in v2.0.
60+
Please use Function.Properties.Equivalence.mkEquivalence instead."
61+
#-}
5062

5163
------------------------------------------------------------------------
5264
-- Equivalence is an equivalence relation
@@ -58,6 +70,11 @@ id {x = S} = record
5870
{ to = F.id
5971
; from = F.id
6072
}
73+
{-# WARNING_ON_USAGE id
74+
"Warning: id was deprecated in v2.0.
75+
Please use Function.Properties.Equivalence.refl or
76+
Function.Construct.Identity.equivalence instead."
77+
#-}
6178

6279
infixr 9 _∘_
6380

@@ -69,6 +86,11 @@ f ∘ g = record
6986
{ to = to f ⟪∘⟫ to g
7087
; from = from g ⟪∘⟫ from f
7188
} where open Equivalence
89+
{-# WARNING_ON_USAGE _∘_
90+
"Warning: _∘_ was deprecated in v2.0.
91+
Please use Function.Properties.Equivalence.trans or
92+
Function.Construct.Composition.equivalence instead."
93+
#-}
7294

7395
-- Symmetry.
7496

@@ -79,6 +101,11 @@ sym eq = record
79101
{ from = to
80102
; to = from
81103
} where open Equivalence eq
104+
{-# WARNING_ON_USAGE sym
105+
"Warning: sym was deprecated in v2.0.
106+
Please use Function.Properties.Equivalence.sym or
107+
Function.Construct.Symmetry.equivalence instead."
108+
#-}
82109

83110
-- For fixed universe levels we can construct setoids.
84111

@@ -92,6 +119,10 @@ setoid s₁ s₂ = record
92119
; trans = flip _∘_
93120
}
94121
}
122+
{-# WARNING_ON_USAGE setoid
123+
"Warning: setoid was deprecated in v2.0.
124+
Please use Function.Properties.Equivalence.setoid instead."
125+
#-}
95126

96127
⇔-setoid : (ℓ : Level) Setoid (suc ℓ) ℓ
97128
⇔-setoid ℓ = record
@@ -103,6 +134,10 @@ setoid s₁ s₂ = record
103134
; trans = flip _∘_
104135
}
105136
}
137+
{-# WARNING_ON_USAGE ⇔-setoid
138+
"Warning: ⇔-setoid was deprecated in v2.0.
139+
Please use Function.Properties.Equivalence.⇔-setoid instead."
140+
#-}
106141

107142
------------------------------------------------------------------------
108143
-- Transformations

src/Function/Injection.agda

+26-5
Original file line numberDiff line numberDiff line change
@@ -7,11 +7,6 @@
77
{-# OPTIONS --cubical-compatible --safe #-}
88
{-# OPTIONS --warn=noUserWarning #-}
99

10-
-- Note: use of the standard function hierarchy is encouraged. The
11-
-- module `Function` re-exports `Injective`, `IsInjection` and
12-
-- `Injection`. The alternative definitions found in this file will
13-
-- eventually be deprecated.
14-
1510
module Function.Injection where
1611

1712
{-# WARNING_ON_IMPORT
@@ -35,6 +30,10 @@ Injective {A = A} {B} f = ∀ {x y} → f ⟨$⟩ x ≈₂ f ⟨$⟩ y → x ≈
3530
where
3631
open Setoid A renaming (_≈_ to _≈₁_)
3732
open Setoid B renaming (_≈_ to _≈₂_)
33+
{-# WARNING_ON_USAGE Injective
34+
"Warning: Injective was deprecated in v2.0.
35+
Please use Function.(Definitions.)Injective instead."
36+
#-}
3837

3938
------------------------------------------------------------------------
4039
-- The set of all injections between two setoids
@@ -47,6 +46,10 @@ record Injection {f₁ f₂ t₁ t₂}
4746
injective : Injective to
4847

4948
open Π to public
49+
{-# WARNING_ON_USAGE Injection
50+
"Warning: Injection was deprecated in v2.0.
51+
Please use Function.(Bundles.)Injection instead."
52+
#-}
5053

5154
------------------------------------------------------------------------
5255
-- The set of all injections from one set to another (i.e. injections
@@ -56,6 +59,10 @@ infix 3 _↣_
5659

5760
_↣_ : {f t} Set f Set t Set _
5861
From ↣ To = Injection (P.setoid From) (P.setoid To)
62+
{-# WARNING_ON_USAGE _↣_
63+
"Warning: _↣_ was deprecated in v2.0.
64+
Please use Function.(Bundles.)_↣_ instead."
65+
#-}
5966

6067
injection : {f t} {From : Set f} {To : Set t} (to : From To)
6168
( {x y} to x ≡ to y x ≡ y) From ↣ To
@@ -66,6 +73,10 @@ injection to injective = record
6673
}
6774
; injective = injective
6875
}
76+
{-# WARNING_ON_USAGE injection
77+
"Warning: injection was deprecated in v2.0.
78+
Please use Function.(Bundles.)mk↣ instead."
79+
#-}
6980

7081
------------------------------------------------------------------------
7182
-- Identity and composition.
@@ -77,6 +88,11 @@ id = record
7788
{ to = F.id
7889
; injective = Fun.id
7990
}
91+
{-# WARNING_ON_USAGE id
92+
"Warning: id was deprecated in v2.0.
93+
Please use Function.Properties.Injection.refl or
94+
Function.Construct.Identity.injection instead."
95+
#-}
8096

8197
_∘_ : {f₁ f₂ m₁ m₂ t₁ t₂}
8298
{F : Setoid f₁ f₂} {M : Setoid m₁ m₂} {T : Setoid t₁ t₂}
@@ -85,3 +101,8 @@ f ∘ g = record
85101
{ to = to f ⟪∘⟫ to g
86102
; injective = (λ {_} injective g) ⟨∘⟩ injective f
87103
} where open Injection
104+
{-# WARNING_ON_USAGE _∘_
105+
"Warning: _∘_ was deprecated in v2.0.
106+
Please use Function.Properties.Injection.trans or
107+
Function.Construct.Composition.injection instead."
108+
#-}

0 commit comments

Comments
 (0)