Skip to content

Commit

Permalink
FVarIdSubst: eliminate Std.HashMap
Browse files Browse the repository at this point in the history
... to reduce dependency on Std.
  • Loading branch information
JLimperg committed Aug 6, 2024
1 parent 04f9302 commit 8b31db8
Show file tree
Hide file tree
Showing 3 changed files with 19 additions and 16 deletions.
11 changes: 6 additions & 5 deletions Aesop/RuleTac/FVarIdSubst.lean
Original file line number Diff line number Diff line change
Expand Up @@ -4,14 +4,14 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Jannis Limperg
-/

import Batteries.Data.HashMap
import Lean

open Lean Lean.Meta

namespace Aesop

structure FVarIdSubst where
map : Batteries.HashMap FVarId FVarId
map : HashMap FVarId FVarId
deriving Inhabited

namespace FVarIdSubst
Expand Down Expand Up @@ -53,15 +53,16 @@ instance : EmptyCollection FVarIdSubst :=
⟨.empty⟩

def insert (s : FVarIdSubst) (old new : FVarId) : FVarIdSubst :=
let map : Batteries.HashMap _ _ := s.map.mapVal λ _ v =>
if v == old then new else v
let map : HashMap _ _ := s.map.fold (init := ∅) λ map k v =>
map.insert k $ if v == old then new else v
⟨map.insert old new⟩

def erase (s : FVarIdSubst) (fvarId : FVarId) : FVarIdSubst :=
⟨s.map.erase fvarId⟩

def append (s t : FVarIdSubst) : FVarIdSubst :=
let map := s.map.mapVal λ _ v => t.get v
let map : HashMap _ _ := s.map.fold (init := ∅) λ map k v =>
map.insert k $ t.get v
⟨t.map.fold (init := map) λ s k v => s.insert k v⟩

def ofFVarSubstIgnoringNonFVarIds (s : FVarSubst) : FVarIdSubst := .mk $
Expand Down
11 changes: 1 addition & 10 deletions Aesop/RuleTac/GoalDiff.lean
Original file line number Diff line number Diff line change
Expand Up @@ -5,6 +5,7 @@ Authors: Jannis Limperg
-/

import Aesop.RuleTac.FVarIdSubst
import Aesop.Util.Basic

open Lean Lean.Meta

Expand Down Expand Up @@ -88,16 +89,6 @@ def diffGoals (old new : MVarId) (fvarSubst : FVarIdSubst) :
fvarSubst
}

-- TODO upstream
local instance [BEq α] [Hashable α] [Monad m] :
ForM m (Batteries.HashMap α β) (α × β) where
forM | m, f => m.forM λ a b => f (a, b)

-- TODO upstream
local instance [BEq α] [Hashable α] [Monad m] :
ForIn m (Batteries.HashMap α β) (α × β) where
forIn := ForM.forIn

namespace GoalDiff

/--
Expand Down
13 changes: 12 additions & 1 deletion Aesop/Util/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -11,7 +11,7 @@ import Batteries.Data.String
import Batteries.Lean.Expr
import Batteries.Lean.Meta.DiscrTree
import Batteries.Lean.PersistentHashSet
import Lean.Meta.Tactic.TryThis
import Lean

open Lean
open Lean.Meta Lean.Elab.Tactic
Expand Down Expand Up @@ -64,6 +64,17 @@ def toArray [BEq α] [Hashable α] (s : PersistentHashSet α) :

end PersistentHashSet

-- TODO upstream; generalise to {m : Type u → Type v}.
-- Need to generalise `HashMap.forM` first.
scoped instance {m : Type u → Type u} [BEq α] [Hashable α] [Monad m] :
ForM m (HashMap α β) (α × β) where
forM | m, f => m.forM λ a b => f (a, b)

-- TODO upstream; generalise to {m : Type u → Type v}.
scoped instance {m : Type u → Type u} [BEq α] [Hashable α] [Monad m] :
ForIn m (HashMap α β) (α × β) where
forIn := ForM.forIn

section DiscrTree

open DiscrTree
Expand Down

0 comments on commit 8b31db8

Please sign in to comment.