Skip to content

Commit

Permalink
chore(data/set/lattice): a few lemmas, golf (#8795)
Browse files Browse the repository at this point in the history
  • Loading branch information
urkud committed Aug 22, 2021
1 parent d3e20b4 commit f915106
Showing 1 changed file with 18 additions and 12 deletions.
30 changes: 18 additions & 12 deletions src/data/set/lattice.lean
Expand Up @@ -365,6 +365,10 @@ variables {s : ι → set α}
@[simp] lemma nonempty_Union : (⋃ i, s i).nonempty ↔ ∃ i, (s i).nonempty :=
by simp [← ne_empty_iff_nonempty]

lemma Union_nonempty_index (s : set α) (t : s.nonempty → set β) :
(⋃ h, t h) = ⋃ x ∈ s, t ⟨x, ‹_›⟩ :=
supr_exists

end

@[simp] theorem Inter_Inter_eq_left {b : β} {s : Π x : β, x = b → set α} :
Expand Down Expand Up @@ -529,6 +533,12 @@ infi_emptyset
theorem bInter_univ (u : α → set β) : (⋂ x ∈ @univ α, u x) = ⋂ x, u x :=
infi_univ

@[simp] lemma bUnion_self (s : set α) : (⋃ x ∈ s, s) = s :=
subset.antisymm (bUnion_subset $ λ x hx, subset.refl s) (λ x hx, mem_bUnion hx hx)

@[simp] lemma Union_nonempty_self (s : set α) : (⋃ h : s.nonempty, s) = s :=
by rw [Union_nonempty_index, bUnion_self]

-- TODO(Jeremy): here is an artifact of the encoding of bounded intersection:
-- without dsimp, the next theorem fails to type check, because there is a lambda
-- in a type that needs to be contracted. Using simp [eq_of_mem_singleton xa] also works.
Expand Down Expand Up @@ -797,6 +807,10 @@ lemma sUnion_inter_sUnion {s t : set (set α)} :
(⋃₀ s) ∩ (⋃₀ t) = (⋃ p ∈ s.prod t, (p : (set α) × (set α )).1 ∩ p.2) :=
Sup_inf_Sup

lemma bUnion_Union (s : ι → set α) (t : α → set β) :
(⋃ x ∈ ⋃ i, s i, t x) = ⋃ i (x ∈ s i), t x :=
by simp [@Union_comm _ ι]

/-- If `S` is a set of sets, and each `s ∈ S` can be represented as an intersection
of sets `T s hs`, then `⋂₀ S` is the intersection of the union of all `T s hs`. -/
lemma sInter_bUnion {S : set (set α)} {T : Π s ∈ S, set (set α)}
Expand Down Expand Up @@ -855,18 +869,11 @@ end

lemma union_distrib_Inter_right {ι : Type*} (s : ι → set α) (t : set α) :
(⋂ i, s i) ∪ t = (⋂ i, s i ∪ t) :=
begin
ext x,
rw [mem_union_eq, mem_Inter],
split; finish
end
infi_sup_eq _ _

lemma union_distrib_Inter_left {ι : Type*} (s : ι → set α) (t : set α) :
t ∪ (⋂ i, s i) = (⋂ i, t ∪ s i) :=
begin
rw [union_comm, union_distrib_Inter_right],
simp [union_comm]
end
sup_infi_eq _ _

lemma union_distrib_bInter_left {ι : Type*} (s : ι → set α) (u : set ι) (t : set α) :
t ∪ (⋂ i ∈ u, s i) = ⋂ i ∈ u, t ∪ s i :=
Expand Down Expand Up @@ -1047,9 +1054,8 @@ section image

lemma image_Union {f : α → β} {s : ι → set α} : f '' (⋃ i, s i) = (⋃ i, f '' s i) :=
begin
apply set.ext, intro x,
simp [image, exists_and_distrib_right.symm, -exists_and_distrib_right],
exact exists_swap
ext1 x,
simp [image, ← exists_and_distrib_right, @exists_swap α]
end

lemma image_bUnion {f : α → β} {s : ι → set α} {p : ι → Prop} :
Expand Down

0 comments on commit f915106

Please sign in to comment.