Skip to content

Commit

Permalink
chore(algebra/category/Module): simp lemmas for monoidal closed (#12608)
Browse files Browse the repository at this point in the history
I'm worried by the fact that I can't express the coercions here without using `@`. They do turn up in the wild, however!



Co-authored-by: Scott Morrison <scott.morrison@gmail.com>
  • Loading branch information
semorrison and semorrison committed Mar 14, 2022
1 parent 31e60c8 commit a2544de
Showing 1 changed file with 13 additions and 0 deletions.
13 changes: 13 additions & 0 deletions src/algebra/category/Module/monoidal.lean
Expand Up @@ -299,4 +299,17 @@ instance : monoidal_closed (Module.{u} R) :=
adj := adjunction.mk_of_hom_equiv
{ hom_equiv := λ N P, monoidal_closed_hom_equiv M N P, } } } }

-- I can't seem to express the function coercion here without writing `@coe_fn`.
@[simp]
lemma monoidal_closed_curry {M N P : Module.{u} R} (f : M ⊗ N ⟶ P) (x : M) (y : N) :
@coe_fn _ _ linear_map.has_coe_to_fun ((monoidal_closed.curry f : N →ₗ[R] (M →ₗ[R] P)) y) x =
f (x ⊗ₜ[R] y) :=
rfl

@[simp]
lemma monoidal_closed_uncurry {M N P : Module.{u} R}
(f : N ⟶ (M ⟶[Module.{u} R] P)) (x : M) (y : N) :
monoidal_closed.uncurry f (x ⊗ₜ[R] y) = (@coe_fn _ _ linear_map.has_coe_to_fun (f y)) x :=
by { simp only [monoidal_closed.uncurry, ihom.adjunction, is_left_adjoint.adj], simp, }

end Module

0 comments on commit a2544de

Please sign in to comment.