Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit 320da57

Browse files
committed
chore(data/fintype/basic): add fintype instance for is_empty (#7692)
Co-authored-by: Eric <37984851+ericrbg@users.noreply.github.com>
1 parent 6e67536 commit 320da57

File tree

4 files changed

+33
-27
lines changed

4 files changed

+33
-27
lines changed

src/analysis/mean_inequalities.lean

Lines changed: 6 additions & 6 deletions
Original file line numberDiff line numberDiff line change
@@ -197,24 +197,24 @@ by exact_mod_cast real.geom_mean_le_arith_mean_weighted _ _ _ (λ i _, (w i).coe
197197
for two `nnreal` numbers. -/
198198
theorem geom_mean_le_arith_mean2_weighted (w₁ w₂ p₁ p₂ : ℝ≥0) :
199199
w₁ + w₂ = 1 → p₁ ^ (w₁:ℝ) * p₂ ^ (w₂:ℝ) ≤ w₁ * p₁ + w₂ * p₂ :=
200-
by simpa only [fin.prod_univ_succ, fin.sum_univ_succ, fin.prod_univ_zero, fin.sum_univ_zero,
201-
fin.cons_succ, fin.cons_zero, add_zero, mul_one]
200+
by simpa only [fin.prod_univ_succ, fin.sum_univ_succ, finset.prod_empty, finset.sum_empty,
201+
fintype.univ_of_is_empty, fin.cons_succ, fin.cons_zero, add_zero, mul_one]
202202
using geom_mean_le_arith_mean_weighted (univ : finset (fin 2))
203203
(fin.cons w₁ $ fin.cons w₂ fin_zero_elim) (fin.cons p₁ $ fin.cons p₂ $ fin_zero_elim)
204204

205205
theorem geom_mean_le_arith_mean3_weighted (w₁ w₂ w₃ p₁ p₂ p₃ : ℝ≥0) :
206206
w₁ + w₂ + w₃ = 1 → p₁ ^ (w₁:ℝ) * p₂ ^ (w₂:ℝ) * p₃ ^ (w₃:ℝ) ≤ w₁ * p₁ + w₂ * p₂ + w₃ * p₃ :=
207-
by simpa only [fin.prod_univ_succ, fin.sum_univ_succ, fin.prod_univ_zero, fin.sum_univ_zero,
208-
fin.cons_succ, fin.cons_zero, add_zero, mul_one, ← add_assoc, mul_assoc]
207+
by simpa only [fin.prod_univ_succ, fin.sum_univ_succ, finset.prod_empty, finset.sum_empty,
208+
fintype.univ_of_is_empty, fin.cons_succ, fin.cons_zero, add_zero, mul_one, ← add_assoc, mul_assoc]
209209
using geom_mean_le_arith_mean_weighted (univ : finset (fin 3))
210210
(fin.cons w₁ $ fin.cons w₂ $ fin.cons w₃ fin_zero_elim)
211211
(fin.cons p₁ $ fin.cons p₂ $ fin.cons p₃ fin_zero_elim)
212212

213213
theorem geom_mean_le_arith_mean4_weighted (w₁ w₂ w₃ w₄ p₁ p₂ p₃ p₄ : ℝ≥0) :
214214
w₁ + w₂ + w₃ + w₄ = 1 → p₁ ^ (w₁:ℝ) * p₂ ^ (w₂:ℝ) * p₃ ^ (w₃:ℝ)* p₄ ^ (w₄:ℝ) ≤
215215
w₁ * p₁ + w₂ * p₂ + w₃ * p₃ + w₄ * p₄ :=
216-
by simpa only [fin.prod_univ_succ, fin.sum_univ_succ, fin.prod_univ_zero, fin.sum_univ_zero,
217-
fin.cons_succ, fin.cons_zero, add_zero, mul_one, ← add_assoc, mul_assoc]
216+
by simpa only [fin.prod_univ_succ, fin.sum_univ_succ, finset.prod_empty, finset.sum_empty,
217+
fintype.univ_of_is_empty, fin.cons_succ, fin.cons_zero, add_zero, mul_one, ← add_assoc, mul_assoc]
218218
using geom_mean_le_arith_mean_weighted (univ : finset (fin 4))
219219
(fin.cons w₁ $ fin.cons w₂ $ fin.cons w₃ $ fin.cons w₄ fin_zero_elim)
220220
(fin.cons p₁ $ fin.cons p₂ $ fin.cons p₃ $ fin.cons p₄ fin_zero_elim)

src/data/fintype/basic.lean

Lines changed: 17 additions & 13 deletions
Original file line numberDiff line numberDiff line change
@@ -51,6 +51,9 @@ univ_nonempty_iff.2 ‹_›
5151
lemma univ_eq_empty : (univ : finset α) = ∅ ↔ ¬nonempty α :=
5252
by rw [← univ_nonempty_iff, nonempty_iff_ne_empty, ne.def, not_not]
5353

54+
lemma univ_eq_empty' : (univ : finset α) = ∅ ↔ is_empty α :=
55+
univ_eq_empty.trans (not_nonempty_iff)
56+
5457
@[simp] theorem subset_univ (s : finset α) : s ⊆ univ := λ a _, mem_univ a
5558

5659
instance : order_top (finset α) :=
@@ -483,16 +486,27 @@ arbitrary `fintype` instances, use either `fintype.card_le_one_iff_subsingleton`
483486
fintype.card α = 1 :=
484487
subsingleton.elim (of_subsingleton $ default α) h ▸ card_of_subsingleton _
485488

489+
@[priority 100] -- see Note [lower instance priority]
490+
instance of_is_empty [is_empty α] : fintype α := ⟨∅, is_empty_elim⟩
491+
492+
/-- Note: this lemma is specifically about `fintype.of_is_empty`. For a statement about
493+
arbitrary `fintype` instances, use `fintype.univ_eq_empty'`. -/
494+
@[simp] theorem univ_of_is_empty [is_empty α] : @univ α _ = ∅ := rfl
495+
496+
/-- Note: this lemma is specifically about `fintype.of_is_empty`. For a statement about
497+
arbitrary `fintype` instances, use `fintype.card_eq_zero_iff`. -/
498+
@[simp] theorem card_of_is_empty [is_empty α] : fintype.card α = 0 := rfl
499+
486500
open_locale classical
487501
variables (α)
488502

489503
/-- Any subsingleton type is (noncomputably) a fintype (with zero or one terms). -/
490-
@[priority 100]
504+
@[priority 5] -- see Note [lower instance priority]
491505
noncomputable instance of_subsingleton' [subsingleton α] : fintype α :=
492506
if h : nonempty α then
493507
of_subsingleton (nonempty.some h)
494508
else
495-
⟨∅, (λ a, false.elim (h ⟨a⟩))⟩
509+
@fintype.of_is_empty _ $ not_nonempty_iff.mp h
496510

497511
end fintype
498512

@@ -651,14 +665,10 @@ fintype.of_subsingleton (default α)
651665
@[simp] lemma univ_unique {α : Type*} [unique α] [f : fintype α] : @finset.univ α _ = {default α} :=
652666
by rw [subsingleton.elim f (@unique.fintype α _)]; refl
653667

654-
instance : fintype empty := ⟨∅, empty.rec _⟩
655-
656668
@[simp] theorem fintype.univ_empty : @univ empty _ = ∅ := rfl
657669

658670
@[simp] theorem fintype.card_empty : fintype.card empty = 0 := rfl
659671

660-
instance : fintype pempty := ⟨∅, pempty.rec _⟩
661-
662672
@[simp] theorem fintype.univ_pempty : @univ pempty _ = ∅ := rfl
663673

664674
@[simp] theorem fintype.card_pempty : fintype.card pempty = 0 := rfl
@@ -1076,8 +1086,6 @@ by { ext, simp only [set.mem_univ, mem_univ, set.mem_to_finset] }
10761086
s.to_finset = ∅ ↔ s = ∅ :=
10771087
by simp [ext_iff, set.ext_iff]
10781088

1079-
instance : fintype (∅ : set α) := ⟨∅, subtype.property⟩
1080-
10811089
@[simp] lemma set.to_finset_empty :
10821090
(∅ : set α).to_finset = ∅ :=
10831091
set.to_finset_eq_empty_iff.mpr rfl
@@ -1535,7 +1543,7 @@ instance function.embedding.is_empty {α β} [infinite α] [fintype β] : is_emp
15351543
⟨λ f, let ⟨x, y, ne, feq⟩ := fintype.exists_ne_map_eq_of_infinite f in ne $ f.injective feq⟩
15361544

15371545
@[priority 100]
1538-
noncomputable instance function.embedding.fintype'' {α β : Type*} [fintype β] : fintype (α ↪ β) :=
1546+
noncomputable instance function.embedding.fintype' {α β : Type*} [fintype β] : fintype (α ↪ β) :=
15391547
begin
15401548
by_cases h : infinite α,
15411549
{ resetI, apply_instance },
@@ -1571,10 +1579,6 @@ assume (hf : surjective f),
15711579
have H : infinite α := infinite.of_surjective f hf,
15721580
by exactI not_fintype α
15731581

1574-
-- the instance generated by `is_empty → subsingleton → fintype is non-computable
1575-
instance function.embedding.fintype' {α β} [infinite α] [fintype β] : fintype (α ↪ β) :=
1576-
{ elems := finset.empty, complete := is_empty_elim }
1577-
15781582
instance nat.infinite : infinite ℕ :=
15791583
⟨λ ⟨s, hs⟩, finset.not_mem_range_self $ s.subset_range_sup_succ (hs _)⟩
15801584

src/data/fintype/card.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -126,7 +126,7 @@ theorem fin.prod_of_fn [comm_monoid β] {n : ℕ} (f : fin n → β) :
126126
by rw [list.of_fn_eq_map, fin.prod_univ_def]
127127

128128
/-- A product of a function `f : fin 0 → β` is `1` because `fin 0` is empty -/
129-
@[simp, to_additive "A sum of a function `f : fin 0 → β` is `0` because `fin 0` is empty"]
129+
@[to_additive "A sum of a function `f : fin 0 → β` is `0` because `fin 0` is empty"]
130130
theorem fin.prod_univ_zero [comm_monoid β] (f : fin 0 → β) : ∏ i, f i = 1 := rfl
131131

132132
/-- A product of a function `f : fin (n + 1) → β` over all `fin (n + 1)`

test/matrix.lean

Lines changed: 9 additions & 7 deletions
Original file line numberDiff line numberDiff line change
@@ -62,9 +62,10 @@ example {α : Type*} [comm_ring α] {a b c d : α} :
6262
begin
6363
simp [matrix.det_succ_row_zero, fin.sum_univ_succ],
6464
/-
65-
Try this: simp only [matrix.det_succ_row_zero, fin.sum_univ_succ, det_fin_zero,
66-
finset.sum_singleton, fin.sum_univ_zero, minor_apply, cons_val_zero, cons_val_succ,
67-
fin.succ_above_zero, fin.coe_succ, fin.coe_zero],
65+
Try this: simp only [det_succ_row_zero, fin.sum_univ_succ, neg_mul_eq_neg_mul_symm, mul_one,
66+
fin.default_eq_zero, fin.coe_zero, one_mul, cons_val_one, fin.coe_succ, univ_unique, minor_apply,
67+
pow_one, fin.zero_succ_above, fin.succ_succ_above_zero, finset.sum_singleton, cons_val_zero,
68+
cons_val_succ, det_fin_zero, pow_zero]
6869
-/
6970
ring
7071
end
@@ -75,10 +76,11 @@ example {α : Type*} [comm_ring α] (A : matrix (fin 3) (fin 3) α) {a b c d e f
7576
begin
7677
simp [matrix.det_succ_row_zero, fin.sum_univ_succ],
7778
/-
78-
Try this: simp only [matrix.det_succ_row_zero, fin.sum_univ_succ, det_fin_zero,
79-
finset.sum_singleton, fin.sum_univ_zero, minor_apply, cons_val_zero, cons_val_succ,
80-
fin.succ_above_zero, fin.succ_succ_above_zero, fin.succ_succ_above_succ,
81-
fin.coe_zero, fin.coe_succ, pow_zero, pow_add],
79+
Try this: simp only [det_succ_row_zero, fin.sum_univ_succ, neg_mul_eq_neg_mul_symm, cons_append,
80+
mul_one, fin.default_eq_zero, fin.coe_zero, cons_vec_bit0_eq_alt0, one_mul, cons_val_one,
81+
cons_vec_alt0, fin.succ_succ_above_one, fin.coe_succ, univ_unique, minor_apply, pow_one,
82+
fin.zero_succ_above, fin.succ_zero_eq_one, fin.succ_succ_above_zero, nat.neg_one_sq,
83+
finset.sum_singleton, cons_val_zero, cons_val_succ, det_fin_zero, head_cons, pow_zero]
8284
-/
8385
ring
8486
end

0 commit comments

Comments
 (0)