Skip to content

Commit 725f123

Browse files
committed
perf: de-prioritize Subalgebra.algebra' and IntermediateField.algebra' (#9936)
Like in #9655, these instances tend to be slow to fail, so we should assign them a low priority. Zulip discussions: https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/Algebra.2Eid.20for.20IntermediateField https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/Timeout.20in.20Submodule.20.28.F0.9D.93.9E.20K.29.20.28.F0.9D.93.9E.20K.29 Co-authored-by: Anne Baanen <Vierkantor@users.noreply.github.com>
1 parent 4fe0086 commit 725f123

File tree

2 files changed

+16
-2
lines changed

2 files changed

+16
-2
lines changed

Mathlib/Algebra/Algebra/Subalgebra/Basic.lean

Lines changed: 8 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -371,7 +371,14 @@ instance : Module R S :=
371371
instance [Semiring R'] [SMul R' R] [Module R' A] [IsScalarTower R' R A] : IsScalarTower R' R S :=
372372
inferInstanceAs (IsScalarTower R' R (toSubmodule S))
373373

374-
instance algebra' [CommSemiring R'] [SMul R' R] [Algebra R' A] [IsScalarTower R' R A] :
374+
/- More general form of `Subalgebra.algebra`.
375+
376+
This instance should have low priority since it is slow to fail:
377+
before failing, it will cause a search through all `SMul R' R` instances,
378+
which can quickly get expensive.
379+
-/
380+
instance (priority := 500) algebra' [CommSemiring R'] [SMul R' R] [Algebra R' A]
381+
[IsScalarTower R' R A] :
375382
Algebra R' S :=
376383
{ (algebraMap R' A).codRestrict S fun x => by
377384
rw [Algebra.algebraMap_eq_smul_one, ← smul_one_smul R x (1 : A), ←

Mathlib/FieldTheory/IntermediateField.lean

Lines changed: 8 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -374,7 +374,14 @@ theorem coe_smul {R} [Semiring R] [SMul R K] [Module R L] [IsScalarTower R K L]
374374
rfl
375375
#align intermediate_field.coe_smul IntermediateField.coe_smul
376376

377-
instance algebra' {K'} [CommSemiring K'] [SMul K' K] [Algebra K' L] [IsScalarTower K' K L] :
377+
/- More general form of `IntermediateField.algebra`.
378+
379+
This instance should have low priority since it is slow to fail:
380+
before failing, it will cause a search through all `SMul K' K` instances,
381+
which can quickly get expensive.
382+
-/
383+
instance (priority := 500) algebra' {K'} [CommSemiring K'] [SMul K' K] [Algebra K' L]
384+
[IsScalarTower K' K L] :
378385
Algebra K' S :=
379386
S.toSubalgebra.algebra'
380387
#align intermediate_field.algebra' IntermediateField.algebra'

0 commit comments

Comments
 (0)