Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit 0d37fd6

Browse files
committed
feat(data/polynomial/algebra_map): aeval_tower (#9250)
Co-authored-by: Chris Hughes <33847686+ChrisHughes24@users.noreply.github.com>
1 parent f181d81 commit 0d37fd6

File tree

2 files changed

+60
-4
lines changed

2 files changed

+60
-4
lines changed

src/algebra/algebra/tower.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -125,15 +125,15 @@ rfl
125125

126126
variables {R S A B}
127127

128-
@[simp, priority 900] lemma _root_.alg_hom.commutes_of_tower (f : A →ₐ[S] B) (r : R) :
128+
@[simp, priority 900] lemma _root_.alg_hom.map_algebra_map (f : A →ₐ[S] B) (r : R) :
129129
f (algebra_map R A r) = algebra_map R B r :=
130130
by rw [algebra_map_apply R S A r, f.commutes, ← algebra_map_apply R S B]
131131

132132
variables (R)
133133

134134
@[simp, priority 900] lemma _root_.alg_hom.comp_algebra_map_of_tower (f : A →ₐ[S] B) :
135135
(f : A →+* B).comp (algebra_map R A) = algebra_map R B :=
136-
ring_hom.ext f.commutes_of_tower
136+
ring_hom.ext f.map_algebra_map
137137

138138
variables (R) {S A B}
139139

src/data/polynomial/algebra_map.lean

Lines changed: 58 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -19,7 +19,8 @@ open_locale big_operators
1919

2020
namespace polynomial
2121
universes u v w z
22-
variables {R : Type u} {S : Type v} {T : Type w} {A : Type z} {a b : R} {n : ℕ}
22+
variables {R : Type u} {S : Type v} {T : Type w} {A : Type z} {A' B' : Type*} {a b : R} {n : ℕ}
23+
variables [comm_semiring A'] [comm_semiring B']
2324

2425
section comm_semiring
2526
variables [comm_semiring R] {p q r : polynomial R}
@@ -52,10 +53,23 @@ When we have `[comm_ring R]`, the function `C` is the same as `algebra_map R (po
5253
(But note that `C` is defined when `R` is not necessarily commutative, in which case
5354
`algebra_map` is not available.)
5455
-/
55-
lemma C_eq_algebra_map {R : Type*} [comm_semiring R] (r : R) :
56+
lemma C_eq_algebra_map (r : R) :
5657
C r = algebra_map R (polynomial R) r :=
5758
rfl
5859

60+
variables {R}
61+
62+
/--
63+
Extensionality lemma for algebra maps out of `polynomial A'` over a smaller base ring than `A'`
64+
-/
65+
@[ext] lemma alg_hom_ext' [algebra R A'] [algebra R B']
66+
{f g : polynomial A' →ₐ[R] B'}
67+
(h₁ : f.comp (is_scalar_tower.to_alg_hom R A' (polynomial A')) =
68+
g.comp (is_scalar_tower.to_alg_hom R A' (polynomial A')))
69+
(h₂ : f X = g X) : f = g :=
70+
alg_hom.coe_ring_hom_injective (polynomial.ring_hom_ext'
71+
(congr_arg alg_hom.to_ring_hom h₁) h₂)
72+
5973
variable (R)
6074

6175
/-- Algebra isomorphism between `polynomial R` and `add_monoid_algebra R ℕ`. This is just an
@@ -259,6 +273,48 @@ lemma is_root_of_aeval_algebra_map_eq_zero [algebra R S] {p : polynomial R}
259273
{r : R} (hr : aeval (algebra_map R S r) p = 0) : p.is_root r :=
260274
is_root_of_eval₂_map_eq_zero inj hr
261275

276+
section aeval_tower
277+
278+
variables [algebra S R] [algebra S A'] [algebra S B']
279+
280+
/-- Version of `aeval` for defining algebra homs out of `polynomial R` over a smaller base ring
281+
than `R`. -/
282+
def aeval_tower (f : R →ₐ[S] A') (x : A') : polynomial R →ₐ[S] A' :=
283+
{ commutes' := λ r, by simp [algebra_map_apply],
284+
..eval₂_ring_hom ↑f x }
285+
286+
variables (g : R →ₐ[S] A') (y : A')
287+
288+
@[simp] lemma aeval_tower_X : aeval_tower g y X = y := eval₂_X _ _
289+
290+
@[simp] lemma aeval_tower_C (x : R) : aeval_tower g y (C x) = g x := eval₂_C _ _
291+
292+
@[simp] lemma aeval_tower_comp_C : ((aeval_tower g y : polynomial R →+* A').comp C) = g :=
293+
ring_hom.ext $ aeval_tower_C _ _
294+
295+
@[simp] lemma aeval_tower_algebra_map (x : R) :
296+
aeval_tower g y (algebra_map R (polynomial R) x) = g x := eval₂_C _ _
297+
298+
@[simp] lemma aeval_tower_comp_algebra_map :
299+
(aeval_tower g y : polynomial R →+* A').comp (algebra_map R (polynomial R)) = g :=
300+
aeval_tower_comp_C _ _
301+
302+
lemma aeval_tower_to_alg_hom (x : R) :
303+
aeval_tower g y (is_scalar_tower.to_alg_hom S R (polynomial R) x) = g x :=
304+
aeval_tower_algebra_map _ _ _
305+
306+
@[simp] lemma aeval_tower_comp_to_alg_hom :
307+
(aeval_tower g y).comp (is_scalar_tower.to_alg_hom S R (polynomial R)) = g :=
308+
alg_hom.coe_ring_hom_injective $ aeval_tower_comp_algebra_map _ _
309+
310+
@[simp] lemma aeval_tower_id : aeval_tower (alg_hom.id S S) = aeval :=
311+
by { ext, simp only [eval_X, aeval_tower_X, coe_aeval_eq_eval], }
312+
313+
@[simp] lemma aeval_tower_of_id : aeval_tower (algebra.of_id S A') = aeval :=
314+
by { ext, simp only [aeval_X, aeval_tower_X], }
315+
316+
end aeval_tower
317+
262318
end comm_semiring
263319

264320
section comm_ring

0 commit comments

Comments
 (0)