Skip to content

Commit

Permalink
chore(Category/TopCat): cleanup (#13170)
Browse files Browse the repository at this point in the history
This PR removes an instance (`ConcreteCategory.instFunLike`) that was mistakenly made global during the port and starts the clean up. It significantly speeds up CI since simp is more efficient.

Unfortunately, some `simp`'s and `rw`'s needed to be changed to `erw` as a result.

Zulip discussion: [#mathlib4 > ConcreteCategory.instFunLike became a global instance](https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/ConcreteCategory.2EinstFunLike.20became.20a.20.20global.20instance)

Issue: #13342 



Co-authored-by: Matthew Ballard <[email protected]>
Co-authored-by: dagurtomas <[email protected]>
Co-authored-by: Kim Morrison <[email protected]>
Co-authored-by: erdOne <[email protected]>
Co-authored-by: Dagur Tómas Ásgeirsson <[email protected]>
  • Loading branch information
5 people committed May 30, 2024
1 parent 9c539b5 commit 6d39f51
Show file tree
Hide file tree
Showing 37 changed files with 263 additions and 120 deletions.
11 changes: 6 additions & 5 deletions Mathlib/AlgebraicGeometry/AffineScheme.lean
Original file line number Diff line number Diff line change
Expand Up @@ -273,9 +273,9 @@ instance isOpenImmersion_fromSpec :
theorem fromSpec_range :
Set.range hU.fromSpec.1.base = (U : Set X) := by
delta IsAffineOpen.fromSpec; dsimp
rw [← Category.assoc, coe_comp, Set.range_comp, Set.range_iff_surjective.mpr, Set.image_univ]
rw [Function.comp.assoc, Set.range_comp, Set.range_iff_surjective.mpr, Set.image_univ]
· exact Subtype.range_coe
rw [← TopCat.epi_iff_surjective]
erw [← coe_comp, ← TopCat.epi_iff_surjective] -- now `erw` after #13170
infer_instance
#align algebraic_geometry.is_affine_open.from_Spec_range AlgebraicGeometry.IsAffineOpen.fromSpec_range

Expand Down Expand Up @@ -306,9 +306,9 @@ theorem _root_.AlgebraicGeometry.Scheme.Hom.isAffineOpen_iff_of_isOpenImmersion
refine ⟨fun hU => @isAffineOfIso _ _
(IsOpenImmersion.isoOfRangeEq (X.ofRestrict U.openEmbedding ≫ f) (Y.ofRestrict _) ?_).hom ?_ hU,
fun hU => hU.imageIsOpenImmersion f⟩
· rw [Scheme.comp_val_base, coe_comp, Set.range_comp]
· erw [Scheme.comp_val_base, coe_comp, Set.range_comp] -- now `erw` after #13170
dsimp [Opens.coe_inclusion, Scheme.restrict]
rw [Subtype.range_coe, Subtype.range_coe]
erw [Subtype.range_coe, Subtype.range_coe] -- now `erw` after #13170
rfl
· infer_instance
#align algebraic_geometry.is_affine_open_iff_of_is_open_immersion AlgebraicGeometry.Scheme.Hom.isAffineOpen_iff_of_isOpenImmersion
Expand Down Expand Up @@ -515,11 +515,12 @@ theorem fromSpec_primeIdealOf (x : U) :
-- Porting note: in the porting note of `Scheme.comp_val_base`, it says that `elementwise` is
-- unnecessary, indeed, the linter did not like it, so I just use `elementwise_of%` instead of
-- adding the corresponding lemma in `Scheme.lean` file
rw [← elementwise_of% Scheme.comp_val_base]
erw [← elementwise_of% Scheme.comp_val_base] -- now `erw` after #13170
simp only [Scheme.Γ_obj, unop_op, Scheme.restrict_presheaf_obj, Category.assoc, ←
Functor.map_comp_assoc, ← op_comp, ← Functor.map_comp, eqToHom_trans, eqToHom_refl, op_id,
CategoryTheory.Functor.map_id, Category.id_comp, Iso.hom_inv_id_assoc,
Scheme.ofRestrict_val_base, Scheme.restrict_carrier, Opens.coe_inclusion]
rfl -- `rfl` was not needed before #13170
#align algebraic_geometry.is_affine_open.from_Spec_prime_ideal_of AlgebraicGeometry.IsAffineOpen.fromSpec_primeIdealOf

set_option backward.isDefEq.lazyWhnfCore false in -- See https://github.com/leanprover-community/mathlib4/issues/12534
Expand Down
13 changes: 7 additions & 6 deletions Mathlib/AlgebraicGeometry/Gluing.lean
Original file line number Diff line number Diff line change
Expand Up @@ -132,9 +132,7 @@ def gluedScheme : Scheme := by
swap
· exact (D.U i).affineCover.map y
constructor
· -- Without removing `Spec.topObj_forget`, we need an `erw` in the following line.
dsimp [-Spec.topObj_forget]
rw [coe_comp, Set.range_comp]
· erw [TopCat.coe_comp, Set.range_comp] -- now `erw` after #13170
refine' Set.mem_image_of_mem _ _
exact (D.U i).affineCover.Covers y
· infer_instance
Expand Down Expand Up @@ -249,7 +247,9 @@ theorem ι_eq_iff (i j : D.J) (x : (D.U i).carrier) (y : (D.U j).carrier) :
D.toLocallyRingedSpaceGlueData.toSheafedSpaceGlueData.toPresheafedSpaceGlueData.toTopGlueData
i j x y)
rw [← ((TopCat.mono_iff_injective D.isoCarrier.inv).mp _).eq_iff]
· simp_rw [← comp_apply, ← D.ι_isoCarrier_inv]; rfl
· erw [← comp_apply] -- now `erw` after #13170
simp_rw [← D.ι_isoCarrier_inv]
rfl -- `rfl` was not needed before #13170
· infer_instance
#align algebraic_geometry.Scheme.glue_data.ι_eq_iff AlgebraicGeometry.Scheme.GlueData.ι_eq_iff

Expand Down Expand Up @@ -372,7 +372,8 @@ theorem fromGlued_injective : Function.Injective 𝒰.fromGlued.1.base := by
intro x y h
obtain ⟨i, x, rfl⟩ := 𝒰.gluedCover.ι_jointly_surjective x
obtain ⟨j, y, rfl⟩ := 𝒰.gluedCover.ι_jointly_surjective y
simp_rw [← comp_apply, ← SheafedSpace.comp_base, ← LocallyRingedSpace.comp_val] at h
erw [← comp_apply, ← comp_apply] at h -- now `erw` after #13170
simp_rw [← SheafedSpace.comp_base, ← LocallyRingedSpace.comp_val] at h
erw [ι_fromGlued, ι_fromGlued] at h
let e :=
(TopCat.pullbackConeIsLimit _ _).conePointUniqueUpToIso
Expand Down Expand Up @@ -427,7 +428,7 @@ instance : Epi 𝒰.fromGlued.val.base := by
intro x
obtain ⟨y, h⟩ := 𝒰.Covers x
use (𝒰.gluedCover.ι (𝒰.f x)).1.base y
rw [← comp_apply]
erw [← comp_apply] -- now `erw` after #13170
rw [← 𝒰.ι_fromGlued (𝒰.f x)] at h
exact h

Expand Down
2 changes: 1 addition & 1 deletion Mathlib/AlgebraicGeometry/Morphisms/QuasiCompact.lean
Original file line number Diff line number Diff line change
Expand Up @@ -68,7 +68,7 @@ instance quasiCompactComp {X Y Z : Scheme} (f : X ⟶ Y) (g : Y ⟶ Z) [QuasiCom
[QuasiCompact g] : QuasiCompact (f ≫ g) := by
constructor
intro U hU hU'
rw [Scheme.comp_val_base, coe_comp, Set.preimage_comp]
rw [Scheme.comp_val_base, TopCat.coe_comp, Set.preimage_comp]
apply QuasiCompact.isCompact_preimage
· exact Continuous.isOpen_preimage (by
-- Porting note: `continuity` failed
Expand Down
27 changes: 13 additions & 14 deletions Mathlib/AlgebraicGeometry/OpenImmersion.lean
Original file line number Diff line number Diff line change
Expand Up @@ -119,11 +119,11 @@ def affineCover (X : Scheme.{u}) : OpenCover X where
apply PresheafedSpace.IsOpenImmersion.ofRestrict
Covers := by
intro x
erw [coe_comp]
erw [TopCat.coe_comp] -- now `erw` after #13170
rw [Set.range_comp, Set.range_iff_surjective.mpr, Set.image_univ]
· erw [Subtype.range_coe_subtype]
exact (X.local_affine x).choose.2
rw [← TopCat.epi_iff_surjective]
erw [← TopCat.epi_iff_surjective] -- now `erw` after #13170
change Epi ((SheafedSpace.forget _).map (LocallyRingedSpace.forgetToSheafedSpace.map _))
infer_instance
#align algebraic_geometry.Scheme.affine_cover AlgebraicGeometry.Scheme.affineCover
Expand All @@ -146,7 +146,7 @@ def OpenCover.bind (f : ∀ x : 𝒰.J, OpenCover (𝒰.obj x)) : OpenCover X wh
change x ∈ Set.range ((f (𝒰.f x)).map ((f (𝒰.f x)).f y) ≫ 𝒰.map (𝒰.f x)).1.base
use z
erw [comp_apply]
rw [hz, hy]
erw [hz, hy] -- now `erw` after #13170
-- Porting note: weirdly, even though no input is needed, `inferInstance` does not work
-- `PresheafedSpace.IsOpenImmersion.comp` is marked as `instance`
IsOpen x := PresheafedSpace.IsOpenImmersion.comp _ _
Expand Down Expand Up @@ -175,10 +175,10 @@ def OpenCover.copy {X : Scheme.{u}} (𝒰 : OpenCover X) (J : Type*) (obj : J
{ J, obj, map
f := fun x => e₁.symm (𝒰.f x)
Covers := fun x => by
rw [e₂, Scheme.comp_val_base, coe_comp, Set.range_comp, Set.range_iff_surjective.mpr,
rw [e₂, Scheme.comp_val_base, TopCat.coe_comp, Set.range_comp, Set.range_iff_surjective.mpr,
Set.image_univ, e₁.rightInverse_symm]
· exact 𝒰.Covers x
· rw [← TopCat.epi_iff_surjective]; infer_instance
· erw [← TopCat.epi_iff_surjective]; infer_instance -- now `erw` after #13170
-- Porting note: weirdly, even though no input is needed, `inferInstance` does not work
-- `PresheafedSpace.IsOpenImmersion.comp` is marked as `instance`
IsOpen := fun i => by rw [e₂]; exact PresheafedSpace.IsOpenImmersion.comp _ _ }
Expand Down Expand Up @@ -567,7 +567,7 @@ theorem range_pullback_snd_of_left :
· erw [TopCat.pullback_snd_image_fst_preimage]
rw [Set.image_univ]
rfl
rw [← TopCat.epi_iff_surjective]
erw [← TopCat.epi_iff_surjective] -- now `erw` after #13170
infer_instance
#align algebraic_geometry.IsOpenImmersion.range_pullback_snd_of_left AlgebraicGeometry.IsOpenImmersion.range_pullback_snd_of_left

Expand All @@ -584,23 +584,24 @@ theorem range_pullback_fst_of_right :
· erw [TopCat.pullback_fst_image_snd_preimage]
rw [Set.image_univ]
rfl
rw [← TopCat.epi_iff_surjective]
erw [← TopCat.epi_iff_surjective] -- now `erw` after #13170
infer_instance
#align algebraic_geometry.IsOpenImmersion.range_pullback_fst_of_right AlgebraicGeometry.IsOpenImmersion.range_pullback_fst_of_right

theorem range_pullback_to_base_of_left :
Set.range (pullback.fst ≫ f : pullback f g ⟶ Z).1.base =
Set.range f.1.base ∩ Set.range g.1.base := by
rw [pullback.condition, Scheme.comp_val_base, coe_comp, Set.range_comp,
rw [pullback.condition, Scheme.comp_val_base, TopCat.coe_comp, Set.range_comp,
range_pullback_snd_of_left, Opens.carrier_eq_coe,
Opens.map_obj, Opens.coe_mk, Set.image_preimage_eq_inter_range]
#align algebraic_geometry.IsOpenImmersion.range_pullback_to_base_of_left AlgebraicGeometry.IsOpenImmersion.range_pullback_to_base_of_left

theorem range_pullback_to_base_of_right :
Set.range (pullback.fst ≫ g : pullback g f ⟶ Z).1.base =
Set.range g.1.base ∩ Set.range f.1.base := by
rw [Scheme.comp_val_base, coe_comp, Set.range_comp, range_pullback_fst_of_right, Opens.map_obj,
Opens.carrier_eq_coe, Opens.coe_mk, Set.image_preimage_eq_inter_range, Set.inter_comm]
rw [Scheme.comp_val_base, TopCat.coe_comp, Set.range_comp, range_pullback_fst_of_right,
Opens.map_obj, Opens.carrier_eq_coe, Opens.coe_mk, Set.image_preimage_eq_inter_range,
Set.inter_comm]
#align algebraic_geometry.IsOpenImmersion.range_pullback_to_base_of_right AlgebraicGeometry.IsOpenImmersion.range_pullback_to_base_of_right

/-- The universal property of open immersions:
Expand Down Expand Up @@ -740,8 +741,7 @@ def Scheme.OpenCover.pullbackCover {X W : Scheme.{u}} (𝒰 : X.OpenCover) (f :
show _ = (pullback.fst : pullback f (𝒰.map (𝒰.f (f.1.base x))) ⟶ _).1.base from
PreservesPullback.iso_hom_fst Scheme.forgetToTop f (𝒰.map (𝒰.f (f.1.base x)))]
-- Porting note: `rw` to `erw` on this single lemma
erw [coe_comp]
rw [Set.range_comp, Set.range_iff_surjective.mpr, Set.image_univ,
rw [TopCat.coe_comp, Set.range_comp, Set.range_iff_surjective.mpr, Set.image_univ,
TopCat.pullback_fst_range]
· obtain ⟨y, h⟩ := 𝒰.Covers (f.1.base x)
exact ⟨y, h.symm⟩
Expand All @@ -763,8 +763,7 @@ def Scheme.OpenCover.pullbackCover' {X W : Scheme.{u}} (𝒰 : X.OpenCover) (f :
show _ = (pullback.snd : pullback (𝒰.map (𝒰.f (f.1.base x))) f ⟶ _).1.base from
PreservesPullback.iso_hom_snd Scheme.forgetToTop (𝒰.map (𝒰.f (f.1.base x))) f]
-- Porting note: `rw` to `erw` on this single lemma
erw [coe_comp]
rw [Set.range_comp, Set.range_iff_surjective.mpr, Set.image_univ,
rw [TopCat.coe_comp, Set.range_comp, Set.range_iff_surjective.mpr, Set.image_univ,
TopCat.pullback_snd_range]
· obtain ⟨y, h⟩ := 𝒰.Covers (f.1.base x)
exact ⟨y, h⟩
Expand Down
6 changes: 3 additions & 3 deletions Mathlib/AlgebraicGeometry/Restrict.lean
Original file line number Diff line number Diff line change
Expand Up @@ -182,7 +182,7 @@ def Scheme.restrictRestrictComm (X : Scheme.{u}) (U V : Opens X.carrier) :
X ∣_ᵤ U ∣_ᵤ ιOpens U ⁻¹ᵁ V ≅ X ∣_ᵤ V ∣_ᵤ ιOpens V ⁻¹ᵁ U := by
refine IsOpenImmersion.isoOfRangeEq (ιOpens _ ≫ ιOpens U) (ιOpens _ ≫ ιOpens V) ?_
simp only [Scheme.restrict_carrier, Scheme.ofRestrict_val_base, Scheme.comp_coeBase,
CategoryTheory.coe_comp, Opens.coe_inclusion, Set.range_comp, Opens.map]
TopCat.coe_comp, Opens.coe_inclusion, Set.range_comp, Opens.map]
rw [Subtype.range_val, Subtype.range_val]
dsimp
rw [Set.image_preimage_eq_inter_range, Set.image_preimage_eq_inter_range,
Expand All @@ -194,7 +194,7 @@ def Scheme.restrictRestrict (X : Scheme.{u}) (U : Opens X.carrier) (V : Opens (X
X ∣_ᵤ U ∣_ᵤ V ≅ X ∣_ᵤ U.openEmbedding.isOpenMap.functor.obj V := by
refine IsOpenImmersion.isoOfRangeEq (ιOpens _ ≫ ιOpens U) (ιOpens _) ?_
simp only [Scheme.restrict_carrier, Scheme.ofRestrict_val_base, Scheme.comp_coeBase,
CategoryTheory.coe_comp, Opens.coe_inclusion, Set.range_comp, Opens.map]
TopCat.coe_comp, Opens.coe_inclusion, Set.range_comp, Opens.map]
rw [Subtype.range_val, Subtype.range_val]
rfl

Expand Down Expand Up @@ -225,7 +225,7 @@ noncomputable abbrev Scheme.restrictMapIso {X Y : Scheme.{u}} (f : X ⟶ Y) [IsI
(H := PresheafedSpace.IsOpenImmersion.comp (hf := inferInstance) (hg := inferInstance))
(Y.ofRestrict _) _
dsimp [restrict]
rw [coe_comp, Set.range_comp, Opens.coe_inclusion, Subtype.range_val, Subtype.range_coe]
rw [Set.range_comp, Subtype.range_val, Subtype.range_coe]
refine' @Set.image_preimage_eq _ _ f.1.base U.1 _
rw [← TopCat.epi_iff_surjective]
infer_instance
Expand Down
1 change: 1 addition & 0 deletions Mathlib/AlgebraicGeometry/Scheme.lean
Original file line number Diff line number Diff line change
Expand Up @@ -5,6 +5,7 @@ Authors: Scott Morrison
-/
import Mathlib.AlgebraicGeometry.Spec
import Mathlib.Algebra.Category.Ring.Constructions
import Mathlib.CategoryTheory.Elementwise

#align_import algebraic_geometry.Scheme from "leanprover-community/mathlib"@"88474d1b5af6d37c2ab728b757771bced7f5194c"

Expand Down
4 changes: 1 addition & 3 deletions Mathlib/AlgebraicTopology/FundamentalGroupoid/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -369,9 +369,7 @@ def fundamentalGroupoidFunctor : TopCat ⥤ CategoryTheory.Grpd where
map_id := fun X => rfl
map_comp := fun {x y z} p q => by
refine Quotient.inductionOn₂ p q fun a b => ?_
simp only [comp_eq, ← Path.Homotopic.map_lift, ← Path.Homotopic.comp_lift, Path.map_trans]
-- This was not needed before leanprover/lean4#2644
erw [ ← Path.Homotopic.comp_lift]; rfl}
simp only [comp_eq, ← Path.Homotopic.map_lift, ← Path.Homotopic.comp_lift, Path.map_trans] }
map_id X := by
simp only
change _ = (⟨_, _, _⟩ : FundamentalGroupoid X ⥤ FundamentalGroupoid X)
Expand Down
1 change: 1 addition & 0 deletions Mathlib/CategoryTheory/ConcreteCategory/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -95,6 +95,7 @@ variable {C : Type u} [Category.{v} C] [ConcreteCategory.{w} C]
-- Porting note: forget_obj_eq_coe has become a syntactic tautology.
#noalign category_theory.forget_obj_eq_coe

/-- In any concrete category, `(forget C).map` is injective. -/
abbrev ConcreteCategory.instFunLike {X Y : C} : FunLike (X ⟶ Y) X Y where
coe f := (forget C).map f
coe_injective' _ _ h := (forget C).map_injective h
Expand Down
3 changes: 2 additions & 1 deletion Mathlib/Condensed/TopComparison.lean
Original file line number Diff line number Diff line change
Expand Up @@ -53,7 +53,8 @@ theorem factorsThrough_of_pullbackCondition {Z B : C} {π : Z ⟶ B} [HasPullbac
have h₂ : ∀ y, G.map pullback.snd ((PreservesPullback.iso G π π).inv y) =
pullback.snd (f := G.map π) (g := G.map π) y := by
simp only [← PreservesPullback.iso_inv_snd]; intro y; rfl
erw [h₁, h₂] at ha'
erw [h₁, h₂, TopCat.pullbackIsoProdSubtype_inv_fst_apply,
TopCat.pullbackIsoProdSubtype_inv_snd_apply] at ha'
simpa using ha'

/--
Expand Down
Original file line number Diff line number Diff line change
Expand Up @@ -215,7 +215,7 @@ theorem imageBasicOpen_image_open :
IsOpen ((coequalizer.π f.1 g.1).base '' (imageBasicOpen f g U s).1) := by
rw [← (TopCat.homeoOfIso (PreservesCoequalizer.iso (SheafedSpace.forget _) f.1
g.1)).isOpen_preimage, TopCat.coequalizer_isOpen_iff, ← Set.preimage_comp]
erw [← coe_comp]
erw [← TopCat.coe_comp]
rw [PreservesCoequalizer.iso_hom, ι_comp_coequalizerComparison]
dsimp only [SheafedSpace.forget]
-- Porting note (#11224): change `rw` to `erw`
Expand Down
30 changes: 22 additions & 8 deletions Mathlib/Geometry/RingedSpace/OpenImmersion.lean
Original file line number Diff line number Diff line change
Expand Up @@ -165,17 +165,18 @@ instance comp {Z : PresheafedSpace C} (f : X ⟶ Y) [hf : IsOpenImmersion f] (g
ext1
dsimp only [IsOpenMap.functor_obj_coe]
-- Porting note: slightly more hand holding here: `g ∘ f` and `fun x => g (f x)`
rw [comp_base, coe_comp, show g.base ∘ f.base = fun x => g.base (f.base x) from rfl,
← Set.image_image]
erw [comp_base, coe_comp, show g.base ∘ f.base = fun x => g.base (f.base x) from rfl,
← Set.image_image] -- now `erw` after #13170
rw [this]
infer_instance
have : IsIso (f.c.app (op <| (Opens.map g.base).obj ((IsOpenMap.functor h).obj U))) := by
have : (Opens.map g.base).obj (h.functor.obj U) = hf.openFunctor.obj U := by
ext1
dsimp only [Opens.map_coe, IsOpenMap.functor_obj_coe, comp_base]
-- Porting note: slightly more hand holding here: `g ∘ f` and `fun x => g (f x)`
rw [coe_comp, show g.base ∘ f.base = fun x => g.base (f.base x) from rfl,
erw [coe_comp, show g.base ∘ f.base = fun x => g.base (f.base x) from rfl,
← Set.image_image g.base f.base, Set.preimage_image_eq _ hg.base_open.inj]
-- now `erw` after #13170
rw [this]
infer_instance
apply IsIso.comp_isIso
Expand Down Expand Up @@ -364,10 +365,22 @@ def pullbackConeOfLeftFst :
· rintro _ ⟨_, h₁, h₂⟩
use (TopCat.pullbackIsoProdSubtype _ _).inv ⟨⟨_, _⟩, h₂⟩
-- Porting note: need a slight hand holding
-- used to be `simpa using h₁` before #13170
change _ ∈ _ ⁻¹' _ ∧ _
simpa using h₁
simp only [TopCat.coe_of, restrict_carrier, Set.preimage_id', Set.mem_preimage,
SetLike.mem_coe]
constructor
· change _ ∈ U.unop at h₁
convert h₁
erw [TopCat.pullbackIsoProdSubtype_inv_fst_apply]
· erw [TopCat.pullbackIsoProdSubtype_inv_snd_apply]
· rintro _ ⟨x, h₁, rfl⟩
exact ⟨_, h₁, ConcreteCategory.congr_hom pullback.condition x⟩))
-- next line used to be
-- `exact ⟨_, h₁, ConcreteCategory.congr_hom pullback.condition x⟩))`
-- before #13170
refine ⟨_, h₁, ?_⟩
change (_ ≫ f.base) _ = (_ ≫ g.base) _
rw [pullback.condition]))
naturality := by
intro U V i
induction U using Opposite.rec'
Expand Down Expand Up @@ -432,7 +445,7 @@ def pullbackConeOfLeftLift : s.pt ⟶ (pullbackConeOfLeft f g).pt where
erw [← this]
dsimp [s']
-- Porting note: need a bit more hand holding here about function composition
rw [coe_comp, show ∀ f g, f ∘ g = fun x => f (g x) from fun _ _ => rfl]
rw [show ∀ f g, f ∘ g = fun x => f (g x) from fun _ _ => rfl]
erw [← Set.preimage_preimage]
erw [Set.preimage_image_eq _
(TopCat.snd_openEmbedding_of_left_openEmbedding hf.base_open g.base).inj]
Expand Down Expand Up @@ -1169,15 +1182,16 @@ theorem lift_range (H' : Set.range g.1.base ⊆ Set.range f.1.base) :
have : _ = (pullback.fst : pullback f g ⟶ _).val.base :=
PreservesPullback.iso_hom_fst
(LocallyRingedSpace.forgetToSheafedSpace ⋙ SheafedSpace.forget _) f g
rw [LocallyRingedSpace.comp_val, SheafedSpace.comp_base, ← this, ← Category.assoc, coe_comp]
erw [LocallyRingedSpace.comp_val, SheafedSpace.comp_base, ← this, ← Category.assoc, coe_comp]
-- now `erw` after #13170
rw [Set.range_comp, Set.range_iff_surjective.mpr, Set.image_univ]
-- Porting note (#11224): change `rw` to `erw` on this lemma
· erw [TopCat.pullback_fst_range]
ext
constructor
· rintro ⟨y, eq⟩; exact ⟨y, eq.symm⟩
· rintro ⟨y, eq⟩; exact ⟨y, eq.symm⟩
· rw [← TopCat.epi_iff_surjective]
· erw [← TopCat.epi_iff_surjective] -- now `erw` after #13170
rw [show (inv (pullback.snd : pullback f g ⟶ _)).val.base = _ from
(LocallyRingedSpace.forgetToSheafedSpace ⋙ SheafedSpace.forget _).map_inv _]
infer_instance
Expand Down
Loading

0 comments on commit 6d39f51

Please sign in to comment.