@@ -51,6 +51,9 @@ univ_nonempty_iff.2 ‹_›
51
51
lemma univ_eq_empty : (univ : finset α) = ∅ ↔ ¬nonempty α :=
52
52
by rw [← univ_nonempty_iff, nonempty_iff_ne_empty, ne.def, not_not]
53
53
54
+ lemma univ_eq_empty' : (univ : finset α) = ∅ ↔ is_empty α :=
55
+ univ_eq_empty.trans (not_nonempty_iff)
56
+
54
57
@[simp] theorem subset_univ (s : finset α) : s ⊆ univ := λ a _, mem_univ a
55
58
56
59
instance : order_top (finset α) :=
@@ -483,16 +486,27 @@ arbitrary `fintype` instances, use either `fintype.card_le_one_iff_subsingleton`
483
486
fintype.card α = 1 :=
484
487
subsingleton.elim (of_subsingleton $ default α) h ▸ card_of_subsingleton _
485
488
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
+
486
500
open_locale classical
487
501
variables (α)
488
502
489
503
/-- Any subsingleton type is (noncomputably) a fintype (with zero or one terms). -/
490
- @[priority 100 ]
504
+ @[priority 5 ] -- see Note [lower instance priority ]
491
505
noncomputable instance of_subsingleton' [subsingleton α] : fintype α :=
492
506
if h : nonempty α then
493
507
of_subsingleton (nonempty.some h)
494
508
else
495
- ⟨∅, (λ a, false.elim (h ⟨a⟩))⟩
509
+ @fintype.of_is_empty _ $ not_nonempty_iff.mp h
496
510
497
511
end fintype
498
512
@@ -651,14 +665,10 @@ fintype.of_subsingleton (default α)
651
665
@[simp] lemma univ_unique {α : Type *} [unique α] [f : fintype α] : @finset.univ α _ = {default α} :=
652
666
by rw [subsingleton.elim f (@unique.fintype α _)]; refl
653
667
654
- instance : fintype empty := ⟨∅, empty.rec _⟩
655
-
656
668
@[simp] theorem fintype.univ_empty : @univ empty _ = ∅ := rfl
657
669
658
670
@[simp] theorem fintype.card_empty : fintype.card empty = 0 := rfl
659
671
660
- instance : fintype pempty := ⟨∅, pempty.rec _⟩
661
-
662
672
@[simp] theorem fintype.univ_pempty : @univ pempty _ = ∅ := rfl
663
673
664
674
@[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] }
1076
1086
s.to_finset = ∅ ↔ s = ∅ :=
1077
1087
by simp [ext_iff, set.ext_iff]
1078
1088
1079
- instance : fintype (∅ : set α) := ⟨∅, subtype.property⟩
1080
-
1081
1089
@[simp] lemma set.to_finset_empty :
1082
1090
(∅ : set α).to_finset = ∅ :=
1083
1091
set.to_finset_eq_empty_iff.mpr rfl
@@ -1535,7 +1543,7 @@ instance function.embedding.is_empty {α β} [infinite α] [fintype β] : is_emp
1535
1543
⟨λ f, let ⟨x, y, ne, feq⟩ := fintype.exists_ne_map_eq_of_infinite f in ne $ f.injective feq⟩
1536
1544
1537
1545
@[priority 100 ]
1538
- noncomputable instance function.embedding.fintype'' {α β : Type *} [fintype β] : fintype (α ↪ β) :=
1546
+ noncomputable instance function.embedding.fintype' {α β : Type *} [fintype β] : fintype (α ↪ β) :=
1539
1547
begin
1540
1548
by_cases h : infinite α,
1541
1549
{ resetI, apply_instance },
@@ -1571,10 +1579,6 @@ assume (hf : surjective f),
1571
1579
have H : infinite α := infinite.of_surjective f hf,
1572
1580
by exactI not_fintype α
1573
1581
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
-
1578
1582
instance nat.infinite : infinite ℕ :=
1579
1583
⟨λ ⟨s, hs⟩, finset.not_mem_range_self $ s.subset_range_sup_succ (hs _)⟩
1580
1584
0 commit comments