Galois connections, insertions and coinsertions #
Galois connections are order theoretic adjoints, i.e. a pair of functions u
and l
,
such that ∀ a b, l a ≤ b ↔ a ≤ u b
.
Main definitions #
GaloisConnection
: A Galois connection is a pair of functionsl
andu
satisfyingl a ≤ b ↔ a ≤ u b
. They are special cases of adjoint functors in category theory, but do not depend on the category theory library in mathlib.GaloisInsertion
: A Galois insertion is a Galois connection wherel ∘ u = id
GaloisCoinsertion
: A Galois coinsertion is a Galois connection whereu ∘ l = id
A Galois connection is a pair of functions l
and u
satisfying
l a ≤ b ↔ a ≤ u b
. They are special cases of adjoint functors in category theory,
but do not depend on the category theory library in mathlib.
Equations
- GaloisConnection l u = ∀ (a : α) (b : β), l a ≤ b ↔ a ≤ u b
Instances For
Makes a Galois connection from an order-preserving bijection.
If (l, u)
is a Galois connection, then the relation x ≤ u (l y)
is a transitive relation.
If l
is a closure operator (Submodule.span
, Subgroup.closure
, ...) and u
is the coercion to
Set
, this reads as "if U
is in the closure of V
and V
is in the closure of W
then U
is
in the closure of W
".
If there exists a b
such that a = u a
, then b = l a
is one such element.
If there exists an a
such that b = l a
, then a = u b
is one such element.
A Galois insertion is a Galois connection where l ∘ u = id
. It also contains a constructive
choice function, to give better definitional equalities when lifting order structures. Dual
to GaloisCoinsertion
- choice (x : α) : u (l x) ≤ x → β
A contructive choice function for images of
l
. - gc : GaloisConnection l u
The Galois connection associated to a Galois insertion.
- le_l_u (x : β) : x ≤ l (u x)
Main property of a Galois insertion.
Property of the choice function.
Instances For
Makes a Galois insertion from an order-preserving bijection.
Equations
Instances For
A constructor for a Galois insertion with the trivial choice
function.
Equations
- GaloisInsertion.monotoneIntro hu hl hul hlu = { choice := fun (x : α) (x_1 : u (l x) ≤ x) => l x, gc := ⋯, le_l_u := ⋯, choice_eq := ⋯ }
Instances For
Make a GaloisInsertion l u
from a GaloisConnection l u
such that ∀ b, b ≤ l (u b)
Equations
Instances For
Lift the bottom along a Galois connection
Equations
- gc.liftOrderBot = OrderBot.mk ⋯
Instances For
A Galois coinsertion is a Galois connection where u ∘ l = id
. It also contains a constructive
choice function, to give better definitional equalities when lifting order structures. Dual to
GaloisInsertion
- choice (x : β) : x ≤ l (u x) → α
A contructive choice function for images of
u
. - gc : GaloisConnection l u
The Galois connection associated to a Galois coinsertion.
- u_l_le (x : α) : u (l x) ≤ x
Main property of a Galois coinsertion.
Property of the choice function.
Instances For
Makes a Galois coinsertion from an order-preserving bijection.
Equations
Instances For
Make a GaloisInsertion
between αᵒᵈ
and βᵒᵈ
from a GaloisCoinsertion
between α
and
β
.
Equations
- x.dual = { choice := x.choice, gc := ⋯, le_l_u := ⋯, choice_eq := ⋯ }
Instances For
Make a GaloisCoinsertion
between αᵒᵈ
and βᵒᵈ
from a GaloisInsertion
between α
and
β
.
Equations
- x.dual = { choice := x.choice, gc := ⋯, u_l_le := ⋯, choice_eq := ⋯ }
Instances For
Make a GaloisInsertion
between α
and β
from a GaloisCoinsertion
between αᵒᵈ
and
βᵒᵈ
.
Equations
- x.ofDual = { choice := x.choice, gc := ⋯, le_l_u := ⋯, choice_eq := ⋯ }
Instances For
Make a GaloisCoinsertion
between α
and β
from a GaloisInsertion
between αᵒᵈ
and
βᵒᵈ
.
Equations
- x.ofDual = { choice := x.choice, gc := ⋯, u_l_le := ⋯, choice_eq := ⋯ }
Instances For
A constructor for a Galois coinsertion with the trivial choice
function.
Equations
- GaloisCoinsertion.monotoneIntro hu hl hlu hul = (GaloisInsertion.monotoneIntro ⋯ ⋯ hlu hul).ofDual
Instances For
Make a GaloisCoinsertion l u
from a GaloisConnection l u
such that ∀ a, u (l a) ≤ a
Equations
Instances For
Lift the top along a Galois connection
Equations
- gc.liftOrderTop = OrderTop.mk ⋯