Skip to content

Commit

Permalink
chore(data/fintype): drop a decidable_pred assumption (#14971)
Browse files Browse the repository at this point in the history
OTOH, now the proof depends on `classical.choice`.
  • Loading branch information
urkud committed Jul 6, 2022
1 parent a95b442 commit d908bc0
Showing 1 changed file with 3 additions and 2 deletions.
5 changes: 3 additions & 2 deletions src/data/fintype/basic.lean
Expand Up @@ -1004,8 +1004,9 @@ begin
end

/-- If the subtype of all-but-one elements is a `fintype` then the type itself is a `fintype`. -/
def fintype_of_fintype_ne (a : α) [decidable_pred (= a)] (h : fintype {b // b ≠ a}) : fintype α :=
fintype.of_equiv _ $ equiv.sum_compl (= a)
def fintype_of_fintype_ne (a : α) (h : fintype {b // b ≠ a}) : fintype α :=
fintype.of_bijective (sum.elim (coe : {b // b = a} → α) (coe : {b // b ≠ a} → α)) $
by { classical, exact (equiv.sum_compl (= a)).bijective }

section finset

Expand Down

0 comments on commit d908bc0

Please sign in to comment.