Documentation

Mathlib.Order.RelIso.Basic

Relation homomorphisms, embeddings, isomorphisms #

This file defines relation homomorphisms, embeddings, isomorphisms and order embeddings and isomorphisms.

Main declarations #

Notation #

structure RelHom {α : Type u_5} {β : Type u_6} (r : ααProp) (s : ββProp) :
Type (max u_5 u_6)

A relation homomorphism with respect to a given pair of relations r and s is a function f : α → β such that r a b → s (f a) (f b).

  • toFun : αβ

    The underlying function of a RelHom

  • map_rel' {a b : α} : r a bs (self.toFun a) (self.toFun b)

    A RelHom sends related elements to related elements

A relation homomorphism with respect to a given pair of relations r and s is a function f : α → β such that r a b → s (f a) (f b).

Equations
class RelHomClass (F : Type u_5) {α : outParam (Type u_6)} {β : outParam (Type u_7)} (r : outParam (ααProp)) (s : outParam (ββProp)) [FunLike F α β] :

RelHomClass F r s asserts that F is a type of functions such that all f : F satisfy r a b → s (f a) (f b).

The relations r and s are outParams since figuring them out from a goal is a higher-order matching problem that Lean usually can't do unaided.

  • map_rel (f : F) {a b : α} : r a bs (f a) (f b)

    A RelHomClass sends related elements to related elements

