Skip to content
New issue

Have a question about this project? Sign up for a free GitHub account to open an issue and contact its maintainers and the community.

By clicking “Sign up for GitHub”, you agree to our terms of service and privacy statement. We’ll occasionally send you account related emails.

Already on GitHub? Sign in to your account

[Merged by Bors] - chore(*Set): golf #12117

Closed
wants to merge 3 commits into from
Closed
Show file tree
Hide file tree
Changes from all commits
Commits
File filter

Filter by extension

Filter by extension

Conversations
Failed to load comments.
Jump to
Jump to file
Failed to load files.
Diff view
Diff view
8 changes: 3 additions & 5 deletions Mathlib/Data/Fintype/Powerset.lean
Original file line number Diff line number Diff line change
Expand Up @@ -59,12 +59,10 @@ theorem Fintype.card_finset_len [Fintype α] (k : ℕ) :
#align fintype.card_finset_len Fintype.card_finset_len

instance Set.fintype [Fintype α] : Fintype (Set α) :=
⟨(@Finset.univ α _).powerset.map ⟨(↑), coe_injective⟩, fun s => by
⟨(@Finset.univ (Finset α) _).map coeEmb.1, fun s => by
classical
refine' mem_map.2 ⟨Finset.univ.filter s, Finset.mem_powerset.2 (Finset.subset_univ _), _⟩
apply (coe_filter _ _).trans
simp
rfl⟩
refine mem_map.2 ⟨Finset.univ.filter (· ∈ s), Finset.mem_univ _, (coe_filter _ _).trans ?_⟩
simp⟩
#align set.fintype Set.fintype

-- Not to be confused with `Set.Finite`, the predicate
Expand Down
8 changes: 4 additions & 4 deletions Mathlib/Data/Set/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -1425,12 +1425,12 @@ theorem sep_false : { x ∈ s | False } = ∅ :=

--Porting note (#10618): removed `simp` attribute because `simp` can prove it
theorem sep_empty (p : α → Prop) : { x ∈ (∅ : Set α) | p x } = ∅ :=
empty_inter p
empty_inter {x | p x}
#align set.sep_empty Set.sep_empty

--Porting note (#10618): removed `simp` attribute because `simp` can prove it
theorem sep_univ : { x ∈ (univ : Set α) | p x } = { x | p x } :=
univ_inter p
univ_inter {x | p x}
#align set.sep_univ Set.sep_univ

@[simp]
Expand All @@ -1440,12 +1440,12 @@ theorem sep_union : { x | (x ∈ s ∨ x ∈ t) ∧ p x } = { x ∈ s | p x }

@[simp]
theorem sep_inter : { x | (x ∈ s ∧ x ∈ t) ∧ p x } = { x ∈ s | p x } ∩ { x ∈ t | p x } :=
inter_inter_distrib_right s t p
inter_inter_distrib_right s t {x | p x}
#align set.sep_inter Set.sep_inter

@[simp]
theorem sep_and : { x ∈ s | p x ∧ q x } = { x ∈ s | p x } ∩ { x ∈ s | q x } :=
inter_inter_distrib_left s p q
inter_inter_distrib_left s {x | p x} {x | q x}
#align set.sep_and Set.sep_and

@[simp]
Expand Down
18 changes: 7 additions & 11 deletions Mathlib/Data/Set/Finite.lean
Original file line number Diff line number Diff line change
Expand Up @@ -1758,17 +1758,13 @@ lemma Set.finite_diff_iUnion_Ioo' (s : Set α) : (s \ ⋃ x : s × s, Ioo x.1 x.
lemma Directed.exists_mem_subset_of_finset_subset_biUnion {α ι : Type*} [Nonempty ι]
{f : ι → Set α} (h : Directed (· ⊆ ·) f) {s : Finset α} (hs : (s : Set α) ⊆ ⋃ i, f i) :
∃ i, (s : Set α) ⊆ f i := by
classical
revert hs
refine s.induction_on ?_ ?_
· simp
intro b t _hbt htc hbtc
obtain ⟨i : ι, hti : (t : Set α) ⊆ f i⟩ := htc (Set.Subset.trans (t.subset_insert b) hbtc)
obtain ⟨j, hbj⟩ : ∃ j, b ∈ f j := by simpa [Set.mem_iUnion₂] using hbtc (t.mem_insert_self b)
rcases h j i with ⟨k, hk, hk'⟩
use k
rw [Finset.coe_insert, Set.insert_subset_iff]
exact ⟨hk hbj, _root_.trans hti hk'⟩
induction s using Finset.cons_induction with
| empty => simp
| cons hbt iht =>
simp only [Finset.coe_cons, Set.insert_subset_iff, Set.mem_iUnion] at hs ⊢
rcases hs.imp_right iht with ⟨⟨i, hi⟩, j, hj⟩
rcases h i j with ⟨k, hik, hjk⟩
exact ⟨k, hik hi, hj.trans hjk⟩
#align directed.exists_mem_subset_of_finset_subset_bUnion Directed.exists_mem_subset_of_finset_subset_biUnion

theorem DirectedOn.exists_mem_subset_of_finset_subset_biUnion {α ι : Type*} {f : ι → Set α}
Expand Down
4 changes: 2 additions & 2 deletions Mathlib/Data/Set/Intervals/Disjoint.lean
Original file line number Diff line number Diff line change
Expand Up @@ -43,12 +43,12 @@ theorem Iio_disjoint_Ici (h : a ≤ b) : Disjoint (Iio a) (Ici b) :=

@[simp]
theorem Iic_disjoint_Ioc (h : a ≤ b) : Disjoint (Iic a) (Ioc b c) :=
(Iic_disjoint_Ioi h).mono le_rfl fun _ => And.left
(Iic_disjoint_Ioi h).mono le_rfl Ioc_subset_Ioi_self
#align set.Iic_disjoint_Ioc Set.Iic_disjoint_Ioc

@[simp]
theorem Ioc_disjoint_Ioc_same : Disjoint (Ioc a b) (Ioc b c) :=
(Iic_disjoint_Ioc (le_refl b)).mono (fun _ => And.right) le_rfl
(Iic_disjoint_Ioc le_rfl).mono Ioc_subset_Iic_self le_rfl
#align set.Ioc_disjoint_Ioc_same Set.Ioc_disjoint_Ioc_same

@[simp]
Expand Down