mathlib documentation

ring_theory.free_comm_ring

Free commutative rings

The theory of the free commutative ring generated by a type α. It is isomorphic to the polynomial ring over ℤ with variables in α

Main definitions

Main results

free_comm_ring has functorial properties (it is an adjoint to the forgetful functor). In this file we have:

Implementation notes

free_comm_ring α is implemented not using mv_polynomial but directly as the free abelian group on multiset α, the type of monomials in this free commutative ring.

Tags

free commutative ring, free ring

def free_comm_ring (α : Type u) :
Type u

free_comm_ring α is the free commutative ring on the type α.

Equations
@[instance]

The structure of a commutative ring on free_comm_ring α.

Equations
@[instance]

Equations
def free_comm_ring.of {α : Type u} (x : α) :

The canonical map from α to the free commutative ring on α.

Equations
theorem free_comm_ring.induction_on {α : Type u} {C : free_comm_ring α → Prop} (z : free_comm_ring α) (hn1 : C (-1)) (hb : ∀ (b : α), C (free_comm_ring.of b)) (ha : ∀ (x y : free_comm_ring α), C xC yC (x + y)) (hm : ∀ (x y : free_comm_ring α), C xC yC (x * y)) :
C z

def free_comm_ring.lift {α : Type u} {R : Type v} [comm_ring R] (f : α → R) :

Lift a map α → R to a additive group homomorphism free_comm_ring α → R. For a version producing a bundled homomorphism, see lift_hom.

Equations
@[simp]
theorem free_comm_ring.lift_of {α : Type u} {R : Type v} [comm_ring R] (f : α → R) (x : α) :

@[simp]
theorem free_comm_ring.lift_comp_of {α : Type u} {R : Type v} [comm_ring R] (f : free_comm_ring α →+* R) :

def free_comm_ring.map {α : Type u} {β : Type v} (f : α → β) :

A map f : α → β produces a ring homomorphism free_comm_ring α →+* free_comm_ring β.

Equations
@[simp]
theorem free_comm_ring.map_of {α : Type u} {β : Type v} (f : α → β) (x : α) :

def free_comm_ring.is_supported {α : Type u} (x : free_comm_ring α) (s : set α) :
Prop

is_supported x s means that all monomials showing up in x have variables in s.

Equations
theorem free_comm_ring.is_supported_upwards {α : Type u} {x : free_comm_ring α} {s t : set α} (hs : x.is_supported s) (hst : s t) :

theorem free_comm_ring.is_supported_add {α : Type u} {x y : free_comm_ring α} {s : set α} (hxs : x.is_supported s) (hys : y.is_supported s) :
(x + y).is_supported s

theorem free_comm_ring.is_supported_neg {α : Type u} {x : free_comm_ring α} {s : set α} (hxs : x.is_supported s) :

theorem free_comm_ring.is_supported_sub {α : Type u} {x y : free_comm_ring α} {s : set α} (hxs : x.is_supported s) (hys : y.is_supported s) :
(x - y).is_supported s

theorem free_comm_ring.is_supported_mul {α : Type u} {x y : free_comm_ring α} {s : set α} (hxs : x.is_supported s) (hys : y.is_supported s) :
(x * y).is_supported s

theorem free_comm_ring.is_supported_zero {α : Type u} {s : set α} :

theorem free_comm_ring.is_supported_one {α : Type u} {s : set α} :

theorem free_comm_ring.is_supported_int {α : Type u} {i : } {s : set α} :

The restriction map from free_comm_ring α to free_comm_ring s where s : set α, defined by sending all variables not in s to zero.

Equations
@[simp]
theorem free_comm_ring.restriction_of {α : Type u} (s : set α) [decidable_pred s] (p : α) :
(free_comm_ring.restriction s) (free_comm_ring.of p) = dite (p s) (λ (H : p s), free_comm_ring.of p, H⟩) (λ (H : p s), 0)

theorem free_comm_ring.is_supported_of {α : Type u} {p : α} {s : set α} :

theorem free_comm_ring.exists_finite_support {α : Type u} (x : free_comm_ring α) :
∃ (s : set α), s.finite x.is_supported s

theorem free_comm_ring.exists_finset_support {α : Type u} (x : free_comm_ring α) :
∃ (s : finset α), x.is_supported s

The canonical ring homomorphism from the free ring generated by α to the free commutative ring generated by α.

Equations
@[instance]
def free_ring.has_coe (α : Type u) :

Equations
@[instance]

@[simp]
theorem free_ring.coe_zero (α : Type u) :
0 = 0

@[simp]
theorem free_ring.coe_one (α : Type u) :
1 = 1

@[simp]
theorem free_ring.coe_of {α : Type u} (a : α) :

@[simp]
theorem free_ring.coe_neg {α : Type u} (x : free_ring α) :

@[simp]
theorem free_ring.coe_add {α : Type u} (x y : free_ring α) :
(x + y) = x + y

@[simp]
theorem free_ring.coe_sub {α : Type u} (x y : free_ring α) :
(x - y) = x - y

@[simp]
theorem free_ring.coe_mul {α : Type u} (x y : free_ring α) :
x * y = (x) * y

theorem free_ring.coe_eq (α : Type u) :
coe = functor.map (λ (l : list α), l)

def free_ring.of' {R : Type u_1} {S : Type u_2} [ring R] [ring S] (e : R S) [is_ring_hom e] :
R ≃+* S

Interpret an equivalence f : R ≃ S as a ring equivalence R ≃+* S.

Equations

If α has size at most 1 then the natural map from the free ring on α to the free commutative ring on α is an isomorphism of rings.

Equations

The free commutative ring on α is isomorphic to the polynomial ring over ℤ with variables in α

Equations