Instances
    theorem RelHomClass.isIrrefl {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} {F : Type u_5} [FunLike F α β] [RelHomClass F r s] (f : F) [IsIrrefl β s] :
    theorem RelHomClass.isAsymm {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} {F : Type u_5} [FunLike F α β] [RelHomClass F r s] (f : F) [IsAsymm β s] :
    IsAsymm α r
    theorem RelHomClass.acc {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} {F : Type u_5} [FunLike F α β] [RelHomClass F r s] (f : F) (a : α) :
    Acc s (f a)Acc r a
    theorem RelHomClass.wellFounded {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} {F : Type u_5} [FunLike F α β] [RelHomClass F r s] (f : F) :
    theorem RelHomClass.isWellFounded {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} {F : Type u_5} [FunLike F α β] [RelHomClass F r s] (f : F) [IsWellFounded β s] :
    instance RelHom.instFunLike {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} :
    FunLike (r →r s) α β
    Equations
    instance RelHom.instRelHomClass {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} :
    RelHomClass (r →r s) r s
    theorem RelHom.map_rel {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} (f : r →r s) {a b : α} :
    r a bs (f a) (f b)
    @[simp]
    theorem RelHom.coe_fn_toFun {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} (f : r →r s) :
    f.toFun = f
    @[simp]
    theorem RelHom.coeFn_mk {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} (f : αβ) (h : ∀ {a b : α}, r a bs (f a) (f b)) :
    { toFun := f, map_rel' := h } = f
    theorem RelHom.coe_fn_injective {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} :
    Function.Injective fun (f : r →r s) => f

    The map coe_fn : (r →r s) → (α → β) is injective.

    theorem RelHom.ext {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} f g : r →r s (h : ∀ (x : α), f x = g x) :
    f = g
    theorem RelHom.ext_iff {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} {f g : r →r s} :
    f = g ∀ (x : α), f x = g x
    def RelHom.id {α : Type u_1} (r : ααProp) :
    r →r r

    Identity map is a relation homomorphism.

    Equations
    • RelHom.id r = { toFun := fun (x : α) => x, map_rel' := }
    @[simp]
    theorem RelHom.id_apply {α : Type u_1} (r : ααProp) (x : α) :
    (RelHom.id r) x = x
    def RelHom.comp {α : Type u_1} {β : Type u_2} {γ : Type u_3} {r : ααProp} {s : ββProp} {t : γγProp} (g : s →r t) (f : r →r s) :
    r →r t

    Composition of two relation homomorphisms is a relation homomorphism.

    Equations
    • g.comp f = { toFun := fun (x : α) => g (f x), map_rel' := }
    @[simp]
    theorem RelHom.comp_apply {α : Type u_1} {β : Type u_2} {γ : Type u_3} {r : ααProp} {s : ββProp} {t : γγProp} (g : s →r t) (f : r →r s) (x : α) :
    (g.comp f) x = g (f x)
    def RelHom.swap {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} (f : r →r s) :

    A relation homomorphism is also a relation homomorphism between dual relations.

    Equations
    • f.swap = { toFun := f, map_rel' := }
    @[simp]
    theorem RelHom.swap_apply {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} (f : r →r s) (a : α) :
    f.swap a = f a
    def RelHom.preimage {α : Type u_1} {β : Type u_2} (f : αβ) (s : ββProp) :

    A function is a relation homomorphism from the preimage relation of s to s.

    Equations
    @[simp]
    theorem RelHom.preimage_apply {α : Type u_1} {β : Type u_2} (f : αβ) (s : ββProp) (a✝ : α) :
    (preimage f s) a✝ = f a✝
    theorem injective_of_increasing {α : Type u_1} {β : Type u_2} (r : ααProp) (s : ββProp) [IsTrichotomous α r] [IsIrrefl β s] (f : αβ) (hf : ∀ {x y : α}, r x ys (f x) (f y)) :

    An increasing function is injective

    theorem RelHom.injective_of_increasing {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} [IsTrichotomous α r] [IsIrrefl β s] (f : r →r s) :

    An increasing function is injective

    theorem Function.Surjective.wellFounded_iff {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} {f : αβ} (hf : Surjective f) (o : ∀ {a b : α}, r a b s (f a) (f b)) :
    structure RelEmbedding {α : Type u_5} {β : Type u_6} (r : ααProp) (s : ββProp) extends α β :
    Type (max u_5 u_6)

    A relation embedding with respect to a given pair of relations r and s is an embedding f : α ↪ β such that r a b ↔ s (f a) (f b).

    A relation embedding with respect to a given pair of relations r and s is an embedding f : α ↪ β such that r a b ↔ s (f a) (f b).

    Equations
    theorem preimage_equivalence {α : Type u_5} {β : Type u_6} (f : αβ) {s : ββProp} (hs : Equivalence s) :
    def RelEmbedding.toRelHom {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} (f : r ↪r s) :
    r →r s

    A relation embedding is also a relation homomorphism

    Equations
    instance RelEmbedding.instCoeRelHom {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} :
    Coe (r ↪r s) (r →r s)
    Equations
    instance RelEmbedding.instFunLike {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} :
    FunLike (r ↪r s) α β
    Equations
    instance RelEmbedding.instRelHomClass {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} :
    RelHomClass (r ↪r s) r s
    instance RelEmbedding.instEmbeddingLike {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} :
    EmbeddingLike (r ↪r s) α β
    @[simp]
    theorem RelEmbedding.coe_toEmbedding {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} {f : r ↪r s} :
    f.toEmbedding = f
    @[simp]
    theorem RelEmbedding.coe_toRelHom {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} {f : r ↪r s} :
    f.toRelHom = f
    theorem RelEmbedding.toEmbedding_injective {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} :
    @[simp]
    theorem RelEmbedding.toEmbedding_inj {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} {f g : r ↪r s} :
    theorem RelEmbedding.injective {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} (f : r ↪r s) :
    theorem RelEmbedding.inj {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} (f : r ↪r s) {a b : α} :
    f a = f b a = b
    theorem RelEmbedding.map_rel_iff {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} (f : r ↪r s) {a b : α} :
    s (f a) (f b) r a b
    @[simp]
    theorem RelEmbedding.coe_mk {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} {f : α β} {h : ∀ {a b : α}, s (f a) (f b) r a b} :
    { toEmbedding := f, map_rel_iff' := h } = f
    theorem RelEmbedding.coe_fn_injective {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} :
    Function.Injective fun (f : r ↪r s) => f

    The map coe_fn : (r ↪r s) → (α → β) is injective.

    theorem RelEmbedding.ext {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} f g : r ↪r s (h : ∀ (x : α), f x = g x) :
    f = g
    theorem RelEmbedding.ext_iff {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} {f g : r ↪r s} :
    f = g ∀ (x : α), f x = g x
    def RelEmbedding.refl {α : Type u_1} (r : ααProp) :
    r ↪r r

    Identity map is a relation embedding.

    Equations
    @[simp]
    theorem RelEmbedding.refl_apply {α : Type u_1} (r : ααProp) (a : α) :
    def RelEmbedding.trans {α : Type u_1} {β : Type u_2} {γ : Type u_3} {r : ααProp} {s : ββProp} {t : γγProp} (f : r ↪r s) (g : s ↪r t) :
    r ↪r t

    Composition of two relation embeddings is a relation embedding.

    Equations
    instance RelEmbedding.instInhabited {α : Type u_1} (r : ααProp) :
    Equations
    theorem RelEmbedding.trans_apply {α : Type u_1} {β : Type u_2} {γ : Type u_3} {r : ααProp} {s : ββProp} {t : γγProp} (f : r ↪r s) (g : s ↪r t) (a : α) :
    (f.trans g) a = g (f a)
    @[simp]
    theorem RelEmbedding.coe_trans {α : Type u_1} {β : Type u_2} {γ : Type u_3} {r : ααProp} {s : ββProp} {t : γγProp} (f : r ↪r s) (g : s ↪r t) :
    (f.trans g) = g f
    def RelEmbedding.swap {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} (f : r ↪r s) :

    A relation embedding is also a relation embedding between dual relations.

    Equations
    @[simp]
    theorem RelEmbedding.swap_apply {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} (f : r ↪r s) (a : α) :
    f.swap a = f a
    def RelEmbedding.preimage {α : Type u_1} {β : Type u_2} (f : α β) (s : ββProp) :
    f ⁻¹'o s ↪r s

    If f is injective, then it is a relation embedding from the preimage relation of s to s.

    Equations
    @[simp]
    theorem RelEmbedding.preimage_apply {α : Type u_1} {β : Type u_2} (f : α β) (s : ββProp) (a : α) :
    (preimage f s) a = f a
    theorem RelEmbedding.eq_preimage {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} (f : r ↪r s) :
    r = f ⁻¹'o s
    theorem RelEmbedding.isIrrefl {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} (f : r ↪r s) [IsIrrefl β s] :
    theorem RelEmbedding.isRefl {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} (f : r ↪r s) [IsRefl β s] :
    IsRefl α r
    theorem RelEmbedding.isSymm {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} (f : r ↪r s) [IsSymm β s] :
    IsSymm α r
    theorem RelEmbedding.isAsymm {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} (f : r ↪r s) [IsAsymm β s] :
    IsAsymm α r
    theorem RelEmbedding.isAntisymm {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} :
    ∀ (x : r ↪r s) [IsAntisymm β s], IsAntisymm α r
    theorem RelEmbedding.isTrans {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} :
    ∀ (x : r ↪r s) [IsTrans β s], IsTrans α r
    theorem RelEmbedding.isTotal {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} :
    ∀ (x : r ↪r s) [IsTotal β s], IsTotal α r
    theorem RelEmbedding.isPreorder {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} :
    ∀ (x : r ↪r s) [IsPreorder β s], IsPreorder α r
    theorem RelEmbedding.isPartialOrder {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} :
    ∀ (x : r ↪r s) [IsPartialOrder β s], IsPartialOrder α r
    theorem RelEmbedding.isLinearOrder {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} :
    ∀ (x : r ↪r s) [IsLinearOrder β s], IsLinearOrder α r
    theorem RelEmbedding.isStrictOrder {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} :
    ∀ (x : r ↪r s) [IsStrictOrder β s], IsStrictOrder α r
    theorem RelEmbedding.isTrichotomous {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} :
    ∀ (x : r ↪r s) [IsTrichotomous β s], IsTrichotomous α r
    theorem RelEmbedding.isStrictTotalOrder {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} :
    ∀ (x : r ↪r s) [IsStrictTotalOrder β s], IsStrictTotalOrder α r
    theorem RelEmbedding.acc {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} (f : r ↪r s) (a : α) :
    Acc s (f a)Acc r a
    theorem RelEmbedding.wellFounded {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} :
    ∀ (x : r ↪r s), WellFounded sWellFounded r
    theorem RelEmbedding.isWellFounded {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} (f : r ↪r s) [IsWellFounded β s] :
    theorem RelEmbedding.isWellOrder {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} :
    ∀ (x : r ↪r s) [IsWellOrder β s], IsWellOrder α r
    def Subtype.relEmbedding {X : Type u_5} (r : XXProp) (p : XProp) :

    The induced relation on a subtype is an embedding under the natural inclusion.

    Equations
    @[simp]
    theorem Subtype.relEmbedding_apply {X : Type u_5} (r : XXProp) (p : XProp) (self : Subtype p) :
    (relEmbedding r p) self = self
    instance Subtype.wellFoundedLT {α : Type u_1} [LT α] [WellFoundedLT α] (p : αProp) :
    instance Subtype.wellFoundedGT {α : Type u_1} [LT α] [WellFoundedGT α] (p : αProp) :
    def Quotient.mkRelHom {α : Type u_1} {x✝ : Setoid α} {r : ααProp} (H : ∀ (a₁ b₁ a₂ b₂ : α), a₁ a₂b₁ b₂r a₁ b₁ = r a₂ b₂) :

    Quotient.mk as a relation homomorphism between the relation and the lift of a relation.

    Equations
    @[simp]
    theorem Quotient.mkRelHom_apply {α : Type u_1} {x✝ : Setoid α} {r : ααProp} (H : ∀ (a₁ b₁ a₂ b₂ : α), a₁ a₂b₁ b₂r a₁ b₁ = r a₂ b₂) (a : α) :
    noncomputable def Quotient.outRelEmbedding {α : Type u_1} {x✝ : Setoid α} {r : ααProp} (H : ∀ (a₁ b₁ a₂ b₂ : α), a₁ a₂b₁ b₂r a₁ b₁ = r a₂ b₂) :

    Quotient.out as a relation embedding between the lift of a relation and the relation.

    Equations
    @[simp]
    theorem Quotient.outRelEmbedding_apply {α : Type u_1} {x✝ : Setoid α} {r : ααProp} (H : ∀ (a₁ b₁ a₂ b₂ : α), a₁ a₂b₁ b₂r a₁ b₁ = r a₂ b₂) (a✝ : Quotient x✝) :
    (outRelEmbedding H) a✝ = a✝.out
    @[simp]
    theorem acc_lift₂_iff {α : Type u_1} {x✝ : Setoid α} {r : ααProp} {H : ∀ (a₁ b₁ a₂ b₂ : α), a₁ a₂b₁ b₂r a₁ b₁ = r a₂ b₂} {a : α} :
    @[simp]
    theorem acc_liftOn₂'_iff {α : Type u_1} {s : Setoid α} {r : ααProp} {H : ∀ (a₁ a₂ b₁ b₂ : α), s a₁ b₁s a₂ b₂r a₁ a₂ = r b₁ b₂} {a : α} :
    Acc (fun (x y : Quotient s) => x.liftOn₂' y r H) (Quotient.mk'' a) Acc r a
    @[simp]
    theorem wellFounded_lift₂_iff {α : Type u_1} {x✝ : Setoid α} {r : ααProp} {H : ∀ (a₁ b₁ a₂ b₂ : α), a₁ a₂b₁ b₂r a₁ b₁ = r a₂ b₂} :

    A relation is well founded iff its lift to a quotient is.

    theorem WellFounded.quotient_lift₂ {α : Type u_1} {x✝ : Setoid α} {r : ααProp} {H : ∀ (a₁ b₁ a₂ b₂ : α), a₁ a₂b₁ b₂r a₁ b₁ = r a₂ b₂} :

    Alias of the reverse direction of wellFounded_lift₂_iff.


    A relation is well founded iff its lift to a quotient is.

    theorem WellFounded.of_quotient_lift₂ {α : Type u_1} {x✝ : Setoid α} {r : ααProp} {H : ∀ (a₁ b₁ a₂ b₂ : α), a₁ a₂b₁ b₂r a₁ b₁ = r a₂ b₂} :

    Alias of the forward direction of wellFounded_lift₂_iff.


    A relation is well founded iff its lift to a quotient is.

    @[simp]
    theorem wellFounded_liftOn₂'_iff {α : Type u_1} {s : Setoid α} {r : ααProp} {H : ∀ (a₁ a₂ b₁ b₂ : α), s a₁ b₁s a₂ b₂r a₁ a₂ = r b₁ b₂} :
    (WellFounded fun (x y : Quotient s) => x.liftOn₂' y r H) WellFounded r
    theorem WellFounded.of_quotient_liftOn₂' {α : Type u_1} {s : Setoid α} {r : ααProp} {H : ∀ (a₁ a₂ b₁ b₂ : α), s a₁ b₁s a₂ b₂r a₁ a₂ = r b₁ b₂} :
    (WellFounded fun (x y : Quotient s) => x.liftOn₂' y r H)WellFounded r

    Alias of the forward direction of wellFounded_liftOn₂'_iff.

    theorem WellFounded.quotient_liftOn₂' {α : Type u_1} {s : Setoid α} {r : ααProp} {H : ∀ (a₁ a₂ b₁ b₂ : α), s a₁ b₁s a₂ b₂r a₁ a₂ = r b₁ b₂} :
    WellFounded rWellFounded fun (x y : Quotient s) => x.liftOn₂' y r H

    Alias of the reverse direction of wellFounded_liftOn₂'_iff.

    def RelEmbedding.ofMapRelIff {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} (f : αβ) [IsAntisymm α r] [IsRefl β s] (hf : ∀ (a b : α), s (f a) (f b) r a b) :
    r ↪r s

    To define a relation embedding from an antisymmetric relation r to a reflexive relation s it suffices to give a function together with a proof that it satisfies s (f a) (f b) ↔ r a b.

    Equations
    @[simp]
    theorem RelEmbedding.ofMapRelIff_coe {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} (f : αβ) [IsAntisymm α r] [IsRefl β s] (hf : ∀ (a b : α), s (f a) (f b) r a b) :
    (ofMapRelIff f hf) = f
    def RelEmbedding.ofMonotone {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} [IsTrichotomous α r] [IsAsymm β s] (f : αβ) (H : ∀ (a b : α), r a bs (f a) (f b)) :
    r ↪r s

    It suffices to prove f is monotone between strict relations to show it is a relation embedding.

    Equations
    @[simp]
    theorem RelEmbedding.ofMonotone_coe {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} [IsTrichotomous α r] [IsAsymm β s] (f : αβ) (H : ∀ (a b : α), r a bs (f a) (f b)) :
    (ofMonotone f H) = f
    def RelEmbedding.ofIsEmpty {α : Type u_1} {β : Type u_2} (r : ααProp) (s : ββProp) [IsEmpty α] :
    r ↪r s

    A relation embedding from an empty type.

    Equations
    def RelEmbedding.sumLiftRelInl {α : Type u_1} {β : Type u_2} (r : ααProp) (s : ββProp) :

    Sum.inl as a relation embedding into Sum.LiftRel r s.

    Equations
    @[simp]
    theorem RelEmbedding.sumLiftRelInl_apply {α : Type u_1} {β : Type u_2} (r : ααProp) (s : ββProp) (val : α) :
    (sumLiftRelInl r s) val = Sum.inl val
    def RelEmbedding.sumLiftRelInr {α : Type u_1} {β : Type u_2} (r : ααProp) (s : ββProp) :

    Sum.inr as a relation embedding into Sum.LiftRel r s.

    Equations
    @[simp]
    theorem RelEmbedding.sumLiftRelInr_apply {α : Type u_1} {β : Type u_2} (r : ααProp) (s : ββProp) (val : β) :
    (sumLiftRelInr r s) val = Sum.inr val
    def RelEmbedding.sumLiftRelMap {α : Type u_1} {β : Type u_2} {γ : Type u_3} {δ : Type u_4} {r : ααProp} {s : ββProp} {t : γγProp} {u : δδProp} (f : r ↪r s) (g : t ↪r u) :

    Sum.map as a relation embedding between Sum.LiftRel relations.

    Equations
    @[simp]
    theorem RelEmbedding.sumLiftRelMap_apply {α : Type u_1} {β : Type u_2} {γ : Type u_3} {δ : Type u_4} {r : ααProp} {s : ββProp} {t : γγProp} {u : δδProp} (f : r ↪r s) (g : t ↪r u) (a✝ : α γ) :
    (f.sumLiftRelMap g) a✝ = Sum.map (⇑f) (⇑g) a✝
    def RelEmbedding.sumLexInl {α : Type u_1} {β : Type u_2} (r : ααProp) (s : ββProp) :

    Sum.inl as a relation embedding into Sum.Lex r s.

    Equations
    @[simp]
    theorem RelEmbedding.sumLexInl_apply {α : Type u_1} {β : Type u_2} (r : ααProp) (s : ββProp) (val : α) :
    (sumLexInl r s) val = Sum.inl val
    def RelEmbedding.sumLexInr {α : Type u_1} {β : Type u_2} (r : ααProp) (s : ββProp) :

    Sum.inr as a relation embedding into Sum.Lex r s.

    Equations
    @[simp]
    theorem RelEmbedding.sumLexInr_apply {α : Type u_1} {β : Type u_2} (r : ααProp) (s : ββProp) (val : β) :
    (sumLexInr r s) val = Sum.inr val
    def RelEmbedding.sumLexMap {α : Type u_1} {β : Type u_2} {γ : Type u_3} {δ : Type u_4} {r : ααProp} {s : ββProp} {t : γγProp} {u : δδProp} (f : r ↪r s) (g : t ↪r u) :

    Sum.map as a relation embedding between Sum.Lex relations.

    Equations
    @[simp]
    theorem RelEmbedding.sumLexMap_apply {α : Type u_1} {β : Type u_2} {γ : Type u_3} {δ : Type u_4} {r : ααProp} {s : ββProp} {t : γγProp} {u : δδProp} (f : r ↪r s) (g : t ↪r u) (a✝ : α γ) :
    (f.sumLexMap g) a✝ = Sum.map (⇑f) (⇑g) a✝
    def RelEmbedding.prodLexMkLeft {α : Type u_1} {β : Type u_2} {r : ααProp} (s : ββProp) {a : α} (h : ¬r a a) :

    fun b ↦ Prod.mk a b as a relation embedding.

    Equations
    @[simp]
    theorem RelEmbedding.prodLexMkLeft_apply {α : Type u_1} {β : Type u_2} {r : ααProp} (s : ββProp) {a : α} (h : ¬r a a) (snd : β) :
    (prodLexMkLeft s h) snd = (a, snd)
    def RelEmbedding.prodLexMkRight {α : Type u_1} {β : Type u_2} {s : ββProp} (r : ααProp) {b : β} (h : ¬s b b) :

    fun a ↦ Prod.mk a b as a relation embedding.

    Equations
    @[simp]
    theorem RelEmbedding.prodLexMkRight_apply {α : Type u_1} {β : Type u_2} {s : ββProp} (r : ααProp) {b : β} (h : ¬s b b) (a : α) :
    (prodLexMkRight r h) a = (a, b)
    def RelEmbedding.prodLexMap {α : Type u_1} {β : Type u_2} {γ : Type u_3} {δ : Type u_4} {r : ααProp} {s : ββProp} {t : γγProp} {u : δδProp} (f : r ↪r s) (g : t ↪r u) :

    Prod.map as a relation embedding.

    Equations
    @[simp]
    theorem RelEmbedding.prodLexMap_apply {α : Type u_1} {β : Type u_2} {γ : Type u_3} {δ : Type u_4} {r : ααProp} {s : ββProp} {t : γγProp} {u : δδProp} (f : r ↪r s) (g : t ↪r u) (a✝ : α × γ) :
    (f.prodLexMap g) a✝ = Prod.map (⇑f) (⇑g) a✝
    structure RelIso {α : Type u_5} {β : Type u_6} (r : ααProp) (s : ββProp) extends α β :
    Type (max u_5 u_6)

    A relation isomorphism is an equivalence that is also a relation embedding.

    A relation isomorphism is an equivalence that is also a relation embedding.

    Equations
    def RelIso.toRelEmbedding {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} (f : r ≃r s) :
    r ↪r s

    Convert a RelIso to a RelEmbedding. This function is also available as a coercion but often it is easier to write f.toRelEmbedding than to write explicitly r and s in the target type.

    Equations
    theorem RelIso.toEquiv_injective {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} :
    instance RelIso.instCoeOutRelEmbedding {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} :
    CoeOut (r ≃r s) (r ↪r s)
    Equations
    instance RelIso.instFunLike {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} :
    FunLike (r ≃r s) α β
    Equations
    instance RelIso.instRelHomClass {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} :
    RelHomClass (r ≃r s) r s
    instance RelIso.instEquivLike {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} :
    EquivLike (r ≃r s) α β
    Equations
    @[simp]
    theorem RelIso.coe_toRelEmbedding {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} (f : r ≃r s) :
    f.toRelEmbedding = f
    @[simp]
    theorem RelIso.coe_toEmbedding {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} (f : r ≃r s) :
    f.toEmbedding = f
    theorem RelIso.map_rel_iff {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} (f : r ≃r s) {a b : α} :
    s (f a) (f b) r a b
    @[simp]
    theorem RelIso.coe_fn_mk {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} (f : α β) (o : ∀ ⦃a b : α⦄, s (f a) (f b) r a b) :
    { toEquiv := f, map_rel_iff' := o } = f
    @[simp]
    theorem RelIso.coe_fn_toEquiv {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} (f : r ≃r s) :
    f.toEquiv = f
    theorem RelIso.coe_fn_injective {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} :
    Function.Injective fun (f : r ≃r s) => f

    The map DFunLike.coe : (r ≃r s) → (α → β) is injective.

    theorem RelIso.ext {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} f g : r ≃r s (h : ∀ (x : α), f x = g x) :
    f = g
    theorem RelIso.ext_iff {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} {f g : r ≃r s} :
    f = g ∀ (x : α), f x = g x
    def RelIso.symm {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} (f : r ≃r s) :
    s ≃r r

    Inverse map of a relation isomorphism is a relation isomorphism.

    Equations
    def RelIso.Simps.apply {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} (h : r ≃r s) :
    αβ

    See Note [custom simps projection]. We need to specify this projection explicitly in this case, because RelIso defines custom coercions other than the ones given by DFunLike.

    Equations
    def RelIso.Simps.symm_apply {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} (h : r ≃r s) :
    βα

    See Note [custom simps projection].

    Equations
    def RelIso.refl {α : Type u_1} (r : ααProp) :
    r ≃r r

    Identity map is a relation isomorphism.

    Equations
    @[simp]
    theorem RelIso.refl_apply {α : Type u_1} (r : ααProp) (a : α) :
    (RelIso.refl r) a = a
    def RelIso.trans {α : Type u_1} {β : Type u_2} {γ : Type u_3} {r : ααProp} {s : ββProp} {t : γγProp} (f₁ : r ≃r s) (f₂ : s ≃r t) :
    r ≃r t

    Composition of two relation isomorphisms is a relation isomorphism.

    Equations
    @[simp]
    theorem RelIso.trans_apply {α : Type u_1} {β : Type u_2} {γ : Type u_3} {r : ααProp} {s : ββProp} {t : γγProp} (f₁ : r ≃r s) (f₂ : s ≃r t) (a✝ : α) :
    (f₁.trans f₂) a✝ = f₂ (f₁ a✝)
    instance RelIso.instInhabited {α : Type u_1} (r : ααProp) :
    Equations
    @[simp]
    theorem RelIso.default_def {α : Type u_1} (r : ααProp) :
    @[simp]
    theorem RelIso.apply_symm_apply {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} (e : r ≃r s) (x : β) :
    e (e.symm x) = x
    @[simp]
    theorem RelIso.symm_apply_apply {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} (e : r ≃r s) (x : α) :
    e.symm (e x) = x
    @[simp]
    theorem RelIso.symm_comp_self {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} (e : r ≃r s) :
    e.symm e = id
    @[simp]
    theorem RelIso.self_comp_symm {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} (e : r ≃r s) :
    e e.symm = id
    @[simp]
    theorem RelIso.symm_trans_apply {α : Type u_1} {β : Type u_2} {γ : Type u_3} {r : ααProp} {s : ββProp} {t : γγProp} (f : r ≃r s) (g : s ≃r t) (a : γ) :
    (f.trans g).symm a = f.symm (g.symm a)
    theorem RelIso.symm_symm_apply {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} (f : r ≃r s) (b : α) :
    f.symm.symm b = f b
    theorem RelIso.apply_eq_iff_eq {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} (f : r ≃r s) {x y : α} :
    f x = f y x = y
    theorem RelIso.apply_eq_iff_eq_symm_apply {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} {x : α} {y : β} (f : r ≃r s) :
    f x = y x = f.symm y
    theorem RelIso.symm_apply_eq {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} (e : r ≃r s) {x : β} {y : α} :
    e.symm x = y x = e y
    theorem RelIso.eq_symm_apply {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} (e : r ≃r s) {x : β} {y : α} :
    y = e.symm x e y = x
    @[simp]
    theorem RelIso.symm_symm {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} (e : r ≃r s) :
    e.symm.symm = e
    theorem RelIso.symm_bijective {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} :
    @[simp]
    theorem RelIso.refl_symm {α : Type u_1} {r : ααProp} :
    @[simp]
    theorem RelIso.trans_refl {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} (e : r ≃r s) :
    @[simp]
    theorem RelIso.refl_trans {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} (e : r ≃r s) :
    @[simp]
    theorem RelIso.symm_trans_self {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} (e : r ≃r s) :
    @[simp]
    theorem RelIso.self_trans_symm {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} (e : r ≃r s) :
    theorem RelIso.trans_assoc {α : Type u_1} {β : Type u_2} {γ : Type u_3} {r : ααProp} {s : ββProp} {t : γγProp} {δ : Type u_5} {u : δδProp} (ab : r ≃r s) (bc : s ≃r t) (cd : t ≃r u) :
    (ab.trans bc).trans cd = ab.trans (bc.trans cd)
    def RelIso.cast {α β : Type u} {r : ααProp} {s : ββProp} (h₁ : α = β) (h₂ : r s) :
    r ≃r s

    A relation isomorphism between equal relations on equal types.

    Equations
    @[simp]
    theorem RelIso.cast_toEquiv {α β : Type u} {r : ααProp} {s : ββProp} (h₁ : α = β) (h₂ : r s) :
    (RelIso.cast h₁ h₂).toEquiv = Equiv.cast h₁
    @[simp]
    theorem RelIso.cast_apply {α β : Type u} {r : ααProp} {s : ββProp} (h₁ : α = β) (h₂ : r s) (a : α) :
    (RelIso.cast h₁ h₂) a = cast h₁ a
    @[simp]
    theorem RelIso.cast_symm {α β : Type u} {r : ααProp} {s : ββProp} (h₁ : α = β) (h₂ : r s) :
    (RelIso.cast h₁ h₂).symm = RelIso.cast
    @[simp]
    theorem RelIso.cast_refl {α : Type u} {r : ααProp} (h₁ : α = α := ) (h₂ : r r := ) :
    @[simp]
    theorem RelIso.cast_trans {α β γ : Type u} {r : ααProp} {s : ββProp} {t : γγProp} (h₁ : α = β) (h₁' : β = γ) (h₂ : r s) (h₂' : s t) :
    (RelIso.cast h₁ h₂).trans (RelIso.cast h₁' h₂') = RelIso.cast
    def RelIso.swap {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} (f : r ≃r s) :

    A relation isomorphism is also a relation isomorphism between dual relations.

    Equations
    • f.swap = { toEquiv := f, map_rel_iff' := }
    def RelIso.compl {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} (f : r ≃r s) :

    A relation isomorphism is also a relation isomorphism between complemented relations.

    Equations
    • f.compl = { toEquiv := f, map_rel_iff' := }
    @[simp]
    theorem RelIso.compl_symm_apply {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} (f : r ≃r s) (a✝ : β) :
    f.compl.symm a✝ = EquivLike.inv f a✝
    @[simp]
    theorem RelIso.compl_apply {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} (f : r ≃r s) (a : α) :
    f.compl a = f a
    @[simp]
    theorem RelIso.coe_fn_symm_mk {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} (f : α β) (o : ∀ {a b : α}, s (f a) (f b) r a b) :
    { toEquiv := f, map_rel_iff' := o }.symm = f.symm
    theorem RelIso.rel_symm_apply {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} (e : r ≃r s) {x : α} {y : β} :
    r x (e.symm y) s (e x) y
    theorem RelIso.symm_apply_rel {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} (e : r ≃r s) {x : β} {y : α} :
    r (e.symm x) y s x (e y)
    theorem RelIso.bijective {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} (e : r ≃r s) :
    theorem RelIso.injective {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} (e : r ≃r s) :
    theorem RelIso.surjective {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} (e : r ≃r s) :
    theorem RelIso.eq_iff_eq {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} (f : r ≃r s) {a b : α} :
    f a = f b a = b
    def RelIso.copy {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} (e : r ≃r s) (f : αβ) (g : βα) (hf : f = e) (hg : g = e.symm) :
    r ≃r s

    Copy of a RelIso with a new toFun and invFun equal to the old ones. Useful to fix definitional equalities.

    Equations
    • e.copy f g hf hg = { toFun := f, invFun := g, left_inv := , right_inv := , map_rel_iff' := }
    @[simp]
    theorem RelIso.coe_copy {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} (e : r ≃r s) (f : αβ) (g : βα) (hf : f = e) (hg : g = e.symm) :
    (e.copy f g hf hg) = f
    theorem RelIso.copy_eq {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} (e : r ≃r s) (f : αβ) (g : βα) (hf : f = e) (hg : g = e.symm) :
    e.copy f g hf hg = e
    def RelIso.preimage {α : Type u_1} {β : Type u_2} (f : α β) (s : ββProp) :
    f ⁻¹'o s ≃r s

    Any equivalence lifts to a relation isomorphism between s and its preimage.

    Equations
    @[simp]
    theorem RelIso.preimage_apply {α : Type u_1} {β : Type u_2} (f : α β) (s : ββProp) (a : α) :
    (RelIso.preimage f s) a = f a
    @[simp]
    theorem RelIso.preimage_symm_apply {α : Type u_1} {β : Type u_2} (f : α β) (s : ββProp) (a : β) :
    instance RelIso.IsWellOrder.preimage {β : Type u_2} {α : Type u} (r : ααProp) [IsWellOrder α r] (f : β α) :
    IsWellOrder β (f ⁻¹'o r)
    instance RelIso.IsWellOrder.ulift {α : Type u} (r : ααProp) [IsWellOrder α r] :
    noncomputable def RelIso.ofSurjective {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} (f : r ↪r s) (H : Function.Surjective f) :
    r ≃r s

    A surjective relation embedding is a relation isomorphism.

    Equations
    @[simp]
    theorem RelIso.ofSurjective_apply {α : Type u_1} {β : Type u_2} {r : ααProp} {s : ββProp} (f : r ↪r s) (H : Function.Surjective f) (a : α) :
    (ofSurjective f H) a = f a
    def RelIso.relHomCongr {α₁ : Type u_5} {β₁ : Type u_6} {α₂ : Type u_7} {β₂ : Type u_8} {r₁ : α₁α₁Prop} {s₁ : β₁β₁Prop} {r₂ : α₂α₂Prop} {s₂ : β₂β₂Prop} (e₁ : r₁ ≃r r₂) (e₂ : s₁ ≃r s₂) :
    r₁ →r s₁ (r₂ →r s₂)

    Transport a RelHom across a pair of RelIsos, by pre- and post-composition.

    This is Equiv.arrowCongr for RelHom.

    Equations
    • One or more equations did not get rendered due to their size.
    @[simp]
    theorem RelIso.relHomCongr_apply {α₁ : Type u_5} {β₁ : Type u_6} {α₂ : Type u_7} {β₂ : Type u_8} {r₁ : α₁α₁Prop} {s₁ : β₁β₁Prop} {r₂ : α₂α₂Prop} {s₂ : β₂β₂Prop} (e₁ : r₁ ≃r r₂) (e₂ : s₁ ≃r s₂) (f₁ : r₁ →r s₁) :
    @[simp]
    theorem RelIso.relHomCongr_symm_apply {α₁ : Type u_5} {β₁ : Type u_6} {α₂ : Type u_7} {β₂ : Type u_8} {r₁ : α₁α₁Prop} {s₁ : β₁β₁Prop} {r₂ : α₂α₂Prop} {s₂ : β₂β₂Prop} (e₁ : r₁ ≃r r₂) (e₂ : s₁ ≃r s₂) (f₂ : r₂ →r s₂) :
    def RelIso.sumLexCongr {α₁ : Type u_5} {α₂ : Type u_6} {β₁ : Type u_7} {β₂ : Type u_8} {r₁ : α₁α₁Prop} {r₂ : α₂α₂Prop} {s₁ : β₁β₁Prop} {s₂ : β₂β₂Prop} (e₁ : r₁ ≃r s₁) (e₂ : r₂ ≃r s₂) :
    Sum.Lex r₁ r₂ ≃r Sum.Lex s₁ s₂

    Given relation isomorphisms r₁ ≃r s₁ and r₂ ≃r s₂, construct a relation isomorphism for the lexicographic orders on the sum.

    Equations
    def RelIso.prodLexCongr {α₁ : Type u_5} {α₂ : Type u_6} {β₁ : Type u_7} {β₂ : Type u_8} {r₁ : α₁α₁Prop} {r₂ : α₂α₂Prop} {s₁ : β₁β₁Prop} {s₂ : β₂β₂Prop} (e₁ : r₁ ≃r s₁) (e₂ : r₂ ≃r s₂) :
    Prod.Lex r₁ r₂ ≃r Prod.Lex s₁ s₂

    Given relation isomorphisms r₁ ≃r s₁ and r₂ ≃r s₂, construct a relation isomorphism for the lexicographic orders on the product.

    Equations
    def RelIso.relIsoOfIsEmpty {α : Type u_1} {β : Type u_2} (r : ααProp) (s : ββProp) [IsEmpty α] [IsEmpty β] :
    r ≃r s

    Two relations on empty types are isomorphic.

    Equations
    def RelIso.ofUniqueOfIrrefl {α : Type u_1} {β : Type u_2} (r : ααProp) (s : ββProp) [IsIrrefl α r] [IsIrrefl β s] [Unique α] [Unique β] :
    r ≃r s

    Two irreflexive relations on a unique type are isomorphic.

    Equations
    @[deprecated RelIso.ofUniqueOfIrrefl (since := "2024-12-26")]
    def RelIso.relIsoOfUniqueOfIrrefl {α : Type u_1} {β : Type u_2} (r : ααProp) (s : ββProp) [IsIrrefl α r] [IsIrrefl β s] [Unique α] [Unique β] :
    r ≃r s

    Alias of RelIso.ofUniqueOfIrrefl.


    Two irreflexive relations on a unique type are isomorphic.

    Equations
    def RelIso.ofUniqueOfRefl {α : Type u_1} {β : Type u_2} (r : ααProp) (s : ββProp) [IsRefl α r] [IsRefl β s] [Unique α] [Unique β] :
    r ≃r s

    Two reflexive relations on a unique type are isomorphic.

    Equations
    @[deprecated RelIso.ofUniqueOfRefl (since := "2024-12-26")]
    def RelIso.relIsoOfUniqueOfRefl {α : Type u_1} {β : Type u_2} (r : ααProp) (s : ββProp) [IsRefl α r] [IsRefl β s] [Unique α] [Unique β] :
    r ≃r s

    Alias of RelIso.ofUniqueOfRefl.


    Two reflexive relations on a unique type are isomorphic.

    Equations