Skip to content

Commit

Permalink
chore: remove a subsingleton instance which leads to unreasonable ins…
Browse files Browse the repository at this point in the history
…tance search paths (#12445)
  • Loading branch information
sgouezel committed Apr 26, 2024
1 parent 5450e92 commit 3877d4f
Show file tree
Hide file tree
Showing 2 changed files with 7 additions and 4 deletions.
10 changes: 6 additions & 4 deletions Mathlib/Algebra/CharP/Basic.lean
Expand Up @@ -466,8 +466,9 @@ section NonAssocSemiring

variable {R} [NonAssocSemiring R]

-- see Note [lower instance priority]
instance (priority := 100) CharOne.subsingleton [CharP R 1] : Subsingleton R :=
-- This lemma is not an instance, to make sure that trying to prove `α` is a subsingleton does
-- not try to find a ring structure on `α`, which can be expensive.
lemma CharOne.subsingleton [CharP R 1] : Subsingleton R :=
Subsingleton.intro <|
suffices ∀ r : R, r = 0 from fun a b => show a = b by rw [this a, this b]
fun r =>
Expand All @@ -477,8 +478,9 @@ instance (priority := 100) CharOne.subsingleton [CharP R 1] : Subsingleton R :=
_ = 0 * r := by rw [CharP.cast_eq_zero]
_ = 0 := by rw [zero_mul]

theorem false_of_nontrivial_of_char_one [Nontrivial R] [CharP R 1] : False :=
false_of_nontrivial_of_subsingleton R
theorem false_of_nontrivial_of_char_one [Nontrivial R] [CharP R 1] : False := by
have : Subsingleton R := CharOne.subsingleton
exact false_of_nontrivial_of_subsingleton R
#align char_p.false_of_nontrivial_of_char_one CharP.false_of_nontrivial_of_char_one

theorem ringChar_ne_one [Nontrivial R] : ringChar R ≠ 1 := by
Expand Down
1 change: 1 addition & 0 deletions Mathlib/Data/ZMod/Basic.lean
Expand Up @@ -297,6 +297,7 @@ theorem cast_one (h : m ∣ n) : (cast (1 : ZMod n) : R) = 1 := by
cases n;
· rw [Nat.dvd_one] at h
subst m
have : Subsingleton R := CharP.CharOne.subsingleton
apply Subsingleton.elim
rw [Nat.mod_eq_of_lt]
· exact Nat.cast_one
Expand Down

0 comments on commit 3877d4f

Please sign in to comment.