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

Commit f30200e

Browse files
committed
refactor(algebra/free_algebra): Make lift an equiv (#4908)
This can probably lead to some API cleanup down the line
1 parent c20ecef commit f30200e

File tree

7 files changed

+218
-182
lines changed

7 files changed

+218
-182
lines changed

src/algebra/category/Algebra/basic.lean

Lines changed: 5 additions & 10 deletions
Original file line numberDiff line numberDiff line change
@@ -82,17 +82,12 @@ def free : Type* ⥤ Algebra R :=
8282
map := λ S T f, free_algebra.lift _ $ (free_algebra.ι _) ∘ f }
8383

8484
/-- The free/forget ajunction for `R`-algebras. -/
85-
@[simps]
8685
def adj : free R ⊣ forget (Algebra R) :=
87-
{ hom_equiv := λ X A,
88-
{ to_fun := λ f, f ∘ (free_algebra.ι _),
89-
inv_fun := λ f, free_algebra.lift _ f,
90-
left_inv := by tidy,
91-
right_inv := by tidy },
92-
unit := { app := λ S, free_algebra.ι _ },
93-
counit :=
94-
{ app := λ S, free_algebra.lift _ $ id,
95-
naturality' := by {intros, ext, simp} } } -- tidy times out :(
86+
adjunction.mk_of_hom_equiv
87+
{ hom_equiv := λ X A, (free_algebra.lift _).symm,
88+
-- Relying on `obviously` to fill out these proofs is very slow :(
89+
hom_equiv_naturality_left_symm' := by {intros, ext, simp},
90+
hom_equiv_naturality_right' := by {intros, ext, simp} }
9691

9792
end Algebra
9893

src/algebra/free_algebra.lean

Lines changed: 37 additions & 24 deletions
Original file line numberDiff line numberDiff line change
@@ -186,11 +186,8 @@ def ι : X → free_algebra R X := λ m, quot.mk _ m
186186

187187
variables {A : Type*} [semiring A] [algebra R A]
188188

189-
/--
190-
Given a function `f : X → A` where `A` is an `R`-algebra, `lift R f` is the unique lift
191-
of `f` to a morphism of `R`-algebras `free_algebra R X → A`.
192-
-/
193-
def lift (f : X → A) : free_algebra R X →ₐ[R] A :=
189+
/-- Internal definition used to define `lift` -/
190+
private def lift_aux (f : X → A) : (free_algebra R X →ₐ[R] A) :=
194191
{ to_fun := λ a, quot.lift_on a (lift_fun _ _ f) $ λ a b h,
195192
begin
196193
induction h,
@@ -230,6 +227,36 @@ def lift (f : X → A) : free_algebra R X →ₐ[R] A :=
230227
map_add' := by { rintros ⟨⟩ ⟨⟩, refl },
231228
commutes' := by tauto }
232229

230+
/--
231+
Given a function `f : X → A` where `A` is an `R`-algebra, `lift R f` is the unique lift
232+
of `f` to a morphism of `R`-algebras `free_algebra R X → A`.
233+
-/
234+
def lift : (X → A) ≃ (free_algebra R X →ₐ[R] A) :=
235+
{ to_fun := lift_aux R,
236+
inv_fun := λ F, F ∘ (ι R),
237+
left_inv := λ f, by {ext, refl},
238+
right_inv := λ F, by {
239+
ext x,
240+
rcases x,
241+
induction x,
242+
case pre.of : {
243+
change ((F : free_algebra R X → A) ∘ (ι R)) _ = _,
244+
refl },
245+
case pre.of_scalar : {
246+
change algebra_map _ _ x = F (algebra_map _ _ x),
247+
rw alg_hom.commutes F x, },
248+
case pre.add : a b ha hb {
249+
change lift_aux R (F ∘ ι R) (quot.mk _ _ + quot.mk _ _) = F (quot.mk _ _ + quot.mk _ _),
250+
rw [alg_hom.map_add, alg_hom.map_add, ha, hb], },
251+
case pre.mul : a b ha hb {
252+
change lift_aux R (F ∘ ι R) (quot.mk _ _ * quot.mk _ _) = F (quot.mk _ _ * quot.mk _ _),
253+
rw [alg_hom.map_mul, alg_hom.map_mul, ha, hb], }, }, }
254+
255+
@[simp] lemma lift_aux_eq (f : X → A) : lift_aux R f = lift R f := rfl
256+
257+
@[simp]
258+
lemma lift_symm_apply (F : free_algebra R X →ₐ[R] A) : (lift R).symm F = F ∘ (ι R) := rfl
259+
233260
variables {R X}
234261

235262
@[simp]
@@ -243,22 +270,7 @@ theorem lift_ι_apply (f : X → A) (x) :
243270
@[simp]
244271
theorem lift_unique (f : X → A) (g : free_algebra R X →ₐ[R] A) :
245272
(g : free_algebra R X → A) ∘ (ι R) = f ↔ g = lift R f :=
246-
begin
247-
refine ⟨λ hyp, _, λ hyp, by rw [hyp, ι_comp_lift]⟩,
248-
ext,
249-
rcases x,
250-
induction x,
251-
{ change ((g : free_algebra R X → A) ∘ (ι R)) _ = _,
252-
rw hyp,
253-
refl },
254-
{ exact alg_hom.commutes g x },
255-
{ change g (quot.mk _ _ + quot.mk _ _) = _,
256-
simp only [alg_hom.map_add, *],
257-
refl },
258-
{ change g (quot.mk _ _ * quot.mk _ _) = _,
259-
simp only [alg_hom.map_mul, *],
260-
refl },
261-
end
273+
(lift R).symm_apply_eq
262274

263275
/-!
264276
At this stage we set the basic definitions as `@[irreducible]`, so from this point onwards one should only use the universal properties of the free algebra, and consider the actual implementation as a quotient of an inductive type as completely hidden.
@@ -272,14 +284,15 @@ attribute [irreducible] ι lift
272284

273285
@[simp]
274286
theorem lift_comp_ι (g : free_algebra R X →ₐ[R] A) :
275-
lift R ((g : free_algebra R X → A) ∘ (ι R)) = g := by {symmetry, rw ←lift_unique}
287+
lift R ((g : free_algebra R X → A) ∘ (ι R)) = g :=
288+
by { rw ←lift_symm_apply, exact (lift R).apply_symm_apply g }
276289

277290
@[ext]
278291
theorem hom_ext {f g : free_algebra R X →ₐ[R] A}
279292
(w : ((f : free_algebra R X → A) ∘ (ι R)) = ((g : free_algebra R X → A) ∘ (ι R))) : f = g :=
280293
begin
281-
have : g = lift R ((g : free_algebra R X → A) ∘ (ι R)), by rw ←lift_unique,
282-
rw [this, ←lift_unique, w],
294+
rw [←lift_symm_apply, ←lift_symm_apply] at w,
295+
exact (lift R).symm.injective w,
283296
end
284297

285298
/--

src/algebra/lie/universal_enveloping.lean

Lines changed: 26 additions & 26 deletions
Original file line numberDiff line numberDiff line change
@@ -77,43 +77,43 @@ variables {A : Type u₃} [ring A] [algebra R A] (f : L →ₗ⁅R⁆ A)
7777

7878
/-- The universal property of the universal enveloping algebra: Lie algebra morphisms into
7979
associative algebras lift to associative algebra morphisms from the universal enveloping algebra. -/
80-
def lift : universal_enveloping_algebra R L →ₐ[R] A :=
81-
ring_quot.lift_alg_hom R (tensor_algebra.lift R (f : L →ₗ[R] A))
82-
begin
83-
intros a b h, induction h with x y,
84-
simp [lie_ring.of_associative_ring_bracket],
85-
end
80+
def lift : (L →ₗ⁅R⁆ A) ≃ (universal_enveloping_algebra R L →ₐ[R] A) :=
81+
{ to_fun := λ f,
82+
ring_quot.lift_alg_hom R ⟨tensor_algebra.lift R (f : L →ₗ[R] A),
83+
begin
84+
intros a b h, induction h with x y,
85+
simp [lie_ring.of_associative_ring_bracket],
86+
end⟩,
87+
inv_fun := λ F, (lie_algebra.of_associative_algebra_hom F).comp (ι R),
88+
left_inv := λ f, by { ext, simp [ι, mk_alg_hom], },
89+
right_inv := λ F, by { ext, simp [ι, mk_alg_hom], } }
90+
91+
@[simp] lemma lift_symm_apply (F : universal_enveloping_algebra R L →ₐ[R] A) :
92+
(lift R).symm F = (lie_algebra.of_associative_algebra_hom F).comp (ι R) :=
93+
rfl
94+
95+
@[simp] lemma ι_comp_lift : (lift R f) ∘ (ι R) = f :=
96+
funext $ lie_algebra.morphism.ext_iff.mp $ (lift R).symm_apply_apply f
8697

8798
@[simp] lemma lift_ι_apply (x : L) : lift R f (ι R x) = f x :=
88-
begin
89-
have : ι R x = ring_quot.mk_alg_hom R (rel R L) (ιₜ x), by refl,
90-
simp [this, lift],
91-
end
92-
93-
lemma ι_comp_lift : (lift R f) ∘ (ι R) = f :=
94-
by { ext, simp, }
99+
by rw [←function.comp_apply (lift R f) (ι R) x, ι_comp_lift]
95100

96101
lemma lift_unique (g : universal_enveloping_algebra R L →ₐ[R] A) :
97102
g ∘ (ι R) = f ↔ g = lift R f :=
98103
begin
99-
split; intros h,
100-
{ apply ring_quot.lift_alg_hom_unique,
101-
rw ← tensor_algebra.lift_unique,
102-
ext x,
103-
change _ = f x, rw ← congr h rfl,
104-
refl, },
105-
{ subst h, apply ι_comp_lift, },
104+
refine iff.trans _ (lift R).symm_apply_eq,
105+
split; {intro h, ext, simp [←h] },
106106
end
107107

108108
@[ext] lemma hom_ext {g₁ g₂ : universal_enveloping_algebra R L →ₐ[R] A}
109109
(h : g₁ ∘ (ι R) = g₂ ∘ (ι R)) : g₁ = g₂ :=
110110
begin
111-
let f₁ := (lie_algebra.of_associative_algebra_hom g₁).comp (ι R),
112-
let f₂ := (lie_algebra.of_associative_algebra_hom g₂).comp (ι R),
113-
have h' : f₁ = f₂, { ext, exact congr h rfl, },
114-
have h₁ : g₁ = lift R f₁, { rw ← lift_unique, refl, },
115-
have h₂ : g₂ = lift R f₂, { rw ← lift_unique, refl, },
116-
rw [h₁, h₂, h'],
111+
apply (lift R).symm.injective,
112+
rw [lift_symm_apply, lift_symm_apply],
113+
ext x,
114+
have := congr h (rfl : x = x),
115+
rw function.comp_apply at this,
116+
simp [this],
117117
end
118118

119119
end universal_enveloping_algebra

src/algebra/ring_quot.lean

Lines changed: 59 additions & 49 deletions
Original file line numberDiff line numberDiff line change
@@ -129,36 +129,41 @@ variables {T : Type u₄} [semiring T]
129129

130130
/--
131131
Any ring homomorphism `f : R →+* T` which respects a relation `r : R → R → Prop`
132-
factors through a morphism `ring_quot r →+* T`.
132+
factors uniquely through a morphism `ring_quot r →+* T`.
133133
-/
134-
def lift (f : R →+* T) {r : R → R → Prop} (w : ∀ ⦃x y⦄, r x y → f x = f y) :
135-
ring_quot r →+* T :=
136-
{ to_fun := quot.lift f
137-
begin
138-
rintros _ _ r,
139-
induction r,
140-
case of : _ _ r { exact w r, },
141-
case add_left : _ _ _ _ r' { simp [r'], },
142-
case mul_left : _ _ _ _ r' { simp [r'], },
143-
case mul_right : _ _ _ _ r' { simp [r'], },
144-
end,
145-
map_zero' := f.map_zero,
146-
map_add' := by { rintros ⟨x⟩ ⟨y⟩, exact f.map_add x y, },
147-
map_one' := f.map_one,
148-
map_mul' := by { rintros ⟨x⟩ ⟨y⟩, exact f.map_mul x y, }, }
134+
def lift {r : R → R → Prop} :
135+
{f : R →+* T // ∀ ⦃x y⦄, r x y → f x = f y} ≃ (ring_quot r →+* T) :=
136+
{ to_fun := λ f', let f := (f' : R →+* T) in
137+
{ to_fun := quot.lift f
138+
begin
139+
rintros _ _ r,
140+
induction r,
141+
case of : _ _ r { exact f'.prop r, },
142+
case add_left : _ _ _ _ r' { simp [r'], },
143+
case mul_left : _ _ _ _ r' { simp [r'], },
144+
case mul_right : _ _ _ _ r' { simp [r'], },
145+
end,
146+
map_zero' := f.map_zero,
147+
map_add' := by { rintros ⟨x⟩ ⟨y⟩, exact f.map_add x y, },
148+
map_one' := f.map_one,
149+
map_mul' := by { rintros ⟨x⟩ ⟨y⟩, exact f.map_mul x y, }, },
150+
inv_fun := λ F, ⟨F.comp (mk_ring_hom r), λ x y h, by { dsimp, rw mk_ring_hom_rel h, }⟩,
151+
left_inv := λ f, by { ext, simp, refl },
152+
right_inv := λ F, by { ext, simp, refl } }
149153

