Skip to content

Commit aa6ef23

Browse files
MatthewDaggittjamesmckinna
authored andcommitted
Change definition of IsStrictTotalOrder (agda#2137)
1 parent a843d0d commit aa6ef23

File tree

25 files changed

+164
-96
lines changed

25 files changed

+164
-96
lines changed

CHANGELOG.md

+40
Original file line numberDiff line numberDiff line change
@@ -918,6 +918,46 @@ Non-backwards compatible changes
918918

919919
* The old names (and the names of all proofs about these functions) have been deprecated appropriately.
920920

921+
### Changes to definition of `IsStrictTotalOrder`
922+
923+
* The previous definition of the record `IsStrictTotalOrder` did not build upon `IsStrictPartialOrder`
924+
as would be expected. Instead it omitted several fields like irreflexivity as they were derivable from the
925+
proof of trichotomy. However, this led to problems further up the hierarchy where bundles such as `StrictTotalOrder`
926+
which contained multiple distinct proofs of `IsStrictPartialOrder`.
927+
928+
* To remedy this the definition of `IsStrictTotalOrder` has been changed to so that it builds upon
929+
`IsStrictPartialOrder` as would be expected.
930+
931+
* To aid migration, the old record definition has been moved to `Relation.Binary.Structures.Biased`
932+
which contains the `isStrictTotalOrderᶜ` smart constructor (which is re-exported by `Relation.Binary`) .
933+
Therefore the old code:
934+
```agda
935+
<-isStrictTotalOrder : IsStrictTotalOrder _≡_ _<_
936+
<-isStrictTotalOrder = record
937+
{ isEquivalence = isEquivalence
938+
; trans = <-trans
939+
; compare = <-cmp
940+
}
941+
```
942+
can be migrated either by updating to the new record fields if you have a proof of `IsStrictPartialOrder`
943+
available:
944+
```agda
945+
<-isStrictTotalOrder : IsStrictTotalOrder _≡_ _<_
946+
<-isStrictTotalOrder = record
947+
{ isStrictPartialOrder = <-isStrictPartialOrder
948+
; compare = <-cmp
949+
}
950+
```
951+
or simply applying the smart constructor to the record definition as follows:
952+
```agda
953+
<-isStrictTotalOrder : IsStrictTotalOrder _≡_ _<_
954+
<-isStrictTotalOrder = isStrictTotalOrderᶜ record
955+
{ isEquivalence = isEquivalence
956+
; trans = <-trans
957+
; compare = <-cmp
958+
}
959+
```
960+
921961
### Changes to triple reasoning interface
922962

923963
* The module `Relation.Binary.Reasoning.Base.Triple` now takes an extra proof that the strict

src/Data/Bool/Properties.agda

+2-3
Original file line numberDiff line numberDiff line change
@@ -210,9 +210,8 @@ true <? _ = no (λ())
210210

211211
<-isStrictTotalOrder : IsStrictTotalOrder _≡_ _<_
212212
<-isStrictTotalOrder = record
213-
{ isEquivalence = isEquivalence
214-
; trans = <-trans
215-
; compare = <-cmp
213+
{ isStrictPartialOrder = <-isStrictPartialOrder
214+
; compare = <-cmp
216215
}
217216

218217
-- Bundles

src/Data/Char/Properties.agda

+2-3
Original file line numberDiff line numberDiff line change
@@ -123,9 +123,8 @@ _<?_ = On.decidable toℕ ℕ._<_ ℕₚ._<?_
123123

124124
<-isStrictTotalOrder : IsStrictTotalOrder _≡_ _<_
125125
<-isStrictTotalOrder = record
126-
{ isEquivalence = PropEq.isEquivalence
127-
; trans = λ {a} {b} {c} <-trans {a} {b} {c}
128-
; compare = <-cmp
126+
{ isStrictPartialOrder = <-isStrictPartialOrder
127+
; compare = <-cmp
129128
}
130129

131130
<-strictPartialOrder : StrictPartialOrder _ _ _

src/Data/Fin/Properties.agda

+2-3
Original file line numberDiff line numberDiff line change
@@ -406,9 +406,8 @@ m <? n = suc (toℕ m) ℕₚ.≤? toℕ n
406406

407407
<-isStrictTotalOrder : IsStrictTotalOrder {A = Fin n} _≡_ _<_
408408
<-isStrictTotalOrder = record
409-
{ isEquivalence = P.isEquivalence
410-
; trans = <-trans
411-
; compare = <-cmp
409+
{ isStrictPartialOrder = <-isStrictPartialOrder
410+
; compare = <-cmp
412411
}
413412

414413
------------------------------------------------------------------------

src/Data/Integer/Properties.agda

+2-3
Original file line numberDiff line numberDiff line change
@@ -331,9 +331,8 @@ _<?_ : Decidable _<_
331331

332332
<-isStrictTotalOrder : IsStrictTotalOrder _≡_ _<_
333333
<-isStrictTotalOrder = record
334-
{ isEquivalence = isEquivalence
335-
; trans = <-trans
336-
; compare = <-cmp
334+
{ isStrictPartialOrder = <-isStrictPartialOrder
335+
; compare = <-cmp
337336
}
338337

339338
------------------------------------------------------------------------

src/Data/List/Relation/Binary/Lex/Strict.agda

+2-3
Original file line numberDiff line numberDiff line change
@@ -121,9 +121,8 @@ module _ {a ℓ₁ ℓ₂} {A : Set a} where
121121
<-isStrictTotalOrder : IsStrictTotalOrder _≈_ _≺_
122122
IsStrictTotalOrder _≋_ _<_
123123
<-isStrictTotalOrder sto = record
124-
{ isEquivalence = Pointwise.isEquivalence isEquivalence
125-
; trans = <-transitive isEquivalence <-resp-≈ trans
126-
; compare = <-compare Eq.sym compare
124+
{ isStrictPartialOrder = <-isStrictPartialOrder isStrictPartialOrder
125+
; compare = <-compare Eq.sym compare
127126
} where open IsStrictTotalOrder sto
128127

129128
<-strictPartialOrder : {a ℓ₁ ℓ₂} StrictPartialOrder a ℓ₁ ℓ₂

src/Data/Nat/Binary/Properties.agda

+2-3
Original file line numberDiff line numberDiff line change
@@ -414,9 +414,8 @@ _<?_ = tri⇒dec< <-cmp
414414

415415
<-isStrictTotalOrder : IsStrictTotalOrder _≡_ _<_
416416
<-isStrictTotalOrder = record
417-
{ isEquivalence = isEquivalence
418-
; trans = <-trans
419-
; compare = <-cmp
417+
{ isStrictPartialOrder = <-isStrictPartialOrder
418+
; compare = <-cmp
420419
}
421420

422421
------------------------------------------------------------------------

src/Data/Nat/Properties.agda

+2-7
Original file line numberDiff line numberDiff line change
@@ -32,12 +32,7 @@ open import Level using (0ℓ)
3232
open import Relation.Unary as U using (Pred)
3333
open import Relation.Binary.Core
3434
using (_⇒_; _Preserves_⟶_; _Preserves₂_⟶_⟶_)
35-
open import Relation.Binary.Bundles
36-
using (DecSetoid; Preorder; TotalPreorder; Poset; TotalOrder; DecTotalOrder; StrictPartialOrder; StrictTotalOrder)
37-
open import Relation.Binary.Structures
38-
using (IsDecEquivalence; IsPreorder; IsTotalPreorder; IsPartialOrder; IsTotalOrder; IsDecTotalOrder; IsStrictPartialOrder; IsStrictTotalOrder)
39-
open import Relation.Binary.Definitions
40-
using (DecidableEquality; Irrelevant; Reflexive; Antisymmetric; Transitive; Total; Decidable; Connex; Irreflexive; Asymmetric; LeftTrans; RightTrans; Trichotomous; tri≈; tri>; tri<; _Respects₂_)
35+
open import Relation.Binary
4136
open import Relation.Binary.Consequences using (flip-Connex)
4237
open import Relation.Binary.PropositionalEquality
4338
open import Relation.Nullary hiding (Irrelevant)
@@ -402,7 +397,7 @@ _>?_ = flip _<?_
402397
}
403398

404399
<-isStrictTotalOrder : IsStrictTotalOrder _≡_ _<_
405-
<-isStrictTotalOrder = record
400+
<-isStrictTotalOrder = isStrictTotalOrderᶜ record
406401
{ isEquivalence = isEquivalence
407402
; trans = <-trans
408403
; compare = <-cmp

src/Data/Product/Relation/Binary/Lex/Strict.agda

+3-6
Original file line numberDiff line numberDiff line change
@@ -269,12 +269,9 @@ module _ {_≈₁_ : Rel A ℓ₁} {_<₁_ : Rel A ℓ₂}
269269
IsStrictTotalOrder _≋_ _<ₗₑₓ_
270270
×-isStrictTotalOrder spo₁ spo₂ =
271271
record
272-
{ isEquivalence = Pointwise.×-isEquivalence
273-
(isEquivalence spo₁) (isEquivalence spo₂)
274-
; trans = ×-transitive {_<₁_ = _<₁_} {_<₂_ = _<₂_}
275-
(isEquivalence spo₁)
276-
(<-resp-≈ spo₁) (trans spo₁)
277-
(trans spo₂)
272+
{ isStrictPartialOrder = ×-isStrictPartialOrder
273+
(isStrictPartialOrder spo₁)
274+
(isStrictPartialOrder spo₂)
278275
; compare = ×-compare (Eq.sym spo₁) (compare spo₁)
279276
(compare spo₂)
280277
}

src/Data/Rational/Properties.agda

+2-3
Original file line numberDiff line numberDiff line change
@@ -702,9 +702,8 @@ _>?_ = flip _<?_
702702

703703
<-isStrictTotalOrder : IsStrictTotalOrder _≡_ _<_
704704
<-isStrictTotalOrder = record
705-
{ isEquivalence = isEquivalence
706-
; trans = <-trans
707-
; compare = <-cmp
705+
{ isStrictPartialOrder = <-isStrictPartialOrder
706+
; compare = <-cmp
708707
}
709708

710709
<-isDenseLinearOrder : IsDenseLinearOrder _≡_ _<_

src/Data/Rational/Unnormalised/Properties.agda

+2-3
Original file line numberDiff line numberDiff line change
@@ -542,9 +542,8 @@ _>?_ = flip _<?_
542542

543543
<-isStrictTotalOrder : IsStrictTotalOrder _≃_ _<_
544544
<-isStrictTotalOrder = record
545-
{ isEquivalence = ≃-isEquivalence
546-
; trans = <-trans
547-
; compare = <-cmp
545+
{ isStrictPartialOrder = <-isStrictPartialOrder
546+
; compare = <-cmp
548547
}
549548

550549
<-isDenseLinearOrder : IsDenseLinearOrder _≃_ _<_

src/Data/Sum/Relation/Binary/LeftOrder.agda

+2-3
Original file line numberDiff line numberDiff line change
@@ -192,9 +192,8 @@ module _ {a₁ a₂} {A₁ : Set a₁} {A₂ : Set a₂}
192192
IsStrictTotalOrder ≈₂ ∼₂
193193
IsStrictTotalOrder (Pointwise ≈₁ ≈₂) (∼₁ ⊎-< ∼₂)
194194
⊎-<-isStrictTotalOrder sto₁ sto₂ = record
195-
{ isEquivalence = PW.⊎-isEquivalence (isEquivalence sto₁) (isEquivalence sto₂)
196-
; trans = ⊎-<-transitive (trans sto₁) (trans sto₂)
197-
; compare = ⊎-<-trichotomous (compare sto₁) (compare sto₂)
195+
{ isStrictPartialOrder = ⊎-<-isStrictPartialOrder (isStrictPartialOrder sto₁) (isStrictPartialOrder sto₂)
196+
; compare = ⊎-<-trichotomous (compare sto₁) (compare sto₂)
198197
}
199198
where open IsStrictTotalOrder
200199

src/Data/Vec/Relation/Binary/Lex/Strict.agda

+1-2
Original file line numberDiff line numberDiff line change
@@ -168,8 +168,7 @@ module _ {_≈_ : Rel A ℓ₁} {_≺_ : Rel A ℓ₂} where
168168
<-isStrictTotalOrder : IsStrictTotalOrder _≈_ _≺_
169169
{n} IsStrictTotalOrder (_≋_ {n} {n}) _<_
170170
<-isStrictTotalOrder ≺-isStrictTotalOrder {n} = record
171-
{ isEquivalence = Pointwise.isEquivalence O.isEquivalence n
172-
; trans = <-trans O.Eq.isPartialEquivalence O.<-resp-≈ O.trans
171+
{ isStrictPartialOrder = <-isStrictPartialOrder O.isStrictPartialOrder
173172
; compare = <-cmp O.Eq.sym O.compare
174173
} where module O = IsStrictTotalOrder ≺-isStrictTotalOrder
175174

src/Induction/WellFounded.agda

+2-2
Original file line numberDiff line numberDiff line change
@@ -14,7 +14,7 @@ open import Induction
1414
open import Level using (Level; _⊔_)
1515
open import Relation.Binary.Core using (Rel)
1616
open import Relation.Binary.Definitions
17-
using (Symmetric; Asymmetric; Irreflexive; _Respects₂_;
17+
using (Symmetric; Asymmetric; Irreflexive; _Respects₂_;
1818
_Respectsʳ_; _Respects_)
1919
open import Relation.Binary.PropositionalEquality.Core using (_≡_; refl)
2020
open import Relation.Binary.Consequences using (asym⇒irr)
@@ -130,7 +130,7 @@ module _ {_<_ : Rel A r} where
130130
wf⇒asym : WellFounded _<_ Asymmetric _<_
131131
wf⇒asym wf = acc⇒asym (wf _)
132132

133-
wf⇒irrefl : {_≈_ : Rel A ℓ} _<_ Respects₂ _≈_
133+
wf⇒irrefl : {_≈_ : Rel A ℓ} _<_ Respects₂ _≈_
134134
Symmetric _≈_ WellFounded _<_ Irreflexive _≈_ _<_
135135
wf⇒irrefl r s wf = asym⇒irr r s (wf⇒asym wf)
136136

src/Relation/Binary.agda

+1
Original file line numberDiff line numberDiff line change
@@ -14,4 +14,5 @@ module Relation.Binary where
1414
open import Relation.Binary.Core public
1515
open import Relation.Binary.Definitions public
1616
open import Relation.Binary.Structures public
17+
open import Relation.Binary.Structures.Biased public
1718
open import Relation.Binary.Bundles public

src/Relation/Binary/Bundles.agda

+1-1
Original file line numberDiff line numberDiff line change
@@ -100,7 +100,7 @@ record Preorder c ℓ₁ ℓ₂ : Set (suc (c ⊔ ℓ₁ ⊔ ℓ₂)) where
100100
_≳_ = flip _≲_
101101

102102
-- Deprecated.
103-
infix 4 _∼_
103+
infix 4 _∼_
104104
_∼_ = _≲_
105105
{-# WARNING_ON_USAGE _∼_
106106
"Warning: _∼_ was deprecated in v2.0.

src/Relation/Binary/Construct/Add/Infimum/Strict.agda

+4-6
Original file line numberDiff line numberDiff line change
@@ -153,9 +153,8 @@ module _ {e} {_≈_ : Rel A e} where
153153
<₋-isStrictTotalOrder-≡ : IsStrictTotalOrder _≡_ _<_
154154
IsStrictTotalOrder _≡_ _<₋_
155155
<₋-isStrictTotalOrder-≡ strictot = record
156-
{ isEquivalence = P.isEquivalence
157-
; trans = <₋-trans trans
158-
; compare = <₋-cmp-≡ compare
156+
{ isStrictPartialOrder = <₋-isStrictPartialOrder-≡ isStrictPartialOrder
157+
; compare = <₋-cmp-≡ compare
159158
} where open IsStrictTotalOrder strictot
160159

161160
------------------------------------------------------------------------
@@ -185,7 +184,6 @@ module _ {e} {_≈_ : Rel A e} where
185184
<₋-isStrictTotalOrder : IsStrictTotalOrder _≈_ _<_
186185
IsStrictTotalOrder _≈₋_ _<₋_
187186
<₋-isStrictTotalOrder strictot = record
188-
{ isEquivalence = ≈₋-isEquivalence isEquivalence
189-
; trans = <₋-trans trans
190-
; compare = <₋-cmp compare
187+
{ isStrictPartialOrder = <₋-isStrictPartialOrder isStrictPartialOrder
188+
; compare = <₋-cmp compare
191189
} where open IsStrictTotalOrder strictot

src/Relation/Binary/Construct/Add/Supremum/Strict.agda

+4-6
Original file line numberDiff line numberDiff line change
@@ -154,9 +154,8 @@ module _ {e} {_≈_ : Rel A e} where
154154
<⁺-isStrictTotalOrder-≡ : IsStrictTotalOrder _≡_ _<_
155155
IsStrictTotalOrder _≡_ _<⁺_
156156
<⁺-isStrictTotalOrder-≡ strictot = record
157-
{ isEquivalence = P.isEquivalence
158-
; trans = <⁺-trans trans
159-
; compare = <⁺-cmp-≡ compare
157+
{ isStrictPartialOrder = <⁺-isStrictPartialOrder-≡ isStrictPartialOrder
158+
; compare = <⁺-cmp-≡ compare
160159
} where open IsStrictTotalOrder strictot
161160

162161
------------------------------------------------------------------------
@@ -186,7 +185,6 @@ module _ {e} {_≈_ : Rel A e} where
186185
<⁺-isStrictTotalOrder : IsStrictTotalOrder _≈_ _<_
187186
IsStrictTotalOrder _≈⁺_ _<⁺_
188187
<⁺-isStrictTotalOrder strictot = record
189-
{ isEquivalence = ≈⁺-isEquivalence isEquivalence
190-
; trans = <⁺-trans trans
191-
; compare = <⁺-cmp compare
188+
{ isStrictPartialOrder = <⁺-isStrictPartialOrder isStrictPartialOrder
189+
; compare = <⁺-cmp compare
192190
} where open IsStrictTotalOrder strictot

src/Relation/Binary/Construct/Flip/EqAndOrd.agda

+2-3
Original file line numberDiff line numberDiff line change
@@ -145,9 +145,8 @@ isStrictPartialOrder {< = <} O = record
145145
isStrictTotalOrder : IsStrictTotalOrder ≈ <
146146
IsStrictTotalOrder ≈ (flip <)
147147
isStrictTotalOrder {< = <} O = record
148-
{ isEquivalence = O.isEquivalence
149-
; trans = trans < O.trans
150-
; compare = compare _ O.compare
148+
{ isStrictPartialOrder = isStrictPartialOrder O.isStrictPartialOrder
149+
; compare = compare _ O.compare
151150
} where module O = IsStrictTotalOrder O
152151

153152
------------------------------------------------------------------------

src/Relation/Binary/Construct/Flip/Ord.agda

+2-3
Original file line numberDiff line numberDiff line change
@@ -139,9 +139,8 @@ isStrictPartialOrder {≈ = ≈} {< = <} O = record
139139
isStrictTotalOrder : IsStrictTotalOrder ≈ <
140140
IsStrictTotalOrder (flip ≈) (flip <)
141141
isStrictTotalOrder {≈ = ≈} {< = <} O = record
142-
{ isEquivalence = isEquivalence O.isEquivalence
143-
; trans = transitive < O.trans
144-
; compare = trichotomous ≈ < O.compare
142+
{ isStrictPartialOrder = isStrictPartialOrder O.isStrictPartialOrder
143+
; compare = trichotomous ≈ < O.compare
145144
} where module O = IsStrictTotalOrder O
146145

147146
------------------------------------------------------------------------

src/Relation/Binary/Construct/NonStrictToStrict.agda

+2-3
Original file line numberDiff line numberDiff line change
@@ -138,9 +138,8 @@ x < y = x ≤ y × x ≉ y
138138
<-isStrictTotalOrder₁ : Decidable _≈_ IsTotalOrder _≈_ _≤_
139139
IsStrictTotalOrder _≈_ _<_
140140
<-isStrictTotalOrder₁ ≟ tot = record
141-
{ isEquivalence = isEquivalence
142-
; trans = <-trans isPartialOrder
143-
; compare = <-trichotomous Eq.sym ≟ antisym total
141+
{ isStrictPartialOrder = <-isStrictPartialOrder isPartialOrder
142+
; compare = <-trichotomous Eq.sym ≟ antisym total
144143
} where open IsTotalOrder tot
145144

146145
<-isStrictTotalOrder₂ : IsDecTotalOrder _≈_ _≤_

src/Relation/Binary/Construct/On.agda

+2-3
Original file line numberDiff line numberDiff line change
@@ -150,9 +150,8 @@ module _ (f : B → A) {≈ : Rel A ℓ₁} {∼ : Rel A ℓ₂} where
150150
isStrictTotalOrder : IsStrictTotalOrder ≈ ∼
151151
IsStrictTotalOrder (≈ on f) (∼ on f)
152152
isStrictTotalOrder sto = record
153-
{ isEquivalence = isEquivalence f Sto.isEquivalence
154-
; trans = transitive f ∼ Sto.trans
155-
; compare = trichotomous f ≈ ∼ Sto.compare
153+
{ isStrictPartialOrder = isStrictPartialOrder Sto.isStrictPartialOrder
154+
; compare = trichotomous _ _ _ Sto.compare
156155
} where module Sto = IsStrictTotalOrder sto
157156

158157
------------------------------------------------------------------------

src/Relation/Binary/Morphism/OrderMonomorphism.agda

+2-3
Original file line numberDiff line numberDiff line change
@@ -105,7 +105,6 @@ isStrictPartialOrder O = record
105105
isStrictTotalOrder : IsStrictTotalOrder _≈₂_ _≲₂_
106106
IsStrictTotalOrder _≈₁_ _≲₁_
107107
isStrictTotalOrder O = record
108-
{ isEquivalence = EqM.isEquivalence O.isEquivalence
109-
; trans = trans O.trans
110-
; compare = compare O.compare
108+
{ isStrictPartialOrder = isStrictPartialOrder O.isStrictPartialOrder
109+
; compare = compare O.compare
111110
} where module O = IsStrictTotalOrder O

0 commit comments

Comments
 (0)