Skip to content

Commit

Permalink
perf (Functor.Flat): restructure proof to reduce simp time (#8065)
Browse files Browse the repository at this point in the history
This file recently jumped almost 200% in #8051. Changing a `simpa using X` call into a `convert X <;> simp` cuts the clock time for the file in half.
  • Loading branch information
mattrobball committed Oct 31, 2023
1 parent 95e9bc1 commit ac9fbe8
Showing 1 changed file with 2 additions and 2 deletions.
4 changes: 2 additions & 2 deletions Mathlib/CategoryTheory/Functor/Flat.lean
Original file line number Diff line number Diff line change
Expand Up @@ -184,7 +184,7 @@ theorem fac (x : J) : lift F hc s ≫ (F.mapCone c).π.app x = s.π.app x := by
simp [lift, ← Functor.map_comp]
#align category_theory.preserves_finite_limits_of_flat.fac CategoryTheory.PreservesFiniteLimitsOfFlat.fac

attribute [local simp] eqToHom_map
-- attribute [local simp] eqToHom_map

theorem uniq {K : J ⥤ C} {c : Cone K} (hc : IsLimit c) (s : Cone (K ⋙ F))
(f₁ f₂ : s.pt ⟶ F.obj c.pt) (h₁ : ∀ j : J, f₁ ≫ (F.mapCone c).π.app j = s.π.app j)
Expand All @@ -208,7 +208,7 @@ theorem uniq {K : J ⥤ C} {c : Cone K} (hc : IsLimit c) (s : Cone (K ⋙ F))
intro j
injection c₀.π.naturality (BiconeHom.left j) with _ e₁
injection c₀.π.naturality (BiconeHom.right j) with _ e₂
simpa using e₁.symm.trans e₂
convert e₁.symm.trans e₂ <;> simp
have : c.extend g₁.right = c.extend g₂.right := by
unfold Cone.extend
congr 1
Expand Down

0 comments on commit ac9fbe8

Please sign in to comment.