The OmegaM state monad. #
We keep track of the linear atoms (up to defeq) that have been encountered so far, and also generate new facts as new atoms are recorded.
The main functions are:
- atoms : OmegaM (List Expr)which returns the atoms recorded so far
- lookup (e : Expr) : OmegaM (Nat × Option (HashSet Expr))which checks if an- Exprhas already been recorded as an atom, and records it.- lookupreturn the index in- atomsfor this- Expr. The- Option (HashSet Expr)return value is- noneis the expression has been previously recorded, and otherwise contains new facts that should be added to the- omegaproblem.- for each new atom aof the form((x : Nat) : Int), the fact that0 ≤ a
- for each new atom aof the formx / k, forka positive numeral, the facts thatk * a ≤ x < k * a + k
- for each new atom of the form ((a - b : Nat) : Int), the fact:b ≤ a ∧ ((a - b : Nat) : Int) = a - b ∨ a < b ∧ ((a - b : Nat) : Int) = 0
- for each new atom of the form min a b, the factsmin a b ≤ aandmin a b ≤ b(and similarly formax)
- for each new atom of the form if P then a else b, the disjunction:(P ∧ (if P then a else b) = a) ∨ (¬ P ∧ (if P then a else b) = b)TheOmegaMmonad also keeps an internal cache of visited expressions (not necessarily atoms, but arbitrary subexpressions of one side of a linear relation) to reduce duplication. The cache mapsExprs to pairs consisting of aLinearCombo, and proof that the expression is equal to the evaluation of theLinearComboat the atoms.
 
- for each new atom 
Context for the OmegaM monad, containing the user configurable options.
- User configurable options for - omega.
Instances For
The internal state for the OmegaM monad, recording previously encountered atoms.
- atoms : Lean.HashMap Lean.Expr NatThe atoms up-to-defeq encountered so far. 
Instances For
Cache of expressions that have been visited, and their reflection as a linear combination.
Equations
Instances For
The OmegaM monad maintains two pieces of state:
- the linear atoms discovered while processing hypotheses
- a cache mapping subexpressions of one side of a linear inequality to LinearCombos (and a proof that theLinearComboevaluates at the atoms to the original expression).
Equations
Instances For
Run a computation in the OmegaM monad, starting with no recorded atoms.
Equations
- m.run cfg = (StateRefT'.run' (StateRefT'.run' m Lean.HashMap.empty) { atoms := ∅ } { cfg := cfg }).run' Lean.Meta.TransparencyMode.instances
Instances For
Retrieve the user-specified configuration options.
Equations
- Lean.Elab.Tactic.Omega.cfg = do let __do_lift ← read pure __do_lift.cfg
Instances For
Retrieve the list of atoms.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Return the Expr representing the list of atoms.
Equations
- Lean.Elab.Tactic.Omega.atomsList = do let __do_lift ← Lean.Elab.Tactic.Omega.atoms liftM (Lean.Meta.mkListLit (Lean.Expr.const `Int []) __do_lift.toList)
Instances For
Return the Expr representing the list of atoms as a Coeffs.
Equations
- Lean.Elab.Tactic.Omega.atomsCoeffs = do let __do_lift ← Lean.Elab.Tactic.Omega.atomsList pure ((Lean.Expr.const `Lean.Omega.Coeffs.ofList []).app __do_lift)
Instances For
Run an OmegaM computation, restoring the state afterwards depending on the result.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Run an OmegaM computation, restoring the state afterwards.
Equations
- Lean.Elab.Tactic.Omega.withoutModifyingState t = Lean.Elab.Tactic.Omega.commitWhen do let __do_lift ← t pure (__do_lift, false)
Instances For
Wrapper around Expr.nat? that also allows Nat.cast.
Equations
- Lean.Elab.Tactic.Omega.natCast? n = match n.getAppFnArgs with | (`Nat.cast, #[head, head_1, n]) => n.nat? | x => n.nat?
Instances For
Wrapper around Expr.int? that also allows Nat.cast.
Equations
- Lean.Elab.Tactic.Omega.intCast? n = match n.getAppFnArgs with | (`Nat.cast, #[head, head_1, n]) => do let a ← n.nat? pure ↑a | x => n.int?
Instances For
If groundNat? e = some n, then e is definitionally equal to OfNat.ofNat n.
If groundInt? e = some i,
then e is definitionally equal to the standard expression for i.
Construct the term with type hint (Eq.refl a : a = b)
Equations
- Lean.Elab.Tactic.Omega.mkEqReflWithExpectedType a b = do let __do_lift ← Lean.Meta.mkEqRefl a let __do_lift_1 ← Lean.Meta.mkEq a b Lean.Meta.mkExpectedTypeHint __do_lift __do_lift_1
Instances For
Analyzes a newly recorded atom, returning a collection of interesting facts about it that should be added to the context.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Look up an expression in the atoms, recording it if it has not previously appeared.
Return its index, and, if it is new, a collection of interesting facts about the atom.
- for each new atom aof the form((x : Nat) : Int), the fact that0 ≤ a
- for each new atom aof the formx / k, forka positive numeral, the facts thatk * a ≤ x < k * a + k
- for each new atom of the form ((a - b : Nat) : Int), the fact:b ≤ a ∧ ((a - b : Nat) : Int) = a - b ∨ a < b ∧ ((a - b : Nat) : Int) = 0
Equations
- One or more equations did not get rendered due to their size.