Documentation

Mathlib.Order.SemiconjSup

Semiconjugate by supₛ #

In this file we prove two facts about semiconjugate (families of) functions.

First, if an order isomorphism fa : α → α→ α is semiconjugate to an order embedding fb : β → β→ β by g : α → β→ β, then fb is semiconjugate to fa by y ↦ supₛ {x | g x ≤ y}↦ supₛ {x | g x ≤ y}≤ y}, see Semiconj.symm_adjoint.

Second, consider two actions f₁ f₂ : G → α → α→ α → α→ α of a group on a complete lattice by order isomorphisms. Then the map x ↦ ⨆ g : G, (f₁ g)⁻¹ (f₂ g x)↦ ⨆ g : G, (f₁ g)⁻¹ (f₂ g x)⨆ g : G, (f₁ g)⁻¹ (f₂ g x)⁻¹ (f₂ g x) semiconjugates each f₁ g' to f₂ g', see Function.supₛ_div_semiconj. In the case of a conditionally complete lattice, a similar statement holds true under an additional assumption that each set {(f₁ g)⁻¹ (f₂ g x) | g : G}⁻¹ (f₂ g x) | g : G} is bounded above, see Function.csupₛ_div_semiconj.

The lemmas come from [Étienne Ghys, Groupes d'homéomorphismes du cercle et cohomologie bornée][ghys87:groupes], Proposition 2.1 and 5.4 respectively. In the paper they are formulated for homeomorphisms of the circle, so in order to apply results from this file one has to lift these homeomorphisms to the real line first.

def IsOrderRightAdjoint {α : Type u_1} {β : Type u_2} [inst : Preorder α] [inst : Preorder β] (f : αβ) (g : βα) :

We say that g : β → α→ α is an order right adjoint function for f : α → β→ β if it sends each y to a least upper bound for {x | f x ≤ y}≤ y}. If α is a partial order, and f : α → β→ β has a right adjoint, then this right adjoint is unique.

Equations
theorem isOrderRightAdjoint_supₛ {α : Type u_1} {β : Type u_2} [inst : CompleteLattice α] [inst : Preorder β] (f : αβ) :
IsOrderRightAdjoint f fun y => supₛ { x | f x y }
theorem isOrderRightAdjoint_csupₛ {α : Type u_1} {β : Type u_2} [inst : ConditionallyCompleteLattice α] [inst : Preorder β] (f : αβ) (hne : ∀ (y : β), x, f x y) (hbdd : ∀ (y : β), BddAbove { x | f x y }) :
IsOrderRightAdjoint f fun y => supₛ { x | f x y }
theorem IsOrderRightAdjoint.unique {α : Type u_1} {β : Type u_2} [inst : PartialOrder α] [inst : Preorder β] {f : αβ} {g₁ : βα} {g₂ : βα} (h₁ : IsOrderRightAdjoint f g₁) (h₂ : IsOrderRightAdjoint f g₂) :
g₁ = g₂
theorem IsOrderRightAdjoint.right_mono {α : Type u_1} {β : Type u_2} [inst : Preorder α] [inst : Preorder β] {f : αβ} {g : βα} (h : IsOrderRightAdjoint f g) :
theorem IsOrderRightAdjoint.orderIso_comp {α : Type u_1} {β : Type u_2} {γ : Type u_3} [inst : Preorder α] [inst : Preorder β] [inst : Preorder γ] {f : αβ} {g : βα} (h : IsOrderRightAdjoint f g) (e : β ≃o γ) :
theorem IsOrderRightAdjoint.comp_orderIso {α : Type u_1} {β : Type u_2} {γ : Type u_3} [inst : Preorder α] [inst : Preorder β] [inst : Preorder γ] {f : αβ} {g : βα} (h : IsOrderRightAdjoint f g) (e : γ ≃o α) :
theorem Function.Semiconj.symm_adjoint {α : Type u_1} {β : Type u_2} [inst : PartialOrder α] [inst : Preorder β] {fa : α ≃o α} {fb : β ↪o β} {g : αβ} (h : Function.Semiconj g (RelIso.toRelEmbedding fa).toEmbedding fb.toEmbedding) {g' : βα} (hg' : IsOrderRightAdjoint g g') :
Function.Semiconj g' fb.toEmbedding (RelIso.toRelEmbedding fa).toEmbedding

If an order automorphism fa is semiconjugate to an order embedding fb by a function g and g' is an order right adjoint of g (i.e. g' y = supₛ {x | f x ≤ y}≤ y}), then fb is semiconjugate to fa by g'.