150154
@[simp]
151155
lemma lift_mk_ring_hom_apply (f : R →+* T) {r : R → R → Prop} (w : ∀ ⦃x y⦄, r x y → f x = f y) (x) :
152-
(lift f w) (mk_ring_hom r x) = f x :=
156+
lift ⟨f, w⟩ (mk_ring_hom r x) = f x :=
153157
rfl
154158

159+
-- note this is essentially `lift.symm_apply_eq.mp h`
155160
lemma lift_unique (f : R →+* T) {r : R → R → Prop} (w : ∀ ⦃x y⦄, r x y → f x = f y)
156-
(g : ring_quot r →+* T) (h : g.comp (mk_ring_hom r) = f) : g = lift f w :=
161+
(g : ring_quot r →+* T) (h : g.comp (mk_ring_hom r) = f) : g = lift ⟨f, w⟩ :=
157162
by { ext, simp [h], }
158163

159164
lemma eq_lift_comp_mk_ring_hom {r : R → R → Prop} (f : ring_quot r →+* T) :
160-
f = lift (f.comp (mk_ring_hom r)) (λ x y h, by { dsimp, rw mk_ring_hom_rel h, }) :=
161-
by { ext, simp, }
165+
f = lift f.comp (mk_ring_hom r), λ x y h, by { dsimp, rw mk_ring_hom_rel h, } :=
166+
(lift.apply_symm_apply f).symm
162167

