A map whose keys are expressions. Keys are identified up to defeq. We use
reducible transparency and treat metavariables as rigid (i.e.,
unassignable).
- rep : Lean.PArray (Option (Lean.Expr × α))
The mappings stored in the map. Defeq expressions are identified, so for each equivalence class of defeq expressions we only store one representative. Missing values indicate expressions that were removed from the map.
- idx : Lean.Meta.DiscrTree Nat
Instances For
Instances For
Equations
- Aesop.instInhabitedEMap = { default := Aesop.instInhabitedEMap.default }
Equations
- Aesop.EMap.empty = { rep := Lean.PersistentArray.empty, idx := Lean.Meta.DiscrTree.empty }
Instances For
Equations
- Aesop.EMap.instEmptyCollection = { emptyCollection := Aesop.EMap.empty }
def
Aesop.EMap.foldl
{σ : Type}
{α : Type u_1}
(init : σ)
(f : σ → Lean.Expr → α → σ)
(map : EMap α)
:
σ
Equations
- Aesop.EMap.foldl init f map = inline (Aesop.EMap.foldlM init f map)
Instances For
@[specialize #[]]
def
Aesop.EMap.alterM
{m : Type → Type u_1}
[Monad m]
[MonadLiftT Lean.MetaM m]
{α β : Type}
(e : Lean.Expr)
(f : α → m (Option α × β))
(map : EMap α)
:
Equations
- One or more equations did not get rendered due to their size.
Instances For
def
Aesop.EMap.alter
{α β : Type}
(e : Lean.Expr)
(f : α → Option α × β)
(map : EMap α)
:
Lean.MetaM (EMap α × Option β)
Equations
- Aesop.EMap.alter e f map = inline (Aesop.EMap.alterM e (fun (a : α) => pure (f a)) map)
Instances For
@[specialize #[]]
def
Aesop.EMap.insertNew
{m : Type → Type u_1}
[Monad m]
[MonadLiftT Lean.MetaM m]
{α : Type}
(e : Lean.Expr)
(a : α)
(map : EMap α)
:
m (EMap α)
Equations
Instances For
@[specialize #[]]
def
Aesop.EMap.insertWithM
{m : Type → Type u_1}
[Monad m]
[MonadLiftT Lean.MetaM m]
{α : Type}
(e : Lean.Expr)
(f : Option α → m α)
(map : EMap α)
:
Equations
- One or more equations did not get rendered due to their size.
Instances For
def
Aesop.EMap.insertWith
{α : Type}
(e : Lean.Expr)
(f : Option α → α)
(map : EMap α)
:
Lean.MetaM (EMap α × Option α × α)
Equations
- Aesop.EMap.insertWith e f map = inline (Aesop.EMap.insertWithM e (fun (a? : Option α) => pure (f a?)) map)
Instances For
Equations
- Aesop.EMap.insert e a map = (fun (x : Aesop.EMap α × Option α × α) => x.fst) <$> inline (Aesop.EMap.insertWithM e (fun (x : Option α) => pure a) map)
Instances For
def
Aesop.EMap.singleton
{m : Type → Type u_1}
[Monad m]
[MonadLiftT Lean.MetaM m]
{α : Type}
(e : Lean.Expr)
(a : α)
:
m (EMap α)
Equations
Instances For
def
Aesop.EMap.findWithKey?
{α : Type}
(e : Lean.Expr)
(map : EMap α)
:
Lean.MetaM (Option (Lean.Expr × α))
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- Aesop.EMap.find? e map = do let __do_lift ← inline (Aesop.EMap.findWithKey? e map) pure (Option.map (fun (x : Lean.Expr × α) => x.snd) __do_lift)
Instances For
Equations
- Aesop.EMap.map f map = Aesop.EMap.mapM f map