Running AtomM
metaprograms recursively #
Tactics such as ring
and abel
are implemented using the AtomM
monad, which tracks "atoms" --
expressions which cannot be further parsed according to the arithmetic operations they handle --
to allow for consistent normalization relative to these atoms.
This file provides methods to allow for such normalization to run recursively: the atoms themselves
will have the normalization run on any of their subterms for which this makes sense. For example,
given an implementation of ring-normalization, the methods in this file implement the bottom-to-top
recursive ring-normalization in which sin (x + y) + sin (y + x)
is normalized first to
sin (x + y) + sin (x + y)
and then to 2 * sin (x + y)
.
Main declarations #
Mathlib.Tactic.AtomM.RecurseM.run
: run a metaprogram (inAtomM
or its slight extensionAtomM.RecurseM
), with atoms normalized according to a provided normalization operation (inAtomM
), run recursively.Mathlib.Tactic.AtomM.recurse
: run a normalization operation (inAtomM
) recursively on an expression.
Configuration for AtomM.Recurse
.
the reducibility setting to use when comparing atoms for defeq
- zetaDelta : Bool
if true, local let variables can be unfolded
Instances For
Equations
The read-only state of the AtomM.Recurse
monad.
- ctx : Lean.Meta.Simp.Context
A basically empty simp context, passed to the
simp
traversal inAtomM.onSubexpressions
. A cleanup routine, which simplifies evaluation results to a more human-friendly format.
Instances For
The monad for AtomM.Recurse
contains, in addition to the AtomM
state,
a simp context for the main traversal and a cleanup function to simplify evaluation results.
Equations
Instances For
A tactic in the AtomM.RecurseM
monad which will simplify expression parent
to a normal form, by
running a core operation eval
(in the AtomM
monad) on the maximal subexpression(s) on which
eval
does not fail.
There is also a subsequent clean-up operation, governed by the context from the AtomM.RecurseM
monad.
root
: true if this is a direct call to the function.AtomM.RecurseM.run
sets this tofalse
in recursive mode.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Runs a tactic in the AtomM.RecurseM
monad, given initial data:
s
: a reference to the mutableAtomM
state, for persisting across calls. This ensures that atom ordering is used consistently.cfg
: the configuration optionseval
: a normalization operation which will be run recursively, potentially dependent on a known atom orderingsimp
: a cleanup operation which will be used to post-process expressionsx
: the tactic to run
Equations
- One or more equations did not get rendered due to their size.
Instances For
The recursive context.
The atom evaluator calls AtomM.onSubexpressions
recursively.
Normalizes an expression, given initial data:
s
: a reference to the mutableAtomM
state, for persisting across calls. This ensures that atom ordering is used consistently.cfg
: the configuration optionseval
: a normalization operation which will be run recursively, potentially dependent on a known atom orderingsimp
: a cleanup operation which will be used to post-process expressionstgt
: the expression to normalize
Equations
- Mathlib.Tactic.AtomM.recurse s cfg eval simp tgt = Mathlib.Tactic.AtomM.RecurseM.run s cfg eval simp (Mathlib.Tactic.AtomM.onSubexpressions eval tgt)