Skip to content

Commit 63cdfe0

Browse files
[ new ] symmetric core of a binary relation (#2071)
* [ new ] symmetric core of a binary relation * [ fix ] name clashes * [ more ] respond to McKinna's comments * [ rename ] fields to lhs≤rhs and rhs≤lhs * [ more ] incorporate parts of McKinna's review * [ minor ] remove implicit argument application from transitive * [ edit ] pull out isEquivalence and do some renaming * [ minor ] respond to easy comments * [ refactor ] remove IsProset, and move Proset to main hierarchy * Eliminate Proset * Fixed CHANGELOG typo --------- Co-authored-by: MatthewDaggitt <[email protected]>
1 parent 67bed00 commit 63cdfe0

File tree

6 files changed

+102
-5
lines changed

6 files changed

+102
-5
lines changed

CHANGELOG.md

Lines changed: 8 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -49,6 +49,11 @@ Deprecated names
4949
New modules
5050
-----------
5151

52+
* Symmetric interior of a binary relation
53+
```
54+
Relation.Binary.Construct.Interior.Symmetric
55+
```
56+
5257
* `Algebra.Module.Bundles.Raw`: raw bundles for module-like algebraic structures
5358

5459
* Nagata's construction of the "idealization of a module":
@@ -205,9 +210,10 @@ Additions to existing modules
205210
* In `Function.Bundles`, added `_⟶ₛ_` as a synonym for `Func` that can
206211
be used infix.
207212

208-
* Added new definitions in `Relation.Binary`
213+
* Added new definitions in `Relation.Binary.Definitions`
209214
```
210-
Stable : Pred A ℓ → Set _
215+
Stable _∼_ = ∀ x y → Nullary.Stable (x ∼ y)
216+
Empty _∼_ = ∀ {x y} → x ∼ y → ⊥
211217
```
212218

213219
* Added new proofs in `Relation.Binary.Properties.Setoid`:

src/Data/Fin/Subset/Properties.agda

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -35,7 +35,7 @@ open import Relation.Binary.Structures
3535
using (IsPreorder; IsPartialOrder; IsStrictPartialOrder; IsDecStrictPartialOrder)
3636
open import Relation.Binary.Bundles
3737
using (Preorder; Poset; StrictPartialOrder; DecStrictPartialOrder)
38-
open import Relation.Binary.Definitions as B hiding (Decidable)
38+
open import Relation.Binary.Definitions as B hiding (Decidable; Empty)
3939
open import Relation.Binary.PropositionalEquality
4040
open import Relation.Nullary.Decidable as Dec using (Dec; yes; no; _⊎-dec_)
4141
open import Relation.Nullary.Negation using (contradiction)
Lines changed: 83 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,83 @@
1+
------------------------------------------------------------------------
2+
-- The Agda standard library
3+
--
4+
-- Symmetric interior of a binary relation
5+
------------------------------------------------------------------------
6+
7+
{-# OPTIONS --cubical-compatible --safe #-}
8+
9+
module Relation.Binary.Construct.Interior.Symmetric where
10+
11+
open import Function.Base using (flip)
12+
open import Level
13+
open import Relation.Binary
14+
15+
private
16+
variable
17+
a b c ℓ r s t : Level
18+
A : Set a
19+
R S T : Rel A r
20+
21+
------------------------------------------------------------------------
22+
-- Definition
23+
24+
record SymInterior (R : Rel A ℓ) (x y : A) : Setwhere
25+
constructor _,_
26+
field
27+
lhs≤rhs : R x y
28+
rhs≤lhs : R y x
29+
30+
open SymInterior public
31+
32+
------------------------------------------------------------------------
33+
-- Properties
34+
35+
-- The symmetric interior is symmetric.
36+
symmetric : Symmetric (SymInterior R)
37+
symmetric (r , r′) = r′ , r
38+
39+
-- The symmetric interior of R is greater than (or equal to) any other symmetric
40+
-- relation contained by R.
41+
unfold : Symmetric S S ⇒ R S ⇒ SymInterior R
42+
unfold sym f s = f s , f (sym s)
43+
44+
-- SymInterior preserves various properties.
45+
reflexive : Reflexive R Reflexive (SymInterior R)
46+
reflexive refl = refl , refl
47+
48+
trans : Trans R S T Trans S R T
49+
Trans (SymInterior R) (SymInterior S) (SymInterior T)
50+
trans trans-rs trans-sr (r , r′) (s , s′) = trans-rs r s , trans-sr s′ r′
51+
52+
transitive : Transitive R Transitive (SymInterior R)
53+
transitive tr = trans tr tr
54+
55+
-- The symmetric interior of a strict relation is empty.
56+
asymmetric⇒empty : Asymmetric R Empty (SymInterior R)
57+
asymmetric⇒empty asym (r , r′) = asym r r′
58+
59+
-- A reflexive transitive relation _≤_ gives rise to a poset in which the
60+
-- equivalence relation is SymInterior _≤_.
61+
62+
isEquivalence : Reflexive R Transitive R IsEquivalence (SymInterior R)
63+
isEquivalence refl trans = record
64+
{ refl = reflexive refl
65+
; sym = symmetric
66+
; trans = transitive trans
67+
}
68+
69+
isPartialOrder : Reflexive R Transitive R IsPartialOrder (SymInterior R) R
70+
isPartialOrder refl trans = record
71+
{ isPreorder = record
72+
{ isEquivalence = isEquivalence refl trans
73+
; reflexive = lhs≤rhs
74+
; trans = trans
75+
}
76+
; antisym = _,_
77+
}
78+
79+
poset : {a} {A : Set a} {R : Rel A ℓ} Reflexive R Transitive R Poset a ℓ ℓ
80+
poset {R = R} refl trans = record
81+
{ _≤_ = R
82+
; isPartialOrder = isPartialOrder refl trans
83+
}

src/Relation/Binary/Definitions.agda

Lines changed: 7 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -12,6 +12,8 @@ module Relation.Binary.Definitions where
1212

1313
open import Agda.Builtin.Equality using (_≡_)
1414

15+
open import Data.Empty using (⊥)
16+
open import Data.Maybe.Base using (Maybe)
1517
open import Data.Product.Base using (_×_; ∃-syntax)
1618
open import Data.Sum.Base using (_⊎_)
1719
open import Function.Base using (_on_; flip)
@@ -243,6 +245,11 @@ DecidableEquality A = Decidable {A = A} _≡_
243245
Universal : REL A B ℓ Set _
244246
Universal _∼_ = x y x ∼ y
245247

248+
-- Empty - no elements are related
249+
250+
Empty : REL A B ℓ Set _
251+
Empty _∼_ = {x y} x ∼ y
252+
246253
-- Non-emptiness - at least one pair of elements are related.
247254

248255
record NonEmpty {A : Set a} {B : Set b}

src/Relation/Unary/Polymorphic/Properties.agda

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -10,7 +10,7 @@
1010
module Relation.Unary.Polymorphic.Properties where
1111

1212
open import Level using (Level)
13-
open import Relation.Binary.Definitions hiding (Decidable; Universal)
13+
open import Relation.Binary.Definitions hiding (Decidable; Universal; Empty)
1414
open import Relation.Nullary.Decidable using (yes; no)
1515
open import Relation.Unary hiding (∅; U)
1616
open import Relation.Unary.Polymorphic

src/Relation/Unary/Properties.agda

Lines changed: 2 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -13,7 +13,8 @@ open import Data.Sum.Base using (inj₁; inj₂)
1313
open import Data.Unit.Base using (tt)
1414
open import Level using (Level)
1515
open import Relation.Binary.Core as Binary
16-
open import Relation.Binary.Definitions hiding (Decidable; Universal; Irrelevant)
16+
open import Relation.Binary.Definitions
17+
hiding (Decidable; Universal; Irrelevant; Empty)
1718
open import Relation.Binary.PropositionalEquality.Core using (refl)
1819
open import Relation.Unary
1920
open import Relation.Nullary.Decidable using (yes; no; _⊎-dec_; _×-dec_; ¬?)

0 commit comments

Comments
 (0)