Skip to content
Draft
Show file tree
Hide file tree
Changes from all commits
Commits
File filter

Filter by extension

Filter by extension

Conversations
Failed to load comments.
Loading
Jump to
Jump to file
Failed to load files.
Loading
Diff view
Diff view
17 changes: 17 additions & 0 deletions src/Lean/Data/NameMap/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -52,6 +52,7 @@ instance : EmptyCollection NameSet := ⟨empty⟩
instance : Inhabited NameSet := ⟨empty⟩
def insert (s : NameSet) (n : Name) : NameSet := Std.TreeSet.insert s n
def contains (s : NameSet) (n : Name) : Bool := Std.TreeSet.contains s n
instance : Membership Name NameSet := ⟨fun (s : NameSet) (n : Name) => s.contains n = true⟩
instance [Monad m] : ForIn m NameSet Name :=
inferInstanceAs (ForIn _ (Std.TreeSet _ _) ..)

Expand Down Expand Up @@ -91,6 +92,8 @@ instance : EmptyCollection NameSSet := ⟨empty⟩
instance : Inhabited NameSSet := ⟨empty⟩
abbrev insert (s : NameSSet) (n : Name) : NameSSet := SSet.insert s n
abbrev contains (s : NameSSet) (n : Name) : Bool := SSet.contains s n
instance : Membership Name NameSSet := ⟨fun (s : NameSSet) (n : Name) => s.contains n = true⟩
instance [Monad m] : ForIn m NameSSet Name := inferInstanceAs (ForIn _ (SSet _) ..)
end NameSSet

@[expose] def NameHashSet := Std.HashSet Name
Expand All @@ -101,9 +104,23 @@ instance : EmptyCollection NameHashSet := ⟨empty⟩
instance : Inhabited NameHashSet := ⟨{}⟩
def insert (s : NameHashSet) (n : Name) := Std.HashSet.insert s n
def contains (s : NameHashSet) (n : Name) : Bool := Std.HashSet.contains s n
instance : Membership Name NameHashSet := ⟨fun (s : NameHashSet) (n : Name) => s.contains n = true⟩

/-- `filter f s` returns the `NameHashSet` consisting of all `x` in `s` where `f x` returns `true`. -/
def filter (f : Name → Bool) (s : NameHashSet) : NameHashSet := Std.HashSet.filter f s

/-- The union of two `NameHashSet`s. -/
def append (s t : NameHashSet) : NameHashSet :=
t.fold (init := s) fun s n => s.insert n

instance : Append NameHashSet := ⟨NameHashSet.append⟩

instance : Singleton Name NameHashSet where
singleton := fun n => (∅ : NameHashSet).insert n

instance : Union NameHashSet where
union := NameHashSet.append

end NameHashSet

def MacroScopesView.isPrefixOf (v₁ v₂ : MacroScopesView) : Bool :=
Expand Down
3 changes: 3 additions & 0 deletions src/Lean/Data/PersistentHashSet.lean
Original file line number Diff line number Diff line change
Expand Up @@ -51,6 +51,9 @@ variable {_ : BEq α} {_ : Hashable α}
@[inline] def contains (s : PersistentHashSet α) (a : α) : Bool :=
s.set.contains a

instance {_ : BEq α} {_ : Hashable α} : Membership α (PersistentHashSet α) :=
⟨fun (s : PersistentHashSet α) (a : α) => s.contains a = true⟩

@[inline] def foldM {β : Type v} {m : Type v → Type w} [Monad m] (f : β → α → m β) (init : β) (s : PersistentHashSet α) : m β :=
s.set.foldlM (init := init) fun d a _ => f d a

Expand Down
13 changes: 13 additions & 0 deletions src/Lean/Data/Position.lean
Original file line number Diff line number Diff line change
Expand Up @@ -22,6 +22,19 @@ namespace Position
protected def lt : Position → Position → Bool
| ⟨l₁, c₁⟩, ⟨l₂, c₂⟩ => Prod.lexLt (l₁, c₁) (l₂, c₂)

instance : LT Position := ⟨fun a b => a.lt b⟩
instance (a b : Position) : Decidable (a < b) := inferInstanceAs (Decidable (a.lt b))

protected def compare : Position → Position → Ordering
| ⟨l₁, c₁⟩, ⟨l₂, c₂⟩ =>
match Ord.compare l₁ l₂ with
| .eq => Ord.compare c₁ c₂
| ord => ord

instance : Ord Position := ⟨Position.compare⟩

instance : LE Position := leOfOrd

instance : ToFormat Position :=
⟨fun ⟨l, c⟩ => "⟨" ++ format l ++ ", " ++ format c ++ "⟩"⟩

Expand Down
14 changes: 14 additions & 0 deletions src/Lean/Data/RBTree.lean
Original file line number Diff line number Diff line change
Expand Up @@ -91,6 +91,8 @@ instance [Repr α] : Repr (RBTree α cmp) where
@[inline] def contains (t : RBTree α cmp) (a : α) : Bool :=
(t.find? a).isSome

instance : Membership α (RBTree α cmp) := ⟨fun (t : RBTree α cmp) (a : α) => t.contains a = true⟩

def fromList (l : List α) (cmp : α → α → Ordering) : RBTree α cmp :=
l.foldl insert (mkRBTree α cmp)