163168
section comm_ring
164169
/-!
@@ -172,8 +177,8 @@ variables {B : Type u₁} [comm_ring B]
172177
def ring_quot_to_ideal_quotient (r : B → B → Prop) :
173178
ring_quot r →+* (ideal.of_rel r).quotient :=
174179
lift
175-
(ideal.quotient.mk (ideal.of_rel r))
176-
(λ x y h, quot.sound (submodule.mem_Inf.mpr (λ p w, w ⟨x, y, h, rfl⟩)))
180+
ideal.quotient.mk (ideal.of_rel r),
181+
λ x y h, quot.sound (submodule.mem_Inf.mpr (λ p w, w ⟨x, y, h, rfl⟩))
177182

178183
@[simp] lemma ring_quot_to_ideal_quotient_apply (r : B → B → Prop) (x : B) :
179184
ring_quot_to_ideal_quotient r (mk_ring_hom r x) = ideal.quotient.mk _ x := rfl
@@ -240,43 +245,48 @@ end
240245

241246
/--
242247
Any `S`-algebra homomorphism `f : A →ₐ[S] B` which respects a relation `s : A → A → Prop`
243-
factors through a morphism `ring_quot s →ₐ[S] B`.
248+
factors uniquely through a morphism `ring_quot s →ₐ[S] B`.
244249
-/
245-
def lift_alg_hom (f : A →ₐ[S] B) {s : A → A → Prop} (w : ∀ ⦃x y⦄, s x y → f x = f y) :
246-
ring_quot s →ₐ[S] B :=
247-
{ to_fun := quot.lift f
248-
begin
249-
rintros _ _ r,
250-
induction r,
251-
case of : _ _ r { exact w r, },
252-
case add_left : _ _ _ _ r' { simp [r'], },
253-
case mul_left : _ _ _ _ r' { simp [r'], },
254-
case mul_right : _ _ _ _ r' { simp [r'], },
255-
end,
256-
map_zero' := f.map_zero,
257-
map_add' := by { rintros ⟨x⟩ ⟨y⟩, exact f.map_add x y, },
258-
map_one' := f.map_one,
259-
map_mul' := by { rintros ⟨x⟩ ⟨y⟩, exact f.map_mul x y, },
260-
commutes' :=
261-
begin
262-
rintros x,
263-
conv_rhs { rw [algebra.algebra_map_eq_smul_one, ←f.map_one, ←f.map_smul], },
264-
rw algebra.algebra_map_eq_smul_one,
265-
exact quot.lift_beta f @w (x • 1),
266-
end, }
250+
def lift_alg_hom {s : A → A → Prop} :
251+
{ f : A →ₐ[S] B // ∀ ⦃x y⦄, s x y → f x = f y} ≃ (ring_quot s →ₐ[S] B) :=
252+
{ to_fun := λ f', let f := (f' : A →ₐ[S] B) in
253+
{ to_fun := quot.lift f
254+
begin
255+
rintros _ _ r,
256+
induction r,
257+
case of : _ _ r { exact f'.prop r, },
258+
case add_left : _ _ _ _ r' { simp [r'], },
259+
case mul_left : _ _ _ _ r' { simp [r'], },
260+
case mul_right : _ _ _ _ r' { simp [r'], },
261+
end,
262+
map_zero' := f.map_zero,
263+
map_add' := by { rintros ⟨x⟩ ⟨y⟩, exact f.map_add x y, },
264+
map_one' := f.map_one,
265+
map_mul' := by { rintros ⟨x⟩ ⟨y⟩, exact f.map_mul x y, },
266+
commutes' :=
267+
begin
268+
rintros x,
269+
conv_rhs { rw [algebra.algebra_map_eq_smul_one, ←f.map_one, ←f.map_smul], },
270+
rw algebra.algebra_map_eq_smul_one,
271+
exact quot.lift_beta f f'.prop (x • 1),
272+
end, },
273+
inv_fun := λ F, ⟨F.comp (mk_alg_hom S s), λ _ _ h, by { dsimp, erw mk_alg_hom_rel S h }⟩,
274+
left_inv := λ f, by { ext, simp, refl },
275+
right_inv := λ F, by { ext, simp, refl } }
267276

