mathlib documentation

data.​opposite

data.​opposite

def opposite  :
Sort uSort u

The type of objects of the opposite of α; used to defined opposite category/group/...

In order to avoid confusion between α and its opposite type, we set up the type of objects opposite α using the following pattern, which will be repeated later for the morphisms.

  1. Define opposite α := α.
  2. Define the isomorphisms op : α → opposite α, unop : opposite α → α.
  3. Make the definition opposite irreducible.

    This has the following consequences.

  4. opposite α and α are distinct types in the elaborator, so you must use op and unop explicitly to convert between them.

  5. Both unop (op X) = X and op (unop X) = X are definitional equalities. Notably, every object of the opposite category is definitionally of the form op X, which greatly simplifies the definition of the structure of the opposite category, for example.

    (If Lean supported definitional eta equality for records, we could achieve the same goals using a structure with one field.)

Equations
def opposite.​op {α : Sort u} :
α → αᵒᵖ

Equations
def opposite.​unop {α : Sort u} :
αᵒᵖ → α

Equations
@[simp]
theorem opposite.​op_inj_iff {α : Sort u} (x y : α) :

@[simp]
theorem opposite.​unop_inj_iff {α : Sort u} (x y : αᵒᵖ) :

@[simp]
theorem opposite.​op_unop {α : Sort u} (x : αᵒᵖ) :

@[simp]
theorem opposite.​unop_op {α : Sort u} (x : α) :

def opposite.​equiv_to_opposite {α : Sort u} :
α αᵒᵖ

The type-level equivalence between a type and its opposite.

Equations
theorem opposite.​op_eq_iff_eq_unop {α : Sort u} {x : α} {y : αᵒᵖ} :

theorem opposite.​unop_eq_iff_eq_op {α : Sort u} {x : αᵒᵖ} {y : α} :

def opposite.​op_induction {α : Sort u} {F : αᵒᵖSort v} (h : Π (X : α), F (opposite.op X)) (X : αᵒᵖ) :
F X

Equations