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
free_comm_ring α: the free commutative ring on a type αlift_hom (f : α → R): the ring homfree_comm_ring α →+* Rinduced by functoriality fromf.map (f : α → β): the ring homfree_comm_ring α →*+ free_comm_ring βinduced by functoriality from f.
Main results
free_comm_ring has functorial properties (it is an adjoint to the forgetful functor).
In this file we have:
of : α → free_comm_ring αlift_hom (f : α → R) : free_comm_ring α →+* Rmap (f : α → β) : free_comm_ring α →+* free_comm_ring βfree_comm_ring_equiv_mv_polynomial_int : free_comm_ring α ≃+* mv_polynomial α ℤ:free_comm_ring αis isomorphic to a polynomial ring.
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
free_comm_ring α is the free commutative ring on the type α.
Equations
Equations
- free_comm_ring.inhabited α = {default := 0}
The canonical map from α to the free commutative ring on α.
Equations
Lift a map α → R to a additive group homomorphism free_comm_ring α → R.
For a version producing a bundled homomorphism, see lift_hom.
Equations
- free_comm_ring.lift f = {to_fun := (free_abelian_group.lift (λ (s : multiplicative (multiset α)), (multiset.map f s.to_add).prod)).to_fun, map_one' := _, map_mul' := _, map_zero' := _, map_add' := _}
A map f : α → β produces a ring homomorphism free_comm_ring α →+* free_comm_ring β.
Equations
is_supported x s means that all monomials showing up in x have variables in s.
Equations
- x.is_supported s = (x ∈ ring.closure (free_comm_ring.of '' s))
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
- free_comm_ring.restriction s = free_comm_ring.lift (λ (p : α), dite (p ∈ s) (λ (H : p ∈ s), free_comm_ring.of ⟨p, H⟩) (λ (H : p ∉ s), 0))
The canonical ring homomorphism from the free ring generated by α to the free commutative ring
generated by α.
Equations
Interpret an equivalence f : R ≃ S as a ring equivalence R ≃+* S.
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
- free_ring.comm_ring α = {add := ring.add (free_ring.ring α), add_assoc := _, zero := ring.zero (free_ring.ring α), zero_add := _, add_zero := _, neg := ring.neg (free_ring.ring α), add_left_neg := _, add_comm := _, mul := ring.mul (free_ring.ring α), mul_assoc := _, one := ring.one (free_ring.ring α), one_mul := _, mul_one := _, left_distrib := _, right_distrib := _, mul_comm := _}
The free commutative ring on α is isomorphic to the polynomial ring over ℤ with
variables in α
Equations
- free_comm_ring_equiv_mv_polynomial_int α = {to_fun := ⇑(free_comm_ring.lift (λ (a : α), mv_polynomial.X a)), inv_fun := mv_polynomial.eval₂ (int.cast_ring_hom (free_comm_ring α)) free_comm_ring.of, left_inv := _, right_inv := _, map_mul' := _, map_add' := _}
The free commutative ring on the empty type is isomorphic to ℤ.
The free commutative ring on a type with one term is isomorphic to ℤ[X].
The free ring on the empty type is isomorphic to ℤ.
The free ring on a type with one term is isomorphic to ℤ[X].