Commit 2023-12-19 08:26 21ad2c6e
View on Github →feat: Mixins for monotonicity of scalar multiplication (#8869) This PR introduces eight typeclasses for monotonicity of left/right scalar multiplication by nonnegative elements:
PosSMulMono
: Ifa ≥ 0
, thenb₁ ≤ b₂
impliesa • b₁ ≤ a • b₂
.PosSMulStrictMono
: Ifa > 0
, thenb₁ < b₂
impliesa • b₁ < a • b₂
.PosSMulReflectLT
: Ifa ≥ 0
, thena • b₁ < a • b₂
impliesb₁ < b₂
.PosSMulReflectLE
: Ifa > 0
, thena • b₁ ≤ a • b₂
impliesb₁ ≤ b₂
.SMulPosMono
: Ifb ≥ 0
, thena₁ ≤ a₂
impliesa₁ • b ≤ a₂ • b
.SMulPosStrictMono
: Ifb > 0
, thena₁ < a₂
impliesa₁ • b < a₂ • b
.SMulPosReflectLT
: Ifb ≥ 0
, thena₁ • b < a₂ • b
impliesa₁ < a₂
.SMulPosReflectLE
: Ifb > 0
, thena₁ • b ≤ a₂ • b
impliesa₁ ≤ a₂
. The design is heavily inspired to the corresponding one for multiplication (seeAlgebra.Order.Ring.Lemmas
). Note however the following differences:- The new typeclasses are custom typeclasses instead of abbreviations for the correct
CovariantClass
/ContravariantClass
invokations. This has the following benefits:- They get displayed as classes in the docs. In particular, one can see their list of instances, instead of their instances being invariably dumped to the
CovariantClass
/ContravariantClass
list. - They don't pollute other typeclass searches. Having many abbreviations of the same typeclass for different purposes always felt like a performance issue to me (more instances with the same key, for no added benefit), and indeed a previous version of this PR hit timeouts due to the higher number of
CovariantClass
/ContravariantClass
instances. SMulPosReflectLT
/SMulPosReflectLE
did not fit in the framework since they relate≤
on two different types. So I would have had to generaliseCovariantClass
/ContravariantClass
to three types and two relations.- Very minor, but the constructors let you work with
a : α
,h : 0 ≤ a
instead ofa : {a : α // 0 ≤ a}
. This actually makes some instances surprisingly cleaner to prove. - The
CovariantClass
/ContravariantClass
framework was only used to automate very simple logic anyway. It was easily copied over.
- They get displayed as classes in the docs. In particular, one can see their list of instances, instead of their instances being invariably dumped to the
- We replace undocumented lemmas stating the equivalence of the four typeclasses mentioning nonnegativity with their positivity version by motivated constructors.
- We abandon series of lemmas of dubious utility. Those were already marked as such in
Algebra.Order.Ring.Lemmas
(by myself). - Some lemmas about commutativity of multiplication don't make sense for scalar multiplication.
This PR links the new typeclasses to
OrderedSMul
and makes all old lemmas inAlgebra.Order.SMul
one-liners in terms of the new lemmas (except when they have the same name, in which case the lemma is simply moved) but doesn't delete the old ones to reduce churn. What remains to be done afterwards is thus: - finish the transition by deleting the duplicate lemmas from
Algebra.Order.SMul
- rearrange the non-duplicate lemmas into new files
- generalise (most of) the lemmas from
Algebra.Order.Module
toAlgebra.Order.Module.Defs
- rethink
OrderedSMul