8
8
9
9
module Relation.Binary.PropositionalEquality where
10
10
11
- import Axiom.Extensionality.Propositional as Ext
12
11
open import Axiom.UniquenessOfIdentityProofs
13
12
open import Function.Base using (id; _∘_)
14
13
import Function.Dependent.Bundles as Dependent
15
14
open import Function.Indexed.Relation.Binary.Equality using (≡-setoid)
16
15
open import Level using (Level; _⊔_)
17
- open import Data.Product.Base using (∃)
18
-
16
+ open import Relation.Nullary using (Irrelevant)
19
17
open import Relation.Nullary.Decidable using (yes; no; dec-yes-irr; dec-no)
20
18
open import Relation.Binary.Bundles using (Setoid)
21
19
open import Relation.Binary.Definitions using (DecidableEquality)
@@ -60,12 +58,6 @@ _≗_ {A = A} {B = B} = Setoid._≈_ (A →-setoid B)
60
58
Dependent.Func (setoid A) (Trivial.indexedSetoid B)
61
59
→-to-⟶ = :→-to-Π
62
60
63
- ------------------------------------------------------------------------
64
- -- Propositionality
65
-
66
- isPropositional : Set a → Set a
67
- isPropositional A = (a b : A) → a ≡ b
68
-
69
61
------------------------------------------------------------------------
70
62
-- More complex rearrangement lemmas
71
63
@@ -113,7 +105,7 @@ module _ (_≟_ : DecidableEquality A) {x y : A} where
113
105
-- See README.Inspect for an explanation of how/why to use this.
114
106
115
107
-- 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
117
109
-- can be used instead."
118
110
119
111
record Reveal_·_is_ {A : Set a} {B : A → Set b}
@@ -125,3 +117,22 @@ record Reveal_·_is_ {A : Set a} {B : A → Set b}
125
117
inspect : ∀ {A : Set a} {B : A → Set b}
126
118
(f : (x : A) → B x) (x : A) → Reveal f · x is f x
127
119
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