This is a version of Proposition 2.1 from [Étienne Ghys, Groupes d'homéomorphismes du cercle et cohomologie bornée][ghys87:groupes].

theorem Function.semiconj_of_isLUB {α : Type u_1} {G : Type u_2} [inst : PartialOrder α] [inst : Group G] (f₁ : G →* α ≃o α) (f₂ : G →* α ≃o α) {h : αα} (H : ∀ (x : α), IsLUB (Set.range fun g' => (RelIso.toRelEmbedding (f₁ g')⁻¹).toEmbedding ((RelIso.toRelEmbedding (f₂ g')).toEmbedding x)) (h x)) (g : G) :
Function.Semiconj h (RelIso.toRelEmbedding (f₂ g)).toEmbedding (RelIso.toRelEmbedding (f₁ g)).toEmbedding
theorem Function.supₛ_div_semiconj {α : Type u_1} {G : Type u_2} [inst : CompleteLattice α] [inst : Group G] (f₁ : G →* α ≃o α) (f₂ : G →* α ≃o α) (g : G) :
Function.Semiconj (fun x => g', (RelIso.toRelEmbedding (f₁ g')⁻¹).toEmbedding ((RelIso.toRelEmbedding (f₂ g')).toEmbedding x)) (RelIso.toRelEmbedding (f₂ g)).toEmbedding (RelIso.toRelEmbedding (f₁ g)).toEmbedding

Consider two actions f₁ f₂ : G → α → α→ α → α→ α of a group on a complete lattice by order isomorphisms. Then the map x ↦ ⨆ g : G, (f₁ g)⁻¹ (f₂ g x)↦ ⨆ g : G, (f₁ g)⁻¹ (f₂ g x)⨆ g : G, (f₁ g)⁻¹ (f₂ g x)⁻¹ (f₂ g x) semiconjugates each f₁ g' to f₂ g'.

This is a version of Proposition 5.4 from [Étienne Ghys, Groupes d'homéomorphismes du cercle et cohomologie bornée][ghys87:groupes].

theorem Function.csupₛ_div_semiconj {α : Type u_1} {G : Type u_2} [inst : ConditionallyCompleteLattice α] [inst : Group G] (f₁ : G →* α ≃o α) (f₂ : G →* α ≃o α) (hbdd : ∀ (x : α), BddAbove (Set.range fun g => (RelIso.toRelEmbedding (f₁ g)⁻¹).toEmbedding ((RelIso.toRelEmbedding (f₂ g)).toEmbedding x))) (g : G) :
Function.Semiconj (fun x => g', (RelIso.toRelEmbedding (f₁ g')⁻¹).toEmbedding ((RelIso.toRelEmbedding (f₂ g')).toEmbedding x)) (RelIso.toRelEmbedding (f₂ g)).toEmbedding (RelIso.toRelEmbedding (f₁ g)).toEmbedding

Consider two actions f₁ f₂ : G → α → α→ α → α→ α of a group on a conditionally complete lattice by order isomorphisms. Suppose that each set $s(x)={f_1(g)^{-1} (f_2(g)(x)) | g \in G}$ is bounded above. Then the map x ↦ supₛ s(x)↦ supₛ s(x) semiconjugates each f₁ g' to f₂ g'.

This is a version of Proposition 5.4 from [Étienne Ghys, Groupes d'homéomorphismes du cercle et cohomologie bornée][ghys87:groupes].