Skip to content
New issue

Have a question about this project? Sign up for a free GitHub account to open an issue and contact its maintainers and the community.

By clicking “Sign up for GitHub”, you agree to our terms of service and privacy statement. We’ll occasionally send you account related emails.

Already on GitHub? Sign in to your account

[Merged by Bors] - perf(linear_algebra): speed up graded_algebra instances #14967

Closed
wants to merge 8 commits into from
9 changes: 5 additions & 4 deletions src/linear_algebra/clifford_algebra/grading.lean
Expand Up @@ -96,7 +96,7 @@ begin
{ rw [alg_hom.commutes, direct_sum.algebra_map_apply], refl },
{ rw [alg_hom.map_add, ihx, ihy, ←map_add], refl },
{ obtain ⟨_, rfl⟩ := hm,
rw [alg_hom.map_mul, ih, lift_ι_apply, graded_algebra.ι_apply, direct_sum.of_mul_of],
rw [alg_hom.map_mul, ih, lift_ι_apply, graded_algebra.ι_apply Q, direct_sum.of_mul_of],
refine direct_sum.of_eq_of_graded_monoid_eq (sigma.subtype_ext _ _);
dsimp only [graded_monoid.mk, subtype.coe_mk],
{ rw [nat.succ_eq_add_one, add_comm, nat.cast_add, nat.cast_one] },
Expand All @@ -110,16 +110,17 @@ end
/-- The clifford algebra is graded by the even and odd parts. -/
instance graded_algebra : graded_algebra (even_odd Q) :=
graded_algebra.of_alg_hom (even_odd Q)
(lift _ ⟨graded_algebra.ι Q, graded_algebra.ι_sq_scalar Q⟩)
-- while not necessary, the `by apply` makes this elaborate faster
(lift Q ⟨by apply graded_algebra.ι Q, by apply graded_algebra.ι_sq_scalar Q⟩)
-- the proof from here onward is mostly similar to the `tensor_algebra` case, with some extra
-- handling for the `supr` in `even_odd`.
(begin
ext m,
dsimp only [linear_map.comp_apply, alg_hom.to_linear_map_apply, alg_hom.comp_apply,
alg_hom.id_apply],
rw [lift_ι_apply, graded_algebra.ι_apply, direct_sum.coe_alg_hom_of, subtype.coe_mk],
rw [lift_ι_apply, graded_algebra.ι_apply Q, direct_sum.coe_alg_hom_of, subtype.coe_mk],
end)
(by exact lift_ι_eq Q)
(by apply lift_ι_eq Q)

lemma supr_ι_range_eq_top : (⨆ i : ℕ, (ι Q).range ^ i) = ⊤ :=
begin
Expand Down
18 changes: 10 additions & 8 deletions src/linear_algebra/exterior_algebra/grading.lean
Expand Up @@ -42,13 +42,13 @@ end
primarily an auxiliary construction used to provide `exterior_algebra.graded_algebra`. -/
def lift_ι : exterior_algebra R M →ₐ[R]
⨁ (i : ℕ), ↥((ι R).range ^ i : submodule R (exterior_algebra R M)) :=
lift R ⟨graded_algebra.ι R M, graded_algebra.ι_sq_zero R M⟩
lift R ⟨by apply graded_algebra.ι R M, by apply graded_algebra.ι_sq_zero R M⟩

variables {R M}

lemma lift_ι_eq (i : ℕ) (x : ((ι R).range ^ i : submodule R (exterior_algebra R M))) :
lift_ι R M x =
direct_sum.of (λ i, (((ι R).range ^ i : submodule R (exterior_algebra R M)) : Type*)) i x :=
direct_sum.of (λ i, ↥((ι R).range ^ i : submodule R (exterior_algebra R M))) i x :=
begin
cases x with x hx,
dsimp only [subtype.coe_mk, direct_sum.lof_eq_of],
Expand All @@ -57,22 +57,24 @@ begin
{ rw [alg_hom.commutes, direct_sum.algebra_map_apply], refl },
{ rw [alg_hom.map_add, ihx, ihy, ←map_add], refl },
{ obtain ⟨_, rfl⟩ := hm,
rw [alg_hom.map_mul, ih, lift_ι, lift_ι_apply, graded_algebra.ι_apply, direct_sum.of_mul_of],
rw [alg_hom.map_mul, ih, lift_ι, lift_ι_apply,
graded_algebra.ι_apply R M, direct_sum.of_mul_of],
exact direct_sum.of_eq_of_graded_monoid_eq (sigma.subtype_ext (add_comm _ _) rfl) }
end

/-- The exterior algebra is graded by the powers of the submodule `(exterior_algebra.ι R).range`. -/
instance graded_algebra :
graded_algebra
((^) (ι R : M →ₗ[R] exterior_algebra R M).range : ℕ → submodule R (exterior_algebra R M)) :=
graded_algebra.of_alg_hom _ (lift_ι R M)
graded_algebra ((^) (ι R : M →ₗ[R] exterior_algebra R M).range : ℕ → submodule R _) :=
graded_algebra.of_alg_hom _
-- while not necessary, the `by apply` makes this elaborate faster
(by apply lift_ι R M)
-- the proof from here onward is identical to the `tensor_algebra` case
(begin
ext m,
dsimp only [linear_map.comp_apply, alg_hom.to_linear_map_apply, alg_hom.comp_apply,
alg_hom.id_apply, lift_ι],
rw [lift_ι_apply, graded_algebra.ι_apply, direct_sum.coe_alg_hom_of, subtype.coe_mk],
rw [lift_ι_apply, graded_algebra.ι_apply R M, direct_sum.coe_alg_hom_of, subtype.coe_mk],
end)
(by exact lift_ι_eq)
(by apply lift_ι_eq)
alreadydone marked this conversation as resolved.
Show resolved Hide resolved

end exterior_algebra
7 changes: 4 additions & 3 deletions src/linear_algebra/tensor_algebra/grading.lean
Expand Up @@ -37,12 +37,13 @@ variables {R M}
instance graded_algebra :
graded_algebra ((^) (ι R : M →ₗ[R] tensor_algebra R M).range : ℕ → submodule R _) :=
graded_algebra.of_alg_hom _
(lift _ $ graded_algebra.ι R M)
-- while not necessary, the `by apply` makes this elaborate faster
(lift R $ by apply graded_algebra.ι R M)
alreadydone marked this conversation as resolved.
Show resolved Hide resolved
(begin
ext m,
dsimp only [linear_map.comp_apply, alg_hom.to_linear_map_apply, alg_hom.comp_apply,
alg_hom.id_apply],
rw [lift_ι_apply, graded_algebra.ι_apply, direct_sum.coe_alg_hom_of, subtype.coe_mk],
rw [lift_ι_apply, graded_algebra.ι_apply R M, direct_sum.coe_alg_hom_of, subtype.coe_mk],
end)
(λ i x, begin
cases x with x hx,
Expand All @@ -52,7 +53,7 @@ graded_algebra.of_alg_hom _
{ rw [alg_hom.commutes, direct_sum.algebra_map_apply], refl },
{ rw [alg_hom.map_add, ihx, ihy, ←map_add], refl },
{ obtain ⟨_, rfl⟩ := hm,
rw [alg_hom.map_mul, ih, lift_ι_apply, graded_algebra.ι_apply, direct_sum.of_mul_of],
rw [alg_hom.map_mul, ih, lift_ι_apply, graded_algebra.ι_apply R M, direct_sum.of_mul_of],
exact direct_sum.of_eq_of_graded_monoid_eq (sigma.subtype_ext (add_comm _ _) rfl) }
end)

Expand Down