Documentation

Mathlib.Data.Rel.Separated

Uniform separation #

This file defines a notion of separation of a set relative to an relation.

For a relation R, a R-separated set s is a set such that every pair of elements of s is R-unrelated.

The concept of uniformly separated sets is used to define two further notions of separation:

TODO #

def SetRel.IsSeparated {X : Type u_1} (R : SetRel X X) (s : Set X) :

Given a relation R, a set s is R-separated if its elements are pairwise R-unrelated from each other.

Equations
Instances For
    theorem SetRel.IsSeparated.singleton {X : Type u_1} {R : SetRel X X} {x : X} :
    @[simp]
    theorem SetRel.IsSeparated.of_subsingleton {X : Type u_1} {R : SetRel X X} {s : Set X} (hs : s.Subsingleton) :
    theorem SetRel.IsSeparated.mono_left {X : Type u_1} {R S : SetRel X X} {s : Set X} (hUV : R S) (hs : S.IsSeparated s) :
    theorem SetRel.IsSeparated.mono_right {X : Type u_1} {R : SetRel X X} {s t : Set X} (hst : s t) (ht : R.IsSeparated t) :
    theorem SetRel.isSeparated_insert' {X : Type u_1} {R : SetRel X X} {s : Set X} {x : X} :
    R.IsSeparated (insert x s) R.IsSeparated s (∀ ys, (x, y) Rx = y) ys, (y, x) Rx = y
    theorem SetRel.isSeparated_insert {X : Type u_1} {R : SetRel X X} {s : Set X} {x : X} [R.IsSymm] :
    R.IsSeparated (insert x s) R.IsSeparated s ys, (x, y) Rx = y
    theorem SetRel.isSeparated_insert_of_notMem {X : Type u_1} {R : SetRel X X} {s : Set X} {x : X} [R.IsSymm] (hx : xs) :
    R.IsSeparated (insert x s) R.IsSeparated s ys, (x, y)R
    theorem SetRel.IsSeparated.insert {X : Type u_1} {R : SetRel X X} {s : Set X} {x : X} [R.IsSymm] (hs : R.IsSeparated s) (h : ys, (x, y) Rx = y) :