mathlib3 documentation

algebra.group.with_one.defs

Adjoining a zero/one to semigroups and related algebraic structures #

THIS FILE IS SYNCHRONIZED WITH MATHLIB4. Any changes to this file require a corresponding PR to mathlib4.

This file contains different results about adjoining an element to an algebraic structure which then behaves like a zero or a one. An example is adjoining a one to a semigroup to obtain a monoid. That this provides an example of an adjunction is proved in algebra.category.Mon.adjunctions.

Another result says that adjoining to a group an element zero gives a group_with_zero. For more information about these structures (which are not that standard in informal mathematics, see algebra.group_with_zero.basic)

Implementation notes #

At various points in this file, id $ is used in at the start of a proof field in a structure. This ensures that the generated _proof_1 lemmas are stated in terms of the algebraic operations and not option.map, as the latter does not typecheck once with_zero/with_one is irreducible.

@[irreducible]
def with_zero (α : Type u_1) :
Type u_1

Add an extra element 0 to a type

Equations
Instances for with_zero
@[protected, instance]
Equations
@[protected, instance]
def with_one.has_repr {α : Type u} [has_repr α] :
Equations
@[protected, instance]
def with_zero.has_repr {α : Type u} [has_repr α] :
Equations
@[protected, instance]
Equations
@[protected, instance]
Equations
@[protected, instance]
def with_one.has_one {α : Type u} :
Equations
@[protected, instance]
def with_zero.has_zero {α : Type u} :
Equations
@[protected, instance]
def with_zero.has_add {α : Type u} [has_add α] :
Equations
@[protected, instance]
def with_one.has_mul {α : Type u} [has_mul α] :
Equations
@[protected, instance]
def with_one.has_inv {α : Type u} [has_inv α] :
Equations
@[protected, instance]
def with_zero.has_neg {α : Type u} [has_neg α] :
Equations
@[protected, instance]
Equations
@[protected, instance]
Equations
@[protected, instance]
def with_one.inhabited {α : Type u} :
Equations
@[protected, instance]
Equations
@[protected, instance]
@[protected, instance]
@[protected, instance]
def with_zero.has_coe_t {α : Type u} :
Equations
@[protected, instance]
def with_one.has_coe_t {α : Type u} :
Equations
def with_zero.rec_zero_coe {α : Type u} {C : with_zero α Sort u_1} (h₁ : C 0) (h₂ : Π (a : α), C a) (n : with_zero α) :
C n

Recursor for with_zero using the preferred forms 0 and ↑a.

Equations
def with_one.rec_one_coe {α : Type u} {C : with_one α Sort u_1} (h₁ : C 1) (h₂ : Π (a : α), C a) (n : with_one α) :
C n

Recursor for with_one using the preferred forms 1 and ↑a.

Equations
def with_zero.unzero {α : Type u} {x : with_zero α} (hx : x 0) :
α

Deconstruct a x : with_zero α to the underlying value in α, given a proof that x ≠ 0.

Equations
def with_one.unone {α : Type u} {x : with_one α} (hx : x 1) :
α

Deconstruct a x : with_one α to the underlying value in α, given a proof that x ≠ 1.

Equations
@[simp]
theorem with_zero.unzero_coe {α : Type u} {x : α} (hx : x 0) :
@[simp]
theorem with_one.unone_coe {α : Type u} {x : α} (hx : x 1) :
@[simp]
theorem with_zero.coe_unzero {α : Type u} {x : with_zero α} (hx : x 0) :
@[simp]
theorem with_one.coe_unone {α : Type u} {x : with_one α} (hx : x 1) :
theorem with_one.some_eq_coe {α : Type u} {a : α} :
theorem with_zero.some_eq_coe {α : Type u} {a : α} :
@[simp]
theorem with_zero.coe_ne_zero {α : Type u} {a : α} :
a 0
@[simp]
theorem with_one.coe_ne_one {α : Type u} {a : α} :
a 1
@[simp]
theorem with_one.one_ne_coe {α : Type u} {a : α} :
1 a
@[simp]
theorem with_zero.zero_ne_coe {α : Type u} {a : α} :
0 a
theorem with_one.ne_one_iff_exists {α : Type u} {x : with_one α} :
x 1 (a : α), a = x
theorem with_zero.ne_zero_iff_exists {α : Type u} {x : with_zero α} :
x 0 (a : α), a = x
@[protected, instance]
def with_zero.can_lift {α : Type u} :
can_lift (with_zero α) α coe (λ (a : with_zero α), a 0)
Equations
@[protected, instance]
def with_one.can_lift {α : Type u} :
can_lift (with_one α) α coe (λ (a : with_one α), a 1)
Equations
@[simp, norm_cast]
theorem with_one.coe_inj {α : Type u} {a b : α} :
a = b a = b
@[simp, norm_cast]
theorem with_zero.coe_inj {α : Type u} {a b : α} :
a = b a = b
@[protected]
theorem with_one.cases_on {α : Type u} {P : with_one α Prop} (x : with_one α) :
P 1 ( (a : α), P a) P x
@[protected]
theorem with_zero.cases_on {α : Type u} {P : with_zero α Prop} (x : with_zero α) :
P 0 ( (a : α), P a) P x
@[protected, instance]
Equations
@[protected, instance]
Equations
@[simp, norm_cast]
theorem with_one.coe_mul {α : Type u} [has_mul α] (a b : α) :
(a * b) = a * b
@[simp, norm_cast]
theorem with_zero.coe_add {α : Type u} [has_add α] (a b : α) :
(a + b) = a + b
@[simp, norm_cast]
theorem with_one.coe_inv {α : Type u} [has_inv α] (a : α) :
@[simp, norm_cast]
theorem with_zero.coe_neg {α : Type u} [has_neg α] (a : α) :
@[protected, instance]
def with_zero.has_one {α : Type u} [one : has_one α] :
Equations
@[simp, norm_cast]
theorem with_zero.coe_one {α : Type u} [has_one α] :
1 = 1
@[protected, instance]
Equations
@[simp, norm_cast]
theorem with_zero.coe_mul {α : Type u} [has_mul α] {a b : α} :
(a * b) = a * b
@[protected, instance]
@[protected, instance]
Equations
@[simp, norm_cast]
theorem with_zero.coe_pow {α : Type u} [has_one α] [has_pow α ] {a : α} (n : ) :
(a ^ n) = a ^ n
@[protected, instance]
def with_zero.has_inv {α : Type u} [has_inv α] :

Given an inverse operation on α there is an inverse operation on with_zero α sending 0 to 0

Equations
@[simp, norm_cast]
theorem with_zero.coe_inv {α : Type u} [has_inv α] (a : α) :
@[simp]
theorem with_zero.inv_zero {α : Type u} [has_inv α] :
0⁻¹ = 0
@[protected, instance]
def with_zero.has_div {α : Type u} [has_div α] :
Equations
@[norm_cast]
theorem with_zero.coe_div {α : Type u} [has_div α] (a b : α) :
(a / b) = a / b
@[protected, instance]
Equations
@[simp, norm_cast]
theorem with_zero.coe_zpow {α : Type u} [div_inv_monoid α] {a : α} (n : ) :
(a ^ n) = a ^ n