Skip to content

Commit

Permalink
feat(data/finset/basic): more lemmas on finsets of subtypes (#3575)
Browse files Browse the repository at this point in the history
Add two more lemmas related to `not_mem_map_subtype_of_not_property`.
  • Loading branch information
jsm28 committed Jul 27, 2020
1 parent 3550f4f commit 2ecf70f
Showing 1 changed file with 20 additions and 3 deletions.
23 changes: 20 additions & 3 deletions src/data/finset/basic.lean
Expand Up @@ -1436,15 +1436,32 @@ lemma subtype_map_of_mem {p : α → Prop} [decidable_pred p] (h : ∀ x ∈ s,
(s.subtype p).map (function.embedding.subtype _) = s :=
by rw [subtype_map, filter_true_of_mem h]

/-- If a `finset` of a subtype is converted to the main type with
`embedding.subtype`, all elements of the result have the property of
the subtype. -/
lemma property_of_mem_map_subtype {p : α → Prop} (s : finset {x // p x}) {a : α}
(h : a ∈ s.map (function.embedding.subtype _)) : p a :=
begin
rcases mem_map.1 h with ⟨x, hx, rfl⟩,
exact x.2
end

/-- If a `finset` of a subtype is converted to the main type with
`embedding.subtype`, the result does not contain any value that does
not satisfy the property of the subtype. -/
lemma not_mem_map_subtype_of_not_property {p : α → Prop} (s : finset {x // p x})
{a : α} (h : ¬ p a) : a ∉ (s.map (function.embedding.subtype _)) :=
mt s.property_of_mem_map_subtype h

/-- If a `finset` of a subtype is converted to the main type with
`embedding.subtype`, the result is a subset of the set giving the
subtype. -/
lemma map_subtype_subset {t : set α} (s : finset t) :
↑(s.map (function.embedding.subtype _)) ⊆ t :=
begin
rw mem_map,
push_neg,
exact λ x, λ hxs hx, h (hx ▸ x.property)
intros a ha,
rw mem_coe at ha,
convert property_of_mem_map_subtype s ha
end

lemma subset_image_iff {f : α → β}
Expand Down

0 comments on commit 2ecf70f

Please sign in to comment.