Conjugacy of group elements #
THIS FILE IS SYNCHRONIZED WITH MATHLIB4. Any changes to this file require a corresponding PR to mathlib4.
See also mul_aut.conj and quandle.conj.
The quotient type of conjugacy classes of a group.
Equations
- conj_classes α = quotient (is_conj.setoid α)
Instances for conj_classes
The canonical quotient map from a monoid α into the conj_classes of α
Equations
- conj_classes.mk a = ⟦a⟧
A monoid_hom maps conjugacy classes of one group to conjugacy classes of another.
Equations
- conj_classes.map f = quotient.lift (conj_classes.mk ∘ ⇑f) _
Certain instances trigger further searches when they are considered as candidate instances; these instances should be assigned a priority lower than the default of 1000 (for example, 900).
The conditions for this rule are as follows:
- a class
Chas instancesinstT : C TandinstT' : C T' - types
TandT'are both specializations of another typeS - the parameters supplied to
Sto produceTare not (fully) determined byinstT, instead they have to be found by instance search If those conditions hold, the instanceinstTshould be assigned lower priority.
For example, suppose the search for an instance of decidable_eq (multiset α) tries the
candidate instance con.quotient.decidable_eq (c : con M) : decidable_eq c.quotient.
Since multiset and con.quotient are both quotient types, unification will check
that the relations list.perm and c.to_setoid.r unify. However, c.to_setoid depends on
a has_mul M instance, so this unification triggers a search for has_mul (list α);
this will traverse all subclasses of has_mul before failing.
On the other hand, the search for an instance of decidable_eq (con.quotient c) for c : con M
can quickly reject the candidate instance multiset.has_decidable_eq because the type of
list.perm : list ?m_1 → list ?m_1 → Prop does not unify with M → M → Prop.
Therefore, we should assign con.quotient.decidable_eq a lower priority because it fails slowly.
(In terms of the rules above, C := decidable_eq, T := con.quotient,
instT := con.quotient.decidable_eq, T' := multiset, instT' := multiset.has_decidable_eq,
and S := quot.)
If the type involved is a free variable (rather than an instantiation of some type S),
the instance priority should be even lower, see Note [lower instance priority].
Equations
The bijection between a comm_group and its conj_classes.
Equations
- conj_classes.mk_equiv = {to_fun := conj_classes.mk (comm_monoid.to_monoid α), inv_fun := quotient.lift id conj_classes.mk_equiv._proof_1, left_inv := _, right_inv := _}
Given an element a, conjugates a is the set of conjugates.
Equations
- conjugates_of a = {b : α | is_conj a b}
Instances for ↥conjugates_of
Given a conjugacy class a, carrier a is the set it represents.
Equations
- conj_classes.carrier = quotient.lift conjugates_of conj_classes.carrier._proof_1