Skip to content

Commit

Permalink
feat(data/setoid/partition): some lemmas about partitions (#2937)
Browse files Browse the repository at this point in the history
  • Loading branch information
ChrisHughes24 committed Jun 5, 2020
1 parent 80a52e9 commit 2131382
Showing 1 changed file with 12 additions and 0 deletions.
12 changes: 12 additions & 0 deletions src/data/setoid/partition.lean
Expand Up @@ -128,6 +128,18 @@ lemma nonempty_of_mem_partition {c : set (set α)} (hc : is_partition c) {s} (h
s.nonempty :=
set.ne_empty_iff_nonempty.1 $ λ hs0, hc.1 $ hs0 ▸ h

lemma is_partition_classes (r : setoid α) : is_partition r.classes :=
⟨empty_not_mem_classes, classes_eqv_classes⟩

lemma is_partition.pairwise_disjoint {c : set (set α)} (hc : is_partition c) :
c.pairwise_disjoint :=
eqv_classes_disjoint hc.2

lemma is_partition.sUnion_eq_univ {c : set (set α)} (hc : is_partition c) :
⋃₀ c = set.univ :=
set.eq_univ_of_forall $ λ x, set.mem_sUnion.2 $
let ⟨t, ht⟩ := hc.2 x in ⟨t, by clear_aux_decl; finish⟩

/-- All elements of a partition of α are the equivalence class of some y ∈ α. -/
lemma exists_of_mem_partition {c : set (set α)} (hc : is_partition c) {s} (hs : s ∈ c) :
∃ y, s = {x | (mk_classes c hc.2).rel x y} :=
Expand Down

0 comments on commit 2131382

Please sign in to comment.