Tools to reformulate category-theoretic lemmas in concrete categories #
The elementwise
attribute #
The elementwise
attribute can be applied to a lemma
@[elementwise]
lemma some_lemma {C : Type*} [category C]
{X Y Z : C} (f : X ⟶ Y) (g : Y ⟶ Z) (h : X ⟶ Z) (w : ...) : f ≫ g = h := ...
and will produce
lemma some_lemma_apply {C : Type*} [category C] [concrete_category C]
{X Y Z : C} (f : X ⟶ Y) (g : Y ⟶ Z) (h : X ⟶ Z) (w : ...) (x : X) : g (f x) = h x := ...
Here X
is being coerced to a type via concrete_category.has_coe_to_sort
and
f
, g
, and h
are being coerced to functions via concrete_category.has_coe_to_fun
.
Further, we simplify the type using concrete_category.coe_id : ((𝟙 X) : X → X) x = x
and
concrete_category.coe_comp : (f ≫ g) x = g (f x)
,
replacing morphism composition with function composition.
The name of the produced lemma can be specified with @[elementwise other_lemma_name]
.
If simp
is added first, the generated lemma will also have the simp
attribute.
Implementation #
This closely follows the implementation of the @[reassoc]
attribute, due to Simon Hudon.
Thanks to Gabriel Ebner for help diagnosing universe issues.
The elementwise
attribute can be applied to a lemma
@[elementwise]
lemma some_lemma {C : Type*} [category C]
{X Y Z : C} (f : X ⟶ Y) (g : Y ⟶ Z) (h : X ⟶ Z) (w : ...) : f ≫ g = h := ...
and will produce
lemma some_lemma_apply {C : Type*} [category C] [concrete_category C]
{X Y Z : C} (f : X ⟶ Y) (g : Y ⟶ Z) (h : X ⟶ Z) (w : ...) (x : X) : g (f x) = h x := ...
Here X
is being coerced to a type via concrete_category.has_coe_to_sort
and
f
, g
, and h
are being coerced to functions via concrete_category.has_coe_to_fun
.
Further, we simplify the type using concrete_category.coe_id : ((𝟙 X) : X → X) x = x
and
concrete_category.coe_comp : (f ≫ g) x = g (f x)
,
replacing morphism composition with function composition.
The [concrete_category C]
argument will be omitted if it is possible to synthesize an instance.
The name of the produced lemma can be specified with @[elementwise other_lemma_name]
.
If simp
is added first, the generated lemma will also have the simp
attribute.
With w : ∀ ..., f ≫ g = h
(with universal quantifiers tolerated),
elementwise_of w : ∀ ... (x : X), g (f x) = h x
.
The type and proof of elementwise_of h
is generated by tactic.derive_elementwise_proof
which makes elementwise_of
meta-programming adjacent. It is not called as a tactic but as
an expression. The goal is to avoid creating assumptions that are dismissed after one use:
example (M N K : Mon.{u}) (f : M ⟶ N) (g : N ⟶ K) (h : M ⟶ K) (w : f ≫ g = h) (m : M) :
g (f m) = h m :=
begin
rw elementwise_of w,
end
With w : ∀ ..., f ≫ g = h
(with universal quantifiers tolerated),
elementwise_of w : ∀ ... (x : X), g (f x) = h x
.
Although elementwise_of
is not a tactic or a meta program, its type is generated
through meta-programming to make it usable inside normal expressions.