Skip to content

Commit

Permalink
feat(data/set): add lemmas about powerset of insert and `singleto…
Browse files Browse the repository at this point in the history
…n` (#18356)

Add lemmas about `𝒫 {x}` and `𝒫 (insert x A)`. 

Mathlib4 pair : leanprover-community/mathlib4/pull/2000



Co-authored-by: Eric Wieser <wieser.eric@gmail.com>
  • Loading branch information
joneugster and eric-wieser committed Feb 22, 2023
1 parent b875cbb commit 4550138
Show file tree
Hide file tree
Showing 2 changed files with 43 additions and 0 deletions.
20 changes: 20 additions & 0 deletions src/data/set/basic.lean
Expand Up @@ -698,6 +698,22 @@ begin
exacts [(ha hx).elim, hxt]
end

theorem subset_insert_iff_of_not_mem {s t : set α} {a : α} (h : a ∉ s) :
s ⊆ insert a t ↔ s ⊆ t :=
begin
constructor,
{ intros g y hy,
specialize g hy,
rw [mem_insert_iff] at g,
rcases g with g | g,
{ rw [g] at hy,
contradiction },
{ assumption }},
{ intros g y hy,
specialize g hy,
exact mem_insert_of_mem _ g }
end

theorem ssubset_iff_insert {s t : set α} : s ⊂ t ↔ ∃ a ∉ s, insert a s ⊆ t :=
begin
simp only [insert_subset, exists_and_distrib_right, ssubset_def, not_subset],
Expand Down Expand Up @@ -1301,6 +1317,10 @@ ext $ λ s, subset_empty_iff
@[simp] theorem powerset_univ : 𝒫 (univ : set α) = univ :=
eq_univ_of_forall subset_univ

/- The powerset of a singleton contains only `∅` and the singleton itself. -/
theorem powerset_singleton (x : α) : 𝒫 ({x} : set α) = {∅, {x}} :=
by { ext y, rw [mem_powerset_iff, subset_singleton_iff_eq, mem_insert_iff, mem_singleton_iff] }

/-! ### Sets defined as an if-then-else -/

lemma mem_dite_univ_right (p : Prop) [decidable p] (t : p → set α) (x : α) :
Expand Down
23 changes: 23 additions & 0 deletions src/data/set/image.lean
Expand Up @@ -471,6 +471,29 @@ end

end image

/-! ### Lemmas about the powerset and image. -/

/-- The powerset of `{a} ∪ s` is `𝒫 s` together with `{a} ∪ t` for each `t ∈ 𝒫 s`. -/
theorem powerset_insert (s : set α) (a : α) :
𝒫 (insert a s) = 𝒫 s ∪ (insert a '' 𝒫 s) :=
begin
ext t,
simp_rw [mem_union, mem_image, mem_powerset_iff],
split,
{ intro h,
by_cases hs : a ∈ t,
{ right,
refine ⟨t \ {a}, _, _⟩,
{ rw [diff_singleton_subset_iff],
assumption },
{ rw [insert_diff_singleton, insert_eq_of_mem hs] }},
{ left,
exact (subset_insert_iff_of_not_mem hs).mp h}},
{ rintros (h | ⟨s', h₁, rfl⟩),
{ exact subset_trans h (subset_insert a s) },
{ exact insert_subset_insert h₁ }}
end

/-! ### Lemmas about range of a function. -/
section range
variables {f : ι → α} {s t : set α}
Expand Down

0 comments on commit 4550138

Please sign in to comment.