Skip to content

Commit c1056b0

Browse files
committed
chore: rm @[eqns] in Ordmap (#11645)
Co-authored-by: Scott Morrison <scott.morrison@gmail.com>
1 parent 31188bc commit c1056b0

File tree

2 files changed

+9
-7
lines changed

2 files changed

+9
-7
lines changed

Mathlib/Data/Ordmap/Ordnode.lean

Lines changed: 7 additions & 5 deletions
Original file line numberDiff line numberDiff line change
@@ -138,14 +138,16 @@ def size : Ordnode α → ℕ
138138
| node sz _ _ _ => sz
139139
#align ordnode.size Ordnode.size
140140

141-
theorem size_nil : size (nil : Ordnode α) = 0 :=
141+
-- Adaptation note:
142+
-- During the port we marked these lemmas with `@[eqns]` to emulate the old Lean 3 behaviour.
143+
-- See https://github.com/leanprover-community/mathlib4/issues/11647
144+
145+
@[simp] theorem size_nil : size (nil : Ordnode α) = 0 :=
142146
rfl
143-
theorem size_node (sz : ℕ) (l : Ordnode α) (x : α) (r : Ordnode α) : size (node sz l x r) = sz :=
147+
@[simp] theorem size_node (sz : ℕ) (l : Ordnode α) (x : α) (r : Ordnode α) :
148+
size (node sz l x r) = sz :=
144149
rfl
145150

146-
attribute [eqns size_nil size_node] size
147-
attribute [simp] size
148-
149151
/-- O(1). Is the set empty?
150152
151153
empty ∅ = tt

Mathlib/Data/Ordmap/Ordset.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -808,7 +808,7 @@ theorem Raised.add_right (k) {n m} (H : Raised n m) : Raised (n + k) (m + k) :=
808808

809809
theorem Raised.right {l x₁ x₂ r₁ r₂} (H : Raised (size r₁) (size r₂)) :
810810
Raised (size (@node' α l x₁ r₁)) (size (@node' α l x₂ r₂)) := by
811-
rw [node', size, size]; generalize size r₂ = m at H ⊢
811+
rw [node', size_node, size_node]; generalize size r₂ = m at H ⊢
812812
rcases H with (rfl | rfl)
813813
· exact Or.inl rfl
814814
· exact Or.inr rfl
@@ -1395,7 +1395,7 @@ theorem Valid'.eraseMax_aux {s l x r o₁ o₂} (H : Valid' o₁ (.node s l x r)
13951395
rcases IHrr H.right with ⟨h, e⟩
13961396
refine' ⟨Valid'.balanceL H.left h (Or.inr ⟨_, Or.inr e, H.3.1⟩), _⟩
13971397
rw [eraseMax, size_balanceL H.3.2.1 h.3 H.2.2.1 h.2 (Or.inr ⟨_, Or.inr e, H.3.1⟩)]
1398-
rw [size, e]; rfl
1398+
rw [size_node, e]; rfl
13991399
#align ordnode.valid'.erase_max_aux Ordnode.Valid'.eraseMax_aux
14001400

14011401
theorem Valid'.eraseMin_aux {s l} {x : α} {r o₁ o₂} (H : Valid' o₁ (.node s l x r) o₂) :

0 commit comments

Comments
 (0)