mathlib documentation

order.directed

Directed indexed families and sets #

This file defines directed indexed families and directed sets. An indexed family/set is directed iff each pair of elements has a shared upper bound.

Main declarations #

def directed {α : Type u} {ι : Sort w} (r : α → α → Prop) (f : ι → α) :
Prop

A family of elements of α is directed (with respect to a relation on α) if there is a member of the family -above any pair in the family.

Equations
  • directed r f = ∀ (x y : ι), ∃ (z : ι), r (f x) (f z) r (f y) (f z)
def directed_on {α : Type u} (r : α → α → Prop) (s : set α) :
Prop

A subset of α is directed if there is an element of the set -above any pair of elements in the set.

Equations
theorem directed_on_iff_directed {α : Type u} {r : α → α → Prop} {s : set α} :
theorem directed_on.directed_coe {α : Type u} {r : α → α → Prop} {s : set α} :

Alias of the forward direction of directed_on_iff_directed.

theorem directed_on_image {α : Type u} {β : Type v} {r : α → α → Prop} {s : set β} {f : β → α} :
theorem directed_on.mono {α : Type u} {r : α → α → Prop} {s : set α} (h : directed_on r s) {r' : α → α → Prop} (H : ∀ {a b : α}, r a br' a b) :
theorem directed_comp {α : Type u} {β : Type v} {r : α → α → Prop} {ι : Sort u_1} {f : ι → β} {g : β → α} :
directed r (g f) directed (g ⁻¹'o r) f
theorem directed.mono {α : Type u} {r s : α → α → Prop} {ι : Sort u_1} {f : ι → α} (H : ∀ (a b : α), r a bs a b) (h : directed r f) :
theorem directed.mono_comp {α : Type u} {β : Type v} (r : α → α → Prop) {ι : Sort u_1} {rb : β → β → Prop} {g : α → β} {f : ι → α} (hg : ∀ ⦃x y : α⦄, r x yrb (g x) (g y)) (hf : directed r f) :
directed rb (g f)
theorem directed_of_sup {α : Type u} {β : Type v} [semilattice_sup α] {f : α → β} {r : β → β → Prop} (H : ∀ ⦃i j : α⦄, i jr (f i) (f j)) :

A monotone function on a sup-semilattice is directed.

theorem monotone.directed_le {α : Type u} {β : Type v} [semilattice_sup α] [preorder β] {f : α → β} :
theorem directed_on_of_sup_mem {α : Type u} [semilattice_sup α] {S : set α} (H : ∀ ⦃i j : α⦄, i Sj Si j S) :

A set stable by supremum is -directed.

theorem directed.extend_bot {α : Type u} {β : Type v} {ι : Sort w} [preorder α] [order_bot α] {e : ι → β} {f : ι → α} (hf : directed has_le.le f) (he : function.injective e) :
theorem directed_of_inf {α : Type u} {β : Type v} [semilattice_inf α] {r : β → β → Prop} {f : α → β} (hf : ∀ (a₁ a₂ : α), a₁ a₂r (f a₂) (f a₁)) :

An antitone function on an inf-semilattice is directed.

theorem directed_on_of_inf_mem {α : Type u} [semilattice_inf α] {S : set α} (H : ∀ ⦃i j : α⦄, i Sj Si j S) :

A set stable by infimum is -directed.

@[class]
structure is_directed (α : Type u_1) (r : α → α → Prop) :
Prop
  • directed : ∀ (a b : α), ∃ (c : α), r a c r b c

is_directed α r states that for any elements a, b there exists an element c such that r a c and r b c.

Instances of this typeclass
theorem directed_of {α : Type u} (r : α → α → Prop) [is_directed α r] (a b : α) :
∃ (c : α), r a c r b c
theorem directed_id {α : Type u} {r : α → α → Prop} [is_directed α r] :
theorem directed_id_iff {α : Type u} {r : α → α → Prop} :
theorem directed_on_univ {α : Type u} {r : α → α → Prop} [is_directed α r] :
theorem directed_on_univ_iff {α : Type u} {r : α → α → Prop} :
@[protected, instance]
def is_total.to_is_directed {α : Type u} {r : α → α → Prop} [is_total α r] :
theorem is_directed_mono {α : Type u} {r : α → α → Prop} (s : α → α → Prop) [is_directed α r] (h : ∀ ⦃a b : α⦄, r a bs a b) :
theorem exists_ge_ge {α : Type u} [has_le α] [is_directed α has_le.le] (a b : α) :
∃ (c : α), a c b c
theorem exists_le_le {α : Type u} [has_le α] [is_directed α ge] (a b : α) :
∃ (c : α), c a c b
@[protected, instance]
@[protected, instance]
@[protected]
theorem is_min.is_bot {α : Type u} [preorder α] {a : α} [is_directed α ge] (h : is_min a) :
@[protected]
theorem is_max.is_top {α : Type u} [preorder α] {a : α} [is_directed α has_le.le] (h : is_max a) :
theorem is_top_or_exists_gt {α : Type u} [preorder α] [is_directed α has_le.le] (a : α) :
is_top a ∃ (b : α), a < b
theorem is_bot_or_exists_lt {α : Type u} [preorder α] [is_directed α ge] (a : α) :
is_bot a ∃ (b : α), b < a
theorem is_bot_iff_is_min {α : Type u} [preorder α] {a : α} [is_directed α ge] :
theorem is_top_iff_is_max {α : Type u} [preorder α] {a : α} [is_directed α has_le.le] :
theorem exists_lt_of_directed_ge (β : Type v) [partial_order β] [is_directed β ge] [nontrivial β] :
∃ (a b : β), a < b
theorem exists_lt_of_directed_le (β : Type v) [partial_order β] [is_directed β has_le.le] [nontrivial β] :
∃ (a b : β), a < b
@[protected, instance]
@[protected, instance]
@[protected, instance]
@[protected, instance]
def order_bot.to_is_directed_ge {α : Type u} [has_le α] [order_bot α] :