Skip to content

Commit 6078b64

Browse files
authored
Add (Is)DecPreorder to Relation.Binary.* (#2488)
* fix issue #2482 * add `isDecPreorder` manifest fields to `IsDec*Order` structures * refactored derived sub-`Bundles` * fixed knock-on consequences
1 parent 22bfd05 commit 6078b64

File tree

4 files changed

+90
-31
lines changed

4 files changed

+90
-31
lines changed

CHANGELOG.md

Lines changed: 18 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -322,17 +322,35 @@ Additions to existing modules
322322
_≡?_ : DecidableEquality (Vec A n)
323323
```
324324

325+
* In `Relation.Binary.Bundles`:
326+
```agda
327+
record DecPreorder c ℓ₁ ℓ₂ : Set (suc (c ⊔ ℓ₁ ⊔ ℓ₂))
328+
```
329+
plus associated sub-bundles.
330+
325331
* In `Relation.Binary.Construct.Interior.Symmetric`:
326332
```agda
327333
decidable : Decidable R → Decidable (SymInterior R)
328334
```
329335
and for `Reflexive` and `Transitive` relations `R`:
330336
```agda
331337
isDecEquivalence : Decidable R → IsDecEquivalence (SymInterior R)
338+
isDecPreorder : Decidable R → IsDecPreorder (SymInterior R) R
332339
isDecPartialOrder : Decidable R → IsDecPartialOrder (SymInterior R) R
340+
decPreorder : Decidable R → DecPreorder _ _ _
333341
decPoset : Decidable R → DecPoset _ _ _
334342
```
335343

344+
* In `Relation.Binary.Structures`:
345+
```agda
346+
record IsDecPreorder (_≲_ : Rel A ℓ₂) : Set (a ⊔ ℓ ⊔ ℓ₂) where
347+
field
348+
isPreorder : IsPreorder _≲_
349+
_≟_ : Decidable _≈_
350+
_≲?_ : Decidable _≲_
351+
```
352+
plus associated `isDecPreorder` fields in each higher `IsDec*Order` structure.
353+
336354
* In `Relation.Nullary.Decidable`:
337355
```agda
338356
does-⇔ : A ⇔ B → (a? : Dec A) → (b? : Dec B) → does a? ≡ does b?

src/Relation/Binary/Bundles.agda

Lines changed: 35 additions & 7 deletions
Original file line numberDiff line numberDiff line change
@@ -132,6 +132,36 @@ record TotalPreorder c ℓ₁ ℓ₂ : Set (suc (c ⊔ ℓ₁ ⊔ ℓ₂)) where
132132
open Preorder preorder public
133133
hiding (Carrier; _≈_; _≲_; isPreorder)
134134

135+
136+
record DecPreorder c ℓ₁ ℓ₂ : Set (suc (c ⊔ ℓ₁ ⊔ ℓ₂)) where
137+
field
138+
Carrier : Set c
139+
_≈_ : Rel Carrier ℓ₁ -- The underlying equality.
140+
_≲_ : Rel Carrier ℓ₂ -- The relation.
141+
isDecPreorder : IsDecPreorder _≈_ _≲_
142+
143+
private module DPO = IsDecPreorder isDecPreorder
144+
145+
open DPO public
146+
using (_≟_; _≲?_; isPreorder)
147+
148+
preorder : Preorder c ℓ₁ ℓ₂
149+
preorder = record
150+
{ isPreorder = isPreorder
151+
}
152+
153+
open Preorder preorder public
154+
hiding (Carrier; _≈_; _≲_; isPreorder; module Eq)
155+
156+
module Eq where
157+
decSetoid : DecSetoid c ℓ₁
158+
decSetoid = record
159+
{ isDecEquivalence = DPO.Eq.isDecEquivalence
160+
}
161+
162+
open DecSetoid decSetoid public
163+
164+
135165
------------------------------------------------------------------------
136166
-- Partial orders
137167
------------------------------------------------------------------------
@@ -173,7 +203,7 @@ record DecPoset c ℓ₁ ℓ₂ : Set (suc (c ⊔ ℓ₁ ⊔ ℓ₂)) where
173203
private module DPO = IsDecPartialOrder isDecPartialOrder
174204

175205
open DPO public
176-
using (_≟_; _≤?_; isPartialOrder)
206+
using (_≟_; _≤?_; isPartialOrder; isDecPreorder)
177207

178208
poset : Poset c ℓ₁ ℓ₂
179209
poset = record
@@ -183,13 +213,11 @@ record DecPoset c ℓ₁ ℓ₂ : Set (suc (c ⊔ ℓ₁ ⊔ ℓ₂)) where
183213
open Poset poset public
184214
hiding (Carrier; _≈_; _≤_; isPartialOrder; module Eq)
185215

186-
module Eq where
187-
decSetoid : DecSetoid c ℓ₁
188-
decSetoid = record
189-
{ isDecEquivalence = DPO.Eq.isDecEquivalence
190-
}
216+
decPreorder : DecPreorder c ℓ₁ ℓ₂
217+
decPreorder = record { isDecPreorder = isDecPreorder }
191218

192-
open DecSetoid decSetoid public
219+
open DecPreorder decPreorder public
220+
using (module Eq)
193221

194222

195223
record StrictPartialOrder c ℓ₁ ℓ₂ : Set (suc (c ⊔ ℓ₁ ⊔ ℓ₂)) where

src/Relation/Binary/Construct/Interior/Symmetric.agda

Lines changed: 7 additions & 6 deletions
Original file line numberDiff line numberDiff line change
@@ -100,12 +100,6 @@ module _ {R : Rel A ℓ} (refl : Reflexive R) (trans : Transitive R) where
100100

101101
module _ (R? : Decidable R) where
102102

103-
isDecEquivalence : IsDecEquivalence (SymInterior R)
104-
isDecEquivalence = record
105-
{ isEquivalence = isEquivalence
106-
; _≟_ = decidable R?
107-
}
108-
109103
isDecPartialOrder : IsDecPartialOrder (SymInterior R) R
110104
isDecPartialOrder = record
111105
{ isPartialOrder = isPartialOrder
@@ -115,3 +109,10 @@ module _ {R : Rel A ℓ} (refl : Reflexive R) (trans : Transitive R) where
115109

116110
decPoset : DecPoset _ ℓ ℓ
117111
decPoset = record { isDecPartialOrder = isDecPartialOrder }
112+
113+
open DecPoset public
114+
using (isDecPreorder; decPreorder)
115+
116+
isDecEquivalence : IsDecEquivalence (SymInterior R)
117+
isDecEquivalence = DecPoset.Eq.isDecEquivalence decPoset
118+

src/Relation/Binary/Structures.agda

Lines changed: 30 additions & 18 deletions
Original file line numberDiff line numberDiff line change
@@ -119,6 +119,26 @@ record IsTotalPreorder (_≲_ : Rel A ℓ₂) : Set (a ⊔ ℓ ⊔ ℓ₂) where
119119
open IsPreorder isPreorder public
120120

121121

122+
record IsDecPreorder (_≲_ : Rel A ℓ₂) : Set (a ⊔ ℓ ⊔ ℓ₂) where
123+
field
124+
isPreorder : IsPreorder _≲_
125+
_≟_ : Decidable _≈_
126+
_≲?_ : Decidable _≲_
127+
128+
open IsPreorder isPreorder public
129+
hiding (module Eq)
130+
131+
module Eq where
132+
133+
isDecEquivalence : IsDecEquivalence
134+
isDecEquivalence = record
135+
{ isEquivalence = isEquivalence
136+
; _≟_ = _≟_
137+
}
138+
139+
open IsDecEquivalence isDecEquivalence public
140+
141+
122142
------------------------------------------------------------------------
123143
-- Partial orders
124144
------------------------------------------------------------------------
@@ -146,15 +166,15 @@ record IsDecPartialOrder (_≤_ : Rel A ℓ₂) : Set (a ⊔ ℓ ⊔ ℓ₂) whe
146166
open IsPartialOrder isPartialOrder public
147167
hiding (module Eq)
148168

149-
module Eq where
150-
151-
isDecEquivalence : IsDecEquivalence
152-
isDecEquivalence = record
153-
{ isEquivalence = isEquivalence
154-
; _≟_ = _≟_
155-
}
169+
isDecPreorder : IsDecPreorder _≤_
170+
isDecPreorder = record
171+
{ isPreorder = isPreorder
172+
; _≟_ = _≟_
173+
; _≲?_ = _≤?_
174+
}
156175

157-
open IsDecEquivalence isDecEquivalence public
176+
open IsDecPreorder isDecPreorder public
177+
using (module Eq)
158178

159179

160180
record IsStrictPartialOrder (_<_ : Rel A ℓ₂) : Set (a ⊔ ℓ ⊔ ℓ₂) where
@@ -234,16 +254,8 @@ record IsDecTotalOrder (_≤_ : Rel A ℓ₂) : Set (a ⊔ ℓ ⊔ ℓ₂) where
234254
; _≤?_ = _≤?_
235255
}
236256

237-
module Eq where
238-
239-
isDecEquivalence : IsDecEquivalence
240-
isDecEquivalence = record
241-
{ isEquivalence = isEquivalence
242-
; _≟_ = _≟_
243-
}
244-
245-
open IsDecEquivalence isDecEquivalence public
246-
257+
open IsDecPartialOrder isDecPartialOrder public
258+
using (isDecPreorder; module Eq)
247259

248260
-- Note that these orders are decidable. The current implementation
249261
-- of `Trichotomous` subsumes irreflexivity and asymmetry. See

0 commit comments

Comments
 (0)