Skip to content

Commit 4220f83

Browse files
jamesmckinnaandreasabel
authored andcommitted
Deprecating Relation.Binary.PropositionalEquality.isPropositional (#2205)
* deprecating `isPropositional` * tighten `import`s * bumped Agda version number in comment
1 parent 586bca1 commit 4220f83

File tree

2 files changed

+26
-10
lines changed

2 files changed

+26
-10
lines changed

CHANGELOG.md

+5
Original file line numberDiff line numberDiff line change
@@ -1727,6 +1727,11 @@ Deprecated names
17271727
invPreorder ↦ converse-preorder
17281728
```
17291729

1730+
* In `Relation.Binary.PropositionalEquality`:
1731+
```agda
1732+
isPropositional ↦ Relation.Nullary.Irrelevant
1733+
```
1734+
17301735
* In `Relation.Unary.Consequences`:
17311736
```agda
17321737
dec⟶recomputable ↦ dec⇒recomputable

src/Relation/Binary/PropositionalEquality.agda

+21-10
Original file line numberDiff line numberDiff line change
@@ -8,14 +8,12 @@
88

99
module Relation.Binary.PropositionalEquality where
1010

11-
import Axiom.Extensionality.Propositional as Ext
1211
open import Axiom.UniquenessOfIdentityProofs
1312
open import Function.Base using (id; _∘_)
1413
import Function.Dependent.Bundles as Dependent
1514
open import Function.Indexed.Relation.Binary.Equality using (≡-setoid)
1615
open import Level using (Level; _⊔_)
17-
open import Data.Product.Base using (∃)
18-
16+
open import Relation.Nullary using (Irrelevant)
1917
open import Relation.Nullary.Decidable using (yes; no; dec-yes-irr; dec-no)
2018
open import Relation.Binary.Bundles using (Setoid)
2119
open import Relation.Binary.Definitions using (DecidableEquality)
@@ -60,12 +58,6 @@ _≗_ {A = A} {B = B} = Setoid._≈_ (A →-setoid B)
6058
Dependent.Func (setoid A) (Trivial.indexedSetoid B)
6159
→-to-⟶ = :→-to-Π
6260

63-
------------------------------------------------------------------------
64-
-- Propositionality
65-
66-
isPropositional : Set a Set a
67-
isPropositional A = (a b : A) a ≡ b
68-
6961
------------------------------------------------------------------------
7062
-- More complex rearrangement lemmas
7163

@@ -113,7 +105,7 @@ module _ (_≟_ : DecidableEquality A) {x y : A} where
113105
-- See README.Inspect for an explanation of how/why to use this.
114106

115107
-- Normally (but not always) the new `with ... in` syntax described at
116-
-- https://agda.readthedocs.io/en/v2.6.3/language/with-abstraction.html#with-abstraction-equality
108+
-- https://agda.readthedocs.io/en/v2.6.4/language/with-abstraction.html#with-abstraction-equality
117109
-- can be used instead."
118110

119111
record Reveal_·_is_ {A : Set a} {B : A Set b}
@@ -125,3 +117,22 @@ record Reveal_·_is_ {A : Set a} {B : A → Set b}
125117
inspect : {A : Set a} {B : A Set b}
126118
(f : (x : A) B x) (x : A) Reveal f · x is f x
127119
inspect f x = [ refl ]
120+
121+
122+
------------------------------------------------------------------------
123+
-- DEPRECATED NAMES
124+
------------------------------------------------------------------------
125+
-- Please use the new names as continuing support for the old names is
126+
-- not guaranteed.
127+
128+
-- Version 2.0
129+
130+
isPropositional : Set a Set a
131+
isPropositional = Irrelevant
132+
133+
{-# WARNING_ON_USAGE isPropositional
134+
"Warning: isPropositional was deprecated in v2.0.
135+
Please use Relation.Nullary.Irrelevant instead. "
136+
#-}
137+
138+

0 commit comments

Comments
 (0)