Typeclasses for algebraic operations #
Notation typeclass for Inv
, the multiplicative analogue of Neg
.
We also introduce notation classes SMul
and VAdd
for multiplicative and additive
actions.
SMul
is typically, but not exclusively, used for scalar multiplication-like operators.
See the module Algebra.AddTorsor
for a motivating example for the name VAdd
(vector addition).
Notation #
a • b
is used as notation forHSMul.hSMul a b
.a +ᵥ b
is used as notation forHVAdd.hVAdd a b
.
The notation typeclass for heterogeneous additive actions.
This enables the notation a +ᵥ b : γ
where a : α
, b : β
.
- hVAdd : α → β → γ
a +ᵥ b
computes the sum ofa
andb
. The meaning of this notation is type-dependent.
Instances
The notation typeclass for heterogeneous scalar multiplication.
This enables the notation a • b : γ
where a : α
, b : β
.
It is assumed to represent a left action in some sense.
The notation a • b
is augmented with a macro (below) to have it elaborate as a left action.
Only the b
argument participates in the elaboration algorithm: the algorithm uses the type of b
when calculating the type of the surrounding arithmetic expression
and it tries to insert coercions into b
to get some b'
such that a • b'
has the same type as b'
.
See the module documentation near the macro for more details.
- hSMul : α → β → γ
a • b
computes the product ofa
andb
. The meaning of this notation is type-dependent, but it is intended to be used for left actions.
Instances
a +ᵥ b
computes the sum of a
and b
.
The meaning of this notation is type-dependent.
Equations
- «term_+ᵥ_» = Lean.ParserDescr.trailingNode `«term_+ᵥ_» 65 66 (Lean.ParserDescr.binary `andthen (Lean.ParserDescr.symbol " +ᵥ ") (Lean.ParserDescr.cat `term 65))
Instances For
a -ᵥ b
computes the difference of a
and b
. The meaning of this notation is
type-dependent, but it is intended to be used for additive torsors.
Equations
- «term_-ᵥ_» = Lean.ParserDescr.trailingNode `«term_-ᵥ_» 65 65 (Lean.ParserDescr.binary `andthen (Lean.ParserDescr.symbol " -ᵥ ") (Lean.ParserDescr.cat `term 66))
Instances For
a • b
computes the product of a
and b
.
The meaning of this notation is type-dependent, but it is intended to be used for left actions.
Equations
- «term_•_» = Lean.ParserDescr.trailingNode `«term_•_» 73 74 (Lean.ParserDescr.binary `andthen (Lean.ParserDescr.symbol " • ") (Lean.ParserDescr.cat `term 73))
Instances For
We have a macro to make x • y
notation participate in the expression tree elaborator,
like other arithmetic expressions such as +
, *
, /
, ^
, =
, inequalities, etc.
The macro is using the leftact%
elaborator introduced in
this RFC.
As a concrete example of the effect of this macro, consider
variable [Ring R] [AddCommMonoid M] [Module R M] (r : R) (N : Submodule R M) (m : M) (n : N)
#check m + r • n
Without the macro, the expression would elaborate as m + ↑(r • n : ↑N) : M
.
With the macro, the expression elaborates as m + r • (↑n : M) : M
.
To get the first interpretation, one can write m + (r • n :)
.
Here is a quick review of the expression tree elaborator:
- It builds up an expression tree of all the immediately accessible operations
that are marked with
binop%
,unop%
,leftact%
,rightact%
,binrel%
, etc. - It elaborates every leaf term of this tree
(without an expected type, so as if it were temporarily wrapped in
(... :)
). - Using the types of each elaborated leaf, it computes a supremum type they can all be coerced to, if such a supremum exists.
- It inserts coercions around leaf terms wherever needed.
The hypothesis is that individual expression trees tend to be calculations with respect to a single algebraic structure.
Note(kmill): If we were to remove HSMul
and switch to using SMul
directly,
then the expression tree elaborator would not be able to insert coercions within the right operand;
they would likely appear as ↑(x • y)
rather than x • ↑y
, unlike other arithmetic operations.
Invert an element of α.
Equations
- «term_⁻¹» = Lean.ParserDescr.trailingNode `«term_⁻¹» 1024 1024 (Lean.ParserDescr.symbol "⁻¹")