Skip to content

Commit f2dc172

Browse files
Saransh-cppandreasabel
authored andcommitted
Rename Relation.Binary.Construct.(Flip/Converse) (#1979)
1 parent 8dd8c59 commit f2dc172

File tree

13 files changed

+246
-204
lines changed

13 files changed

+246
-204
lines changed

CHANGELOG.md

Lines changed: 8 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -918,6 +918,14 @@ Major improvements
918918
Deprecated modules
919919
------------------
920920
921+
### Moving `Relation.Binary.Construct.(Converse/Flip)`
922+
923+
* The following files have been moved:
924+
```agda
925+
Relation.Binary.Construct.Converse ↦ Relation.Binary.Construct.Flip.EqAndOrd
926+
Relation.Binary.Construct.Flip ↦ Relation.Binary.Construct.Flip.Ord
927+
```
928+
921929
### Deprecation of old function hierarchy
922930

923931
* The module `Function.Related` has been deprecated in favour of `Function.Related.Propositional`

src/Algebra/Construct/NaturalChoice/Base.agda

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -11,7 +11,7 @@ open import Algebra.Core
1111
open import Level as L hiding (_⊔_)
1212
open import Function.Base using (flip)
1313
open import Relation.Binary
14-
open import Relation.Binary.Construct.Converse using ()
14+
open import Relation.Binary.Construct.Flip.EqAndOrd using ()
1515
renaming (totalPreorder to flipOrder)
1616
import Relation.Binary.Properties.TotalOrder as TotalOrderProperties
1717

src/Algebra/Construct/NaturalChoice/Max.agda

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -14,7 +14,7 @@ module Algebra.Construct.NaturalChoice.Max
1414
open import Algebra.Core
1515
open import Algebra.Definitions
1616
open import Algebra.Construct.NaturalChoice.Base
17-
open import Relation.Binary.Construct.Converse using ()
17+
open import Relation.Binary.Construct.Flip.EqAndOrd using ()
1818
renaming (totalOrder to flip)
1919

2020
open TotalOrder totalOrder renaming (Carrier to A)

src/Algebra/Construct/NaturalChoice/MaxOp.agda

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -12,7 +12,7 @@ open import Algebra.Construct.NaturalChoice.Base
1212
import Algebra.Construct.NaturalChoice.MinOp as MinOp
1313
open import Function.Base using (flip)
1414
open import Relation.Binary
15-
open import Relation.Binary.Construct.Converse using ()
15+
open import Relation.Binary.Construct.Flip.EqAndOrd using ()
1616
renaming (totalPreorder to flipOrder)
1717

1818
module Algebra.Construct.NaturalChoice.MaxOp

src/Data/Fin/Induction.agda

Lines changed: 4 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -23,8 +23,8 @@ open import Induction.WellFounded as WF
2323
open import Level using (Level)
2424
open import Relation.Binary
2525
using (Rel; Decidable; IsPartialOrder; IsStrictPartialOrder; StrictPartialOrder)
26-
import Relation.Binary.Construct.Converse as Converse
27-
import Relation.Binary.Construct.Flip as Flip
26+
import Relation.Binary.Construct.Flip.EqAndOrd as EqAndOrd
27+
import Relation.Binary.Construct.Flip.Ord as Ord
2828
import Relation.Binary.Construct.NonStrictToStrict as ToStrict
2929
import Relation.Binary.Construct.On as On
3030
open import Relation.Binary.Definitions using (Tri; tri<; tri≈; tri>)
@@ -123,7 +123,7 @@ module _ {_≈_ : Rel (Fin n) ℓ} where
123123
pigeon : (xs : Vec (Fin n) n) Linked (flip _⊏_) (i ∷ xs) WellFounded _⊏_
124124
pigeon xs i∷xs↑ =
125125
let (i₁ , i₂ , i₁<i₂ , xs[i₁]≡xs[i₂]) = pigeonhole (n<1+n n) (Vec.lookup (i ∷ xs)) in
126-
let xs[i₁]⊏xs[i₂] = Linkedₚ.lookup⁺ (Flip.transitive _⊏_ ⊏.trans) {xs = i ∷ xs} i∷xs↑ i₁<i₂ in
126+
let xs[i₁]⊏xs[i₂] = Linkedₚ.lookup⁺ (Ord.transitive _⊏_ ⊏.trans) {xs = i ∷ xs} i∷xs↑ i₁<i₂ in
127127
let xs[i₁]⊏xs[i₁] = ⊏.<-respʳ-≈ (⊏.Eq.reflexive xs[i₁]≡xs[i₂]) xs[i₁]⊏xs[i₂] in
128128
contradiction xs[i₁]⊏xs[i₁] (⊏.irrefl ⊏.Eq.refl)
129129

@@ -137,7 +137,7 @@ module _ {_≈_ : Rel (Fin n) ℓ} where
137137

138138
spo-noetherian : {r} {_⊏_ : Rel (Fin n) r}
139139
IsStrictPartialOrder _≈_ _⊏_ WellFounded (flip _⊏_)
140-
spo-noetherian isSPO = spo-wellFounded (Converse.isStrictPartialOrder isSPO)
140+
spo-noetherian isSPO = spo-wellFounded (EqAndOrd.isStrictPartialOrder isSPO)
141141

142142
po-noetherian : {r} {_⊑_ : Rel (Fin n) r} IsPartialOrder _≈_ _⊑_
143143
WellFounded (flip (ToStrict._<_ _≈_ _⊑_))
Lines changed: 7 additions & 185 deletions
Original file line numberDiff line numberDiff line change
@@ -1,195 +1,17 @@
11
------------------------------------------------------------------------
22
-- The Agda standard library
33
--
4-
-- Many properties which hold for `∼` also hold for `flip ∼`. Unlike
5-
-- the module `Relation.Binary.Construct.Flip` this module does not
6-
-- flip the underlying equality.
4+
-- This module is DEPRECATED. Please use
5+
-- `Relation.Binary.Construct.Flip.EqAndOrd` instead.
76
------------------------------------------------------------------------
87

98
{-# OPTIONS --cubical-compatible --safe #-}
109

11-
open import Relation.Binary
12-
1310
module Relation.Binary.Construct.Converse where
1411

15-
open import Data.Product
16-
open import Function.Base using (flip; _∘_)
17-
open import Level using (Level)
18-
19-
private
20-
variable
21-
a b p ℓ ℓ₁ ℓ₂ : Level
22-
A B : Set a
23-
≈ ∼ ≤ < : Rel A ℓ
24-
25-
------------------------------------------------------------------------
26-
-- Properties
27-
28-
module _ (∼ : Rel A ℓ) where
29-
30-
refl : Reflexive ∼ Reflexive (flip ∼)
31-
refl refl = refl
32-
33-
sym : Symmetric ∼ Symmetric (flip ∼)
34-
sym sym = sym
35-
36-
trans : Transitive ∼ Transitive (flip ∼)
37-
trans trans = flip trans
38-
39-
asym : Asymmetric ∼ Asymmetric (flip ∼)
40-
asym asym = asym
41-
42-
total : Total ∼ Total (flip ∼)
43-
total total x y = total y x
44-
45-
resp : {p} (P : A Set p) Symmetric ∼
46-
P Respects ∼ P Respects (flip ∼)
47-
resp _ sym resp ∼ = resp (sym ∼)
48-
49-
max : {⊥} Minimum ∼ ⊥ Maximum (flip ∼) ⊥
50-
max min = min
51-
52-
min : {⊤} Maximum ∼ ⊤ Minimum (flip ∼) ⊤
53-
min max = max
54-
55-
module _ {≈ : Rel A ℓ₁} (∼ : Rel A ℓ₂) where
56-
57-
reflexive : Symmetric ≈ (≈ ⇒ ∼) (≈ ⇒ flip ∼)
58-
reflexive sym impl = impl ∘ sym
59-
60-
irrefl : Symmetric ≈ Irreflexive ≈ ∼ Irreflexive ≈ (flip ∼)
61-
irrefl sym irrefl x≈y y∼x = irrefl (sym x≈y) y∼x
62-
63-
antisym : Antisymmetric ≈ ∼ Antisymmetric ≈ (flip ∼)
64-
antisym antisym = flip antisym
65-
66-
compare : Trichotomous ≈ ∼ Trichotomous ≈ (flip ∼)
67-
compare cmp x y with cmp x y
68-
... | tri< x<y x≉y y≮x = tri> y≮x x≉y x<y
69-
... | tri≈ x≮y x≈y y≮x = tri≈ y≮x x≈y x≮y
70-
... | tri> x≮y x≉y y<x = tri< y<x x≉y x≮y
71-
72-
module _ (∼₁ : Rel A ℓ₁) (∼₂ : Rel A ℓ₂) where
73-
74-
resp₂ : ∼₁ Respects₂ ∼₂ (flip ∼₁) Respects₂ ∼₂
75-
resp₂ (resp₁ , resp₂) = resp₂ , resp₁
76-
77-
module _ (∼ : REL A B ℓ) where
78-
79-
dec : Decidable ∼ Decidable (flip ∼)
80-
dec dec = flip dec
81-
82-
------------------------------------------------------------------------
83-
-- Structures
84-
85-
isEquivalence : IsEquivalence ≈ IsEquivalence (flip ≈)
86-
isEquivalence {≈ = ≈} eq = record
87-
{ refl = refl ≈ Eq.refl
88-
; sym = sym ≈ Eq.sym
89-
; trans = trans ≈ Eq.trans
90-
} where module Eq = IsEquivalence eq
91-
92-
isDecEquivalence : IsDecEquivalence ≈ IsDecEquivalence (flip ≈)
93-
isDecEquivalence {≈ = ≈} eq = record
94-
{ isEquivalence = isEquivalence Dec.isEquivalence
95-
; _≟_ = dec ≈ Dec._≟_
96-
} where module Dec = IsDecEquivalence eq
97-
98-
isPreorder : IsPreorder ≈ ∼ IsPreorder ≈ (flip ∼)
99-
isPreorder {≈ = ≈} {∼ = ∼} O = record
100-
{ isEquivalence = O.isEquivalence
101-
; reflexive = reflexive ∼ O.Eq.sym O.reflexive
102-
; trans = trans ∼ O.trans
103-
} where module O = IsPreorder O
104-
105-
isTotalPreorder : IsTotalPreorder ≈ ∼ IsTotalPreorder ≈ (flip ∼)
106-
isTotalPreorder O = record
107-
{ isPreorder = isPreorder O.isPreorder
108-
; total = total _ O.total
109-
} where module O = IsTotalPreorder O
110-
111-
isPartialOrder : IsPartialOrder ≈ ≤ IsPartialOrder ≈ (flip ≤)
112-
isPartialOrder {≤ = ≤} O = record
113-
{ isPreorder = isPreorder O.isPreorder
114-
; antisym = antisym ≤ O.antisym
115-
} where module O = IsPartialOrder O
116-
117-
isTotalOrder : IsTotalOrder ≈ ≤ IsTotalOrder ≈ (flip ≤)
118-
isTotalOrder O = record
119-
{ isPartialOrder = isPartialOrder O.isPartialOrder
120-
; total = total _ O.total
121-
} where module O = IsTotalOrder O
122-
123-
isDecTotalOrder : IsDecTotalOrder ≈ ≤ IsDecTotalOrder ≈ (flip ≤)
124-
isDecTotalOrder O = record
125-
{ isTotalOrder = isTotalOrder O.isTotalOrder
126-
; _≟_ = O._≟_
127-
; _≤?_ = dec _ O._≤?_
128-
} where module O = IsDecTotalOrder O
129-
130-
isStrictPartialOrder : IsStrictPartialOrder ≈ <
131-
IsStrictPartialOrder ≈ (flip <)
132-
isStrictPartialOrder {< = <} O = record
133-
{ isEquivalence = O.isEquivalence
134-
; irrefl = irrefl < O.Eq.sym O.irrefl
135-
; trans = trans < O.trans
136-
; <-resp-≈ = resp₂ _ _ O.<-resp-≈
137-
} where module O = IsStrictPartialOrder O
138-
139-
isStrictTotalOrder : IsStrictTotalOrder ≈ <
140-
IsStrictTotalOrder ≈ (flip <)
141-
isStrictTotalOrder {< = <} O = record
142-
{ isEquivalence = O.isEquivalence
143-
; trans = trans < O.trans
144-
; compare = compare _ O.compare
145-
} where module O = IsStrictTotalOrder O
146-
147-
------------------------------------------------------------------------
148-
-- Bundles
149-
150-
setoid : Setoid a ℓ Setoid a ℓ
151-
setoid S = record
152-
{ isEquivalence = isEquivalence S.isEquivalence
153-
} where module S = Setoid S
154-
155-
decSetoid : DecSetoid a ℓ DecSetoid a ℓ
156-
decSetoid S = record
157-
{ isDecEquivalence = isDecEquivalence S.isDecEquivalence
158-
} where module S = DecSetoid S
159-
160-
preorder : Preorder a ℓ₁ ℓ₂ Preorder a ℓ₁ ℓ₂
161-
preorder O = record
162-
{ isPreorder = isPreorder O.isPreorder
163-
} where module O = Preorder O
164-
165-
totalPreorder : TotalPreorder a ℓ₁ ℓ₂ TotalPreorder a ℓ₁ ℓ₂
166-
totalPreorder O = record
167-
{ isTotalPreorder = isTotalPreorder O.isTotalPreorder
168-
} where module O = TotalPreorder O
169-
170-
poset : Poset a ℓ₁ ℓ₂ Poset a ℓ₁ ℓ₂
171-
poset O = record
172-
{ isPartialOrder = isPartialOrder O.isPartialOrder
173-
} where module O = Poset O
174-
175-
totalOrder : TotalOrder a ℓ₁ ℓ₂ TotalOrder a ℓ₁ ℓ₂
176-
totalOrder O = record
177-
{ isTotalOrder = isTotalOrder O.isTotalOrder
178-
} where module O = TotalOrder O
179-
180-
decTotalOrder : DecTotalOrder a ℓ₁ ℓ₂ DecTotalOrder a ℓ₁ ℓ₂
181-
decTotalOrder O = record
182-
{ isDecTotalOrder = isDecTotalOrder O.isDecTotalOrder
183-
} where module O = DecTotalOrder O
184-
185-
strictPartialOrder : StrictPartialOrder a ℓ₁ ℓ₂
186-
StrictPartialOrder a ℓ₁ ℓ₂
187-
strictPartialOrder O = record
188-
{ isStrictPartialOrder = isStrictPartialOrder O.isStrictPartialOrder
189-
} where module O = StrictPartialOrder O
12+
open import Relation.Binary.Construct.Flip.EqAndOrd public
19013

191-
strictTotalOrder : StrictTotalOrder a ℓ₁ ℓ₂
192-
StrictTotalOrder a ℓ₁ ℓ₂
193-
strictTotalOrder O = record
194-
{ isStrictTotalOrder = isStrictTotalOrder O.isStrictTotalOrder
195-
} where module O = StrictTotalOrder O
14+
{-# WARNING_ON_IMPORT
15+
"Relation.Binary.Construct.Converse was deprecated in v2.0.
16+
Use Relation.Binary.Construct.Flip.EqAndOrd instead."
17+
#-}
Lines changed: 17 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,17 @@
1+
------------------------------------------------------------------------
2+
-- The Agda standard library
3+
--
4+
-- This module is DEPRECATED. Please use
5+
-- `Relation.Binary.Construct.Flip.Ord` instead.
6+
------------------------------------------------------------------------
7+
8+
{-# OPTIONS --cubical-compatible --safe #-}
9+
10+
module Relation.Binary.Construct.Flip where
11+
12+
open import Relation.Binary.Construct.Flip.Ord public
13+
14+
{-# WARNING_ON_IMPORT
15+
"Relation.Binary.Construct.Flip was deprecated in v2.0.
16+
Use Relation.Binary.Construct.Flip.Ord instead."
17+
#-}

0 commit comments

Comments
 (0)