268277
@[simp]
269278
lemma lift_alg_hom_mk_alg_hom_apply (f : A →ₐ[S] B) {s : A → A → Prop} (w : ∀ ⦃x y⦄, s x y → f x = f y) (x) :
270-
(lift_alg_hom S f w) ((mk_alg_hom S s) x) = f x :=
279+
(lift_alg_hom S ⟨f, w⟩) ((mk_alg_hom S s) x) = f x :=
271280
rfl
272281

282+
-- note this is essentially `(lift_alg_hom S).symm_apply_eq.mp h`
273283
lemma lift_alg_hom_unique (f : A →ₐ[S] B) {s : A → A → Prop} (w : ∀ ⦃x y⦄, s x y → f x = f y)
274-
(g : ring_quot s →ₐ[S] B) (h : g.comp (mk_alg_hom S s) = f) : g = lift_alg_hom S f w :=
284+
(g : ring_quot s →ₐ[S] B) (h : g.comp (mk_alg_hom S s) = f) : g = lift_alg_hom S ⟨f, w⟩ :=
275285
by { ext, simp [h], }
276286

277287
lemma eq_lift_alg_hom_comp_mk_alg_hom {s : A → A → Prop} (f : ring_quot s →ₐ[S] B) :
278-
f = lift_alg_hom S (f.comp (mk_alg_hom S s)) (λ x y h, by { dsimp, erw mk_alg_hom_rel S h, }) :=
279-
by { ext, simp, }
288+
f = lift_alg_hom S f.comp (mk_alg_hom S s), λ x y h, by { dsimp, erw mk_alg_hom_rel S h, } :=
289+
((lift_alg_hom S).apply_symm_apply f).symm
280290

281291
end algebra
282292

0 commit comments

Comments
 (0)