Open
Description
There has been some discussion of what tactics we would like in a few threads on Zulip. I am making this tracking issue to try to organize these ideas and put them in a visible place. Anyone with repo permission should feel free to edit or add to the list.
TODO: Integrate suggestions from this thread.
New Tactics
- Tendsto
- Should find limits in "simple" cases (e.g., no
0/0
etc). - Simplest version is
Continuous.tendsto' (by continuity) _ _ (by simp)
. - More advanced version should know that
exp(-1/x)
tends to𝓝[>] 0
asx
tends toatTop
.
*Should know that 1/x tends to cobounded asx
tends to𝓝[≠] 0
and vice versa (for any normed field, not only reals)
- Should find limits in "simple" cases (e.g., no
- Exists positive from limit
- Prove statements like
(δ : Real) (h : 0 < δ) ⊢ ∃ ε > 0, ε ^ 2 + 5 * ε + sin ε < δ ∧ 3 * ε < δ
by provingTendsto (fun ε ↦ ε ^ 2 + 5 * ε + sin ε) (nhds 0) (nhds 0)
, similarly for3 * ε
, then using this fact to get a witness
- Prove statements like
- Nonzero A tactic that proves
expr ≠ 0
.- at least as powerful as whatever
field_simp
currently uses (it tries several tactics); - knows lemmas like
a ≠ 0 → -a ≠ 0, a ≠ 0 → b ≠ 0 → a * b ≠ 0 and a ≠ 0 → a ^ n ≠ 0
; - fallbacks to positivity if it can't deal with the head symbol (e.g.,
+
) and there is aPartialOrder
instance
- at least as powerful as whatever
- Tactic that can compute things about polynomials
- Polynomial equality
- Polynomial disequality
- n-th coefficient, leading coefficient
- degree (
compute_degree
) - monicity (is that the word for the property of being monic?)
- ring in characteristic n
- More tactics resolving statements in decidable first-order theories
- Such as:
- Presburger arithmetic. (Is this just
omega
?) - Reals.
- Lists.
- Presburger arithmetic. (Is this just
- Often these are hard to write fast tactics for.
- Nevertheless, I think they could often be helpful.
- Maybe the best version of these tactics is to use a look-up table for common instances.
- Such as:
- Tactic to decide the order on logarithmico-exponential functions
- Tactic to provide approximations of real-valued expressions (i.e prove 3.83 < ln(2) + pi < 3.84)
- See (https://github.com/Timeroot/ComputableReal/tree/main) which does this, albeit with some feature yet-to-be-added:
-
pi
-
exp
-
log
- ... see the link for more
-
- See (https://github.com/Timeroot/ComputableReal/tree/main) which does this, albeit with some feature yet-to-be-added:
-
push
tactic that generalizespush_neg
fromneg
to any def. #21841- Pushes any instance of the def it is fed to the leaves of the syntax tree. For example,
push Real.log
would applylog_mul
.
- Pushes any instance of the def it is fed to the leaves of the syntax tree. For example,
-
field
field tactic #4837 -
module_nf
-
recommend
- Priority mechanism for
hint
tactic #25302- Preferably in order of typical completion time, though we could also use exponential backoff to avoid spending too much time on expensive tactics.
- Have a
?
version that gives aTry This
on success.
- a tactic that rewrites modulo associativity and commutativity:#general > "Missing Tactics" list @ 💬
Enhancements of existing Tactics
- Improved positivity
- more extensions
- an attribute that automatically produces correct positivity extension in a simple case (e.g., we only know how to prove
0 ≤ f a
, never0 < f a
, or we always can prove0 < f a
) - support for goals
expr < 0
, so that it can prove-3 * ε - δ < 0
- Ability to supply terms
-
zify!
zify!
tactic #7450 rify
- There should be a
by_contra!
similar tocontrapose!
- A preprocessor for linarith which identifies calls to natural division/modulo and introduces the facts
0 ≤ n%d
n%d < d
d * n/d + n%d = n
- Potential problem: if
d = 0
these do not all hold.
- Make
abel
work for multiplicative monoids/groups. - Make
linear_combination
work for groups. - Reimplement
polyrith
in pure Lean, removing the dependence on sage and an internet connection. - Boost
omega
into a full decision procedure for Presburger arithmetic. clear_unneeded
:- A
clear_unneeded
tactic #25319 - If
(a : α)
is a hypothesis andα
is Nonempty, buta
is never used elsewhere, deletea
. - If
(h : a = ...)
if a hypothesis anda
is never used elsewhere, deleteh
.
- A
- a version of
simp_rw
that can create new goals for side-conditions - a version of
field_simp
that collects non-zeroness as side conditions. - More
?
versions of tactics-
unfold?
to list possible unfold applications. ([Merged by Bors] - feat: interactiveunfold?
tactic #12016) -
assumption?
(replaces byexact
) -
field_simp?
-
linarith?/nlinarith?
#25320- (replaces by
linarith only
) - Or even better, by a call to
linear_combination
- (replaces by
-
- Let
norm_cast
useFin.cast_val_eq_self
, soNat.cast (Fin.val n)
simplifies ton
. -
cases (r:ℚ)
should replacer
with_ / _
instead of an application ofRat.mk'
- push_neg should be aware of Infinite, Finite, Set.Infinite and Set.Finite, Nat.Odd, Nat.Even
- a ring_nf option to normalize a ring equality a = b to a - b = 0
- Let
linear_combination
be used nonterminally. -
suffices
should work with incremental elaboration - in conv mode, simp, unfold, etc should throw errors when doing nothing (like they usually do)
- Version of
plausible
that produces a counterexample demonstration. - A option for the
simp?
tactic that, instead of outputting asimp only
call, outputs asimp_rw
call.- This would make it easier to step through the transformations of simp call when debugging it.
- Make
norm_num
convert Gaussian rationals to the forma_n / a_d + b_n / b_d * I
- A variant of
extract_goal
calledextract_mwe
that provides the same things, but with additional import and opened namespace information. -
ext at
#general > "Missing Tactics" list @ 💬
simproc
s
- A simproc that simplifies expressions of the form
∃ a', ... ∧ a' = a ∧ ...
(Zulip). ([Merged by Bors] - feat(Simproc): simproc for∃ a', ... ∧ a' = a ∧ ...
#22273) - TODO: The contents of #mathlib4 > potential simp_procs @ 💬
As-Yet Unported mathlib3 Tactics
- TODO identify these