Lemmas for the linear_combination tactic #
These should not be used directly in user code.
Addition #
Multiplication #
Division #
Lemmas constructing the reduction of a goal to a specified built-up hypothesis #
Alias of the forward direction of sub_eq_zero
.
Lookup functions for lemmas by operation and relation(s) #
Given two (in)equalities, look up the lemma to add them.
Equations
- Mathlib.Ineq.eq.addRelRelData Mathlib.Ineq.eq = `Mathlib.Tactic.LinearCombination.add_eq_eq
- Mathlib.Ineq.eq.addRelRelData Mathlib.Ineq.le = `Mathlib.Tactic.LinearCombination.add_eq_le
- Mathlib.Ineq.eq.addRelRelData Mathlib.Ineq.lt = `Mathlib.Tactic.LinearCombination.add_eq_lt
- Mathlib.Ineq.le.addRelRelData Mathlib.Ineq.eq = `Mathlib.Tactic.LinearCombination.add_le_eq
- Mathlib.Ineq.le.addRelRelData Mathlib.Ineq.le = `add_le_add
- Mathlib.Ineq.le.addRelRelData Mathlib.Ineq.lt = `add_lt_add_of_le_of_lt
- Mathlib.Ineq.lt.addRelRelData Mathlib.Ineq.eq = `Mathlib.Tactic.LinearCombination.add_lt_eq
- Mathlib.Ineq.lt.addRelRelData Mathlib.Ineq.le = `add_lt_add_of_lt_of_le
- Mathlib.Ineq.lt.addRelRelData Mathlib.Ineq.lt = `add_lt_add
Instances For
Finite inductive type extending Mathlib.Ineq
: a type of inequality (eq
, le
or lt
),
together with, in the case of lt
, a boolean, typically representing the strictness (< or ≤) of
some other inequality.
- eq: Mathlib.Ineq.WithStrictness
- le: Mathlib.Ineq.WithStrictness
- lt: Bool → Mathlib.Ineq.WithStrictness
Instances For
Given an (in)equality, look up the lemma to left-multiply it by a constant. If relevant, also take into account the degree of positivity which can be proved of the constant: strict or non-strict.
Equations
- Mathlib.Ineq.mulRelConstData Mathlib.Ineq.WithStrictness.eq = `Mathlib.Tactic.LinearCombination.mul_eq_const
- Mathlib.Ineq.mulRelConstData Mathlib.Ineq.WithStrictness.le = `Mathlib.Tactic.LinearCombination.mul_le_const
- Mathlib.Ineq.mulRelConstData (Mathlib.Ineq.WithStrictness.lt true) = `Mathlib.Tactic.LinearCombination.mul_lt_const
- Mathlib.Ineq.mulRelConstData (Mathlib.Ineq.WithStrictness.lt false) = `Mathlib.Tactic.LinearCombination.mul_lt_const_weak
Instances For
Given an (in)equality, look up the lemma to right-multiply it by a constant. If relevant, also take into account the degree of positivity which can be proved of the constant: strict or non-strict.
Equations
- Mathlib.Ineq.mulConstRelData Mathlib.Ineq.WithStrictness.eq = `Mathlib.Tactic.LinearCombination.mul_const_eq
- Mathlib.Ineq.mulConstRelData Mathlib.Ineq.WithStrictness.le = `Mathlib.Tactic.LinearCombination.mul_const_le
- Mathlib.Ineq.mulConstRelData (Mathlib.Ineq.WithStrictness.lt true) = `Mathlib.Tactic.LinearCombination.mul_const_lt
- Mathlib.Ineq.mulConstRelData (Mathlib.Ineq.WithStrictness.lt false) = `Mathlib.Tactic.LinearCombination.mul_const_lt_weak
Instances For
Given an (in)equality, look up the lemma to divide it by a constant. If relevant, also take into account the degree of positivity which can be proved of the constant: strict or non-strict.
Equations
- Mathlib.Ineq.divRelConstData Mathlib.Ineq.WithStrictness.eq = `Mathlib.Tactic.LinearCombination.div_eq_const
- Mathlib.Ineq.divRelConstData Mathlib.Ineq.WithStrictness.le = `Mathlib.Tactic.LinearCombination.div_le_const
- Mathlib.Ineq.divRelConstData (Mathlib.Ineq.WithStrictness.lt true) = `Mathlib.Tactic.LinearCombination.div_lt_const
- Mathlib.Ineq.divRelConstData (Mathlib.Ineq.WithStrictness.lt false) = `Mathlib.Tactic.LinearCombination.div_lt_const_weak
Instances For
Given two (in)equalities P
and Q
, look up the lemma to deduce Q
from P
, and the relation
appearing in the side condition produced by this lemma.
Equations
- Mathlib.Ineq.eq.relImpRelData Mathlib.Ineq.eq = some (`Mathlib.Tactic.LinearCombination.eq_of_eq, Mathlib.Ineq.eq)
- Mathlib.Ineq.eq.relImpRelData Mathlib.Ineq.le = some (`Mathlib.Tactic.LinearCombination.le_of_eq, Mathlib.Ineq.le)
- Mathlib.Ineq.eq.relImpRelData Mathlib.Ineq.lt = some (`Mathlib.Tactic.LinearCombination.lt_of_eq, Mathlib.Ineq.lt)
- Mathlib.Ineq.le.relImpRelData Mathlib.Ineq.eq = none
- Mathlib.Ineq.le.relImpRelData Mathlib.Ineq.le = some (`Mathlib.Tactic.LinearCombination.le_of_le, Mathlib.Ineq.le)
- Mathlib.Ineq.le.relImpRelData Mathlib.Ineq.lt = some (`Mathlib.Tactic.LinearCombination.lt_of_le, Mathlib.Ineq.lt)
- Mathlib.Ineq.lt.relImpRelData Mathlib.Ineq.eq = none
- Mathlib.Ineq.lt.relImpRelData Mathlib.Ineq.le = some (`Mathlib.Tactic.LinearCombination.le_of_lt, Mathlib.Ineq.le)
- Mathlib.Ineq.lt.relImpRelData Mathlib.Ineq.lt = some (`Mathlib.Tactic.LinearCombination.lt_of_lt, Mathlib.Ineq.le)
Instances For
Given an (in)equality, look up the lemma to move everything to the LHS.
Equations
- Mathlib.Ineq.eq.rearrangeData = `Mathlib.Tactic.LinearCombination.eq_rearrange
- Mathlib.Ineq.le.rearrangeData = `Mathlib.Tactic.LinearCombination.le_rearrange
- Mathlib.Ineq.lt.rearrangeData = `Mathlib.Tactic.LinearCombination.lt_rearrange