Expand All @@ -115,9 +117,21 @@ def union (t₁ t₂ : RBTree α cmp) : RBTree α cmp :=
else
t₂.fold .insert t₁

instance : Append (RBTree α cmp) := ⟨union⟩

instance : Union (RBTree α cmp) := ⟨union⟩

def diff (t₁ t₂ : RBTree α cmp) : RBTree α cmp :=
t₂.fold .erase t₁

instance : SDiff (RBTree α cmp) := ⟨diff⟩

instance : Singleton α (RBTree α cmp) where
singleton := fun a => (mkRBTree α cmp).insert a

instance : Inter (RBTree α cmp) where
inter := fun s t => s.fold (fun r a => if t.contains a then r.insert a else r) (mkRBTree α cmp)

/--
`filter f m` returns the `RBTree` consisting of all
`x` in `m` where `f x` returns `true`.
Expand Down
9 changes: 9 additions & 0 deletions src/Lean/Data/SSet.lean
Original file line number Diff line number Diff line change
Expand Up @@ -28,9 +28,18 @@ abbrev insert (s : SSet α) (a : α) : SSet α :=
abbrev contains (s : SSet α) (a : α) : Bool :=
SMap.contains s a

instance : Membership α (SSet α) := ⟨fun (s : SSet α) (a : α) => s.contains a = true⟩

abbrev forM [Monad m] (s : SSet α) (f : α → m PUnit) : m PUnit :=
SMap.forM s fun a _ => f a

protected def forIn [Monad m] (s : SSet α) (init : σ) (f : α → σ → m (ForInStep σ)) : m σ := do
let inst : ForIn m (SMap α Unit) (α × Unit) := inferInstance
inst.forIn s init fun (a, _) acc => f a acc

instance [Monad m] : ForIn m (SSet α) α where
forIn := SSet.forIn

/-- Move from stage 1 into stage 2. -/
abbrev switch (s : SSet α) : SSet α :=
SMap.switch s
Expand Down
5 changes: 5 additions & 0 deletions src/Lean/Declaration.lean
Original file line number Diff line number Diff line change
Expand Up @@ -61,6 +61,11 @@ def lt : ReducibilityHints → ReducibilityHints → Bool
| .regular _, .opaque => true
| _, _ => false

instance : LT ReducibilityHints := ⟨fun a b => a.lt b⟩
instance (a b : ReducibilityHints) : Decidable (a < b) := inferInstanceAs (Decidable (a.lt b))
instance : LE ReducibilityHints := ⟨fun a b => a < b ∨ a = b⟩
instance (a b : ReducibilityHints) : Decidable (a ≤ b) := inferInstanceAs (Decidable (a < b ∨ a = b))

protected def compare : ReducibilityHints → ReducibilityHints → Ordering
| .abbrev, .abbrev => .eq
| .abbrev, _ => .lt
Expand Down
2 changes: 2 additions & 0 deletions src/Lean/Meta/Tactic/FVarSubst.lean
Original file line number Diff line number Diff line change
Expand Up @@ -74,6 +74,8 @@ This satisfies `(s.append t).apply e = t.apply (s.apply e)`
def append (s t : FVarSubst) : FVarSubst :=
s.1.foldl (fun s' k v => s'.insert k (t.apply v)) t

instance : Append FVarSubst := ⟨append⟩

end FVarSubst
end Meta

Expand Down
22 changes: 22 additions & 0 deletions src/Lean/Meta/TransparencyMode.lean
Original file line number Diff line number Diff line change
Expand Up @@ -32,6 +32,28 @@ def lt : TransparencyMode → TransparencyMode → Bool
| default, all => true
| _, _ => false

instance : LT TransparencyMode := ⟨fun a b => a.lt b⟩
instance (a b : TransparencyMode) : Decidable (a < b) := inferInstanceAs (Decidable (a.lt b))

instance : LE TransparencyMode := ⟨fun a b => a < b ∨ a = b⟩
instance (a b : TransparencyMode) : Decidable (a ≤ b) := inferInstanceAs (Decidable (a < b ∨ a = b))

protected def beq : TransparencyMode → TransparencyMode → Bool
| none, none => true
| reducible, reducible => true
| instances, instances => true
| default, default => true
| all, all => true
| _, _ => false

instance : BEq TransparencyMode := ⟨TransparencyMode.beq⟩

protected def compare (a b : TransparencyMode) : Ordering :=
if a.lt b then .lt else if b.lt a then .gt else .eq

instance : Ord TransparencyMode where
compare := TransparencyMode.compare

end TransparencyMode

example (a b c : TransparencyMode) : a.lt b → b.lt c → a.lt c := by
Expand Down
2 changes: 2 additions & 0 deletions src/Lean/SubExpr.lean
Original file line number Diff line number Diff line change
Expand Up @@ -81,6 +81,8 @@ def all (pred : Nat → Bool) (p : Pos) : Bool :=

def append : Pos → Pos → Pos := foldl push

instance : Append Pos := ⟨append⟩

/-- Creates a subexpression `Pos` from an array of 'coordinates'.
Each coordinate is a number {0,1,2} expressing which child subexpression should be explored.
The first coordinate in the array corresponds to the root of the expression tree. -/
Expand Down
Loading