Tools to reformulate category-theoretic axioms in a more associativity-friendly way
The reassoc
attribute
The reassoc
attribute can be applied to a lemma
@[reassoc]
lemma some_lemma : foo ≫ bar = baz := ...
and produce
lemma some_lemma_assoc {Y : C} (f : X ⟶ Y) : foo ≫ bar ≫ f = baz ≫ f := ...
The name of the produced lemma can be specified with @[reassoc other_lemma_name]
. If
simp
is added first, the generated lemma will also have the simp
attribute.
The reassoc_axiom
command
When declaring a class of categories, the axioms can be reformulated to be more amenable to manipulation in right associated expressions:
class some_class (C : Type) [category C] :=
(foo : Π X : C, X ⟶ X)
(bar : ∀ {X Y : C} (f : X ⟶ Y), foo X ≫ f = f ≫ foo Y)
reassoc_axiom some_class.bar
Here too, the reassoc
attribute can be used instead. It works well when combined with
simp
:
attribute [simp, reassoc] some_class.bar
From an expression f ≫ g
, extract the expression representing the category instance.
(implementation for @[reassoc]
)
Given a declaration named n
of the form f ≫ g = h
, proves a new lemma named n'
of the form ∀ {W} (k), f ≫ (g ≫ k) = h ≫ k
.
The reassoc
attribute can be applied to a lemma
@[reassoc]
lemma some_lemma : foo ≫ bar = baz := ...
to produce
lemma some_lemma_assoc {Y : C} (f : X ⟶ Y) : foo ≫ bar ≫ f = baz ≫ f := ...
The name of the produced lemma can be specified with @[reassoc other_lemma_name]
. If
simp
is added first, the generated lemma will also have the simp
attribute.
When declaring a class of categories, the axioms can be reformulated to be more amenable to manipulation in right associated expressions:
class some_class (C : Type) [category C] :=
(foo : Π X : C, X ⟶ X)
(bar : ∀ {X Y : C} (f : X ⟶ Y), foo X ≫ f = f ≫ foo Y)
reassoc_axiom some_class.bar
The above will produce:
lemma some_class.bar_assoc {Z : C} (g : Y ⟶ Z) :
foo X ≫ f ≫ g = f ≫ foo Y ≫ g := ...
Here too, the reassoc
attribute can be used instead. It works well when combined with
simp
:
attribute [simp, reassoc] some_class.bar
reassoc h
, for assumption h : x ≫ y = z
, creates a new assumption
h : ∀ {W} (f : Z ⟶ W), x ≫ y ≫ f = z ≫ f
.
reassoc! h
, does the same but deletes the initial h
assumption.
(You can also add the attribute @[reassoc]
to lemmas to generate new declarations generalized
in this way.)
Equations
- tactic.calculated_Prop β hh = β
With h : x ≫ y ≫ z = x
(with universal quantifiers tolerated),
reassoc_of h : ∀ {X'} (f : W ⟶ X'), x ≫ y ≫ z ≫ f = x ≫ f
.
The type and proof of reassoc_of h
is generated by tactic.derive_reassoc_proof
which make reassoc_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: