Skip to content

Commit

Permalink
chore: update Std (#11858)
Browse files Browse the repository at this point in the history
Co-authored-by: Scott Morrison <scott.morrison@gmail.com>
  • Loading branch information
semorrison and semorrison committed Apr 3, 2024
1 parent 8d5f6cf commit f28589d
Show file tree
Hide file tree
Showing 6 changed files with 9 additions and 7 deletions.
2 changes: 1 addition & 1 deletion Mathlib/Data/List/Intervals.lean
Expand Up @@ -112,7 +112,7 @@ theorem inter_consecutive (n m l : ℕ) : Ico n m ∩ Ico m l = [] := by

@[simp]
theorem bagInter_consecutive (n m l : Nat) : @List.bagInter ℕ instBEq (Ico n m) (Ico m l) = [] :=
(bagInter_nil_iff_inter_nil _ _).2 (inter_consecutive n m l)
(bagInter_nil_iff_inter_nil _ _).2 (by convert inter_consecutive n m l)
#align list.Ico.bag_inter_consecutive List.Ico.bagInter_consecutive

@[simp]
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Data/List/Lattice.lean
Expand Up @@ -181,7 +181,7 @@ theorem forall_mem_inter_of_forall_right (l₁ : List α) (h : ∀ x ∈ l₂, p

@[simp]
theorem inter_reverse {xs ys : List α} : xs.inter ys.reverse = xs.inter ys := by
simp only [List.inter, mem_reverse]
simp only [List.inter, elem_eq_mem, mem_reverse]
#align list.inter_reverse List.inter_reverse

end Inter
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Data/List/Nodup.lean
Expand Up @@ -469,7 +469,7 @@ theorem Nodup.pairwise_coe [IsSymm α r] (hl : l.Nodup) :

-- Porting note (#10756): new theorem
theorem Nodup.take_eq_filter_mem [DecidableEq α] :
∀ {l : List α} {n : ℕ} (_ : l.Nodup), l.take n = l.filter (· ∈ l.take n)
∀ {l : List α} {n : ℕ} (_ : l.Nodup), l.take n = l.filter (l.take n).elem
| [], n, _ => by simp
| b::l, 0, _ => by simp
| b::l, n+1, hl => by
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Data/List/Perm.lean
Expand Up @@ -607,7 +607,7 @@ theorem Perm.erasep (f : α → Prop) [DecidablePred f] {l₁ l₂ : List α}
theorem Perm.take_inter {α : Type*} [DecidableEq α] {xs ys : List α} (n : ℕ) (h : xs ~ ys)
(h' : ys.Nodup) : xs.take n ~ ys.inter (xs.take n) := by
simp only [List.inter]
exact Perm.trans (show xs.take n ~ xs.filter (· ∈ xs.take n) by
exact Perm.trans (show xs.take n ~ xs.filter (xs.take n).elem by
conv_lhs => rw [Nodup.take_eq_filter_mem ((Perm.nodup_iff h).2 h')])
(Perm.filter _ h)
#align list.perm.take_inter List.Perm.take_inter
Expand Down
6 changes: 4 additions & 2 deletions Mathlib/Data/Multiset/FinsetOps.lean
Expand Up @@ -217,8 +217,10 @@ def ndinter (s t : Multiset α) : Multiset α :=
#align multiset.ndinter Multiset.ndinter

@[simp]
theorem coe_ndinter (l₁ l₂ : List α) : @ndinter α _ l₁ l₂ = (l₁ ∩ l₂ : List α) :=
rfl
theorem coe_ndinter (l₁ l₂ : List α) : @ndinter α _ l₁ l₂ = (l₁ ∩ l₂ : List α) := by
simp only [ndinter, mem_coe, filter_coe, coe_eq_coe, ← elem_eq_mem]
apply Perm.refl

#align multiset.coe_ndinter Multiset.coe_ndinter

@[simp, nolint simpNF] -- Porting note (#10675): dsimp can not prove this
Expand Down
2 changes: 1 addition & 1 deletion lake-manifest.json
Expand Up @@ -4,7 +4,7 @@
[{"url": "https://github.com/leanprover/std4",
"type": "git",
"subDir": null,
"rev": "67c40f89be2c17c6f4b0b6056d2de8591c0f92d3",
"rev": "d4161291e2a4c1c92d710bf670570aa79bf0d6ef",
"name": "std",
"manifestFile": "lake-manifest.json",
"inputRev": "main",
Expand Down

0 comments on commit f28589d

Please sign in to comment.