Skip to content

Commit

Permalink
chore(data/polynomial/eval): add map_ring_hom_{id,comp} lemmas (#7019)
Browse files Browse the repository at this point in the history
  • Loading branch information
eric-wieser committed Apr 3, 2021
1 parent 76a3b82 commit fc87598
Showing 1 changed file with 7 additions and 0 deletions.
7 changes: 7 additions & 0 deletions src/data/polynomial/eval.lean
Expand Up @@ -560,6 +560,13 @@ def map_ring_hom (f : R →+* S) : polynomial R →+* polynomial S :=

@[simp] lemma coe_map_ring_hom (f : R →+* S) : ⇑(map_ring_hom f) = map f := rfl

@[simp] lemma map_ring_hom_id : map_ring_hom (ring_hom.id R) = ring_hom.id (polynomial R) :=
ring_hom.ext $ λ x, map_id

@[simp] lemma map_ring_hom_comp [semiring T] (f : S →+* T) (g : R →+* S) :
(map_ring_hom f).comp (map_ring_hom g) = map_ring_hom (f.comp g) :=
ring_hom.ext $ map_map g f

lemma map_list_prod (L : list (polynomial R)) : L.prod.map f = (L.map $ map f).prod :=
eq.symm $ list.prod_hom _ (monoid_hom.of (map f))

Expand Down

0 comments on commit fc87598

Please sign in to comment.