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
galois_connection
: 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.galois_insertion
: A Galois insertion is a Galois connection wherel ∘ u = id
galois_coinsertion
: A Galois coinsertion is a Galois connection whereu ∘ l = id
Implementation details
Galois insertions can be used to lift order structures from one type to another.
For example if α
is a complete lattice, and l : α → β
, and u : β → α
form
a Galois insertion, then β
is also a complete lattice. l
is the lower adjoint and
u
is the upper adjoint.
An example of this is the Galois insertion is in group thery. If G
is a topological space,
then there is a Galois insertion between the set of subsets of G
, set G
, and the
set of subgroups of G
, subgroup G
. The left adjoint is subgroup.closure
,
taking the subgroup
generated by a set
, The right adjoint is the coercion from subgroup G
to
set G
, taking the underlying set of an subgroup.
Naively lifting a lattice structure along this Galois insertion would mean that the definition
of inf
on subgroups would be subgroup.closure (↑S ∩ ↑T)
. This is an undesirable definition
because the intersection of subgroups is already a subgroup, so there is no need to take the
closure. For this reason a choice
function is added as a field to the galois_insertion
structure. It has type Π S : set G, ↑(closure S) ≤ S → subgroup G
. When ↑(closure S) ≤ S
, then
S
is already a subgroup, so this function can be defined using subgroup.mk
and not closure
.
This means the infimum of subgroups will be defined to be the intersection of sets, paired
with a proof that intersection of subgroups is a subgroup, rather than the closure of the
intersection.
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
- galois_connection l u = ∀ (a : α) (b : β), l a ≤ b ↔ a ≤ u b
Makes a Galois connection from an order-preserving bijection.
- choice : Π (x : α), u (l x) ≤ x → β
- gc : galois_connection l u
- le_l_u : ∀ (x : β), x ≤ l (u x)
- choice_eq : ∀ (a : α) (h : u (l a) ≤ a), c.choice a h = l a
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 galois_coinsertion
A constructor for a Galois insertion with the trivial choice
function.
Makes a Galois insertion from an order-preserving bijection.
Make a galois_insertion l u
from a galois_connection l u
such that ∀ b, b ≤ l (u b)
Lift the bottom along a Galois connection
Equations
- gc.lift_order_bot = {bot := l ⊥, le := partial_order.le _inst_2, lt := partial_order.lt _inst_2, le_refl := _, le_trans := _, lt_iff_le_not_le := _, le_antisymm := _, bot_le := _}
Lift the suprema along a Galois insertion
Equations
- gi.lift_semilattice_sup = {sup := λ (a b : β), l (u a ⊔ u b), le := partial_order.le _inst_1, lt := partial_order.lt _inst_1, le_refl := _, le_trans := _, lt_iff_le_not_le := _, le_antisymm := _, le_sup_left := _, le_sup_right := _, sup_le := _}
Lift the infima along a Galois insertion
Equations
- gi.lift_semilattice_inf = {inf := λ (a b : β), gi.choice (u a ⊓ u b) _, le := partial_order.le _inst_1, lt := partial_order.lt _inst_1, le_refl := _, le_trans := _, lt_iff_le_not_le := _, le_antisymm := _, inf_le_left := _, inf_le_right := _, le_inf := _}
Lift the suprema and infima along a Galois insertion
Equations
- gi.lift_lattice = {sup := semilattice_sup.sup gi.lift_semilattice_sup, le := semilattice_sup.le gi.lift_semilattice_sup, lt := semilattice_sup.lt gi.lift_semilattice_sup, le_refl := _, le_trans := _, lt_iff_le_not_le := _, le_antisymm := _, le_sup_left := _, le_sup_right := _, sup_le := _, inf := semilattice_inf.inf gi.lift_semilattice_inf, inf_le_left := _, inf_le_right := _, le_inf := _}
Lift the top along a Galois insertion
Equations
- gi.lift_order_top = {top := gi.choice ⊤ galois_insertion.lift_order_top._proof_1, le := partial_order.le _inst_1, lt := partial_order.lt _inst_1, le_refl := _, le_trans := _, lt_iff_le_not_le := _, le_antisymm := _, le_top := _}
Lift the top, bottom, suprema, and infima along a Galois insertion
Equations
- gi.lift_bounded_lattice = {sup := lattice.sup gi.lift_lattice, le := lattice.le gi.lift_lattice, lt := lattice.lt gi.lift_lattice, le_refl := _, le_trans := _, lt_iff_le_not_le := _, le_antisymm := _, le_sup_left := _, le_sup_right := _, sup_le := _, inf := lattice.inf gi.lift_lattice, inf_le_left := _, inf_le_right := _, le_inf := _, top := order_top.top gi.lift_order_top, le_top := _, bot := order_bot.bot _.lift_order_bot, bot_le := _}
Lift all suprema and infima along a Galois insertion
Equations
- gi.lift_complete_lattice = {sup := bounded_lattice.sup gi.lift_bounded_lattice, le := bounded_lattice.le gi.lift_bounded_lattice, lt := bounded_lattice.lt gi.lift_bounded_lattice, le_refl := _, le_trans := _, lt_iff_le_not_le := _, le_antisymm := _, le_sup_left := _, le_sup_right := _, sup_le := _, inf := bounded_lattice.inf gi.lift_bounded_lattice, inf_le_left := _, inf_le_right := _, le_inf := _, top := bounded_lattice.top gi.lift_bounded_lattice, le_top := _, bot := bounded_lattice.bot gi.lift_bounded_lattice, bot_le := _, Sup := λ (s : set β), l (⨆ (b : β) (H : b ∈ s), u b), Inf := λ (s : set β), gi.choice (⨅ (b : β) (H : b ∈ s), u b) _, le_Sup := _, Sup_le := _, Inf_le := _, le_Inf := _}
- choice : Π (x : β), x ≤ l (u x) → α
- gc : galois_connection l u
- u_l_le : ∀ (x : α), u (l x) ≤ x
- choice_eq : ∀ (a : β) (h : a ≤ l (u a)), c.choice a h = u a
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
galois_insertion
Make a galois_insertion u l
in the order_dual
, from a galois_coinsertion l u
Equations
- galois_coinsertion.dual = λ (x : galois_coinsertion l u), {choice := x.choice, gc := _, le_l_u := _, choice_eq := _}
Make a galois_coinsertion u l
in the order_dual
, from a galois_insertion l u
Equations
- galois_insertion.dual = λ (x : galois_insertion l u), {choice := x.choice, gc := _, u_l_le := _, choice_eq := _}
Make a galois_coinsertion l u
from a galois_insertion l u
in the order_dual
Equations
- galois_coinsertion.of_dual = λ (x : galois_insertion u l), {choice := x.choice, gc := _, u_l_le := _, choice_eq := _}
Make a galois_insertion l u
from a galois_coinsertion l u
in the order_dual
Equations
- galois_insertion.of_dual = λ (x : galois_coinsertion u l), {choice := x.choice, gc := _, le_l_u := _, choice_eq := _}
Makes a Galois coinsertion from an order-preserving bijection.
Equations
- rel_iso.to_galois_coinsertion oi = {choice := λ (b : β) (h : b ≤ ⇑oi (⇑(rel_iso.symm oi) b)), ⇑(rel_iso.symm oi) b, gc := _, u_l_le := _, choice_eq := _}
A constructor for a Galois coinsertion with the trivial choice
function.
Equations
- galois_coinsertion.monotone_intro hu hl hlu hul = galois_coinsertion.of_dual (galois_insertion.monotone_intro _ _ hlu hul)
Make a galois_coinsertion l u
from a galois_connection l u
such that ∀ b, b ≤ l (u b)
Lift the top along a Galois connection
Equations
- gc.lift_order_top = {top := u ⊤, le := partial_order.le _inst_1, lt := partial_order.lt _inst_1, le_refl := _, le_trans := _, lt_iff_le_not_le := _, le_antisymm := _, le_top := _}
Lift the infima along a Galois coinsertion
Equations
- gi.lift_semilattice_inf = {inf := λ (a b : α), u (l a ⊓ l b), le := partial_order.le _inst_1, lt := partial_order.lt _inst_1, le_refl := _, le_trans := _, lt_iff_le_not_le := _, le_antisymm := _, inf_le_left := _, inf_le_right := _, le_inf := _}
Lift the suprema along a Galois coinsertion
Equations
- gi.lift_semilattice_sup = {sup := λ (a b : α), gi.choice (l a ⊔ l b) _, le := partial_order.le _inst_1, lt := partial_order.lt _inst_1, le_refl := _, le_trans := _, lt_iff_le_not_le := _, le_antisymm := _, le_sup_left := _, le_sup_right := _, sup_le := _}
Lift the suprema and infima along a Galois coinsertion
Equations
- gi.lift_lattice = {sup := semilattice_sup.sup gi.lift_semilattice_sup, le := semilattice_sup.le gi.lift_semilattice_sup, lt := semilattice_sup.lt gi.lift_semilattice_sup, le_refl := _, le_trans := _, lt_iff_le_not_le := _, le_antisymm := _, le_sup_left := _, le_sup_right := _, sup_le := _, inf := semilattice_inf.inf gi.lift_semilattice_inf, inf_le_left := _, inf_le_right := _, le_inf := _}
Lift the bot along a Galois coinsertion
Equations
- gi.lift_order_bot = {bot := gi.choice ⊥ galois_coinsertion.lift_order_bot._proof_1, le := partial_order.le _inst_1, lt := partial_order.lt _inst_1, le_refl := _, le_trans := _, lt_iff_le_not_le := _, le_antisymm := _, bot_le := _}
Lift the top, bottom, suprema, and infima along a Galois coinsertion
Equations
- gi.lift_bounded_lattice = {sup := lattice.sup gi.lift_lattice, le := lattice.le gi.lift_lattice, lt := lattice.lt gi.lift_lattice, le_refl := _, le_trans := _, lt_iff_le_not_le := _, le_antisymm := _, le_sup_left := _, le_sup_right := _, sup_le := _, inf := lattice.inf gi.lift_lattice, inf_le_left := _, inf_le_right := _, le_inf := _, top := order_top.top _.lift_order_top, le_top := _, bot := order_bot.bot gi.lift_order_bot, bot_le := _}
Lift all suprema and infima along a Galois coinsertion
Equations
- gi.lift_complete_lattice = {sup := bounded_lattice.sup gi.lift_bounded_lattice, le := bounded_lattice.le gi.lift_bounded_lattice, lt := bounded_lattice.lt gi.lift_bounded_lattice, le_refl := _, le_trans := _, lt_iff_le_not_le := _, le_antisymm := _, le_sup_left := _, le_sup_right := _, sup_le := _, inf := bounded_lattice.inf gi.lift_bounded_lattice, inf_le_left := _, inf_le_right := _, le_inf := _, top := bounded_lattice.top gi.lift_bounded_lattice, le_top := _, bot := bounded_lattice.bot gi.lift_bounded_lattice, bot_le := _, Sup := λ (s : set α), gi.choice (⨆ (a : α) (H : a ∈ s), l a) _, Inf := λ (s : set α), u (⨅ (a : α) (H : a ∈ s), l a), le_Sup := _, Sup_le := _, Inf_le := _, le_Inf := _}