Documentation

Mathlib.Logic.Relation

Relation closures #

This file defines the reflexive, transitive, reflexive transitive and equivalence closures of relations and proves some basic results on them.

Note that this is about unbundled relations, that is terms of types of the form α → β → Prop. For the bundled version, see Rel.

Definitions #

theorem Std.Refl.reflexive {α : Sort u_1} {r : ααProp} [Refl r] :
@[deprecated Std.Refl.reflexive (since := "2026-01-09")]
theorem IsRefl.reflexive {α : Sort u_1} {r : ααProp} [Std.Refl r] :

Alias of Std.Refl.reflexive.

theorem Reflexive.rel_of_ne_imp {α : Sort u_1} {r : ααProp} (h : Reflexive r) {x y : α} (hr : x yr x y) :
r x y

To show a reflexive relation r : α → α → Prop holds over x y : α, it suffices to show it holds when x ≠ y.

theorem Reflexive.ne_imp_iff {α : Sort u_1} {r : ααProp} (h : Reflexive r) {x y : α} :
x yr x y r x y

If a reflexive relation r : α → α → Prop holds over x y : α, then it holds whether or not x ≠ y.

theorem reflexive_ne_imp_iff {α : Sort u_1} {r : ααProp} [Std.Refl r] {x y : α} :
x yr x y r x y

If a reflexive relation r : α → α → Prop holds over x y : α, then it holds whether or not x ≠ y. Unlike Reflexive.ne_imp_iff, this uses [Std.Refl r].

theorem reflexive_iff_subrelation_eq {α : Sort u_1} {r : ααProp} :
theorem irreflexive_iff_subrelation_ne {α : Sort u_1} {r : ααProp} :
theorem Symmetric.iff {α : Sort u_1} {r : ααProp} (H : Symmetric r) (x y : α) :
r x y r y x
theorem Symmetric.flip_eq {α : Sort u_1} {r : ααProp} (h : Symmetric r) :
flip r = r
theorem Symmetric.swap_eq {α : Sort u_1} {r : ααProp} :
theorem flip_eq_iff {α : Sort u_1} {r : ααProp} :
theorem swap_eq_iff {α : Sort u_1} {r : ααProp} :
theorem Reflexive.comap {α : Sort u_1} {β : Sort u_2} {r : ββProp} (h : Reflexive r) (f : αβ) :
theorem Symmetric.comap {α : Sort u_1} {β : Sort u_2} {r : ββProp} (h : Symmetric r) (f : αβ) :
theorem Transitive.comap {α : Sort u_1} {β : Sort u_2} {r : ββProp} (h : Transitive r) (f : αβ) :
theorem Equivalence.comap {α : Sort u_1} {β : Sort u_2} {r : ββProp} (h : Equivalence r) (f : αβ) :
def Relation.Comp {α : Sort u_1} {β : Sort u_2} {γ : Sort u_3} (r : αβProp) (p : βγProp) (a : α) (c : γ) :

The composition of two relations, yielding a new relation. The result relates a term of α and a term of γ if there is an intermediate term of β related to both.

