-
Notifications
You must be signed in to change notification settings - Fork 295
[Merged by Bors] - refactor(data/set/finite): reorganize and put emphasis on fintype instances #14136
[Merged by Bors] - refactor(data/set/finite): reorganize and put emphasis on fintype instances #14136
Changes from 34 commits
5835bcf
0c3ecf3
a4693f5
2e5e43a
b76df6d
eae560a
d8ffb5a
41e6bd9
b572e0b
8774203
ec15466
feed464
4b25d90
84ba4e1
ab44747
b69ec3e
ad7e533
cdcf54b
023cfa5
5799e88
000351e
494939c
5e3d59b
5409e54
c2078a1
0b4b83b
92a8b56
31812fd
6c2898a
9c12a57
3854af7
aaeac4c
1e3c1a8
0497014
a68b38d
File filter
Filter by extension
Conversations
Jump to
Diff view
Diff view
There are no files selected for viewing
Original file line number | Diff line number | Diff line change |
---|---|---|
|
@@ -641,6 +641,13 @@ by simp [to_finset] | |
@[simp] theorem mem_to_finset_val {s : set α} [fintype s] {a : α} : a ∈ s.to_finset.1 ↔ a ∈ s := | ||
mem_to_finset | ||
|
||
/-- Membership of a set with a `fintype` instance is decidable. | ||
|
||
Using this as an instance leads to potential loops with `subtype.fintype` under certain decidability | ||
assumptions, so it should only be declared a local instance. -/ | ||
def decidable_mem_of_fintype [decidable_eq α] (s : set α) [fintype s] (a) : decidable (a ∈ s) := | ||
b-mehta marked this conversation as resolved.
Show resolved
Hide resolved
|
||
decidable_of_iff _ mem_to_finset | ||
|
||
-- We use an arbitrary `[fintype s]` instance here, | ||
-- not necessarily coming from a `[fintype α]`. | ||
@[simp] | ||
|
@@ -667,9 +674,46 @@ by simp only [finset.ssubset_def, to_finset_mono, ssubset_def] | |
disjoint s.to_finset t.to_finset ↔ disjoint s t := | ||
by simp only [disjoint_iff_disjoint_coe, coe_to_finset] | ||
|
||
lemma to_finset_inter {α : Type*} [decidable_eq α] (s t : set α) [fintype (s ∩ t : set α)] | ||
kmill marked this conversation as resolved.
Show resolved
Hide resolved
|
||
[fintype s] [fintype t] : (s ∩ t).to_finset = s.to_finset ∩ t.to_finset := | ||
by { ext, simp } | ||
|
||
lemma to_finset_union {α : Type*} [decidable_eq α] (s t : set α) [fintype (s ∪ t : set α)] | ||
There was a problem hiding this comment. Choose a reason for hiding this commentThe reason will be displayed to describe this comment to others. Learn more. What's the reason for some of these to be There was a problem hiding this comment. Choose a reason for hiding this commentThe reason will be displayed to describe this comment to others. Learn more. I'm not sure, this is how they were, and I didn't want to see if mathlib would compile with them as |
||
[fintype s] [fintype t] : (s ∪ t).to_finset = s.to_finset ∪ t.to_finset := | ||
by { ext, simp } | ||
|
||
lemma to_finset_diff {α : Type*} [decidable_eq α] (s t : set α) [fintype s] [fintype t] | ||
[fintype (s \ t : set α)] : (s \ t).to_finset = s.to_finset \ t.to_finset := | ||
by { ext, simp } | ||
|
||
lemma to_finset_ne_eq_erase {α : Type*} [decidable_eq α] [fintype α] (a : α) | ||
[fintype {x : α | x ≠ a}] : {x : α | x ≠ a}.to_finset = finset.univ.erase a := | ||
by { ext, simp } | ||
|
||
theorem to_finset_compl [decidable_eq α] [fintype α] (s : set α) [fintype s] [fintype ↥sᶜ] : | ||
(sᶜ).to_finset = s.to_finsetᶜ := | ||
by { ext a, simp } | ||
by { ext, simp } | ||
|
||
/- TODO Without the coercion arrow (`↥`) there is an elaboration bug; | ||
it essentially infers `fintype.{v} (set.univ.{u} : set α)` with `v` and `u` distinct. | ||
Reported in leanprover-community/lean#672 -/ | ||
@[simp] lemma to_finset_univ [fintype ↥(set.univ : set α)] [fintype α] : | ||
kmill marked this conversation as resolved.
Show resolved
Hide resolved
|
||
(set.univ : set α).to_finset = finset.univ := | ||
by { ext, simp } | ||
|
||
@[simp] lemma to_finset_range [decidable_eq α] [fintype β] (f : β → α) [fintype (set.range f)] : | ||
(set.range f).to_finset = finset.univ.image f := | ||
by { ext, simp } | ||
|
||
/- TODO The `↥` circumvents an elaboration bug. See comment on `set.to_finset_univ`. -/ | ||
lemma to_finset_singleton (a : α) [fintype ↥({a} : set α)] : ({a} : set α).to_finset = {a} := | ||
by { ext, simp } | ||
|
||
/- TODO The `↥` circumvents an elaboration bug. See comment on `set.to_finset_univ`. -/ | ||
@[simp] lemma to_finset_insert [decidable_eq α] {a : α} {s : set α} | ||
[fintype ↥(insert a s : set α)] [fintype s] : | ||
kmill marked this conversation as resolved.
Show resolved
Hide resolved
|
||
(insert a s).to_finset = insert a s.to_finset := | ||
by { ext, simp } | ||
|
||
lemma filter_mem_univ_eq_to_finset [fintype α] (s : set α) [fintype s] [decidable_pred (∈ s)] : | ||
finset.univ.filter (∈ s) = s.to_finset := | ||
|
@@ -1168,23 +1212,12 @@ instance Prop.fintype : fintype Prop := | |
instance subtype.fintype (p : α → Prop) [decidable_pred p] [fintype α] : fintype {x // p x} := | ||
fintype.subtype (univ.filter p) (by simp) | ||
|
||
/- TODO Without the coercion arrow (`↥`) there is an elaboration bug; | ||
it essentially infers `fintype.{v} (set.univ.{u} : set α)` with `v` and `u` distinct. | ||
Reported in leanprover-community/lean#672 -/ | ||
@[simp] lemma set.to_finset_univ [fintype ↥(set.univ : set α)] [fintype α] : | ||
(set.univ : set α).to_finset = finset.univ := | ||
by { ext, simp only [set.mem_univ, mem_univ, set.mem_to_finset] } | ||
|
||
@[simp] lemma set.to_finset_eq_empty_iff {s : set α} [fintype s] : s.to_finset = ∅ ↔ s = ∅ := | ||
by simp [ext_iff, set.ext_iff] | ||
by simp only [ext_iff, set.ext_iff, set.mem_to_finset, not_mem_empty, set.mem_empty_eq] | ||
|
||
@[simp] lemma set.to_finset_empty : (∅ : set α).to_finset = ∅ := | ||
set.to_finset_eq_empty_iff.mpr rfl | ||
|
||
@[simp] lemma set.to_finset_range [decidable_eq α] [fintype β] (f : β → α) [fintype (set.range f)] : | ||
(set.range f).to_finset = finset.univ.image f := | ||
by simp [ext_iff] | ||
|
||
/-- A set on a fintype, when coerced to a type, is a fintype. -/ | ||
def set_fintype [fintype α] (s : set α) [decidable_pred (∈ s)] : fintype s := | ||
subtype.fintype (λ x, x ∈ s) | ||
|
@@ -1412,6 +1445,26 @@ begin | |
simp | ||
end | ||
|
||
@[simp] | ||
lemma fintype.card_subtype_compl [fintype α] | ||
(p : α → Prop) [fintype {x // p x}] [fintype {x // ¬ p x}] : | ||
fintype.card {x // ¬ p x} = fintype.card α - fintype.card {x // p x} := | ||
begin | ||
classical, | ||
rw [fintype.card_of_subtype (set.to_finset pᶜ), set.to_finset_compl p, finset.card_compl, | ||
fintype.card_of_subtype (set.to_finset p)]; | ||
intros; simp; refl | ||
There was a problem hiding this comment. Choose a reason for hiding this commentThe reason will be displayed to describe this comment to others. Learn more. Are we missing a simp lemma? There was a problem hiding this comment. Choose a reason for hiding this commentThe reason will be displayed to describe this comment to others. Learn more. This is code that had previously been through review, but I went ahead and fixed the non-terminal What's going on is that this proof is that |
||
end | ||
|
||
/-- If two subtypes of a fintype have equal cardinality, so do their complements. -/ | ||
lemma fintype.card_compl_eq_card_compl [fintype α] | ||
(p q : α → Prop) | ||
[fintype {x // p x}] [fintype {x // ¬ p x}] | ||
[fintype {x // q x}] [fintype {x // ¬ q x}] | ||
(h : fintype.card {x // p x} = fintype.card {x // q x}) : | ||
fintype.card {x // ¬ p x} = fintype.card {x // ¬ q x} := | ||
by simp only [fintype.card_subtype_compl, h] | ||
|
||
theorem fintype.card_quotient_le [fintype α] (s : setoid α) [decidable_rel ((≈) : α → α → Prop)] : | ||
fintype.card (quotient s) ≤ fintype.card α := | ||
fintype.card_le_of_surjective _ (surjective_quotient_mk _) | ||
|
There was a problem hiding this comment.
Choose a reason for hiding this comment
The reason will be displayed to describe this comment to others. Learn more.
Does it work to leave this as an underscore?
There was a problem hiding this comment.
Choose a reason for hiding this comment
The reason will be displayed to describe this comment to others. Learn more.
Without it it creates a
fintype {x // x ∈ l.val}
goal. I'm not sure why this happens exactly (see the discussion just above).