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: If a ≥ 0, then b₁ ≤ b₂ implies a • b₁ ≤ a • b₂.
  • PosSMulStrictMono: If a > 0, then b₁ < b₂ implies a • b₁ < a • b₂.
  • PosSMulReflectLT: If a ≥ 0, then a • b₁ < a • b₂ implies b₁ < b₂.
  • PosSMulReflectLE: If a > 0, then a • b₁ ≤ a • b₂ implies b₁ ≤ b₂.
  • SMulPosMono: If b ≥ 0, then a₁ ≤ a₂ implies a₁ • b ≤ a₂ • b.
  • SMulPosStrictMono: If b > 0, then a₁ < a₂ implies a₁ • b < a₂ • b.
  • SMulPosReflectLT: If b ≥ 0, then a₁ • b < a₂ • b implies a₁ < a₂.
  • SMulPosReflectLE: If b > 0, then a₁ • b ≤ a₂ • b implies a₁ ≤ a₂. The design is heavily inspired to the corresponding one for multiplication (see Algebra.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 generalise CovariantClass/ContravariantClass to three types and two relations.
    • Very minor, but the constructors let you work with a : α, h : 0 ≤ a instead of a : {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.
  • 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 in Algebra.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 to Algebra.Order.Module.Defs
  • rethink OrderedSMul

Estimated changes

added theorem PosSMulMono.lift
added theorem PosSMulMono.of_pos
added theorem PosSMulReflectLE.lift
added theorem PosSMulReflectLT.lift
added theorem PosSMulStrictMono.lift
added theorem SMulPosMono.lift
added theorem SMulPosMono.of_pos
added theorem SMulPosReflectLE.lift
added theorem SMulPosReflectLT.lift
added theorem SMulPosStrictMono.lift
added theorem inv_smul_le_iff_of_pos
added theorem inv_smul_lt_iff_of_pos
added theorem le_inv_smul_iff_of_pos
added theorem le_smul_of_one_le_left
added theorem lt_inv_smul_iff_of_pos
added theorem lt_smul_of_one_lt_left
added theorem neg_of_smul_neg_left'
added theorem neg_of_smul_neg_left
added theorem neg_of_smul_neg_right'
added theorem neg_of_smul_neg_right
added theorem neg_of_smul_pos_left
added theorem neg_of_smul_pos_right
added theorem pos_of_smul_pos_left
added theorem pos_of_smul_pos_right
added theorem smul_le_of_le_one_left
added theorem smul_le_smul'
added theorem smul_le_smul
added theorem smul_lt_of_lt_one_left
added theorem smul_lt_smul'
added theorem smul_lt_smul
added theorem smul_max_of_nonneg
added theorem smul_min_of_nonneg
added theorem smul_neg_of_neg_of_pos
added theorem smul_neg_of_pos_of_neg
added theorem smul_nonneg'
added theorem smul_nonneg
added theorem smul_pos'
added theorem smul_pos
modified theorem inv_smul_le_iff
modified theorem inv_smul_lt_iff
modified theorem le_inv_smul_iff
modified theorem lt_inv_smul_iff
modified theorem monotone_smul_left
modified theorem smul_le_smul_of_nonneg
modified theorem smul_lt_smul_of_pos
deleted theorem smul_nonneg
modified theorem smul_pos_iff_of_pos
modified theorem strictMono_smul_left