Equations
Instances For
    @[simp]
    theorem Relation.comp_eq_fun {α : Sort u_1} {β : Sort u_2} {γ : Sort u_3} {r : αβProp} (f : γβ) :
    (Comp r fun (x1 : β) (x2 : γ) => x1 = f x2) = fun (x1 : α) (x2 : γ) => r x1 (f x2)
    @[simp]
    theorem Relation.comp_eq {α : Sort u_1} {β : Sort u_2} {r : αβProp} :
    (Comp r fun (x1 x2 : β) => x1 = x2) = r
    @[simp]
    theorem Relation.fun_eq_comp {α : Sort u_1} {β : Sort u_2} {γ : Sort u_3} {r : αβProp} (f : γα) :
    Comp (fun (x1 : γ) (x2 : α) => f x1 = x2) r = fun (x : γ) => r (f x)
    @[simp]
    theorem Relation.eq_comp {α : Sort u_1} {β : Sort u_2} {r : αβProp} :
    Comp (fun (x1 x2 : α) => x1 = x2) r = r
    @[simp]
    theorem Relation.iff_comp {α : Sort u_1} {r : PropαProp} :
    Comp (fun (x1 x2 : Prop) => x1 x2) r = r
    @[simp]
    theorem Relation.comp_iff {α : Sort u_1} {r : αPropProp} :
    (Comp r fun (x1 x2 : Prop) => x1 x2) = r
    theorem Relation.comp_assoc {α : Sort u_1} {β : Sort u_2} {γ : Sort u_3} {δ : Sort u_4} {r : αβProp} {p : βγProp} {q : γδProp} :
    Comp (Comp r p) q = Comp r (Comp p q)
    theorem Relation.flip_comp {α : Sort u_1} {β : Sort u_2} {γ : Sort u_3} {r : αβProp} {p : βγProp} :
    flip (Comp r p) = Comp (flip p) (flip r)
    def Relation.Fibration {α : Sort u_1} {β : Sort u_2} ( : ααProp) ( : ββProp) (f : αβ) :

    A function f : α → β is a fibration between the relation and if for all a : α and b : β, whenever b : β and f a are related by , b is the image of some a' : α under f, and a' and a are related by .

    Equations
    Instances For
      theorem Acc.of_fibration {α : Sort u_1} {β : Sort u_2} { : ααProp} { : ββProp} (f : αβ) (fib : Relation.Fibration f) {a : α} (ha : Acc a) :
      Acc (f a)

      If f : α → β is a fibration between relations and , and a : α is accessible under , then f a is accessible under .

      theorem Acc.of_downward_closed {α : Sort u_1} {β : Sort u_2} { : ββProp} (f : αβ) (dc : ∀ {a : α} {b : β}, b (f a) (c : α), f c = b) (a : α) (ha : Acc (InvImage f) a) :
      Acc (f a)
      def Relation.Map {α : Sort u_1} {β : Sort u_2} {γ : Sort u_3} {δ : Sort u_4} (r : αβProp) (f : αγ) (g : βδ) :
      γδProp

      The map of a relation r through a pair of functions pushes the relation to the codomains of the functions. The resulting relation is defined by having pairs of terms related if they have preimages related by r.

      Equations
      Instances For
        theorem Relation.map_apply {α : Sort u_1} {β : Sort u_2} {γ : Sort u_3} {δ : Sort u_4} {r : αβProp} {f : αγ} {g : βδ} {c : γ} {d : δ} :
        Relation.Map r f g c d (a : α), (b : β), r a b f a = c g b = d
        @[simp]
        theorem Relation.map_map {α : Sort u_1} {β : Sort u_2} {γ : Sort u_3} {δ : Sort u_4} {ε : Sort u_5} {ζ : Sort u_6} (r : αβProp) (f₁ : αγ) (g₁ : βδ) (f₂ : γε) (g₂ : δζ) :
        Relation.Map (Relation.Map r f₁ g₁) f₂ g₂ = Relation.Map r (f₂ f₁) (g₂ g₁)
        @[simp]
        theorem Relation.map_apply_apply {α : Sort u_1} {β : Sort u_2} {γ : Sort u_3} {δ : Sort u_4} {f : αγ} {g : βδ} (hf : Function.Injective f) (hg : Function.Injective g) (r : αβProp) (a : α) (b : β) :
        Relation.Map r f g (f a) (g b) r a b
        @[simp]
        theorem Relation.map_id_id {α : Sort u_1} {β : Sort u_2} (r : αβProp) :
        instance Relation.instDecidableMapOfExistsAndEq {α : Sort u_1} {β : Sort u_2} {γ : Sort u_3} {δ : Sort u_4} {r : αβProp} {f : αγ} {g : βδ} {c : γ} {d : δ} [Decidable ( (a : α), (b : β), r a b f a = c g b = d)] :
        Equations
        theorem Relation.map_reflexive {α : Sort u_1} {β : Sort u_2} {r : ααProp} (hr : Reflexive r) {f : αβ} (hf : Function.Surjective f) :
        theorem Relation.map_symmetric {α : Sort u_1} {β : Sort u_2} {r : ααProp} (hr : Symmetric r) (f : αβ) :
        theorem Relation.map_transitive {α : Sort u_1} {β : Sort u_2} {r : ααProp} (hr : Transitive r) {f : αβ} (hf : ∀ (x y : α), f x = f yr x y) :
        theorem Relation.map_equivalence {α : Sort u_1} {β : Sort u_2} {r : ααProp} (hr : Equivalence r) (f : αβ) (hf : Function.Surjective f) (hf_ker : ∀ (x y : α), f x = f yr x y) :
        theorem Relation.map_mono {α : Sort u_1} {β : Sort u_2} {γ : Sort u_3} {δ : Sort u_4} {r s : αβProp} {f : αγ} {g : βδ} (h : ∀ (x : α) (y : β), r x ys x y) (x : γ) (y : δ) :
        Relation.Map r f g x yRelation.Map s f g x y
        theorem Relation.le_onFun_map {α : Sort u_1} {β : Sort u_2} {r : ααProp} (f : αβ) :
        theorem Relation.onFun_map_eq_of_injective {α : Sort u_1} {β : Sort u_2} {r : ααProp} {f : αβ} (hinj : Function.Injective f) :
        theorem Relation.map_onFun_le {α : Sort u_1} {β : Sort u_2} {r : ββProp} (f : αβ) :
        theorem Relation.map_onFun_eq_of_surjective {α : Sort u_1} {β : Sort u_2} {r : ββProp} {f : αβ} (hsurj : Function.Surjective f) :
        theorem Relation.map_onFun_map_eq_map {α : Sort u_1} {β : Sort u_2} {r : ααProp} (f : αβ) :
        theorem Relation.onFun_map_onFun_eq_onFun {α : Sort u_1} {β : Sort u_2} {r : ββProp} (f : αβ) :
        theorem Relation.onFun_map_onFun_iff_onFun {α : Sort u_1} {β : Sort u_2} {r : ββProp} (f : αβ) (a₁ a₂ : α) :
        Relation.Map (Function.onFun r f) f f (f a₁) (f a₂) r (f a₁) (f a₂)
        inductive Relation.ReflTransGen {α : Sort u_1} (r : ααProp) (a : α) :
        αProp

        ReflTransGen r: reflexive transitive closure of r

        Instances For
          theorem Relation.ReflTransGen.cases_tail_iff {α : Sort u_1} (r : ααProp) (a a✝ : α) :
          ReflTransGen r a a✝ a✝ = a (b : α), ReflTransGen r a b r b a✝
          inductive Relation.ReflGen {α : Sort u_1} (r : ααProp) (a : α) :
          αProp

          ReflGen r: reflexive closure of r

          • refl {α : Sort u_1} {r : ααProp} {a : α} : ReflGen r a a
          • single {α : Sort u_1} {r : ααProp} {a b : α} : r a bReflGen r a b
          Instances For
            theorem Relation.reflGen_iff {α : Sort u_1} (r : ααProp) (a a✝ : α) :
            ReflGen r a a✝ a✝ = a r a a✝
            inductive Relation.EqvGen {α : Sort u_1} (r : ααProp) :
            ααProp

            EqvGen r: equivalence closure of r.

            Instances For
              theorem Relation.eqvGen_iff {α : Sort u_1} (r : ααProp) (a✝ a✝¹ : α) :
              EqvGen r a✝ a✝¹ r a✝ a✝¹ a✝¹ = a✝ EqvGen r a✝¹ a✝ (y : α), EqvGen r a✝ y EqvGen r y a✝¹
              theorem Relation.transGen_iff {α : Sort u} (r : ααProp) (a✝ a✝¹ : α) :
              TransGen r a✝ a✝¹ r a✝ a✝¹ (b : α), TransGen r a✝ b r b a✝¹
              theorem Relation.ReflGen.to_reflTransGen {α : Sort u_1} {r : ααProp} {a b : α} :
              ReflGen r a bReflTransGen r a b
              theorem Relation.ReflGen.mono {α : Sort u_1} {r p : ααProp} (hp : ∀ (a b : α), r a bp a b) {a b : α} :
              ReflGen r a bReflGen p a b
              instance Relation.ReflGen.instRefl {α : Sort u_1} {r : ααProp} :
              theorem Relation.ReflTransGen.trans {α : Sort u_1} {r : ααProp} {a b c : α} (hab : ReflTransGen r a b) (hbc : ReflTransGen r b c) :
              theorem Relation.ReflTransGen.single {α : Sort u_1} {r : ααProp} {a b : α} (hab : r a b) :
              theorem Relation.ReflTransGen.head {α : Sort u_1} {r : ααProp} {a b c : α} (hab : r a b) (hbc : ReflTransGen r b c) :
              theorem Relation.ReflTransGen.symmetric {α : Sort u_1} {r : ααProp} (h : Symmetric r) :
              theorem Relation.ReflTransGen.cases_tail {α : Sort u_1} {r : ααProp} {a b : α} :
              ReflTransGen r a bb = a (c : α), ReflTransGen r a c r c b
              theorem Relation.ReflTransGen.head_induction_on {α : Sort u_1} {r : ααProp} {b : α} {motive : (a : α) → ReflTransGen r a bProp} {a : α} (h : ReflTransGen r a b) (refl : motive b ) (head : ∀ {a c : α} (h' : r a c) (h : ReflTransGen r c b), motive c hmotive a ) :
              motive a h
              theorem Relation.ReflTransGen.trans_induction_on {α : Sort u_1} {r : ααProp} {motive : {a b : α} → ReflTransGen r a bProp} {a b : α} (h : ReflTransGen r a b) (refl : ∀ (a : α), motive ) (single : ∀ {a b : α} (h : r a b), motive ) (trans : ∀ {a b c : α} (h₁ : ReflTransGen r a b) (h₂ : ReflTransGen r b c), motive h₁motive h₂motive ) :
              motive h
              theorem Relation.ReflTransGen.cases_head {α : Sort u_1} {r : ααProp} {a b : α} (h : ReflTransGen r a b) :
              a = b (c : α), r a c ReflTransGen r c b
              theorem Relation.ReflTransGen.cases_head_iff {α : Sort u_1} {r : ααProp} {a b : α} :
              ReflTransGen r a b a = b (c : α), r a c ReflTransGen r c b
              theorem Relation.ReflTransGen.total_of_right_unique {α : Sort u_1} {r : ααProp} {a b c : α} (U : Relator.RightUnique r) (ab : ReflTransGen r a b) (ac : ReflTransGen r a c) :
              theorem Relation.TransGen.to_reflTransGen {α : Sort u_1} {r : ααProp} {a b : α} (h : TransGen r a b) :
              theorem Relation.TransGen.trans_left {α : Sort u_1} {r : ααProp} {a b c : α} (hab : TransGen r a b) (hbc : ReflTransGen r b c) :
              TransGen r a c
              theorem Relation.TransGen.head' {α : Sort u_1} {r : ααProp} {a b c : α} (hab : r a b) (hbc : ReflTransGen r b c) :
              TransGen r a c
              theorem Relation.TransGen.tail' {α : Sort u_1} {r : ααProp} {a b c : α} (hab : ReflTransGen r a b) (hbc : r b c) :
              TransGen r a c
              theorem Relation.TransGen.head {α : Sort u_1} {r : ααProp} {a b c : α} (hab : r a b) (hbc : TransGen r b c) :
              TransGen r a c
              theorem Relation.TransGen.head_induction_on {α : Sort u_1} {r : ααProp} {b : α} {motive : (a : α) → TransGen r a bProp} {a : α} (h : TransGen r a b) (single : ∀ {a : α} (h : r a b), motive a ) (head : ∀ {a c : α} (h' : r a c) (h : TransGen r c b), motive c hmotive a ) :
              motive a h
              theorem Relation.TransGen.trans_induction_on {α : Sort u_1} {r : ααProp} {motive : {a b : α} → TransGen r a bProp} {a b : α} (h : TransGen r a b) (single : ∀ {a b : α} (h : r a b), motive ) (trans : ∀ {a b c : α} (h₁ : TransGen r a b) (h₂ : TransGen r b c), motive h₁motive h₂motive ) :
              motive h
              theorem Relation.TransGen.trans_right {α : Sort u_1} {r : ααProp} {a b c : α} (hab : ReflTransGen r a b) (hbc : TransGen r b c) :
              TransGen r a c
              theorem Relation.TransGen.tail'_iff {α : Sort u_1} {r : ααProp} {a c : α} :
              TransGen r a c (b : α), ReflTransGen r a b r b c
              theorem Relation.TransGen.head'_iff {α : Sort u_1} {r : ααProp} {a c : α} :
              TransGen r a c (b : α), r a b ReflTransGen r b c
              theorem Relation.TransGen.symmetric {α : Sort u_1} {r : ααProp} (hr : Symmetric r) :
              theorem Relation.reflGen_eq_self {α : Sort u_1} {r : ααProp} (hr : Reflexive r) :
              theorem Relation.reflexive_reflGen {α : Sort u_1} {r : ααProp} :
              theorem Relation.reflGen_minimal {α : Sort u_1} {r r' : ααProp} (hr' : Reflexive r') (h : ∀ (x y : α), r x yr' x y) {x y : α} (hxy : ReflGen r x y) :
              r' x y
              instance Relation.instIsTransTransGen {α : Sort u_1} {r : ααProp} :
              instance Relation.instTransTransGen_mathlib {α : Sort u_1} {r : ααProp} :
              Equations
              instance Relation.instTransTransGen_mathlib_1 {α : Sort u_1} {r : ααProp} :
              Equations
              instance Relation.instTransTransGenReflTransGen {α : Sort u_1} {r : ααProp} :
              Equations
              instance Relation.instTransReflTransGenTransGen {α : Sort u_1} {r : ααProp} :
              Equations
              theorem Relation.transGen_eq_self {α : Sort u_1} {r : ααProp} (trans : Transitive r) :
              theorem Relation.transitive_transGen {α : Sort u_1} {r : ααProp} :
              theorem Relation.transGen_idem {α : Sort u_1} {r : ααProp} :
              theorem Relation.TransGen.lift {α : Sort u_1} {β : Sort u_2} {r : ααProp} {p : ββProp} {a b : α} (f : αβ) (h : ∀ (a b : α), r a bp (f a) (f b)) (hab : TransGen r a b) :
              TransGen p (f a) (f b)
              theorem Relation.TransGen.lift' {α : Sort u_1} {β : Sort u_2} {r : ααProp} {p : ββProp} {a b : α} (f : αβ) (h : ∀ (a b : α), r a bTransGen p (f a) (f b)) (hab : TransGen r a b) :
              TransGen p (f a) (f b)
              theorem Relation.TransGen.closed {α : Sort u_1} {r : ααProp} {a b : α} {p : ααProp} :
              (∀ (a b : α), r a bTransGen p a b)TransGen r a bTransGen p a b
              theorem Relation.TransGen.closed' {α : Sort u_1} {r : ααProp} {P : αProp} (dc : ∀ {a b : α}, r a bP bP a) {a b : α} (h : TransGen r a b) :
              P bP a
              theorem Relation.TransGen.mono {α : Sort u_1} {r : ααProp} {a b : α} {p : ααProp} :
              (∀ (a b : α), r a bp a b)TransGen r a bTransGen p a b
              theorem Relation.transGen_minimal {α : Sort u_1} {r r' : ααProp} (hr' : Transitive r') (h : ∀ (x y : α), r x yr' x y) {x y : α} (hxy : TransGen r x y) :
              r' x y
              theorem Relation.TransGen.swap {α : Sort u_1} {r : ααProp} {a b : α} (h : TransGen r b a) :
              theorem Relation.transGen_swap {α : Sort u_1} {r : ααProp} {a b : α} :
              theorem Relation.reflTransGen_iff_eq {α : Sort u_1} {r : ααProp} {a b : α} (h : ∀ (b : α), ¬r a b) :
              ReflTransGen r a b b = a
              theorem Relation.reflTransGen_iff_eq_or_transGen {α : Sort u_1} {r : ααProp} {a b : α} :
              ReflTransGen r a b b = a TransGen r a b
              theorem Relation.ReflTransGen.lift {α : Sort u_1} {β : Sort u_2} {r : ααProp} {p : ββProp} {a b : α} (f : αβ) (h : ∀ (a b : α), r a bp (f a) (f b)) (hab : ReflTransGen r a b) :
              ReflTransGen p (f a) (f b)
              theorem Relation.ReflTransGen.mono {α : Sort u_1} {r : ααProp} {a b : α} {p : ααProp} :
              (∀ (a b : α), r a bp a b)ReflTransGen r a bReflTransGen p a b
              theorem Relation.reflTransGen_eq_self {α : Sort u_1} {r : ααProp} (refl : Reflexive r) (trans : Transitive r) :
              theorem Relation.reflexive_reflTransGen {α : Sort u_1} {r : ααProp} :
              theorem Relation.transitive_reflTransGen {α : Sort u_1} {r : ααProp} :
              instance Relation.instTransReflTransGen {α : Sort u_1} {r : ααProp} :
              Equations
              instance Relation.instTransReflTransGen_1 {α : Sort u_1} {r : ααProp} :
              Equations
              instance Relation.instReflReflTransGen {α : Sort u_1} {r : ααProp} :
              instance Relation.instIsTransReflTransGen {α : Sort u_1} {r : ααProp} :
              theorem Relation.reflTransGen_idem {α : Sort u_1} {r : ααProp} :
              theorem Relation.ReflTransGen.lift' {α : Sort u_1} {β : Sort u_2} {r : ααProp} {p : ββProp} {a b : α} (f : αβ) (h : ∀ (a b : α), r a bReflTransGen p (f a) (f b)) (hab : ReflTransGen r a b) :
              ReflTransGen p (f a) (f b)
              theorem Relation.reflTransGen_closed {α : Sort u_1} {r : ααProp} {a b : α} {p : ααProp} :
              (∀ (a b : α), r a bReflTransGen p a b)ReflTransGen r a bReflTransGen p a b
              theorem Relation.ReflTransGen.swap {α : Sort u_1} {r : ααProp} {a b : α} (h : ReflTransGen r b a) :
              theorem Relation.reflTransGen_swap {α : Sort u_1} {r : ααProp} {a b : α} :
              @[simp]
              theorem Relation.reflGen_transGen {α : Sort u_1} {r : ααProp} :
              @[simp]
              theorem Relation.transGen_reflGen {α : Sort u_1} {r : ααProp} :
              @[simp]
              theorem Relation.reflTransGen_reflGen {α : Sort u_1} {r : ααProp} :
              @[simp]
              theorem Relation.reflTransGen_transGen {α : Sort u_1} {r : ααProp} :
              theorem Relation.reflTransGen_eq_transGen {α : Sort u_1} {r : ααProp} (hr : Reflexive r) :
              theorem Relation.reflTransGen_eq_reflGen {α : Sort u_1} {r : ααProp} (hr : Transitive r) :
              theorem Relation.EqvGen.is_equivalence {α : Sort u_1} (r : ααProp) :
              def Relation.EqvGen.setoid {α : Sort u_1} (r : ααProp) :

              EqvGen.setoid r is the setoid generated by a relation r.

              The motivation for this definition is that Quot r behaves like Quotient (EqvGen.setoid r), see for example Quot.eqvGen_exact and Quot.eqvGen_sound.

              Equations
              Instances For
                theorem Relation.EqvGen.mono {α : Sort u_1} {a b : α} {r p : ααProp} (hrp : ∀ (a b : α), r a bp a b) (h : EqvGen r a b) :
                EqvGen p a b
                def Relation.Join {α : Sort u_1} (r : ααProp) :
                ααProp

                The join of a relation on a single type is a new relation for which pairs of terms are related if there is a third term they are both related to. For example, if r is a relation representing rewrites in a term rewriting system, then confluence is the property that if a rewrites to both b and c, then join r relates b and c (see Relation.church_rosser).

                Equations
                Instances For
                  theorem Relation.church_rosser {α : Sort u_1} {r : ααProp} {a b c : α} (h : ∀ (a b c : α), r a br a c (d : α), ReflGen r b d ReflTransGen r c d) (hab : ReflTransGen r a b) (hac : ReflTransGen r a c) :

                  A sufficient condition for the Church-Rosser property.

                  theorem Relation.join_of_single {α : Sort u_1} {r : ααProp} {a b : α} (h : Reflexive r) (hab : r a b) :
                  Join r a b
                  theorem Relation.symmetric_join {α : Sort u_1} {r : ααProp} :
                  theorem Relation.reflexive_join {α : Sort u_1} {r : ααProp} (h : Reflexive r) :
                  theorem Relation.transitive_join {α : Sort u_1} {r : ααProp} (ht : Transitive r) (h : ∀ (a b c : α), r a br a cJoin r b c) :
                  theorem Relation.equivalence_join {α : Sort u_1} {r : ααProp} (hr : Reflexive r) (ht : Transitive r) (h : ∀ (a b c : α), r a br a cJoin r b c) :
                  theorem Relation.equivalence_join_reflTransGen {α : Sort u_1} {r : ααProp} (h : ∀ (a b c : α), r a br a c (d : α), ReflGen r b d ReflTransGen r c d) :
                  theorem Relation.join_of_equivalence {α : Sort u_1} {r : ααProp} {a b : α} {r' : ααProp} (hr : Equivalence r) (h : ∀ (a b : α), r' a br a b) :
                  Join r' a br a b
                  theorem Relation.reflTransGen_of_transitive_reflexive {α : Sort u_1} {r : ααProp} {a b : α} {r' : ααProp} (hr : Reflexive r) (ht : Transitive r) (h : ∀ (a b : α), r' a br a b) (h' : ReflTransGen r' a b) :
                  r a b
                  @[deprecated Relation.reflTransGen_of_transitive_reflexive (since := "2025-12-17")]
                  theorem Relation.reflTransGen_minimal {α : Sort u_1} {r : ααProp} {a b : α} {r' : ααProp} (hr : Reflexive r) (ht : Transitive r) (h : ∀ (a b : α), r' a br a b) (h' : ReflTransGen r' a b) :
                  r a b

                  Alias of Relation.reflTransGen_of_transitive_reflexive.

                  theorem Relation.reflTransGen_of_equivalence {α : Sort u_1} {r : ααProp} {a b : α} {r' : ααProp} (hr : Equivalence r) :
                  (∀ (a b : α), r' a br a b)ReflTransGen r' a br a b
                  theorem Quot.eqvGen_exact {α : Sort u_1} {r : ααProp} {a b : α} (H : mk r a = mk r b) :
                  theorem Quot.eqvGen_sound {α : Sort u_1} {r : ααProp} {a b : α} (H : Relation.EqvGen r a b) :
                  mk r a = mk r b
                  theorem Equivalence.eqvGen_iff {α : Sort u_1} {r : ααProp} {a b : α} (h : Equivalence r) :
                  Relation.EqvGen r a b r a b
                  theorem Equivalence.eqvGen_eq {α : Sort u_1} {r : ααProp} (h : Equivalence r) :