From 5eb1c0ba2b6435f7f5dcd338942306d6471e2e99 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Mon, 13 Jun 2022 11:35:33 +0100 Subject: [PATCH 01/84] initial --- src/algebra/category/Group/epi_mono.lean | 47 ++++++++++++++++++++++++ 1 file changed, 47 insertions(+) create mode 100644 src/algebra/category/Group/epi_mono.lean diff --git a/src/algebra/category/Group/epi_mono.lean b/src/algebra/category/Group/epi_mono.lean new file mode 100644 index 0000000000000..e531fe2652c80 --- /dev/null +++ b/src/algebra/category/Group/epi_mono.lean @@ -0,0 +1,47 @@ +import category_theory.epi_mono +import algebra.category.Group.basic + +namespace monoid_hom + +universes u + +variables {A B : Type u} [group A] [group B] + +@[to_additive add_monoid_hom.range_zero_eq_bot] +lemma range_one_eq_bot : + (1 : A →* B).range = ⊥ := +set_like.ext $ λ a, iff.trans monoid_hom.mem_range $ + iff.trans (by simp only [one_apply, exists_const]; split; intros h; subst h) + subgroup.mem_bot.symm + +@[to_additive add_monoid_hom.ker_eq_bot_of_cancel] +lemma ker_eq_bot_of_cancel {f : A →* B} + (h : ∀ (u v : f.ker →* A), f.comp u = f.comp v → u = v) : + f.ker = (⊥ : subgroup A) := +begin + specialize h 1 f.ker.subtype begin + ext1, + simp only [comp_one, one_apply, coe_comp, subgroup.coe_subtype, function.comp_app], + rw [←subtype.val_eq_coe, f.mem_ker.mp x.2], + end, + rw [←subgroup.subtype_range f.ker, ←h, range_one_eq_bot], +end + + +end monoid_hom + +namespace Group + +open category_theory + +universe u + +variables {A B : Group.{u}} (f : A ⟶ B) + +lemma ker_eq_bot_of_mono [mono f] : f.ker = ⊥ := +monoid_hom.ker_eq_bot_of_cancel $ λ u v h, +monoid_hom.ext $ λ x, begin + +end + +end Group From df2506713c24888ad71e9bc605a1189651c3bdfc Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Mon, 13 Jun 2022 21:05:52 +0100 Subject: [PATCH 02/84] finished epi --- src/algebra/category/Group/epi_mono.lean | 102 +++++++++++++++++++++-- 1 file changed, 95 insertions(+), 7 deletions(-) diff --git a/src/algebra/category/Group/epi_mono.lean b/src/algebra/category/Group/epi_mono.lean index e531fe2652c80..5cf76128aed9a 100644 --- a/src/algebra/category/Group/epi_mono.lean +++ b/src/algebra/category/Group/epi_mono.lean @@ -1,11 +1,41 @@ import category_theory.epi_mono +import group_theory.quotient_group import algebra.category.Group.basic +universes u + namespace monoid_hom -universes u +open quotient_group + +variables {A B : Type u} + +section + +variables [group A] [group B] + +@[to_additive add_monoid_hom.ker_eq_bot] +lemma ker_eq_bot (f : A →* B) : + f.ker = ⊥ ↔ function.injective f := +{ mp := λ h1 x y h2, begin + rw [←div_eq_one, ←map_div, ←mem_ker] at h2, + rwa [h1, subgroup.mem_bot, div_eq_one] at h2, + end, + mpr := λ h, set_like.ext $ λ x, iff.trans (mem_ker _) $ + iff.trans begin + rw ←(by rw map_one : f x = f 1 ↔ f x = 1), + exact ⟨λ h', h h', λ h, h ▸ rfl⟩, + end subgroup.mem_bot.symm } -variables {A B : Type u} [group A] [group B] +@[to_additive add_monoid_hom.range_eq_top] +lemma range_eq_top (f : A →* B) : + f.range = ⊤ ↔ function.surjective f := +{ mp := λ h x, begin + rcases show x ∈ f.range, from h.symm ▸ subgroup.mem_top _ with ⟨a, h⟩, + exact ⟨a, h⟩, + end, + mpr := λ h, set_like.ext $ λ x, iff.trans mem_range + ⟨λ _, trivial, λ _, h x⟩ } @[to_additive add_monoid_hom.range_zero_eq_bot] lemma range_one_eq_bot : @@ -13,6 +43,10 @@ lemma range_one_eq_bot : set_like.ext $ λ a, iff.trans monoid_hom.mem_range $ iff.trans (by simp only [one_apply, exists_const]; split; intros h; subst h) subgroup.mem_bot.symm +@[to_additive add_monoid_hom.ker_zero_eq_top] +lemma ker_one_eq_top : + (1 : A →* B).ker = ⊤ := +set_like.ext $ λ a, iff.trans (monoid_hom.mem_ker _) $ by simp @[to_additive add_monoid_hom.ker_eq_bot_of_cancel] lemma ker_eq_bot_of_cancel {f : A →* B} @@ -27,21 +61,75 @@ begin rw [←subgroup.subtype_range f.ker, ←h, range_one_eq_bot], end +end + +section + +variables [comm_group A] [comm_group B] + +@[to_additive add_monoid_hom.range_eq_top_of_cancel] +lemma range_eq_top_of_cancel {f : A →* B} + (h : ∀ (u v : B →* B ⧸ f.range), u.comp f = v.comp f → u = v) : + f.range = ⊤ := +begin + specialize h 1 (quotient_group.mk' _) begin + ext1, + simp only [one_apply, coe_comp, coe_mk', function.comp_app], + rw [show (1 : B ⧸ f.range) = (1 : B), from quotient_group.coe_one _, quotient_group.eq, inv_one, + one_mul], + exact ⟨x, rfl⟩, + end, + replace h : (quotient_group.mk' _).ker = (1 : B →* B ⧸ f.range).ker := by rw h, + rwa [ker_one_eq_top, quotient_group.ker_mk] at h, +end + +end end monoid_hom -namespace Group +section open category_theory -universe u +namespace Group variables {A B : Group.{u}} (f : A ⟶ B) lemma ker_eq_bot_of_mono [mono f] : f.ker = ⊥ := -monoid_hom.ker_eq_bot_of_cancel $ λ u v h, -monoid_hom.ext $ λ x, begin +monoid_hom.ker_eq_bot_of_cancel $ λ u v, + (@cancel_mono _ _ _ _ _ f _ + (show Group.of f.ker ⟶ A, from u) _).1 -end +lemma mono_iff_ker_eq_bot : + mono f ↔ f.ker = ⊥ := +⟨λ h, @@ker_eq_bot_of_mono f h, +λ h, concrete_category.mono_of_injective _ $ (monoid_hom.ker_eq_bot f).1 h⟩ + +lemma mono_iff_injective : + mono f ↔ function.injective f := +iff.trans (mono_iff_ker_eq_bot f) $ monoid_hom.ker_eq_bot f end Group + +namespace CommGroup + +variables {A B : CommGroup.{u}} (f : A ⟶ B) + +@[to_additive] +lemma range_eq_top_of_epi [epi f] : f.range = ⊤ := +monoid_hom.range_eq_top_of_cancel $ λ u v, + (@cancel_epi _ _ _ _ _ f _ + (show CommGroup.of B ⟶ CommGroup.of (B ⧸ f.range), from u) _).1 + +lemma epi_iff_range_eq_top : + epi f ↔ f.range = ⊤ := +⟨λ h, @@range_eq_top_of_epi f h, +λ h, concrete_category.epi_of_surjective _ $ (monoid_hom.range_eq_top f).1 h⟩ + +lemma epi_iff_surjective : + epi f ↔ function.surjective f := +iff.trans (epi_iff_range_eq_top f) $ monoid_hom.range_eq_top f + +end CommGroup + +end From 99553bdaf0d7ff4d9cbf22117e2c4fd0157a3108 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Mon, 13 Jun 2022 21:09:09 +0100 Subject: [PATCH 03/84] add to_additive --- src/algebra/category/Group/epi_mono.lean | 4 +++- 1 file changed, 3 insertions(+), 1 deletion(-) diff --git a/src/algebra/category/Group/epi_mono.lean b/src/algebra/category/Group/epi_mono.lean index 5cf76128aed9a..625de63f53108 100644 --- a/src/algebra/category/Group/epi_mono.lean +++ b/src/algebra/category/Group/epi_mono.lean @@ -115,17 +115,19 @@ namespace CommGroup variables {A B : CommGroup.{u}} (f : A ⟶ B) -@[to_additive] +@[to_additive AddCommGroup.range_eq_top_of_epi] lemma range_eq_top_of_epi [epi f] : f.range = ⊤ := monoid_hom.range_eq_top_of_cancel $ λ u v, (@cancel_epi _ _ _ _ _ f _ (show CommGroup.of B ⟶ CommGroup.of (B ⧸ f.range), from u) _).1 +@[to_additive AddCommGroup.epi_iff_range_eq_top] lemma epi_iff_range_eq_top : epi f ↔ f.range = ⊤ := ⟨λ h, @@range_eq_top_of_epi f h, λ h, concrete_category.epi_of_surjective _ $ (monoid_hom.range_eq_top f).1 h⟩ +@[to_additive AddCommGroup.epi_iff_surjective] lemma epi_iff_surjective : epi f ↔ function.surjective f := iff.trans (epi_iff_range_eq_top f) $ monoid_hom.range_eq_top f From 793b87b1e0a206f5b617bd293acaf13830beaffb Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Mon, 13 Jun 2022 21:11:37 +0100 Subject: [PATCH 04/84] add doc_string --- src/algebra/category/Group/epi_mono.lean | 12 ++++++++++++ 1 file changed, 12 insertions(+) diff --git a/src/algebra/category/Group/epi_mono.lean b/src/algebra/category/Group/epi_mono.lean index 625de63f53108..a96f48f62f8c7 100644 --- a/src/algebra/category/Group/epi_mono.lean +++ b/src/algebra/category/Group/epi_mono.lean @@ -1,7 +1,19 @@ +/- +Copyright (c) 2022 Jujian Zhang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jujian Zhang +-/ import category_theory.epi_mono import group_theory.quotient_group import algebra.category.Group.basic +/-! +# Monomorphisms and epimorphism in `Group` and `CommGroup` +In this file, we prove monomorphisms in category of group are injective homomorphisms and +epimorphisms are surjective homomorphisms. +-/ + + universes u namespace monoid_hom From d80df7fc10a05d62fd46c4a7793fa42a65a15044 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Wed, 15 Jun 2022 19:53:51 +0100 Subject: [PATCH 05/84] generalize --- src/algebra/category/Group/epi_mono.lean | 378 ++++++++++++++++++++++- 1 file changed, 362 insertions(+), 16 deletions(-) diff --git a/src/algebra/category/Group/epi_mono.lean b/src/algebra/category/Group/epi_mono.lean index a96f48f62f8c7..49ba14e7fa893 100644 --- a/src/algebra/category/Group/epi_mono.lean +++ b/src/algebra/category/Group/epi_mono.lean @@ -4,6 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jujian Zhang -/ import category_theory.epi_mono +-- import algebra.category.Group.limits +import group_theory.coset +import group_theory.perm.basic import group_theory.quotient_group import algebra.category.Group.basic @@ -121,29 +124,372 @@ lemma mono_iff_injective : mono f ↔ function.injective f := iff.trans (mono_iff_ker_eq_bot f) $ monoid_hom.ker_eq_bot f -end Group +namespace surjective_of_epi_auxs -namespace CommGroup +local notation `X` := set.range (function.swap left_coset f.range.carrier) -variables {A B : CommGroup.{u}} (f : A ⟶ B) +inductive X_with_infinity +| from_coset : X → X_with_infinity +| infinity : X_with_infinity -@[to_additive AddCommGroup.range_eq_top_of_epi] -lemma range_eq_top_of_epi [epi f] : f.range = ⊤ := -monoid_hom.range_eq_top_of_cancel $ λ u v, - (@cancel_epi _ _ _ _ _ f _ - (show CommGroup.of B ⟶ CommGroup.of (B ⧸ f.range), from u) _).1 +open X_with_infinity equiv.perm +open_locale coset -@[to_additive AddCommGroup.epi_iff_range_eq_top] -lemma epi_iff_range_eq_top : - epi f ↔ f.range = ⊤ := -⟨λ h, @@range_eq_top_of_epi f h, -λ h, concrete_category.epi_of_surjective _ $ (monoid_hom.range_eq_top f).1 h⟩ +local notation `X'` := X_with_infinity f +local notation `⊙` := X_with_infinity.infinity +local notation `SX'` := equiv.perm X' + +noncomputable instance : decidable_eq X' := classical.dec_eq _ +noncomputable def tau : SX' := equiv.swap +(from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩) ⊙ + +local notation `τ` := tau f + +lemma τ_apply_infinity : + τ ⊙ = from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩ := +begin + dunfold tau, + rw [equiv.swap_apply_right], +end + +lemma τ_apply_from_coset : + τ (from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩) = ⊙ := +begin + dunfold tau, + rw [equiv.swap_apply_left], +end + +lemma τ_apply_from_coset' (x : B) (hx : x ∈ f.range) : + τ (from_coset ⟨x *l f.range.carrier, ⟨x, rfl⟩⟩) = ⊙ := +begin + convert τ_apply_from_coset _, + ext b, + simp only [mem_left_coset_iff, subgroup.mem_carrier, monoid_hom.mem_range], + rcases hx with ⟨c, rfl⟩, + split, + { rintros ⟨a, ha⟩, + use c * a, + rw [map_mul, ha, ←mul_assoc, mul_right_inv (f c), one_mul], }, + { rintros ⟨a, rfl⟩, + use c⁻¹ * a, + rw [map_mul, map_inv], }, +end + +lemma τ_symm_apply_from_coset : + (equiv.symm τ) (from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩) = ⊙ := +begin + dunfold tau, + rw [equiv.symm_swap, equiv.swap_apply_left], +end + +lemma τ_symm_apply_infinity : + (equiv.symm τ) ⊙ = from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩ := +begin + dunfold tau, + rw [equiv.symm_swap, equiv.swap_apply_right], +end + +def G : B ⟶ Group.of SX' := +{ to_fun := λ β, + { to_fun := λ x, + match x with + | from_coset y := from_coset ⟨β *l y, begin + rcases y.2 with ⟨b, hb⟩, + rw [subtype.val_eq_coe] at hb, + rw [←hb, set.mem_range, left_coset_assoc], + use β * b, + end⟩ + | ⊙ := ⊙ + end, + inv_fun := λ x, + match x with + | from_coset y := from_coset ⟨β⁻¹ *l y, begin + rcases y.2 with ⟨b, hb⟩, + rw [subtype.val_eq_coe] at hb, + rw [←hb, set.mem_range, left_coset_assoc], + use β⁻¹ * b, + end⟩ + | ⊙ := ⊙ + end, + left_inv := λ x, + match x with + | from_coset y := begin + change from_coset ⟨_, _⟩ = _, + simp + end + | ⊙ := rfl + end, + right_inv := λ x, + match x with + | from_coset y := begin + change from_coset _ = _, + simp + end + | ⊙ := rfl + end }, + map_one' := begin + ext x, + rcases x with ⟨⟨_, ⟨y, rfl⟩⟩⟩, + { simp only [equiv.coe_fn_mk, equiv.perm.coe_one, id.def], + change from_coset _ = _, + simp }, + { refl }, + end, + map_mul' := λ b1 b2, begin + ext x, + rcases x with ⟨⟨_, ⟨y, rfl⟩⟩⟩, + { simp only [equiv.coe_fn_mk, equiv.perm.coe_mul, function.comp_app], + change from_coset _ = from_coset _, + simp only [left_coset_assoc, subtype.coe_mk, subtype.mk_eq_mk, mul_assoc], }, + { refl }, + end } + +local notation `g` := G f + +noncomputable def H : B ⟶ Group.of SX':= +{ to_fun := λ β, (equiv.trans (equiv.symm τ) (g β)).trans τ, + map_one' := begin + ext x, + simp + end, + map_mul' := λ b1 b2, begin + ext x, + simp + end } + +local notation `h` := H f + +lemma g_apply_from_coset (x : B) (y : X) : + (g x).to_fun (from_coset y) = + from_coset ⟨x *l y, begin + rcases y.2 with ⟨b, hb⟩, + rw [subtype.val_eq_coe] at hb, + rw [←hb, set.mem_range, left_coset_assoc], + use x * b, + end⟩ := rfl + +lemma g_apply_infinity (x : B) : + (g x).to_fun ⊙ = ⊙ := rfl + +lemma h_apply_infinity (x : B) (hx : x ∈ f.range) : + (h x).to_fun ⊙ = ⊙ := +begin + dunfold H, + simp only [monoid_hom.coe_mk, equiv.to_fun_as_coe, equiv.coe_trans, function.comp_app], + rw [τ_symm_apply_infinity], + have := g_apply_from_coset f x ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩, + rw [equiv.to_fun_as_coe] at this, + rw this, + simpa only [←subtype.val_eq_coe] using τ_apply_from_coset' f x hx, +end + +lemma h_apply_from_coset (x : B) : + (h x).to_fun (from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩) = + from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩ := +begin + dunfold H, + simp only [monoid_hom.coe_mk, equiv.to_fun_as_coe, equiv.coe_trans, function.comp_app], + rw [τ_symm_apply_from_coset], + have := g_apply_infinity f x, + rw [equiv.to_fun_as_coe] at this, + rw [this, τ_apply_infinity], +end + +lemma h_apply_from_coset' (x : B) (b : B) (hb : b ∈ f.range): + (h x).to_fun (from_coset ⟨b *l f.range.carrier, ⟨b, rfl⟩⟩) = + from_coset ⟨b *l f.range.carrier, ⟨b, rfl⟩⟩ := +begin + have eq1 : b *l (monoid_hom.range f).carrier = (monoid_hom.range f).carrier, + { rcases hb with ⟨a, rfl⟩, + ext y, + simp only [mem_left_coset_iff, subgroup.mem_carrier, monoid_hom.mem_range], + split, + { rintros ⟨a', ha'⟩, + use a * a', + rw [map_mul, ha', ←mul_assoc, mul_right_inv, one_mul], }, + { rintros ⟨a', rfl⟩, + use a⁻¹ * a', + rw [map_mul, map_inv], }, }, + convert h_apply_from_coset f x, +end + +lemma h_apply_from_coset_nin_range (x : B) (hx : x ∈ f.range) (b : B) (hb : b ∉ f.range) : + (h x).to_fun (from_coset ⟨b *l f.range.carrier, ⟨b, rfl⟩⟩) = + from_coset ⟨(x * b) *l f.range.carrier, ⟨x * b, rfl⟩⟩ := +begin + dunfold H tau, + simp only [monoid_hom.coe_mk, equiv.to_fun_as_coe, equiv.coe_trans, function.comp_app], + rw [equiv.symm_swap], + rw @equiv.swap_apply_of_ne_of_ne X' _ + (from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩) ⊙ + (from_coset ⟨b *l f.range.carrier, ⟨b, rfl⟩⟩) begin + intro r, + simp only [subtype.mk_eq_mk] at r, + change b *l f.range = f.range at r, + have eq1 : (f.range : set B) = 1 *l f.range, + { rw one_left_coset, }, + conv_rhs at r { rw eq1 }, + rw [left_coset_eq_iff, mul_one] at r, + apply hb, + replace r : b⁻¹⁻¹ ∈ f.range := subgroup.inv_mem _ r, + rwa [inv_inv] at r, + end begin + simp only [ne.def, not_false_iff], + end, + have := g_apply_from_coset f x, + rw [equiv.to_fun_as_coe] at this, + rw this, + clear this, + simp only [←subtype.val_eq_coe, left_coset_assoc], + refine equiv.swap_apply_of_ne_of_ne begin + intro r, + simp only [subtype.mk_eq_mk] at r, + change x * b *l f.range = f.range at r, + have eq1 : (f.range : set B) = 1 *l f.range, + { rw one_left_coset, }, + conv_rhs at r { rw eq1 }, + rw [left_coset_eq_iff, mul_one] at r, + replace r : (x * b)⁻¹⁻¹ ∈ f.range := subgroup.inv_mem _ r, + rw inv_inv at r, + apply hb, + replace hx : x⁻¹ ∈ f.range := subgroup.inv_mem _ hx, + have := subgroup.mul_mem _ hx r, + rwa [←mul_assoc, mul_left_inv, one_mul] at this, + end begin + simp only [ne.def, not_false_iff], + end, +end + +lemma agree : + f.range.carrier = {x | h x = g x} := +begin + ext b, + split, + { rintros ⟨a, rfl⟩, + change h (f a) = g (f a), + ext, + rcases x with ⟨⟨_, ⟨y, rfl⟩⟩⟩, + { have := g_apply_from_coset f (f a), + rw [equiv.to_fun_as_coe] at this, + rw this, + clear this, + by_cases m : y ∈ f.range, + { have := h_apply_from_coset' f (f a) y m, + rw [equiv.to_fun_as_coe] at this, + simp only [this, ←subtype.val_eq_coe], + congr' 1, + change y *l f.range = f a *l (y *l f.range), + rw [left_coset_assoc, left_coset_eq_iff], + refine subgroup.mul_mem _ (subgroup.inv_mem _ m) + (subgroup.mul_mem _ ⟨a, rfl⟩ m), }, + { have := h_apply_from_coset_nin_range f (f a) ⟨a, rfl⟩ y m, + rw [equiv.to_fun_as_coe] at this, + rw this, + clear this, + simp only [←subtype.val_eq_coe, left_coset_assoc], + refl, }, }, + { have := g_apply_infinity f (f a), + rw [equiv.to_fun_as_coe] at this, + rw this, + clear this, + have := h_apply_infinity f (f a) ⟨a, rfl⟩, + simpa only [equiv.to_fun_as_coe] using this, }, }, + { rintros (hb : h b = g b), + have eq1 : (h b).to_fun (from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩) = + (from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩), + { dunfold H tau, + simp only [equiv.symm_swap, monoid_hom.coe_mk, equiv.to_fun_as_coe, equiv.coe_trans, + function.comp_app, equiv.swap_apply_left], + have := g_apply_infinity f b, + rw [equiv.to_fun_as_coe] at this, + rw [this, equiv.swap_apply_right], }, + have eq2 : (g b).to_fun (from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩) = + (from_coset ⟨b *l f.range.carrier, ⟨b, rfl⟩⟩), + { unfold G, + simp only [monoid_hom.coe_mk], + change from_coset _ = _, + refl, }, + have eq3 : (h b).to_fun (from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩) = + (g b).to_fun (from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩), + { simp only [equiv.to_fun_as_coe, hb], }, + rw [eq1, eq2] at eq3, + simp only [subtype.mk_eq_mk] at eq3, + change (f.range : set B) = b *l f.range at eq3, + rwa [show (f.range : set B) = 1 *l f.range, by rw one_left_coset _, + left_coset_assoc, left_coset_eq_iff, inv_one, one_mul, mul_one] at eq3 } +end + +lemma comp_eq : f ≫ g = f ≫ h := +begin + ext1 a, + rw [comp_apply, comp_apply], + have : f a ∈ f.range.carrier := ⟨a, rfl⟩, + rw agree at this, + simp only [set.mem_set_of_eq] at this, + rw this, +end + +lemma g_eq_h [epi f] : g = h := +(cancel_epi f).1 $ comp_eq f + +lemma g_ne_h [epi f] (x : B) (hx : x ∉ f.range) : + g ≠ h := +begin + intros r, + replace r : ∀ a, (g x).to_fun a = (h x).to_fun a, + { intros a, + simp only [equiv.to_fun_as_coe, r], }, + specialize r (from_coset ⟨f.range, ⟨1, one_left_coset _⟩⟩), + rw [H, g_apply_from_coset, monoid_hom.coe_mk, tau] at r, + simp only [monoid_hom.coe_range, subtype.coe_mk, equiv.symm_swap, + equiv.to_fun_as_coe, equiv.coe_trans, function.comp_app] at r, + generalize_proofs h1 h2 at r, + rw [show from_coset ⟨set.range f, h2⟩ = from_coset ⟨f.range.carrier, h2⟩, + by simpa only [subtype.mk_eq_mk], equiv.swap_apply_left] at r, + have := g_apply_infinity f x, + rw [equiv.to_fun_as_coe] at this, + rw [this, equiv.swap_apply_right] at r, + clear this, + simp only [subtype.mk_eq_mk] at r, + change x *l f.range = f.range at r, + rw [show (f.range : set B) = 1 *l f.range, from (one_left_coset _).symm, + left_coset_assoc, mul_one, left_coset_eq_iff, mul_one] at r, + replace r := subgroup.inv_mem _ r, + rw inv_inv at r, + exact hx r, +end + +end surjective_of_epi_auxs + +lemma surjective_of_epi [epi f] : function.surjective f := +begin + have := surjective_of_epi_auxs.g_eq_h f, + by_contra r, + simp_rw [not_forall, not_exists] at r, + rcases r with ⟨b, hb⟩, + contrapose! this, + apply surjective_of_epi_auxs.g_ne_h f b (λ r, _), + rcases r with ⟨c, hc⟩, + specialize hb c, + exact hb hc, +end -@[to_additive AddCommGroup.epi_iff_surjective] lemma epi_iff_surjective : epi f ↔ function.surjective f := -iff.trans (epi_iff_range_eq_top f) $ monoid_hom.range_eq_top f +⟨λ h, @@surjective_of_epi f h, concrete_category.epi_of_surjective _⟩ + +lemma epi_iff_range_eq_top : + epi f ↔ f.range = ⊤ := +iff.trans (epi_iff_surjective _) begin + rw subgroup.eq_top_iff', + split, + { intros h x, + rcases h x with ⟨a, h⟩, + exact ⟨a, h⟩, }, + { intros h x, + exact h x, }, +end -end CommGroup +end Group end From 9a7d5c9ff8893abcba672174d4335f29eb0e8585 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Wed, 15 Jun 2022 20:15:24 +0100 Subject: [PATCH 06/84] add to_addtive --- src/algebra/category/Group/epi_mono.lean | 36 ++++++++++++++++++++---- 1 file changed, 31 insertions(+), 5 deletions(-) diff --git a/src/algebra/category/Group/epi_mono.lean b/src/algebra/category/Group/epi_mono.lean index 49ba14e7fa893..55d180c843438 100644 --- a/src/algebra/category/Group/epi_mono.lean +++ b/src/algebra/category/Group/epi_mono.lean @@ -4,12 +4,13 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jujian Zhang -/ import category_theory.epi_mono --- import algebra.category.Group.limits import group_theory.coset import group_theory.perm.basic import group_theory.quotient_group import algebra.category.Group.basic +noncomputable theory + /-! # Monomorphisms and epimorphism in `Group` and `CommGroup` In this file, we prove monomorphisms in category of group are injective homomorphisms and @@ -110,16 +111,19 @@ namespace Group variables {A B : Group.{u}} (f : A ⟶ B) +@[to_additive AddGroup.ker_eq_bot_of_mono] lemma ker_eq_bot_of_mono [mono f] : f.ker = ⊥ := monoid_hom.ker_eq_bot_of_cancel $ λ u v, (@cancel_mono _ _ _ _ _ f _ (show Group.of f.ker ⟶ A, from u) _).1 +@[to_additive AddGroup.mono_iff_ker_eq_bot] lemma mono_iff_ker_eq_bot : mono f ↔ f.ker = ⊥ := ⟨λ h, @@ker_eq_bot_of_mono f h, λ h, concrete_category.mono_of_injective _ $ (monoid_hom.ker_eq_bot f).1 h⟩ +@[to_additive AddGroup.mono_iff_injective] lemma mono_iff_injective : mono f ↔ function.injective f := iff.trans (mono_iff_ker_eq_bot f) $ monoid_hom.ker_eq_bot f @@ -128,8 +132,11 @@ namespace surjective_of_epi_auxs local notation `X` := set.range (function.swap left_coset f.range.carrier) +/-- +Define `X'` to be the set of all right cosets with an extra point at "infinity". +-/ inductive X_with_infinity -| from_coset : X → X_with_infinity +| from_coset : set.range (function.swap left_coset f.range.carrier) → X_with_infinity | infinity : X_with_infinity open X_with_infinity equiv.perm @@ -139,8 +146,12 @@ local notation `X'` := X_with_infinity f local notation `⊙` := X_with_infinity.infinity local notation `SX'` := equiv.perm X' -noncomputable instance : decidable_eq X' := classical.dec_eq _ -noncomputable def tau : SX' := equiv.swap +instance : decidable_eq X' := classical.dec_eq _ + +/-- +Let `τ` be the permutation on `X'` exchanging `f.range` and the point at infinity. +-/ +def tau : SX' := equiv.swap (from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩) ⊙ local notation `τ` := tau f @@ -189,6 +200,10 @@ begin rw [equiv.symm_swap, equiv.swap_apply_right], end +/-- +Let `g : B ⟶ S(X')` be defined as as such that, for any `β : B`, `g(β)` is the function sending +point at infinity to point at infinity and sending coset `y` to `β *l y`. +-/ def G : B ⟶ Group.of SX' := { to_fun := λ β, { to_fun := λ x, @@ -246,6 +261,9 @@ def G : B ⟶ Group.of SX' := local notation `g` := G f +/-- +Define `h : B ⟶ S(X')` to be `τ g τ⁻¹` +-/ noncomputable def H : B ⟶ Group.of SX':= { to_fun := λ β, (equiv.trans (equiv.symm τ) (g β)).trans τ, map_one' := begin @@ -259,6 +277,13 @@ noncomputable def H : B ⟶ Group.of SX':= local notation `h` := H f +/-- +The strategy is the following: assuming `epi f` +* prove that `f.range = {x | h x = g x}`; +* thus `f ≫ h = f ≫ g` so that `h = g`; +* but if `f` is not surjective, then some `x ≠ f.range`, then `h x ≠ g x` at the coset `f.range`. +-/ + lemma g_apply_from_coset (x : B) (y : X) : (g x).to_fun (from_coset y) = from_coset ⟨x *l y, begin @@ -432,7 +457,7 @@ end lemma g_eq_h [epi f] : g = h := (cancel_epi f).1 $ comp_eq f -lemma g_ne_h [epi f] (x : B) (hx : x ∉ f.range) : +lemma g_ne_h (x : B) (hx : x ∉ f.range) : g ≠ h := begin intros r, @@ -492,4 +517,5 @@ end end Group + end From 5170d987f846283375c98433a74646efde489857 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Wed, 15 Jun 2022 20:17:39 +0100 Subject: [PATCH 07/84] remove noncomputable --- src/algebra/category/Group/epi_mono.lean | 6 ++---- 1 file changed, 2 insertions(+), 4 deletions(-) diff --git a/src/algebra/category/Group/epi_mono.lean b/src/algebra/category/Group/epi_mono.lean index 55d180c843438..5d2b97662b5c8 100644 --- a/src/algebra/category/Group/epi_mono.lean +++ b/src/algebra/category/Group/epi_mono.lean @@ -9,8 +9,6 @@ import group_theory.perm.basic import group_theory.quotient_group import algebra.category.Group.basic -noncomputable theory - /-! # Monomorphisms and epimorphism in `Group` and `CommGroup` In this file, we prove monomorphisms in category of group are injective homomorphisms and @@ -146,12 +144,12 @@ local notation `X'` := X_with_infinity f local notation `⊙` := X_with_infinity.infinity local notation `SX'` := equiv.perm X' -instance : decidable_eq X' := classical.dec_eq _ +noncomputable instance : decidable_eq X' := classical.dec_eq _ /-- Let `τ` be the permutation on `X'` exchanging `f.range` and the point at infinity. -/ -def tau : SX' := equiv.swap +noncomputable def tau : SX' := equiv.swap (from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩) ⊙ local notation `τ` := tau f From 370c5f323c7c1358cf40939d4b6e1b6f77b68e6f Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Wed, 15 Jun 2022 20:30:03 +0100 Subject: [PATCH 08/84] minor golf --- src/algebra/category/Group/epi_mono.lean | 20 +++++--------------- 1 file changed, 5 insertions(+), 15 deletions(-) diff --git a/src/algebra/category/Group/epi_mono.lean b/src/algebra/category/Group/epi_mono.lean index 5d2b97662b5c8..066be2cd10162 100644 --- a/src/algebra/category/Group/epi_mono.lean +++ b/src/algebra/category/Group/epi_mono.lean @@ -486,15 +486,13 @@ end surjective_of_epi_auxs lemma surjective_of_epi [epi f] : function.surjective f := begin - have := surjective_of_epi_auxs.g_eq_h f, by_contra r, simp_rw [not_forall, not_exists] at r, rcases r with ⟨b, hb⟩, - contrapose! this, - apply surjective_of_epi_auxs.g_ne_h f b (λ r, _), - rcases r with ⟨c, hc⟩, - specialize hb c, - exact hb hc, + refine surjective_of_epi_auxs.g_ne_h f b (λ r, _) _, + { rcases r with ⟨c, hc⟩, + exact hb _ hc, }, + apply surjective_of_epi_auxs.g_eq_h, end lemma epi_iff_surjective : @@ -503,15 +501,7 @@ lemma epi_iff_surjective : lemma epi_iff_range_eq_top : epi f ↔ f.range = ⊤ := -iff.trans (epi_iff_surjective _) begin - rw subgroup.eq_top_iff', - split, - { intros h x, - rcases h x with ⟨a, h⟩, - exact ⟨a, h⟩, }, - { intros h x, - exact h x, }, -end +iff.trans (epi_iff_surjective _) (subgroup.eq_top_iff' f.range).symm end Group From 29808f5e9459245c8591f4f9a58fe8264f5c3968 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Wed, 15 Jun 2022 20:38:03 +0100 Subject: [PATCH 09/84] minimize import --- src/algebra/category/Group/epi_mono.lean | 4 +--- 1 file changed, 1 insertion(+), 3 deletions(-) diff --git a/src/algebra/category/Group/epi_mono.lean b/src/algebra/category/Group/epi_mono.lean index 066be2cd10162..e51192126587b 100644 --- a/src/algebra/category/Group/epi_mono.lean +++ b/src/algebra/category/Group/epi_mono.lean @@ -4,13 +4,11 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jujian Zhang -/ import category_theory.epi_mono -import group_theory.coset -import group_theory.perm.basic import group_theory.quotient_group import algebra.category.Group.basic /-! -# Monomorphisms and epimorphism in `Group` and `CommGroup` +# Monomorphisms and epimorphism in `Group` In this file, we prove monomorphisms in category of group are injective homomorphisms and epimorphisms are surjective homomorphisms. -/ From 4640408f0f53cdfdc6cdc0b91d4d3f1275667c20 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Fri, 17 Jun 2022 01:49:15 +0100 Subject: [PATCH 10/84] golf --- src/algebra/category/Group/epi_mono.lean | 121 ++++++++++------------- 1 file changed, 51 insertions(+), 70 deletions(-) diff --git a/src/algebra/category/Group/epi_mono.lean b/src/algebra/category/Group/epi_mono.lean index e51192126587b..2d5a2389d2079 100644 --- a/src/algebra/category/Group/epi_mono.lean +++ b/src/algebra/category/Group/epi_mono.lean @@ -131,43 +131,76 @@ local notation `X` := set.range (function.swap left_coset f.range.carrier) /-- Define `X'` to be the set of all right cosets with an extra point at "infinity". -/ +@[nolint has_inhabited_instance] inductive X_with_infinity | from_coset : set.range (function.swap left_coset f.range.carrier) → X_with_infinity | infinity : X_with_infinity + open X_with_infinity equiv.perm open_locale coset local notation `X'` := X_with_infinity f -local notation `⊙` := X_with_infinity.infinity +local notation `∞` := X_with_infinity.infinity local notation `SX'` := equiv.perm X' +instance : has_scalar B X' := +{ smul := λ b x, match x with + | from_coset y := from_coset ⟨b *l y, begin + rcases y.2 with ⟨b', hb'⟩, + use b * b', + rw [←subtype.val_eq_coe, ←hb', left_coset_assoc], + end⟩ + | ∞ := ∞ + end } + +lemma mul_smul (b b' : B) (x : X') : + (b * b') • x = b • b' • x := +match x with +| from_coset y := begin + change from_coset _ = from_coset _, + simp only [←subtype.val_eq_coe, left_coset_assoc], +end +| ∞ := rfl +end + +lemma one_smul (x : X') : + (1 : B) • x = x := +match x with +| from_coset y := begin + change from_coset _ = from_coset _, + simp only [←subtype.val_eq_coe, one_left_coset, subtype.ext_iff_val], +end +| ∞ := rfl +end + + noncomputable instance : decidable_eq X' := classical.dec_eq _ /-- Let `τ` be the permutation on `X'` exchanging `f.range` and the point at infinity. -/ noncomputable def tau : SX' := equiv.swap -(from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩) ⊙ +(from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩) ∞ local notation `τ` := tau f lemma τ_apply_infinity : - τ ⊙ = from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩ := + τ ∞ = from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩ := begin dunfold tau, rw [equiv.swap_apply_right], end lemma τ_apply_from_coset : - τ (from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩) = ⊙ := + τ (from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩) = ∞ := begin dunfold tau, rw [equiv.swap_apply_left], end lemma τ_apply_from_coset' (x : B) (hx : x ∈ f.range) : - τ (from_coset ⟨x *l f.range.carrier, ⟨x, rfl⟩⟩) = ⊙ := + τ (from_coset ⟨x *l f.range.carrier, ⟨x, rfl⟩⟩) = ∞ := begin convert τ_apply_from_coset _, ext b, @@ -183,14 +216,14 @@ begin end lemma τ_symm_apply_from_coset : - (equiv.symm τ) (from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩) = ⊙ := + (equiv.symm τ) (from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩) = ∞ := begin dunfold tau, rw [equiv.symm_swap, equiv.swap_apply_left], end lemma τ_symm_apply_infinity : - (equiv.symm τ) ⊙ = from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩ := + (equiv.symm τ) ∞ = from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩ := begin dunfold tau, rw [equiv.symm_swap, equiv.swap_apply_right], @@ -202,58 +235,12 @@ point at infinity to point at infinity and sending coset `y` to `β *l y`. -/ def G : B ⟶ Group.of SX' := { to_fun := λ β, - { to_fun := λ x, - match x with - | from_coset y := from_coset ⟨β *l y, begin - rcases y.2 with ⟨b, hb⟩, - rw [subtype.val_eq_coe] at hb, - rw [←hb, set.mem_range, left_coset_assoc], - use β * b, - end⟩ - | ⊙ := ⊙ - end, - inv_fun := λ x, - match x with - | from_coset y := from_coset ⟨β⁻¹ *l y, begin - rcases y.2 with ⟨b, hb⟩, - rw [subtype.val_eq_coe] at hb, - rw [←hb, set.mem_range, left_coset_assoc], - use β⁻¹ * b, - end⟩ - | ⊙ := ⊙ - end, - left_inv := λ x, - match x with - | from_coset y := begin - change from_coset ⟨_, _⟩ = _, - simp - end - | ⊙ := rfl - end, - right_inv := λ x, - match x with - | from_coset y := begin - change from_coset _ = _, - simp - end - | ⊙ := rfl - end }, - map_one' := begin - ext x, - rcases x with ⟨⟨_, ⟨y, rfl⟩⟩⟩, - { simp only [equiv.coe_fn_mk, equiv.perm.coe_one, id.def], - change from_coset _ = _, - simp }, - { refl }, - end, - map_mul' := λ b1 b2, begin - ext x, - rcases x with ⟨⟨_, ⟨y, rfl⟩⟩⟩, - { simp only [equiv.coe_fn_mk, equiv.perm.coe_mul, function.comp_app], - change from_coset _ = from_coset _, - simp only [left_coset_assoc, subtype.coe_mk, subtype.mk_eq_mk, mul_assoc], }, - { refl }, - end } + { to_fun := λ x, β • x, + inv_fun := λ x, β⁻¹ • x, + left_inv := λ x, by dsimp only; rw [←mul_smul, mul_left_inv, one_smul], + right_inv := λ x, by dsimp only; rw [←mul_smul, mul_right_inv, one_smul] }, + map_one' := by ext; simp [one_smul], + map_mul' := λ b1 b2, by ext x; simp [mul_smul] } local notation `g` := G f @@ -262,14 +249,8 @@ Define `h : B ⟶ S(X')` to be `τ g τ⁻¹` -/ noncomputable def H : B ⟶ Group.of SX':= { to_fun := λ β, (equiv.trans (equiv.symm τ) (g β)).trans τ, - map_one' := begin - ext x, - simp - end, - map_mul' := λ b1 b2, begin - ext x, - simp - end } + map_one' := by ext; simp, + map_mul' := λ b1 b2, by ext; simp } local notation `h` := H f @@ -290,10 +271,10 @@ lemma g_apply_from_coset (x : B) (y : X) : end⟩ := rfl lemma g_apply_infinity (x : B) : - (g x).to_fun ⊙ = ⊙ := rfl + (g x).to_fun ∞ = ∞ := rfl lemma h_apply_infinity (x : B) (hx : x ∈ f.range) : - (h x).to_fun ⊙ = ⊙ := + (h x).to_fun ∞ = ∞ := begin dunfold H, simp only [monoid_hom.coe_mk, equiv.to_fun_as_coe, equiv.coe_trans, function.comp_app], @@ -342,7 +323,7 @@ begin simp only [monoid_hom.coe_mk, equiv.to_fun_as_coe, equiv.coe_trans, function.comp_app], rw [equiv.symm_swap], rw @equiv.swap_apply_of_ne_of_ne X' _ - (from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩) ⊙ + (from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩) ∞ (from_coset ⟨b *l f.range.carrier, ⟨b, rfl⟩⟩) begin intro r, simp only [subtype.mk_eq_mk] at r, From 697ddd95704349c2bc526cf12d0e33b5612ce0d1 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Fri, 17 Jun 2022 01:58:54 +0100 Subject: [PATCH 11/84] golf --- src/algebra/category/Group/epi_mono.lean | 70 ++++++++---------------- 1 file changed, 22 insertions(+), 48 deletions(-) diff --git a/src/algebra/category/Group/epi_mono.lean b/src/algebra/category/Group/epi_mono.lean index 2d5a2389d2079..b1419d96768a4 100644 --- a/src/algebra/category/Group/epi_mono.lean +++ b/src/algebra/category/Group/epi_mono.lean @@ -233,7 +233,7 @@ end Let `g : B ⟶ S(X')` be defined as as such that, for any `β : B`, `g(β)` is the function sending point at infinity to point at infinity and sending coset `y` to `β *l y`. -/ -def G : B ⟶ Group.of SX' := +def G : B →* SX' := { to_fun := λ β, { to_fun := λ x, β • x, inv_fun := λ x, β⁻¹ • x, @@ -247,7 +247,7 @@ local notation `g` := G f /-- Define `h : B ⟶ S(X')` to be `τ g τ⁻¹` -/ -noncomputable def H : B ⟶ Group.of SX':= +noncomputable def H : B →* SX':= { to_fun := λ β, (equiv.trans (equiv.symm τ) (g β)).trans τ, map_one' := by ext; simp, map_mul' := λ b1 b2, by ext; simp } @@ -262,7 +262,7 @@ The strategy is the following: assuming `epi f` -/ lemma g_apply_from_coset (x : B) (y : X) : - (g x).to_fun (from_coset y) = + (g x) (from_coset y) = from_coset ⟨x *l y, begin rcases y.2 with ⟨b, hb⟩, rw [subtype.val_eq_coe] at hb, @@ -271,34 +271,28 @@ lemma g_apply_from_coset (x : B) (y : X) : end⟩ := rfl lemma g_apply_infinity (x : B) : - (g x).to_fun ∞ = ∞ := rfl + (g x) ∞ = ∞ := rfl lemma h_apply_infinity (x : B) (hx : x ∈ f.range) : - (h x).to_fun ∞ = ∞ := + (h x) ∞ = ∞ := begin dunfold H, simp only [monoid_hom.coe_mk, equiv.to_fun_as_coe, equiv.coe_trans, function.comp_app], - rw [τ_symm_apply_infinity], - have := g_apply_from_coset f x ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩, - rw [equiv.to_fun_as_coe] at this, - rw this, + rw [τ_symm_apply_infinity, g_apply_from_coset], simpa only [←subtype.val_eq_coe] using τ_apply_from_coset' f x hx, end lemma h_apply_from_coset (x : B) : - (h x).to_fun (from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩) = + (h x) (from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩) = from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩ := begin dunfold H, simp only [monoid_hom.coe_mk, equiv.to_fun_as_coe, equiv.coe_trans, function.comp_app], - rw [τ_symm_apply_from_coset], - have := g_apply_infinity f x, - rw [equiv.to_fun_as_coe] at this, - rw [this, τ_apply_infinity], + rw [τ_symm_apply_from_coset, g_apply_infinity, τ_apply_infinity], end lemma h_apply_from_coset' (x : B) (b : B) (hb : b ∈ f.range): - (h x).to_fun (from_coset ⟨b *l f.range.carrier, ⟨b, rfl⟩⟩) = + (h x) (from_coset ⟨b *l f.range.carrier, ⟨b, rfl⟩⟩) = from_coset ⟨b *l f.range.carrier, ⟨b, rfl⟩⟩ := begin have eq1 : b *l (monoid_hom.range f).carrier = (monoid_hom.range f).carrier, @@ -316,7 +310,7 @@ begin end lemma h_apply_from_coset_nin_range (x : B) (hx : x ∈ f.range) (b : B) (hb : b ∉ f.range) : - (h x).to_fun (from_coset ⟨b *l f.range.carrier, ⟨b, rfl⟩⟩) = + (h x) (from_coset ⟨b *l f.range.carrier, ⟨b, rfl⟩⟩) = from_coset ⟨(x * b) *l f.range.carrier, ⟨x * b, rfl⟩⟩ := begin dunfold H tau, @@ -338,11 +332,7 @@ begin end begin simp only [ne.def, not_false_iff], end, - have := g_apply_from_coset f x, - rw [equiv.to_fun_as_coe] at this, - rw this, - clear this, - simp only [←subtype.val_eq_coe, left_coset_assoc], + simp only [g_apply_from_coset, ←subtype.val_eq_coe, left_coset_assoc], refine equiv.swap_apply_of_ne_of_ne begin intro r, simp only [subtype.mk_eq_mk] at r, @@ -371,40 +361,27 @@ begin change h (f a) = g (f a), ext, rcases x with ⟨⟨_, ⟨y, rfl⟩⟩⟩, - { have := g_apply_from_coset f (f a), - rw [equiv.to_fun_as_coe] at this, - rw this, - clear this, + { rw [g_apply_from_coset], by_cases m : y ∈ f.range, - { have := h_apply_from_coset' f (f a) y m, - rw [equiv.to_fun_as_coe] at this, - simp only [this, ←subtype.val_eq_coe], + { rw [h_apply_from_coset' _ _ _ m], + simp only [←subtype.val_eq_coe], congr' 1, change y *l f.range = f a *l (y *l f.range), rw [left_coset_assoc, left_coset_eq_iff], refine subgroup.mul_mem _ (subgroup.inv_mem _ m) (subgroup.mul_mem _ ⟨a, rfl⟩ m), }, - { have := h_apply_from_coset_nin_range f (f a) ⟨a, rfl⟩ y m, - rw [equiv.to_fun_as_coe] at this, - rw this, - clear this, + { rw [h_apply_from_coset_nin_range _ _ ⟨_, rfl⟩ _ m], simp only [←subtype.val_eq_coe, left_coset_assoc], refl, }, }, - { have := g_apply_infinity f (f a), - rw [equiv.to_fun_as_coe] at this, - rw this, - clear this, - have := h_apply_infinity f (f a) ⟨a, rfl⟩, - simpa only [equiv.to_fun_as_coe] using this, }, }, + { rw [g_apply_infinity, h_apply_infinity], + exact ⟨_, rfl⟩, }, }, { rintros (hb : h b = g b), have eq1 : (h b).to_fun (from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩) = (from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩), { dunfold H tau, simp only [equiv.symm_swap, monoid_hom.coe_mk, equiv.to_fun_as_coe, equiv.coe_trans, function.comp_app, equiv.swap_apply_left], - have := g_apply_infinity f b, - rw [equiv.to_fun_as_coe] at this, - rw [this, equiv.swap_apply_right], }, + rw [g_apply_infinity, equiv.swap_apply_right], }, have eq2 : (g b).to_fun (from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩) = (from_coset ⟨b *l f.range.carrier, ⟨b, rfl⟩⟩), { unfold G, @@ -421,7 +398,7 @@ begin left_coset_assoc, left_coset_eq_iff, inv_one, one_mul, mul_one] at eq3 } end -lemma comp_eq : f ≫ g = f ≫ h := +lemma comp_eq : f ≫ (show B ⟶ Group.of SX', from g) = f ≫ h := begin ext1 a, rw [comp_apply, comp_apply], @@ -438,7 +415,7 @@ lemma g_ne_h (x : B) (hx : x ∉ f.range) : g ≠ h := begin intros r, - replace r : ∀ a, (g x).to_fun a = (h x).to_fun a, + replace r : ∀ a, (g x) a = (h x) a, { intros a, simp only [equiv.to_fun_as_coe, r], }, specialize r (from_coset ⟨f.range, ⟨1, one_left_coset _⟩⟩), @@ -447,11 +424,8 @@ begin equiv.to_fun_as_coe, equiv.coe_trans, function.comp_app] at r, generalize_proofs h1 h2 at r, rw [show from_coset ⟨set.range f, h2⟩ = from_coset ⟨f.range.carrier, h2⟩, - by simpa only [subtype.mk_eq_mk], equiv.swap_apply_left] at r, - have := g_apply_infinity f x, - rw [equiv.to_fun_as_coe] at this, - rw [this, equiv.swap_apply_right] at r, - clear this, + by simpa only [subtype.mk_eq_mk], equiv.swap_apply_left, g_apply_infinity, + equiv.swap_apply_right] at r, simp only [subtype.mk_eq_mk] at r, change x *l f.range = f.range at r, rw [show (f.range : set B) = 1 *l f.range, from (one_left_coset _).symm, From c605b39085425abe479687a156bf306882654c42 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Fri, 17 Jun 2022 02:12:54 +0100 Subject: [PATCH 12/84] golf --- src/algebra/category/Group/epi_mono.lean | 75 +++++++++++------------- 1 file changed, 35 insertions(+), 40 deletions(-) diff --git a/src/algebra/category/Group/epi_mono.lean b/src/algebra/category/Group/epi_mono.lean index b1419d96768a4..2b15db948713f 100644 --- a/src/algebra/category/Group/epi_mono.lean +++ b/src/algebra/category/Group/epi_mono.lean @@ -174,6 +174,29 @@ end | ∞ := rfl end +lemma from_coset_eq_of_mem_range {b : B} (hb : b ∈ f.range) : + from_coset ⟨b *l f.range.carrier, ⟨b, rfl⟩⟩ = + from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩ := +begin + congr, + change b *l f.range = f.range, + nth_rewrite 1 [show (f.range : set B) = 1 *l f.range, from (one_left_coset _).symm], + rw [left_coset_eq_iff, mul_one], + exact subgroup.inv_mem _ hb, +end + +lemma from_coset_ne_of_nin_range {b : B} (hb : b ∉ f.range) : + from_coset ⟨b *l f.range.carrier, ⟨b, rfl⟩⟩ ≠ + from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩ := +begin + intros r, + simp only [subtype.mk_eq_mk] at r, + change b *l f.range = f.range at r, + nth_rewrite 1 [show (f.range : set B) = 1 *l f.range, from (one_left_coset _).symm] at r, + rw [left_coset_eq_iff, mul_one] at r, + exact hb (inv_inv b ▸ (subgroup.inv_mem _ r)), +end + noncomputable instance : decidable_eq X' := classical.dec_eq _ @@ -318,38 +341,14 @@ begin rw [equiv.symm_swap], rw @equiv.swap_apply_of_ne_of_ne X' _ (from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩) ∞ - (from_coset ⟨b *l f.range.carrier, ⟨b, rfl⟩⟩) begin - intro r, - simp only [subtype.mk_eq_mk] at r, - change b *l f.range = f.range at r, - have eq1 : (f.range : set B) = 1 *l f.range, - { rw one_left_coset, }, - conv_rhs at r { rw eq1 }, - rw [left_coset_eq_iff, mul_one] at r, - apply hb, - replace r : b⁻¹⁻¹ ∈ f.range := subgroup.inv_mem _ r, - rwa [inv_inv] at r, - end begin - simp only [ne.def, not_false_iff], - end, + (from_coset ⟨b *l f.range.carrier, ⟨b, rfl⟩⟩) + (from_coset_ne_of_nin_range _ hb) (by simp), simp only [g_apply_from_coset, ←subtype.val_eq_coe, left_coset_assoc], refine equiv.swap_apply_of_ne_of_ne begin - intro r, - simp only [subtype.mk_eq_mk] at r, - change x * b *l f.range = f.range at r, - have eq1 : (f.range : set B) = 1 *l f.range, - { rw one_left_coset, }, - conv_rhs at r { rw eq1 }, - rw [left_coset_eq_iff, mul_one] at r, - replace r : (x * b)⁻¹⁻¹ ∈ f.range := subgroup.inv_mem _ r, - rw inv_inv at r, - apply hb, - replace hx : x⁻¹ ∈ f.range := subgroup.inv_mem _ hx, - have := subgroup.mul_mem _ hx r, - rwa [←mul_assoc, mul_left_inv, one_mul] at this, - end begin - simp only [ne.def, not_false_iff], - end, + refine from_coset_ne_of_nin_range _ (λ r, hb _), + convert subgroup.mul_mem _ (subgroup.inv_mem _ hx) r, + rw [←mul_assoc, mul_left_inv, one_mul], + end (by simp), end lemma agree : @@ -373,23 +372,19 @@ begin { rw [h_apply_from_coset_nin_range _ _ ⟨_, rfl⟩ _ m], simp only [←subtype.val_eq_coe, left_coset_assoc], refl, }, }, - { rw [g_apply_infinity, h_apply_infinity], - exact ⟨_, rfl⟩, }, }, + { rw [g_apply_infinity, h_apply_infinity _ _ ⟨_, rfl⟩], }, }, { rintros (hb : h b = g b), - have eq1 : (h b).to_fun (from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩) = + have eq1 : (h b) (from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩) = (from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩), { dunfold H tau, simp only [equiv.symm_swap, monoid_hom.coe_mk, equiv.to_fun_as_coe, equiv.coe_trans, function.comp_app, equiv.swap_apply_left], rw [g_apply_infinity, equiv.swap_apply_right], }, - have eq2 : (g b).to_fun (from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩) = + have eq2 : (g b) (from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩) = (from_coset ⟨b *l f.range.carrier, ⟨b, rfl⟩⟩), - { unfold G, - simp only [monoid_hom.coe_mk], - change from_coset _ = _, - refl, }, - have eq3 : (h b).to_fun (from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩) = - (g b).to_fun (from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩), + { refl }, + have eq3 : (h b) (from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩) = + (g b) (from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩), { simp only [equiv.to_fun_as_coe, hb], }, rw [eq1, eq2] at eq3, simp only [subtype.mk_eq_mk] at eq3, From e1d1cf845474d56c80992764dbf6c6b3bce16c7f Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Fri, 17 Jun 2022 02:55:07 +0100 Subject: [PATCH 13/84] golf --- src/algebra/category/Group/epi_mono.lean | 130 ++++++----------------- 1 file changed, 30 insertions(+), 100 deletions(-) diff --git a/src/algebra/category/Group/epi_mono.lean b/src/algebra/category/Group/epi_mono.lean index 2b15db948713f..ff7fe93a03f18 100644 --- a/src/algebra/category/Group/epi_mono.lean +++ b/src/algebra/category/Group/epi_mono.lean @@ -210,47 +210,26 @@ local notation `τ` := tau f lemma τ_apply_infinity : τ ∞ = from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩ := -begin - dunfold tau, - rw [equiv.swap_apply_right], -end +equiv.swap_apply_right _ _ lemma τ_apply_from_coset : τ (from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩) = ∞ := -begin - dunfold tau, - rw [equiv.swap_apply_left], -end +equiv.swap_apply_left _ _ lemma τ_apply_from_coset' (x : B) (hx : x ∈ f.range) : τ (from_coset ⟨x *l f.range.carrier, ⟨x, rfl⟩⟩) = ∞ := begin - convert τ_apply_from_coset _, - ext b, - simp only [mem_left_coset_iff, subgroup.mem_carrier, monoid_hom.mem_range], - rcases hx with ⟨c, rfl⟩, - split, - { rintros ⟨a, ha⟩, - use c * a, - rw [map_mul, ha, ←mul_assoc, mul_right_inv (f c), one_mul], }, - { rintros ⟨a, rfl⟩, - use c⁻¹ * a, - rw [map_mul, map_inv], }, + rw from_coset_eq_of_mem_range _ hx, + exact τ_apply_from_coset _, end lemma τ_symm_apply_from_coset : (equiv.symm τ) (from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩) = ∞ := -begin - dunfold tau, - rw [equiv.symm_swap, equiv.swap_apply_left], -end +by rw [tau, equiv.symm_swap, equiv.swap_apply_left] lemma τ_symm_apply_infinity : (equiv.symm τ) ∞ = from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩ := -begin - dunfold tau, - rw [equiv.symm_swap, equiv.swap_apply_right], -end +by rw [tau, equiv.symm_swap, equiv.swap_apply_right] /-- Let `g : B ⟶ S(X')` be defined as as such that, for any `β : B`, `g(β)` is the function sending @@ -263,7 +242,7 @@ def G : B →* SX' := left_inv := λ x, by dsimp only; rw [←mul_smul, mul_left_inv, one_smul], right_inv := λ x, by dsimp only; rw [←mul_smul, mul_right_inv, one_smul] }, map_one' := by ext; simp [one_smul], - map_mul' := λ b1 b2, by ext x; simp [mul_smul] } + map_mul' := λ b1 b2, by ext; simp [mul_smul] } local notation `g` := G f @@ -271,7 +250,7 @@ local notation `g` := G f Define `h : B ⟶ S(X')` to be `τ g τ⁻¹` -/ noncomputable def H : B →* SX':= -{ to_fun := λ β, (equiv.trans (equiv.symm τ) (g β)).trans τ, +{ to_fun := λ β, ((τ).symm.trans (g β)).trans τ, map_one' := by ext; simp, map_mul' := λ b1 b2, by ext; simp } @@ -288,8 +267,7 @@ lemma g_apply_from_coset (x : B) (y : X) : (g x) (from_coset y) = from_coset ⟨x *l y, begin rcases y.2 with ⟨b, hb⟩, - rw [subtype.val_eq_coe] at hb, - rw [←hb, set.mem_range, left_coset_assoc], + rw [←subtype.val_eq_coe, ←hb, set.mem_range, left_coset_assoc], use x * b, end⟩ := rfl @@ -299,8 +277,7 @@ lemma g_apply_infinity (x : B) : lemma h_apply_infinity (x : B) (hx : x ∈ f.range) : (h x) ∞ = ∞ := begin - dunfold H, - simp only [monoid_hom.coe_mk, equiv.to_fun_as_coe, equiv.coe_trans, function.comp_app], + simp only [H, monoid_hom.coe_mk, equiv.to_fun_as_coe, equiv.coe_trans, function.comp_app], rw [τ_symm_apply_infinity, g_apply_from_coset], simpa only [←subtype.val_eq_coe] using τ_apply_from_coset' f x hx, end @@ -308,43 +285,25 @@ end lemma h_apply_from_coset (x : B) : (h x) (from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩) = from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩ := -begin - dunfold H, - simp only [monoid_hom.coe_mk, equiv.to_fun_as_coe, equiv.coe_trans, function.comp_app], - rw [τ_symm_apply_from_coset, g_apply_infinity, τ_apply_infinity], -end +by simp [H, τ_symm_apply_from_coset, g_apply_infinity, τ_apply_infinity] lemma h_apply_from_coset' (x : B) (b : B) (hb : b ∈ f.range): (h x) (from_coset ⟨b *l f.range.carrier, ⟨b, rfl⟩⟩) = from_coset ⟨b *l f.range.carrier, ⟨b, rfl⟩⟩ := -begin - have eq1 : b *l (monoid_hom.range f).carrier = (monoid_hom.range f).carrier, - { rcases hb with ⟨a, rfl⟩, - ext y, - simp only [mem_left_coset_iff, subgroup.mem_carrier, monoid_hom.mem_range], - split, - { rintros ⟨a', ha'⟩, - use a * a', - rw [map_mul, ha', ←mul_assoc, mul_right_inv, one_mul], }, - { rintros ⟨a', rfl⟩, - use a⁻¹ * a', - rw [map_mul, map_inv], }, }, - convert h_apply_from_coset f x, -end +(from_coset_eq_of_mem_range _ hb).symm ▸ h_apply_from_coset f x lemma h_apply_from_coset_nin_range (x : B) (hx : x ∈ f.range) (b : B) (hb : b ∉ f.range) : (h x) (from_coset ⟨b *l f.range.carrier, ⟨b, rfl⟩⟩) = from_coset ⟨(x * b) *l f.range.carrier, ⟨x * b, rfl⟩⟩ := begin - dunfold H tau, - simp only [monoid_hom.coe_mk, equiv.to_fun_as_coe, equiv.coe_trans, function.comp_app], + simp only [H, tau, monoid_hom.coe_mk, equiv.to_fun_as_coe, equiv.coe_trans, function.comp_app], rw [equiv.symm_swap], rw @equiv.swap_apply_of_ne_of_ne X' _ (from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩) ∞ (from_coset ⟨b *l f.range.carrier, ⟨b, rfl⟩⟩) (from_coset_ne_of_nin_range _ hb) (by simp), simp only [g_apply_from_coset, ←subtype.val_eq_coe, left_coset_assoc], - refine equiv.swap_apply_of_ne_of_ne begin + exact equiv.swap_apply_of_ne_of_ne begin refine from_coset_ne_of_nin_range _ (λ r, hb _), convert subgroup.mul_mem _ (subgroup.inv_mem _ hx) r, rw [←mul_assoc, mul_left_inv, one_mul], @@ -358,48 +317,31 @@ begin split, { rintros ⟨a, rfl⟩, change h (f a) = g (f a), - ext, - rcases x with ⟨⟨_, ⟨y, rfl⟩⟩⟩, + ext ⟨⟨_, ⟨y, rfl⟩⟩⟩, { rw [g_apply_from_coset], by_cases m : y ∈ f.range, - { rw [h_apply_from_coset' _ _ _ m], - simp only [←subtype.val_eq_coe], - congr' 1, - change y *l f.range = f a *l (y *l f.range), - rw [left_coset_assoc, left_coset_eq_iff], - refine subgroup.mul_mem _ (subgroup.inv_mem _ m) - (subgroup.mul_mem _ ⟨a, rfl⟩ m), }, + { rw [h_apply_from_coset' _ _ _ m, from_coset_eq_of_mem_range _ m], + change from_coset _ = from_coset ⟨f a *l (y *l _), _⟩, + simpa only [←from_coset_eq_of_mem_range _ (subgroup.mul_mem _ ⟨a, rfl⟩ m), + left_coset_assoc] }, { rw [h_apply_from_coset_nin_range _ _ ⟨_, rfl⟩ _ m], - simp only [←subtype.val_eq_coe, left_coset_assoc], - refl, }, }, + simpa only [←subtype.val_eq_coe, left_coset_assoc], }, }, { rw [g_apply_infinity, h_apply_infinity _ _ ⟨_, rfl⟩], }, }, { rintros (hb : h b = g b), have eq1 : (h b) (from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩) = - (from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩), - { dunfold H tau, - simp only [equiv.symm_swap, monoid_hom.coe_mk, equiv.to_fun_as_coe, equiv.coe_trans, - function.comp_app, equiv.swap_apply_left], - rw [g_apply_infinity, equiv.swap_apply_right], }, + (from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩) := by simp [H, tau, g_apply_infinity], have eq2 : (g b) (from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩) = - (from_coset ⟨b *l f.range.carrier, ⟨b, rfl⟩⟩), - { refl }, - have eq3 : (h b) (from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩) = - (g b) (from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩), - { simp only [equiv.to_fun_as_coe, hb], }, - rw [eq1, eq2] at eq3, - simp only [subtype.mk_eq_mk] at eq3, - change (f.range : set B) = b *l f.range at eq3, - rwa [show (f.range : set B) = 1 *l f.range, by rw one_left_coset _, - left_coset_assoc, left_coset_eq_iff, inv_one, one_mul, mul_one] at eq3 } + (from_coset ⟨b *l f.range.carrier, ⟨b, rfl⟩⟩) := rfl, + by_contra r, + exact (from_coset_ne_of_nin_range _ r).symm (by rw [←eq1, ←eq2, fun_like.congr_fun hb]), } end lemma comp_eq : f ≫ (show B ⟶ Group.of SX', from g) = f ≫ h := begin ext1 a, - rw [comp_apply, comp_apply], + simp only [comp_apply, comp_apply], have : f a ∈ f.range.carrier := ⟨a, rfl⟩, - rw agree at this, - simp only [set.mem_set_of_eq] at this, + rw [agree, set.mem_set_of_eq] at this, rw this, end @@ -410,24 +352,13 @@ lemma g_ne_h (x : B) (hx : x ∉ f.range) : g ≠ h := begin intros r, - replace r : ∀ a, (g x) a = (h x) a, - { intros a, - simp only [equiv.to_fun_as_coe, r], }, - specialize r (from_coset ⟨f.range, ⟨1, one_left_coset _⟩⟩), + replace r := fun_like.congr_fun (fun_like.congr_fun r x) + ((from_coset ⟨f.range, ⟨1, one_left_coset _⟩⟩)), rw [H, g_apply_from_coset, monoid_hom.coe_mk, tau] at r, simp only [monoid_hom.coe_range, subtype.coe_mk, equiv.symm_swap, equiv.to_fun_as_coe, equiv.coe_trans, function.comp_app] at r, - generalize_proofs h1 h2 at r, - rw [show from_coset ⟨set.range f, h2⟩ = from_coset ⟨f.range.carrier, h2⟩, - by simpa only [subtype.mk_eq_mk], equiv.swap_apply_left, g_apply_infinity, - equiv.swap_apply_right] at r, - simp only [subtype.mk_eq_mk] at r, - change x *l f.range = f.range at r, - rw [show (f.range : set B) = 1 *l f.range, from (one_left_coset _).symm, - left_coset_assoc, mul_one, left_coset_eq_iff, mul_one] at r, - replace r := subgroup.inv_mem _ r, - rw inv_inv at r, - exact hx r, + erw [equiv.swap_apply_left, g_apply_infinity, equiv.swap_apply_right] at r, + exact from_coset_ne_of_nin_range _ hx r, end end surjective_of_epi_auxs @@ -453,5 +384,4 @@ iff.trans (epi_iff_surjective _) (subgroup.eq_top_iff' f.range).symm end Group - end From bf97e8bb1881a99c385dc4039f19f1ccf54c2b3e Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Fri, 17 Jun 2022 13:46:25 +0100 Subject: [PATCH 14/84] golf a little bit more --- src/algebra/category/Group/epi_mono.lean | 106 +++++++++-------------- 1 file changed, 41 insertions(+), 65 deletions(-) diff --git a/src/algebra/category/Group/epi_mono.lean b/src/algebra/category/Group/epi_mono.lean index ff7fe93a03f18..c23b752f1a51c 100644 --- a/src/algebra/category/Group/epi_mono.lean +++ b/src/algebra/category/Group/epi_mono.lean @@ -30,8 +30,7 @@ variables [group A] [group B] lemma ker_eq_bot (f : A →* B) : f.ker = ⊥ ↔ function.injective f := { mp := λ h1 x y h2, begin - rw [←div_eq_one, ←map_div, ←mem_ker] at h2, - rwa [h1, subgroup.mem_bot, div_eq_one] at h2, + rwa [←div_eq_one, ←map_div, ←mem_ker, h1, subgroup.mem_bot, div_eq_one] at h2, end, mpr := λ h, set_like.ext $ λ x, iff.trans (mem_ker _) $ iff.trans begin @@ -42,19 +41,15 @@ lemma ker_eq_bot (f : A →* B) : @[to_additive add_monoid_hom.range_eq_top] lemma range_eq_top (f : A →* B) : f.range = ⊤ ↔ function.surjective f := -{ mp := λ h x, begin - rcases show x ∈ f.range, from h.symm ▸ subgroup.mem_top _ with ⟨a, h⟩, - exact ⟨a, h⟩, - end, - mpr := λ h, set_like.ext $ λ x, iff.trans mem_range - ⟨λ _, trivial, λ _, h x⟩ } +{ mp := λ h x, show x ∈ f.range, from h.symm ▸ subgroup.mem_top _, + mpr := λ h, set_like.ext $ λ x, iff.trans mem_range ⟨λ _, trivial, λ _, h x⟩ } @[to_additive add_monoid_hom.range_zero_eq_bot] lemma range_one_eq_bot : (1 : A →* B).range = ⊥ := set_like.ext $ λ a, iff.trans monoid_hom.mem_range $ - iff.trans (by simp only [one_apply, exists_const]; split; intros h; subst h) - subgroup.mem_bot.symm + iff.trans (by simp only [one_apply, exists_const]; split; intros h; subst h) subgroup.mem_bot.symm + @[to_additive add_monoid_hom.ker_zero_eq_top] lemma ker_one_eq_top : (1 : A →* B).ker = ⊤ := @@ -147,9 +142,8 @@ local notation `SX'` := equiv.perm X' instance : has_scalar B X' := { smul := λ b x, match x with | from_coset y := from_coset ⟨b *l y, begin - rcases y.2 with ⟨b', hb'⟩, - use b * b', - rw [←subtype.val_eq_coe, ←hb', left_coset_assoc], + rw [←subtype.val_eq_coe, ←y.2.some_spec, left_coset_assoc], + use b * y.2.some, end⟩ | ∞ := ∞ end } @@ -218,10 +212,7 @@ equiv.swap_apply_left _ _ lemma τ_apply_from_coset' (x : B) (hx : x ∈ f.range) : τ (from_coset ⟨x *l f.range.carrier, ⟨x, rfl⟩⟩) = ∞ := -begin - rw from_coset_eq_of_mem_range _ hx, - exact τ_apply_from_coset _, -end +(from_coset_eq_of_mem_range _ hx).symm ▸ τ_apply_from_coset _ lemma τ_symm_apply_from_coset : (equiv.symm τ) (from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩) = ∞ := @@ -260,15 +251,14 @@ local notation `h` := H f The strategy is the following: assuming `epi f` * prove that `f.range = {x | h x = g x}`; * thus `f ≫ h = f ≫ g` so that `h = g`; -* but if `f` is not surjective, then some `x ≠ f.range`, then `h x ≠ g x` at the coset `f.range`. +* but if `f` is not surjective, then some `x ∉ f.range`, then `h x ≠ g x` at the coset `f.range`. -/ lemma g_apply_from_coset (x : B) (y : X) : (g x) (from_coset y) = from_coset ⟨x *l y, begin - rcases y.2 with ⟨b, hb⟩, - rw [←subtype.val_eq_coe, ←hb, set.mem_range, left_coset_assoc], - use x * b, + rw [←subtype.val_eq_coe, ←y.2.some_spec, set.mem_range, left_coset_assoc], + use x * y.2.some, end⟩ := rfl lemma g_apply_infinity (x : B) : @@ -300,8 +290,7 @@ begin rw [equiv.symm_swap], rw @equiv.swap_apply_of_ne_of_ne X' _ (from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩) ∞ - (from_coset ⟨b *l f.range.carrier, ⟨b, rfl⟩⟩) - (from_coset_ne_of_nin_range _ hb) (by simp), + (from_coset ⟨b *l f.range.carrier, ⟨b, rfl⟩⟩) (from_coset_ne_of_nin_range _ hb) (by simp), simp only [g_apply_from_coset, ←subtype.val_eq_coe, left_coset_assoc], exact equiv.swap_apply_of_ne_of_ne begin refine from_coset_ne_of_nin_range _ (λ r, hb _), @@ -312,46 +301,35 @@ end lemma agree : f.range.carrier = {x | h x = g x} := -begin - ext b, - split, - { rintros ⟨a, rfl⟩, - change h (f a) = g (f a), - ext ⟨⟨_, ⟨y, rfl⟩⟩⟩, - { rw [g_apply_from_coset], - by_cases m : y ∈ f.range, - { rw [h_apply_from_coset' _ _ _ m, from_coset_eq_of_mem_range _ m], - change from_coset _ = from_coset ⟨f a *l (y *l _), _⟩, - simpa only [←from_coset_eq_of_mem_range _ (subgroup.mul_mem _ ⟨a, rfl⟩ m), - left_coset_assoc] }, - { rw [h_apply_from_coset_nin_range _ _ ⟨_, rfl⟩ _ m], - simpa only [←subtype.val_eq_coe, left_coset_assoc], }, }, - { rw [g_apply_infinity, h_apply_infinity _ _ ⟨_, rfl⟩], }, }, - { rintros (hb : h b = g b), - have eq1 : (h b) (from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩) = - (from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩) := by simp [H, tau, g_apply_infinity], - have eq2 : (g b) (from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩) = - (from_coset ⟨b *l f.range.carrier, ⟨b, rfl⟩⟩) := rfl, - by_contra r, - exact (from_coset_ne_of_nin_range _ r).symm (by rw [←eq1, ←eq2, fun_like.congr_fun hb]), } -end +set.ext $ λ b, +⟨begin + rintros ⟨a, rfl⟩, + change h (f a) = g (f a), + ext ⟨⟨_, ⟨y, rfl⟩⟩⟩, + { rw [g_apply_from_coset], + by_cases m : y ∈ f.range, + { rw [h_apply_from_coset' _ _ _ m, from_coset_eq_of_mem_range _ m], + change from_coset _ = from_coset ⟨f a *l (y *l _), _⟩, + simpa only [←from_coset_eq_of_mem_range _ (subgroup.mul_mem _ ⟨a, rfl⟩ m), + left_coset_assoc] }, + { rw [h_apply_from_coset_nin_range _ _ ⟨_, rfl⟩ _ m], + simpa only [←subtype.val_eq_coe, left_coset_assoc], }, }, + { rw [g_apply_infinity, h_apply_infinity _ _ ⟨_, rfl⟩], }, +end, λ (hb : h b = g b), classical.by_contradiction $ λ r, begin + have eq1 : (h b) (from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩) = + (from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩) := by simp [H, tau, g_apply_infinity], + have eq2 : (g b) (from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩) = + (from_coset ⟨b *l f.range.carrier, ⟨b, rfl⟩⟩) := rfl, + exact (from_coset_ne_of_nin_range _ r).symm (by rw [←eq1, ←eq2, fun_like.congr_fun hb]), +end⟩ lemma comp_eq : f ≫ (show B ⟶ Group.of SX', from g) = f ≫ h := -begin - ext1 a, - simp only [comp_apply, comp_apply], - have : f a ∈ f.range.carrier := ⟨a, rfl⟩, - rw [agree, set.mem_set_of_eq] at this, - rw this, -end - -lemma g_eq_h [epi f] : g = h := -(cancel_epi f).1 $ comp_eq f +fun_like.ext _ _ $ λ a, + by simp only [comp_apply, show h (f a) = _, from (by simp [←agree] : f a ∈ {b | h b = g b})] lemma g_ne_h (x : B) (hx : x ∉ f.range) : g ≠ h := -begin - intros r, +λ r, begin replace r := fun_like.congr_fun (fun_like.congr_fun r x) ((from_coset ⟨f.range, ⟨1, one_left_coset _⟩⟩)), rw [H, g_apply_from_coset, monoid_hom.coe_mk, tau] at r, @@ -364,14 +342,12 @@ end end surjective_of_epi_auxs lemma surjective_of_epi [epi f] : function.surjective f := -begin - by_contra r, - simp_rw [not_forall, not_exists] at r, +classical.by_contradiction $ λ r, begin + push_neg at r, rcases r with ⟨b, hb⟩, - refine surjective_of_epi_auxs.g_ne_h f b (λ r, _) _, - { rcases r with ⟨c, hc⟩, - exact hb _ hc, }, - apply surjective_of_epi_auxs.g_eq_h, + refine surjective_of_epi_auxs.g_ne_h f b _ ((cancel_epi f).1 (surjective_of_epi_auxs.comp_eq f)), + rintros ⟨c, hc⟩, + exact hb _ hc, end lemma epi_iff_surjective : From 4b3cc16e6df818d6b9a26a2ba41e8cd53c80d915 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Sun, 17 Jul 2022 16:26:57 +0100 Subject: [PATCH 15/84] apply suggestion and fix --- src/algebra/category/Group/epi_mono.lean | 39 +++--------------------- src/group_theory/subgroup/basic.lean | 5 +++ 2 files changed, 10 insertions(+), 34 deletions(-) diff --git a/src/algebra/category/Group/epi_mono.lean b/src/algebra/category/Group/epi_mono.lean index c23b752f1a51c..1f1c25706caa7 100644 --- a/src/algebra/category/Group/epi_mono.lean +++ b/src/algebra/category/Group/epi_mono.lean @@ -26,35 +26,6 @@ section variables [group A] [group B] -@[to_additive add_monoid_hom.ker_eq_bot] -lemma ker_eq_bot (f : A →* B) : - f.ker = ⊥ ↔ function.injective f := -{ mp := λ h1 x y h2, begin - rwa [←div_eq_one, ←map_div, ←mem_ker, h1, subgroup.mem_bot, div_eq_one] at h2, - end, - mpr := λ h, set_like.ext $ λ x, iff.trans (mem_ker _) $ - iff.trans begin - rw ←(by rw map_one : f x = f 1 ↔ f x = 1), - exact ⟨λ h', h h', λ h, h ▸ rfl⟩, - end subgroup.mem_bot.symm } - -@[to_additive add_monoid_hom.range_eq_top] -lemma range_eq_top (f : A →* B) : - f.range = ⊤ ↔ function.surjective f := -{ mp := λ h x, show x ∈ f.range, from h.symm ▸ subgroup.mem_top _, - mpr := λ h, set_like.ext $ λ x, iff.trans mem_range ⟨λ _, trivial, λ _, h x⟩ } - -@[to_additive add_monoid_hom.range_zero_eq_bot] -lemma range_one_eq_bot : - (1 : A →* B).range = ⊥ := -set_like.ext $ λ a, iff.trans monoid_hom.mem_range $ - iff.trans (by simp only [one_apply, exists_const]; split; intros h; subst h) subgroup.mem_bot.symm - -@[to_additive add_monoid_hom.ker_zero_eq_top] -lemma ker_one_eq_top : - (1 : A →* B).ker = ⊤ := -set_like.ext $ λ a, iff.trans (monoid_hom.mem_ker _) $ by simp - @[to_additive add_monoid_hom.ker_eq_bot_of_cancel] lemma ker_eq_bot_of_cancel {f : A →* B} (h : ∀ (u v : f.ker →* A), f.comp u = f.comp v → u = v) : @@ -65,7 +36,7 @@ begin simp only [comp_one, one_apply, coe_comp, subgroup.coe_subtype, function.comp_app], rw [←subtype.val_eq_coe, f.mem_ker.mp x.2], end, - rw [←subgroup.subtype_range f.ker, ←h, range_one_eq_bot], + rw [←subgroup.subtype_range f.ker, ←h, range_one], end end @@ -87,7 +58,7 @@ begin exact ⟨x, rfl⟩, end, replace h : (quotient_group.mk' _).ker = (1 : B →* B ⧸ f.range).ker := by rw h, - rwa [ker_one_eq_top, quotient_group.ker_mk] at h, + rwa [ker_one, quotient_group.ker_mk] at h, end end @@ -112,12 +83,12 @@ monoid_hom.ker_eq_bot_of_cancel $ λ u v, lemma mono_iff_ker_eq_bot : mono f ↔ f.ker = ⊥ := ⟨λ h, @@ker_eq_bot_of_mono f h, -λ h, concrete_category.mono_of_injective _ $ (monoid_hom.ker_eq_bot f).1 h⟩ +λ h, concrete_category.mono_of_injective _ $ (monoid_hom.ker_eq_bot_iff f).1 h⟩ @[to_additive AddGroup.mono_iff_injective] lemma mono_iff_injective : mono f ↔ function.injective f := -iff.trans (mono_iff_ker_eq_bot f) $ monoid_hom.ker_eq_bot f +iff.trans (mono_iff_ker_eq_bot f) $ monoid_hom.ker_eq_bot_iff f namespace surjective_of_epi_auxs @@ -139,7 +110,7 @@ local notation `X'` := X_with_infinity f local notation `∞` := X_with_infinity.infinity local notation `SX'` := equiv.perm X' -instance : has_scalar B X' := +instance : has_smul B X' := { smul := λ b x, match x with | from_coset y := from_coset ⟨b *l y, begin rw [←subtype.val_eq_coe, ←y.2.some_spec, left_coset_assoc], diff --git a/src/group_theory/subgroup/basic.lean b/src/group_theory/subgroup/basic.lean index 1163576cc90a1..6e4820d65f988 100644 --- a/src/group_theory/subgroup/basic.lean +++ b/src/group_theory/subgroup/basic.lean @@ -2024,6 +2024,11 @@ lemma range_top_of_surjective {N} [group N] (f : G →* N) (hf : function.surjec f.range = (⊤ : subgroup N) := range_top_iff_surjective.2 hf +@[simp, to_additive] +lemma range_one : (1 : G →* N).range = ⊥ := +set_like.ext $ λ a, iff.trans monoid_hom.mem_range $ + iff.trans (by simp only [one_apply, exists_const]; split; intros h; subst h) subgroup.mem_bot.symm + @[simp, to_additive] lemma _root_.subgroup.subtype_range (H : subgroup G) : H.subtype.range = H := by { rw [range_eq_map, ← set_like.coe_set_eq, coe_map, subgroup.coe_subtype], ext, simp } From f21bd67ebd94b5c760347853bee515c671ed040a Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Mon, 18 Jul 2022 12:34:44 +0100 Subject: [PATCH 16/84] Update src/algebra/category/Group/epi_mono.lean Co-authored-by: Markus Himmel --- src/algebra/category/Group/epi_mono.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/algebra/category/Group/epi_mono.lean b/src/algebra/category/Group/epi_mono.lean index 1f1c25706caa7..38c11c7da94a8 100644 --- a/src/algebra/category/Group/epi_mono.lean +++ b/src/algebra/category/Group/epi_mono.lean @@ -8,7 +8,7 @@ import group_theory.quotient_group import algebra.category.Group.basic /-! -# Monomorphisms and epimorphism in `Group` +# Monomorphisms and epimorphisms in `Group` In this file, we prove monomorphisms in category of group are injective homomorphisms and epimorphisms are surjective homomorphisms. -/ From 7b39fcd01f7ac7420447823b9c8d8468be3d7034 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Mon, 18 Jul 2022 12:34:58 +0100 Subject: [PATCH 17/84] Update src/algebra/category/Group/epi_mono.lean Co-authored-by: Markus Himmel --- src/algebra/category/Group/epi_mono.lean | 9 +++++---- 1 file changed, 5 insertions(+), 4 deletions(-) diff --git a/src/algebra/category/Group/epi_mono.lean b/src/algebra/category/Group/epi_mono.lean index 38c11c7da94a8..b7cde57312867 100644 --- a/src/algebra/category/Group/epi_mono.lean +++ b/src/algebra/category/Group/epi_mono.lean @@ -112,10 +112,11 @@ local notation `SX'` := equiv.perm X' instance : has_smul B X' := { smul := λ b x, match x with - | from_coset y := from_coset ⟨b *l y, begin - rw [←subtype.val_eq_coe, ←y.2.some_spec, left_coset_assoc], - use b * y.2.some, - end⟩ + | from_coset y := from_coset ⟨b *l y, + begin + rw [←subtype.val_eq_coe, ←y.2.some_spec, left_coset_assoc], + use b * y.2.some, + end⟩ | ∞ := ∞ end } From c039b7d92fafd460c13019e48f2824c9ec693838 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Mon, 18 Jul 2022 12:37:02 +0100 Subject: [PATCH 18/84] Update src/algebra/category/Group/epi_mono.lean Co-authored-by: Markus Himmel --- src/algebra/category/Group/epi_mono.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/algebra/category/Group/epi_mono.lean b/src/algebra/category/Group/epi_mono.lean index b7cde57312867..8de6754b5af7d 100644 --- a/src/algebra/category/Group/epi_mono.lean +++ b/src/algebra/category/Group/epi_mono.lean @@ -219,7 +219,7 @@ noncomputable def H : B →* SX':= local notation `h` := H f -/-- +/-! The strategy is the following: assuming `epi f` * prove that `f.range = {x | h x = g x}`; * thus `f ≫ h = f ≫ g` so that `h = g`; From ef6a749b18929df9ec53a08fb2f957dfdfcfb5bf Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Mon, 18 Jul 2022 12:37:11 +0100 Subject: [PATCH 19/84] Update src/algebra/category/Group/epi_mono.lean Co-authored-by: Markus Himmel --- src/algebra/category/Group/epi_mono.lean | 3 +-- 1 file changed, 1 insertion(+), 2 deletions(-) diff --git a/src/algebra/category/Group/epi_mono.lean b/src/algebra/category/Group/epi_mono.lean index 8de6754b5af7d..46d38b5c0f191 100644 --- a/src/algebra/category/Group/epi_mono.lean +++ b/src/algebra/category/Group/epi_mono.lean @@ -233,8 +233,7 @@ lemma g_apply_from_coset (x : B) (y : X) : use x * y.2.some, end⟩ := rfl -lemma g_apply_infinity (x : B) : - (g x) ∞ = ∞ := rfl +lemma g_apply_infinity (x : B) : (g x) ∞ = ∞ := rfl lemma h_apply_infinity (x : B) (hx : x ∈ f.range) : (h x) ∞ = ∞ := From 677d54254040dc752412e837c1a12430c5d3679c Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Mon, 18 Jul 2022 12:37:21 +0100 Subject: [PATCH 20/84] Update src/algebra/category/Group/epi_mono.lean Co-authored-by: Markus Himmel --- src/algebra/category/Group/epi_mono.lean | 44 +++++++++++------------- 1 file changed, 21 insertions(+), 23 deletions(-) diff --git a/src/algebra/category/Group/epi_mono.lean b/src/algebra/category/Group/epi_mono.lean index 46d38b5c0f191..3143144f14557 100644 --- a/src/algebra/category/Group/epi_mono.lean +++ b/src/algebra/category/Group/epi_mono.lean @@ -270,29 +270,27 @@ begin end (by simp), end -lemma agree : - f.range.carrier = {x | h x = g x} := -set.ext $ λ b, -⟨begin - rintros ⟨a, rfl⟩, - change h (f a) = g (f a), - ext ⟨⟨_, ⟨y, rfl⟩⟩⟩, - { rw [g_apply_from_coset], - by_cases m : y ∈ f.range, - { rw [h_apply_from_coset' _ _ _ m, from_coset_eq_of_mem_range _ m], - change from_coset _ = from_coset ⟨f a *l (y *l _), _⟩, - simpa only [←from_coset_eq_of_mem_range _ (subgroup.mul_mem _ ⟨a, rfl⟩ m), - left_coset_assoc] }, - { rw [h_apply_from_coset_nin_range _ _ ⟨_, rfl⟩ _ m], - simpa only [←subtype.val_eq_coe, left_coset_assoc], }, }, - { rw [g_apply_infinity, h_apply_infinity _ _ ⟨_, rfl⟩], }, -end, λ (hb : h b = g b), classical.by_contradiction $ λ r, begin - have eq1 : (h b) (from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩) = - (from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩) := by simp [H, tau, g_apply_infinity], - have eq2 : (g b) (from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩) = - (from_coset ⟨b *l f.range.carrier, ⟨b, rfl⟩⟩) := rfl, - exact (from_coset_ne_of_nin_range _ r).symm (by rw [←eq1, ←eq2, fun_like.congr_fun hb]), -end⟩ +lemma agree : f.range.carrier = {x | h x = g x} := +begin + refine set.ext (λ b, ⟨_, λ (hb : h b = g b), classical.by_contradiction (λ r, _)⟩), + { rintros ⟨a, rfl⟩, + change h (f a) = g (f a), + ext ⟨⟨_, ⟨y, rfl⟩⟩⟩, + { rw [g_apply_from_coset], + by_cases m : y ∈ f.range, + { rw [h_apply_from_coset' _ _ _ m, from_coset_eq_of_mem_range _ m], + change from_coset _ = from_coset ⟨f a *l (y *l _), _⟩, + simpa only [←from_coset_eq_of_mem_range _ (subgroup.mul_mem _ ⟨a, rfl⟩ m), + left_coset_assoc] }, + { rw [h_apply_from_coset_nin_range _ _ ⟨_, rfl⟩ _ m], + simpa only [←subtype.val_eq_coe, left_coset_assoc], }, }, + { rw [g_apply_infinity, h_apply_infinity _ _ ⟨_, rfl⟩], } }, + { have eq1 : (h b) (from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩) = + (from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩) := by simp [H, tau, g_apply_infinity], + have eq2 : (g b) (from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩) = + (from_coset ⟨b *l f.range.carrier, ⟨b, rfl⟩⟩) := rfl, + exact (from_coset_ne_of_nin_range _ r).symm (by rw [←eq1, ←eq2, fun_like.congr_fun hb]) } +end lemma comp_eq : f ≫ (show B ⟶ Group.of SX', from g) = f ≫ h := fun_like.ext _ _ $ λ a, From 02c7601f0637f8d30b309a4f496d1a780c20e17a Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Mon, 18 Jul 2022 12:37:41 +0100 Subject: [PATCH 21/84] Update src/algebra/category/Group/epi_mono.lean Co-authored-by: Markus Himmel --- src/algebra/category/Group/epi_mono.lean | 8 +++----- 1 file changed, 3 insertions(+), 5 deletions(-) diff --git a/src/algebra/category/Group/epi_mono.lean b/src/algebra/category/Group/epi_mono.lean index 3143144f14557..831b2f0832a10 100644 --- a/src/algebra/category/Group/epi_mono.lean +++ b/src/algebra/category/Group/epi_mono.lean @@ -263,11 +263,9 @@ begin (from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩) ∞ (from_coset ⟨b *l f.range.carrier, ⟨b, rfl⟩⟩) (from_coset_ne_of_nin_range _ hb) (by simp), simp only [g_apply_from_coset, ←subtype.val_eq_coe, left_coset_assoc], - exact equiv.swap_apply_of_ne_of_ne begin - refine from_coset_ne_of_nin_range _ (λ r, hb _), - convert subgroup.mul_mem _ (subgroup.inv_mem _ hx) r, - rw [←mul_assoc, mul_left_inv, one_mul], - end (by simp), + refine equiv.swap_apply_of_ne_of_ne (from_coset_ne_of_nin_range _ (λ r, hb _)) (by simp), + convert subgroup.mul_mem _ (subgroup.inv_mem _ hx) r, + rw [←mul_assoc, mul_left_inv, one_mul], end lemma agree : f.range.carrier = {x | h x = g x} := From 99d294501d1eb23b2cb32a2f18ca4e20a83c4d36 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Mon, 18 Jul 2022 12:38:01 +0100 Subject: [PATCH 22/84] Update src/algebra/category/Group/epi_mono.lean Co-authored-by: Markus Himmel --- src/algebra/category/Group/epi_mono.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/src/algebra/category/Group/epi_mono.lean b/src/algebra/category/Group/epi_mono.lean index 831b2f0832a10..7e9ef08d04541 100644 --- a/src/algebra/category/Group/epi_mono.lean +++ b/src/algebra/category/Group/epi_mono.lean @@ -169,8 +169,8 @@ noncomputable instance : decidable_eq X' := classical.dec_eq _ /-- Let `τ` be the permutation on `X'` exchanging `f.range` and the point at infinity. -/ -noncomputable def tau : SX' := equiv.swap -(from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩) ∞ +noncomputable def tau : SX' := +equiv.swap (from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩) ∞ local notation `τ` := tau f From 5fd752913c522369bcc686b0d516c4499fc285a0 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Mon, 18 Jul 2022 12:38:17 +0100 Subject: [PATCH 23/84] Update src/algebra/category/Group/epi_mono.lean Co-authored-by: Markus Himmel --- src/algebra/category/Group/epi_mono.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/algebra/category/Group/epi_mono.lean b/src/algebra/category/Group/epi_mono.lean index 7e9ef08d04541..2a035404a29b3 100644 --- a/src/algebra/category/Group/epi_mono.lean +++ b/src/algebra/category/Group/epi_mono.lean @@ -204,7 +204,7 @@ def G : B →* SX' := inv_fun := λ x, β⁻¹ • x, left_inv := λ x, by dsimp only; rw [←mul_smul, mul_left_inv, one_smul], right_inv := λ x, by dsimp only; rw [←mul_smul, mul_right_inv, one_smul] }, - map_one' := by ext; simp [one_smul], + map_one' := by { ext, simp [one_smul] }, map_mul' := λ b1 b2, by ext; simp [mul_smul] } local notation `g` := G f From 0d8ea158d2327870bc08b304e7e9560c3ec2f038 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Mon, 18 Jul 2022 13:03:21 +0100 Subject: [PATCH 24/84] add `noncomputable` --- src/algebra/category/Group/epi_mono.lean | 39 +++++------------------- 1 file changed, 7 insertions(+), 32 deletions(-) diff --git a/src/algebra/category/Group/epi_mono.lean b/src/algebra/category/Group/epi_mono.lean index 1f1c25706caa7..18540a30da935 100644 --- a/src/algebra/category/Group/epi_mono.lean +++ b/src/algebra/category/Group/epi_mono.lean @@ -13,16 +13,15 @@ In this file, we prove monomorphisms in category of group are injective homomorp epimorphisms are surjective homomorphisms. -/ +noncomputable theory -universes u +universes u v namespace monoid_hom open quotient_group -variables {A B : Type u} - -section +variables {A : Type u} {B : Type v} variables [group A] [group B] @@ -39,30 +38,6 @@ begin rw [←subgroup.subtype_range f.ker, ←h, range_one], end -end - -section - -variables [comm_group A] [comm_group B] - -@[to_additive add_monoid_hom.range_eq_top_of_cancel] -lemma range_eq_top_of_cancel {f : A →* B} - (h : ∀ (u v : B →* B ⧸ f.range), u.comp f = v.comp f → u = v) : - f.range = ⊤ := -begin - specialize h 1 (quotient_group.mk' _) begin - ext1, - simp only [one_apply, coe_comp, coe_mk', function.comp_app], - rw [show (1 : B ⧸ f.range) = (1 : B), from quotient_group.coe_one _, quotient_group.eq, inv_one, - one_mul], - exact ⟨x, rfl⟩, - end, - replace h : (quotient_group.mk' _).ker = (1 : B →* B ⧸ f.range).ker := by rw h, - rwa [ker_one, quotient_group.ker_mk] at h, -end - -end - end monoid_hom section @@ -95,7 +70,7 @@ namespace surjective_of_epi_auxs local notation `X` := set.range (function.swap left_coset f.range.carrier) /-- -Define `X'` to be the set of all right cosets with an extra point at "infinity". +Define `X'` to be the set of all left cosets with an extra point at "infinity". -/ @[nolint has_inhabited_instance] inductive X_with_infinity @@ -163,12 +138,12 @@ begin end -noncomputable instance : decidable_eq X' := classical.dec_eq _ +instance : decidable_eq X' := classical.dec_eq _ /-- Let `τ` be the permutation on `X'` exchanging `f.range` and the point at infinity. -/ -noncomputable def tau : SX' := equiv.swap +def tau : SX' := equiv.swap (from_coset ⟨f.range.carrier, ⟨1, one_left_coset _⟩⟩) ∞ local notation `τ` := tau f @@ -211,7 +186,7 @@ local notation `g` := G f /-- Define `h : B ⟶ S(X')` to be `τ g τ⁻¹` -/ -noncomputable def H : B →* SX':= +def H : B →* SX':= { to_fun := λ β, ((τ).symm.trans (g β)).trans τ, map_one' := by ext; simp, map_mul' := λ b1 b2, by ext; simp } From c199b075ba547f900729ce12a7bcd4770ff1ace2 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Mon, 18 Jul 2022 13:53:52 +0100 Subject: [PATCH 25/84] reformat --- src/algebra/category/Group/epi_mono.lean | 54 +++++++++++++----------- 1 file changed, 29 insertions(+), 25 deletions(-) diff --git a/src/algebra/category/Group/epi_mono.lean b/src/algebra/category/Group/epi_mono.lean index eb02f54836dbf..04bff2688f969 100644 --- a/src/algebra/category/Group/epi_mono.lean +++ b/src/algebra/category/Group/epi_mono.lean @@ -98,20 +98,22 @@ instance : has_smul B X' := lemma mul_smul (b b' : B) (x : X') : (b * b') • x = b • b' • x := match x with -| from_coset y := begin - change from_coset _ = from_coset _, - simp only [←subtype.val_eq_coe, left_coset_assoc], -end +| from_coset y := + begin + change from_coset _ = from_coset _, + simp only [←subtype.val_eq_coe, left_coset_assoc], + end | ∞ := rfl end lemma one_smul (x : X') : (1 : B) • x = x := match x with -| from_coset y := begin - change from_coset _ = from_coset _, - simp only [←subtype.val_eq_coe, one_left_coset, subtype.ext_iff_val], -end +| from_coset y := + begin + change from_coset _ = from_coset _, + simp only [←subtype.val_eq_coe, one_left_coset, subtype.ext_iff_val], + end | ∞ := rfl end @@ -170,17 +172,17 @@ lemma τ_symm_apply_infinity : by rw [tau, equiv.symm_swap, equiv.swap_apply_right] /-- -Let `g : B ⟶ S(X')` be defined as as such that, for any `β : B`, `g(β)` is the function sending +Let `g : B ⟶ S(X')` be defined as such that, for any `β : B`, `g(β)` is the function sending point at infinity to point at infinity and sending coset `y` to `β *l y`. -/ def G : B →* SX' := { to_fun := λ β, { to_fun := λ x, β • x, inv_fun := λ x, β⁻¹ • x, - left_inv := λ x, by dsimp only; rw [←mul_smul, mul_left_inv, one_smul], - right_inv := λ x, by dsimp only; rw [←mul_smul, mul_right_inv, one_smul] }, + left_inv := λ x, by { dsimp only, rw [←mul_smul, mul_left_inv, one_smul] }, + right_inv := λ x, by { dsimp only, rw [←mul_smul, mul_right_inv, one_smul] } }, map_one' := by { ext, simp [one_smul] }, - map_mul' := λ b1 b2, by ext; simp [mul_smul] } + map_mul' := λ b1 b2, by { ext, simp [mul_smul] } } local notation `g` := G f @@ -189,8 +191,8 @@ Define `h : B ⟶ S(X')` to be `τ g τ⁻¹` -/ def H : B →* SX':= { to_fun := λ β, ((τ).symm.trans (g β)).trans τ, - map_one' := by ext; simp, - map_mul' := λ b1 b2, by ext; simp } + map_one' := by { ext, simp }, + map_mul' := λ b1 b2, by { ext, simp } } local notation `h` := H f @@ -203,15 +205,16 @@ The strategy is the following: assuming `epi f` lemma g_apply_from_coset (x : B) (y : X) : (g x) (from_coset y) = - from_coset ⟨x *l y, begin - rw [←subtype.val_eq_coe, ←y.2.some_spec, set.mem_range, left_coset_assoc], - use x * y.2.some, - end⟩ := rfl + from_coset ⟨x *l y, + begin + rwa [←subtype.val_eq_coe, ←y.2.some_spec, set.mem_range, left_coset_assoc], + use x * y.2.some + end⟩ := +rfl lemma g_apply_infinity (x : B) : (g x) ∞ = ∞ := rfl -lemma h_apply_infinity (x : B) (hx : x ∈ f.range) : - (h x) ∞ = ∞ := +lemma h_apply_infinity (x : B) (hx : x ∈ f.range) : (h x) ∞ = ∞ := begin simp only [H, monoid_hom.coe_mk, equiv.to_fun_as_coe, equiv.coe_trans, function.comp_app], rw [τ_symm_apply_infinity, g_apply_from_coset], @@ -271,7 +274,8 @@ fun_like.ext _ _ $ λ a, lemma g_ne_h (x : B) (hx : x ∉ f.range) : g ≠ h := -λ r, begin +begin + intros r, replace r := fun_like.congr_fun (fun_like.congr_fun r x) ((from_coset ⟨f.range, ⟨1, one_left_coset _⟩⟩)), rw [H, g_apply_from_coset, monoid_hom.coe_mk, tau] at r, @@ -284,12 +288,12 @@ end end surjective_of_epi_auxs lemma surjective_of_epi [epi f] : function.surjective f := -classical.by_contradiction $ λ r, begin +begin + by_contra r, push_neg at r, rcases r with ⟨b, hb⟩, - refine surjective_of_epi_auxs.g_ne_h f b _ ((cancel_epi f).1 (surjective_of_epi_auxs.comp_eq f)), - rintros ⟨c, hc⟩, - exact hb _ hc, + exact surjective_of_epi_auxs.g_ne_h f b (λ ⟨c, hc⟩, hb _ hc) + ((cancel_epi f).1 (surjective_of_epi_auxs.comp_eq f)), end lemma epi_iff_surjective : From f82ff81f7e838f7be78b13f94e43b14eda4d09d7 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Mon, 18 Jul 2022 14:31:00 +0100 Subject: [PATCH 26/84] initial --- .../Group/equivalence_Group_AddGroup.lean | 104 ++++++++++++++++++ 1 file changed, 104 insertions(+) create mode 100644 src/algebra/category/Group/equivalence_Group_AddGroup.lean diff --git a/src/algebra/category/Group/equivalence_Group_AddGroup.lean b/src/algebra/category/Group/equivalence_Group_AddGroup.lean new file mode 100644 index 0000000000000..2a0e66944c2d9 --- /dev/null +++ b/src/algebra/category/Group/equivalence_Group_AddGroup.lean @@ -0,0 +1,104 @@ +import algebra.category.Group.basic + +namespace Group + +@[simps] +def to_AddGroup : Group ⥤ AddGroup := +{ obj := λ X, ⟨additive X⟩, + map := λ X Y f, + { to_fun := λ x, f x, + map_zero' := by { erw [map_one], refl }, + map_add' := λ x y, by { erw [map_mul], refl } }, + map_id' := λ X, by { ext, refl }, + map_comp' := λ X Y Z f g, by { ext, refl } } + +end Group + +namespace CommGroup + +@[simps] +def to_AddCommGroup : CommGroup ⥤ AddCommGroup := +{ obj := λ X, ⟨additive X⟩, + map := λ X Y f, + { to_fun := λ x, f x, + map_zero' := by { erw [map_one], refl }, + map_add' := λ x y, by { erw [map_mul], refl } }, + map_id' := λ X, by { ext, refl }, + map_comp' := λ X Y Z f g, by { ext, refl } } + +end CommGroup + +namespace AddGroup + +@[simps] +def to_Group : AddGroup ⥤ Group := +{ obj := λ X, ⟨multiplicative X⟩, + map := λ X Y f, + { to_fun := λ x, f x, + map_one' := by { erw [map_zero], refl }, + map_mul' := λ x y, by { erw [map_add], refl } }, + map_id' := λ X, by { ext, refl }, + map_comp' := λ X Y Z f g, by { ext, refl } } + +end AddGroup + +namespace AddCommGroup + +@[simps] +def to_CommGroup : AddCommGroup ⥤ CommGroup := +{ obj := λ X, ⟨multiplicative X⟩, + map := λ X Y f, + { to_fun := λ x, f x, + map_one' := by { erw [map_zero], refl }, + map_mul' := λ x y, by { erw [map_add], refl } }, + map_id' := λ X, by { ext, refl }, + map_comp' := λ X Y Z f g, by { ext, refl } } + +end AddCommGroup + +@[simps] +def Group_AddGroup_equivalence : Group ≌ AddGroup := +{ functor := Group.to_AddGroup, + inverse := AddGroup.to_Group, + unit_iso := + { hom := + { app := λ X, (𝟙 X), + naturality' := λ _ _ _, by { ext, refl } }, + inv := + { app := λ X, (𝟙 X), + naturality' := λ _ _ _, by { ext, refl } }, + hom_inv_id' := rfl, + inv_hom_id' := rfl }, + counit_iso := + { hom := + { app := λ X, (𝟙 X), + naturality' := λ _ _ _, by { ext, refl } }, + inv := + { app := λ X, (𝟙 X), + naturality' := λ _ _ _, by { ext, refl } }, + hom_inv_id' := rfl, + inv_hom_id' := rfl }, + functor_unit_iso_comp' := λ X, rfl } + +def CommGroup_AddCommGroup_equivalence : CommGroup ≌ AddCommGroup := +{ functor := CommGroup.to_AddCommGroup, + inverse := AddCommGroup.to_CommGroup, + unit_iso := + { hom := + { app := λ X, (𝟙 X), + naturality' := λ _ _ _, by { ext, refl } }, + inv := + { app := λ X, (𝟙 X), + naturality' := λ _ _ _, by { ext, refl } }, + hom_inv_id' := rfl, + inv_hom_id' := rfl }, + counit_iso := + { hom := + { app := λ X, (𝟙 X), + naturality' := λ _ _ _, by { ext, refl } }, + inv := + { app := λ X, (𝟙 X), + naturality' := λ _ _ _, by { ext, refl } }, + hom_inv_id' := rfl, + inv_hom_id' := rfl }, + functor_unit_iso_comp' := λ X, rfl } From dab193b9d01dd5e1f57958d6ed9e9e6170e3aaba Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Mon, 18 Jul 2022 16:07:48 +0100 Subject: [PATCH 27/84] add missing --- src/algebra/category/Group/epi_mono.lean | 97 +++++++++++++++++++++++- 1 file changed, 96 insertions(+), 1 deletion(-) diff --git a/src/algebra/category/Group/epi_mono.lean b/src/algebra/category/Group/epi_mono.lean index 04bff2688f969..765a8c3bcc9f8 100644 --- a/src/algebra/category/Group/epi_mono.lean +++ b/src/algebra/category/Group/epi_mono.lean @@ -5,7 +5,7 @@ Authors: Jujian Zhang -/ import category_theory.epi_mono import group_theory.quotient_group -import algebra.category.Group.basic +import algebra.category.Group.equivalence_Group_AddGroup /-! # Monomorphisms and epimorphisms in `Group` @@ -21,8 +21,11 @@ namespace monoid_hom open quotient_group + variables {A : Type u} {B : Type v} +section + variables [group A] [group B] @[to_additive add_monoid_hom.ker_eq_bot_of_cancel] @@ -38,6 +41,30 @@ begin rw [←subgroup.subtype_range f.ker, ←h, range_one], end +end + +section + + variables [comm_group A] [comm_group B] + + @[to_additive add_monoid_hom.range_eq_top_of_cancel] + lemma range_eq_top_of_cancel {f : A →* B} + (h : ∀ (u v : B →* B ⧸ f.range), u.comp f = v.comp f → u = v) : + f.range = ⊤ := + begin + specialize h 1 (quotient_group.mk' _) begin + ext1, + simp only [one_apply, coe_comp, coe_mk', function.comp_app], + rw [show (1 : B ⧸ f.range) = (1 : B), from quotient_group.coe_one _, quotient_group.eq, inv_one, + one_mul], + exact ⟨x, rfl⟩, + end, + replace h : (quotient_group.mk' _).ker = (1 : B →* B ⧸ f.range).ker := by rw h, + rwa [ker_one, quotient_group.ker_mk] at h, + end + +end + end monoid_hom section @@ -306,4 +333,72 @@ iff.trans (epi_iff_surjective _) (subgroup.eq_top_iff' f.range).symm end Group +namespace AddGroup + +variables {A B : AddGroup.{u}} (f : A ⟶ B) + +lemma epi_iff_surjective : + epi f ↔ function.surjective f := +begin + have i1 : epi f ↔ epi (Group_AddGroup_equivalence.inverse.map f), + { refine ⟨_, Group_AddGroup_equivalence.inverse.epi_of_epi_map⟩, + introsI e', + apply Group_AddGroup_equivalence.inverse.map_epi }, + rwa Group.epi_iff_surjective at i1, +end + +lemma epi_iff_range_eq_top : + epi f ↔ f.range = ⊤ := +iff.trans (epi_iff_surjective _) (add_subgroup.eq_top_iff' f.range).symm + +end AddGroup + +namespace CommGroup + +variables {A B : CommGroup.{u}} (f : A ⟶ B) + +@[to_additive AddCommGroup.ker_eq_bot_of_mono] +lemma ker_eq_bot_of_mono [mono f] : f.ker = ⊥ := +monoid_hom.ker_eq_bot_of_cancel $ λ u v, + (@cancel_mono _ _ _ _ _ f _ + (show CommGroup.of f.ker ⟶ A, from u) _).1 + +@[to_additive AddCommGroup.mono_iff_ker_eq_bot] +lemma mono_iff_ker_eq_bot : + mono f ↔ f.ker = ⊥ := +⟨λ h, @@ker_eq_bot_of_mono f h, +λ h, concrete_category.mono_of_injective _ $ (monoid_hom.ker_eq_bot_iff f).1 h⟩ + +@[to_additive AddCommGroup.mono_iff_injective] +lemma mono_iff_injective : + mono f ↔ function.injective f := +iff.trans (mono_iff_ker_eq_bot f) $ monoid_hom.ker_eq_bot_iff f + +@[to_additive] +instance : functor.reflects_epimorphisms (forget₂ CommGroup Group) := +functor.reflects_epimorphisms_of_faithful _ + +@[to_additive] +lemma range_eq_top_of_epi [epi f] : f.range = ⊤ := +monoid_hom.range_eq_top_of_cancel $ λ u v h, + (@cancel_epi _ _ _ _ _ f _ (show B ⟶ ⟨B ⧸ monoid_hom.range f⟩, from u) v).1 h + +@[to_additive] +lemma epi_iff_range_eq_top : epi f ↔ f.range = ⊤ := +⟨λ hf, by exactI range_eq_top_of_epi _, + λ hf, concrete_category.epi_of_surjective _ $ monoid_hom.range_top_iff_surjective.mp hf⟩ + +@[to_additive] +lemma epi_iff_surjective : epi f ↔ function.surjective f := +by rw [epi_iff_range_eq_top, monoid_hom.range_top_iff_surjective] + +@[to_additive] +instance : functor.preserves_epimorphisms (forget₂ CommGroup Group) := +{ preserves := λ X Y f e, begin + rw [epi_iff_surjective, ←@Group.epi_iff_surjective ⟨X⟩ ⟨Y⟩ f] at e, + exact e, + end } + +end CommGroup + end From 38ba5c9d9dfa795805e824659dd997c7d69938c3 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Mon, 18 Jul 2022 16:14:30 +0100 Subject: [PATCH 28/84] add docs --- .../Group/equivalence_Group_AddGroup.lean | 16 ++++++++++++++++ 1 file changed, 16 insertions(+) diff --git a/src/algebra/category/Group/equivalence_Group_AddGroup.lean b/src/algebra/category/Group/equivalence_Group_AddGroup.lean index 2a0e66944c2d9..a9d5fc2873eaf 100644 --- a/src/algebra/category/Group/equivalence_Group_AddGroup.lean +++ b/src/algebra/category/Group/equivalence_Group_AddGroup.lean @@ -1,5 +1,21 @@ +/- +Copyright (c) 2022 Jujian Zhang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jujian Zhang +-/ import algebra.category.Group.basic +/-! +# Equivalence between `Group` and `AddGroup` + +This file contains two equivalence: +* `Group_AddGroup_equivalence` : the equivalence between `Group` and `AddGroup` by sending + `X : Group` to `additive X` and `Y : AddGroup` to `multiplicative Y`. +* `CommGroup_AddCommGroup_equivlance` : the equivalence between `CommGroup` and `AddCommGroup` by + sending `X : CommGroup` to `additive X` and `Y : AddCommGroup` to `multiplicative Y`. +-/ + + namespace Group @[simps] From 68b3523786b9629166364317d5e9adf3ac132bd8 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Mon, 18 Jul 2022 16:54:00 +0100 Subject: [PATCH 29/84] fix linter error --- src/algebra/category/Group/epi_mono.lean | 11 +++++------ 1 file changed, 5 insertions(+), 6 deletions(-) diff --git a/src/algebra/category/Group/epi_mono.lean b/src/algebra/category/Group/epi_mono.lean index 765a8c3bcc9f8..8794f7c9ebcf5 100644 --- a/src/algebra/category/Group/epi_mono.lean +++ b/src/algebra/category/Group/epi_mono.lean @@ -52,13 +52,12 @@ section (h : ∀ (u v : B →* B ⧸ f.range), u.comp f = v.comp f → u = v) : f.range = ⊤ := begin - specialize h 1 (quotient_group.mk' _) begin - ext1, + specialize h 1 (quotient_group.mk' _) _, + { ext1, simp only [one_apply, coe_comp, coe_mk', function.comp_app], - rw [show (1 : B ⧸ f.range) = (1 : B), from quotient_group.coe_one _, quotient_group.eq, inv_one, - one_mul], - exact ⟨x, rfl⟩, - end, + rw [show (1 : B ⧸ f.range) = (1 : B), from quotient_group.coe_one _, quotient_group.eq, + inv_one, one_mul], + exact ⟨x, rfl⟩, }, replace h : (quotient_group.mk' _).ker = (1 : B →* B ⧸ f.range).ker := by rw h, rwa [ker_one, quotient_group.ker_mk] at h, end From c2f9119df6ba6fa23233a4fc70be2f87cd2d523d Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Mon, 18 Jul 2022 23:07:49 +0100 Subject: [PATCH 30/84] fix doc_blame error --- .../Group/equivalence_Group_AddGroup.lean | 33 +++++++++++++------ 1 file changed, 23 insertions(+), 10 deletions(-) diff --git a/src/algebra/category/Group/equivalence_Group_AddGroup.lean b/src/algebra/category/Group/equivalence_Group_AddGroup.lean index a9d5fc2873eaf..fb2c2b69e974f 100644 --- a/src/algebra/category/Group/equivalence_Group_AddGroup.lean +++ b/src/algebra/category/Group/equivalence_Group_AddGroup.lean @@ -18,8 +18,10 @@ This file contains two equivalence: namespace Group -@[simps] -def to_AddGroup : Group ⥤ AddGroup := +/-- +The functor `Group ⥤ AddGroup` by sending `X ↦ additive X` and `f ↦ f`. +-/ +@[simps] def to_AddGroup : Group ⥤ AddGroup := { obj := λ X, ⟨additive X⟩, map := λ X Y f, { to_fun := λ x, f x, @@ -32,8 +34,10 @@ end Group namespace CommGroup -@[simps] -def to_AddCommGroup : CommGroup ⥤ AddCommGroup := +/-- +The functor `CommGroup ⥤ AddCommGroup` by sending `X ↦ additive X` and `f ↦ f`. +-/ +@[simps] def to_AddCommGroup : CommGroup ⥤ AddCommGroup := { obj := λ X, ⟨additive X⟩, map := λ X Y f, { to_fun := λ x, f x, @@ -46,8 +50,10 @@ end CommGroup namespace AddGroup -@[simps] -def to_Group : AddGroup ⥤ Group := +/-- +The functor `AddGroup ⥤ Group` by sending `X ↦ multiplicative Y` and `f ↦ f`. +-/ +@[simps] def to_Group : AddGroup ⥤ Group := { obj := λ X, ⟨multiplicative X⟩, map := λ X Y f, { to_fun := λ x, f x, @@ -60,8 +66,10 @@ end AddGroup namespace AddCommGroup -@[simps] -def to_CommGroup : AddCommGroup ⥤ CommGroup := +/-- +The functor `AddCommGroup ⥤ CommGroup` by sending `X ↦ multiplicative Y` and `f ↦ f`. +-/ +@[simps] def to_CommGroup : AddCommGroup ⥤ CommGroup := { obj := λ X, ⟨multiplicative X⟩, map := λ X Y f, { to_fun := λ x, f x, @@ -72,8 +80,10 @@ def to_CommGroup : AddCommGroup ⥤ CommGroup := end AddCommGroup -@[simps] -def Group_AddGroup_equivalence : Group ≌ AddGroup := +/-- +The equivalence of categories between `Group` and `AddGroup` +-/ +@[simps] def Group_AddGroup_equivalence : Group ≌ AddGroup := { functor := Group.to_AddGroup, inverse := AddGroup.to_Group, unit_iso := @@ -96,6 +106,9 @@ def Group_AddGroup_equivalence : Group ≌ AddGroup := inv_hom_id' := rfl }, functor_unit_iso_comp' := λ X, rfl } +/-- +The equivalence of categories between `CommGroup` and `AddCommGroup`. +-/ def CommGroup_AddCommGroup_equivalence : CommGroup ≌ AddCommGroup := { functor := CommGroup.to_AddCommGroup, inverse := AddCommGroup.to_CommGroup, From ff66715e330f4d72f12806b598ba5e5dd188fad8 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Tue, 19 Jul 2022 12:59:30 +0100 Subject: [PATCH 31/84] fix after merge master --- src/algebra/category/Group/epi_mono.lean | 36 +++++++++++++++++++----- 1 file changed, 29 insertions(+), 7 deletions(-) diff --git a/src/algebra/category/Group/epi_mono.lean b/src/algebra/category/Group/epi_mono.lean index 45bbc6f5c6ef4..92db1b23e3f8e 100644 --- a/src/algebra/category/Group/epi_mono.lean +++ b/src/algebra/category/Group/epi_mono.lean @@ -23,6 +23,8 @@ open quotient_group variables {A : Type u} {B : Type v} +section + variables [group A] [group B] @[to_additive add_monoid_hom.ker_eq_bot_of_cancel] @@ -30,6 +32,29 @@ lemma ker_eq_bot_of_cancel {f : A →* B} (h : ∀ (u v : f.ker →* A), f.comp f.ker = ⊥ := by simpa using _root_.congr_arg range (h f.ker.subtype 1 (by tidy)) +end + +section + +variables [comm_group A] [comm_group B] + +@[to_additive add_monoid_hom.range_eq_top_of_cancel] +lemma range_eq_top_of_cancel {f : A →* B} + (h : ∀ (u v : B →* B ⧸ f.range), u.comp f = v.comp f → u = v) : + f.range = ⊤ := +begin + specialize h 1 (quotient_group.mk' _) _, + { ext1, + simp only [one_apply, coe_comp, coe_mk', function.comp_app], + rw [show (1 : B ⧸ f.range) = (1 : B), from quotient_group.coe_one _, quotient_group.eq, + inv_one, one_mul], + exact ⟨x, rfl⟩, }, + replace h : (quotient_group.mk' _).ker = (1 : B →* B ⧸ f.range).ker := by rw h, + rwa [ker_one, quotient_group.ker_mk] at h, +end + +end + end monoid_hom section @@ -308,18 +333,15 @@ variables {A B : CommGroup.{u}} (f : A ⟶ B) @[to_additive AddCommGroup.ker_eq_bot_of_mono] lemma ker_eq_bot_of_mono [mono f] : f.ker = ⊥ := monoid_hom.ker_eq_bot_of_cancel $ λ u v, - (@cancel_mono _ _ _ _ _ f _ - (show CommGroup.of f.ker ⟶ A, from u) _).1 + (@cancel_mono _ _ _ _ _ f _ (show Group.of f.ker ⟶ A, from u) _).1 @[to_additive AddCommGroup.mono_iff_ker_eq_bot] -lemma mono_iff_ker_eq_bot : - mono f ↔ f.ker = ⊥ := +lemma mono_iff_ker_eq_bot : mono f ↔ f.ker = ⊥ := ⟨λ h, @@ker_eq_bot_of_mono f h, -λ h, concrete_category.mono_of_injective _ $ (monoid_hom.ker_eq_bot_iff f).1 h⟩ + λ h, concrete_category.mono_of_injective _ $ (monoid_hom.ker_eq_bot_iff f).1 h⟩ @[to_additive AddCommGroup.mono_iff_injective] -lemma mono_iff_injective : - mono f ↔ function.injective f := +lemma mono_iff_injective : mono f ↔ function.injective f := iff.trans (mono_iff_ker_eq_bot f) $ monoid_hom.ker_eq_bot_iff f @[to_additive] From 8af22a99eed6e5c0ed0eb1d73e1cce8eba52f8e5 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Wed, 27 Jul 2022 22:01:10 +0100 Subject: [PATCH 32/84] need cache --- src/algebraic_geometry/skyscraper.lean | 90 ++++++++++++++++++++++++++ src/category_theory/eq_to_hom.lean | 10 +++ 2 files changed, 100 insertions(+) create mode 100644 src/algebraic_geometry/skyscraper.lean diff --git a/src/algebraic_geometry/skyscraper.lean b/src/algebraic_geometry/skyscraper.lean new file mode 100644 index 0000000000000..3f7262a2b1ff3 --- /dev/null +++ b/src/algebraic_geometry/skyscraper.lean @@ -0,0 +1,90 @@ +import algebraic_geometry.sheafed_space +import topology.sheaves.sheaf_condition.unique_gluing + +section + +open topological_space +open category_theory category_theory.limits +open Top +open opposite + +universes u v w + +variables {X : Top.{u}} (p₀ : X) {C : Type v} [category.{w} C] (S : C) +variables {star : C} (ts : is_terminal star) +variable [Π (U : (opens X)ᵒᵖ), decidable (p₀ ∈ unop U)] + +instance : Π (U : opens X), decidable (p₀ ∈ U) := +λ U, (infer_instance : decidable (p₀ ∈ unop (op U))) + +@[simps] +def skyscraper_presheaf : presheaf C X := +{ obj := λ U, ite (p₀ ∈ unop U) S star, + map := λ U V i, dite (p₀ ∈ unop V) + (λ h, eq_to_hom (if_pos (le_of_hom i.unop h)) ≫ 𝟙 S ≫ eq_to_hom (if_pos h).symm) + (λ h, ts.from _ ≫ eq_to_hom (if_neg h).symm), + map_id' := λ U, + begin + split_ifs, + { simp, }, + { rw eq_comp_eq_to_hom, + exact ts.hom_ext _ _, }, + end, + map_comp' := λ U V W iVU iWV, + begin + by_cases hW : p₀ ∈ unop W, + { have hV : p₀ ∈ unop V := le_of_hom iWV.unop hW, + have hU : p₀ ∈ unop U := le_of_hom iVU.unop hV, + split_ifs, + simp }, + { split_ifs; + rw eq_comp_eq_to_hom; + exact ts.hom_ext _ _, } + end } + +section + +variables {p₀} + +lemma skyscraper_presheaf_obj_of_mem {U : opens X} (h : p₀ ∈ U) : + (skyscraper_presheaf p₀ S ts).obj (op U) = S := +if_pos h + +lemma skyscraper_presheaf_obj_of_not_mem {U : opens X} (h : p₀ ∉ U) : + (skyscraper_presheaf p₀ S ts).obj (op U) = star := +if_neg h + +end + +def skyscraper_sheaf : sheaf C X := +⟨skyscraper_presheaf p₀ S ts, λ c U s hs x hx, + ⟨dite (p₀ ∈ U) + (λ h, x (hs p₀ h).some_spec.some (hs p₀ h).some_spec.some_spec.1 ≫ + eq_to_hom ((skyscraper_presheaf_obj_of_mem S ts (hs p₀ h).some_spec.some_spec.2).trans + (skyscraper_presheaf_obj_of_mem S ts h).symm)) + (λ h, ts.from c ≫ (eq_to_hom (skyscraper_presheaf_obj_of_not_mem S ts h).symm)), + λ V inc h, + begin + by_cases hV : p₀ ∈ V, + { have hU : p₀ ∈ U := le_of_hom inc hV, + split_ifs, + generalize_proofs h₁ h₂ h₃ h₄, + dsimp, + split_ifs, + -- generalize_proofs h₅ h₆, + rw [category.id_comp, eq_to_hom_trans, category.assoc, eq_to_hom_trans], + generalize_proofs h₅, + have := hx _ _ h₂.some_spec.1 h rfl, + }, + end, _⟩⟩ + +/- + +dite (p₀ ∈ U) + (λ h, (begin + have := x (𝟙 _) _, + dsimp at this, + end : c ⟶ S) ≫ (eq_to_hom (skyscraper_presheaf_obj_of_mem S ts h).symm)) + (λ h, ts.from c ≫ (eq_to_hom (skyscraper_presheaf_obj_of_not_mem S ts h).symm)) +-/ +end diff --git a/src/category_theory/eq_to_hom.lean b/src/category_theory/eq_to_hom.lean index 65f70dda347ec..4899e382b313a 100644 --- a/src/category_theory/eq_to_hom.lean +++ b/src/category_theory/eq_to_hom.lean @@ -45,6 +45,16 @@ def eq_to_hom {X Y : C} (p : X = Y) : X ⟶ Y := by rw p; exact 𝟙 _ eq_to_hom p ≫ eq_to_hom q = eq_to_hom (p.trans q) := by { cases p, cases q, simp, } +lemma eq_comp_eq_to_hom {X Y Y' : C} (p : Y = Y') (f : X ⟶ Y) (g : X ⟶ Y') : + f ≫ eq_to_hom p = g ↔ f = g ≫ eq_to_hom p.symm := +{ mp := λ h, h ▸ by simp, + mpr := λ h, by simp [eq_whisker h (eq_to_hom p)] } + +lemma eq_eq_to_hom_comp {X X' Y : C} (p : X = X') (f : X ⟶ Y) (g : X' ⟶ Y) : + eq_to_hom p ≫ g = f ↔ g = eq_to_hom p.symm ≫ f := +{ mp := λ h, h ▸ by simp, + mpr := λ h, h ▸ by simp [whisker_eq _ h] } + /-- If we (perhaps unintentionally) perform equational rewriting on the source object of a morphism, From 3f98b7215dc42287e6c1f7e1f910449d2f9ad066 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Thu, 28 Jul 2022 16:51:56 +0100 Subject: [PATCH 33/84] calculating stalk --- src/algebraic_geometry/skyscraper.lean | 282 +++++++++++++++++++++++-- 1 file changed, 268 insertions(+), 14 deletions(-) diff --git a/src/algebraic_geometry/skyscraper.lean b/src/algebraic_geometry/skyscraper.lean index 3f7262a2b1ff3..3399eb5da0adf 100644 --- a/src/algebraic_geometry/skyscraper.lean +++ b/src/algebraic_geometry/skyscraper.lean @@ -1,5 +1,6 @@ import algebraic_geometry.sheafed_space import topology.sheaves.sheaf_condition.unique_gluing +import topology.sheaves.stalks section @@ -12,10 +13,7 @@ universes u v w variables {X : Top.{u}} (p₀ : X) {C : Type v} [category.{w} C] (S : C) variables {star : C} (ts : is_terminal star) -variable [Π (U : (opens X)ᵒᵖ), decidable (p₀ ∈ unop U)] - -instance : Π (U : opens X), decidable (p₀ ∈ U) := -λ U, (infer_instance : decidable (p₀ ∈ unop (op U))) +variable [Π (U : opens X), decidable (p₀ ∈ U)] @[simps] def skyscraper_presheaf : presheaf C X := @@ -71,20 +69,276 @@ def skyscraper_sheaf : sheaf C X := generalize_proofs h₁ h₂ h₃ h₄, dsimp, split_ifs, - -- generalize_proofs h₅ h₆, rw [category.id_comp, eq_to_hom_trans, category.assoc, eq_to_hom_trans], generalize_proofs h₅, have := hx _ _ h₂.some_spec.1 h rfl, - }, - end, _⟩⟩ + work_on_goal 2 { exact h₁.some ⊓ V }, + work_on_goal 2 { exact hom_of_le inf_le_left }, + work_on_goal 2 { exact hom_of_le inf_le_right }, + dsimp at this, + have hV' : p₀ ∈ h₁.some ⊓ V := ⟨h₂.some_spec.2, hV⟩, + split_ifs at this, + rw [category.id_comp, eq_to_hom_trans, eq_to_hom_trans, eq_comp_eq_to_hom, + category.assoc, eq_to_hom_trans] at this, + generalize_proofs h₆ at this, + rw [this, eq_comp_eq_to_hom], }, + { dsimp, + split_ifs, + rw [←category.assoc, eq_comp_eq_to_hom], + exact ts.hom_ext _ _, } + end, + λ y (hy : x.is_amalgamation y), + begin + split_ifs, + { generalize_proofs h₁ h₂ h₃ h₄, + have := hy h₂.some h₂.some_spec.1, + dsimp at this, + split_ifs at this with H, + { rw [category.id_comp, eq_to_hom_trans, eq_comp_eq_to_hom] at this, + rw [this, eq_comp_eq_to_hom, category.assoc, eq_to_hom_trans, eq_to_hom_refl, + category.comp_id], }, + { exfalso, + exact H h₂.some_spec.2, }, }, + { rw [←eq_comp_eq_to_hom], + exact (ts.hom_ext _ _), } + end⟩⟩ + +example : true := trivial + +end + +section + +open topological_space +open category_theory category_theory.limits +open Top +open opposite + +universes u v + +variables {X : Top.{u}} (p₀ : X) {C : Type v} [category.{u} C] (S : C) +variables {star : C} (ts : is_terminal star) +variable [Π (U : opens X), decidable (p₀ ∈ U)] +variable [has_colimits C] + +lemma mem_nhds_of_closure_singleton {y : X} (U : open_nhds y) (h : y ∈ (closure ({p₀} : set X))) : + p₀ ∈ U.1 := +begin + have := U.2, + contrapose this, + change p₀ ∈ U.1.1ᶜ at this, + change y ∈ U.1.1ᶜ, + have h1 : {p₀} ⊆ U.1.1ᶜ := set.singleton_subset_iff.mpr this, + rw ←is_closed.closure_subset_iff at h1, + { exact h1 h }, + rw is_closed_compl_iff, + exact U.1.2 +end + +lemma mem_nhds_of_not_mem_closure_singleton {y : X} (h : y ∉ (closure ({p₀} : set X))) : + ∃ (U : open_nhds y), p₀ ∉ U.1 := +begin + have : ∃ (s : set X), is_closed s ∧ p₀ ∈ s ∧ y ∉ s, + { contrapose h, + push_neg at h ⊢, + rw mem_closure_iff, + intros t ht hy, + specialize h tᶜ (is_closed_compl_iff.mpr ht), + refine ⟨p₀, _⟩, + contrapose! h, + rw [set.mem_inter_iff, not_and_distrib, set.mem_singleton_iff, eq_self_iff_true, + not_true, or_false] at h, + exact ⟨h, λ r, r hy⟩ }, + rcases this with ⟨s, hs, hp₀, hy⟩, + resetI, + exact ⟨⟨⟨sᶜ, is_closed.is_open_compl⟩, hy⟩, λ r, r hp₀⟩, +end + +@[simps] def skyscraper_presheaf_cocone_of_mem_closure₀ {y : X} (h : y ∈ (closure ({p₀} : set X))) : + cocone ((open_nhds.inclusion y).op ⋙ skyscraper_presheaf p₀ S ts) := +{ X := S, + ι := + { app := λ U, eq_to_hom + begin + dsimp, + rw if_pos, + exact mem_nhds_of_closure_singleton _ _ h, + end ≫ 𝟙 S, + naturality' := λ U V inc, + begin + simp only [functor.op_obj, unop_op, functor.comp_map, functor.op_map, skyscraper_presheaf_map, + category.id_comp, eq_to_hom_trans, functor.const_obj_map, category.assoc], + by_cases hV : p₀ ∈ (open_nhds.inclusion y).obj V.unop, + { have hU : p₀ ∈ unop ((open_nhds.inclusion y).op.obj U) := le_of_hom inc.unop hV, + split_ifs, + erw [category.comp_id, category.comp_id, category.comp_id, eq_to_hom_trans], + refl }, + { split_ifs with hU; + erw [category.comp_id, category.comp_id, category.comp_id, eq_comp_eq_to_hom, + eq_comp_eq_to_hom]; + exact ts.hom_ext _ _, }, + end } } + +noncomputable def skyscraper_presheaf_of_mem_closure₀_from + {y : X} (h : y ∈ (closure ({p₀} : set X))) : + S ⟶ (skyscraper_presheaf p₀ S ts).stalk y := +eq_to_hom (skyscraper_presheaf_obj_of_mem S ts (by tauto : p₀ ∈ ⊤)).symm ≫ + (skyscraper_presheaf p₀ S ts).germ (⟨y, trivial⟩ : (⊤ : opens X)) + +noncomputable lemma skyscraper_presheaf_cocone_of_mem_closure₀_is_colimit + {y : X} (h : y ∈ (closure ({p₀} : set X))) : + is_colimit (skyscraper_presheaf_cocone_of_mem_closure₀ p₀ S ts h) := +{ desc := λ c, (skyscraper_presheaf_of_mem_closure₀_from p₀ S ts h ≫ colimit.desc _ _ : S ⟶ c.X), + fac' := λ c U, + begin + dsimp, + simp only [skyscraper_presheaf_of_mem_closure₀_from, presheaf.germ, category.comp_id, + category.assoc, colimit.ι_desc, eq_to_hom_trans_assoc], + have := c.ι.naturality (hom_of_le $ (le_top : unop U ≤ _)).op, + dsimp at this, + have h' : p₀ ∈ (open_nhds.inclusion y).obj (unop U) := mem_nhds_of_closure_singleton _ _ h, + have h'' : p₀ ∈ (open_nhds.inclusion y).obj ⊤ := trivial, + split_ifs at this, + rw [category.comp_id, category.id_comp, eq_to_hom_trans, eq_eq_to_hom_comp] at this, + rw [this, eq_eq_to_hom_comp, ←category.assoc, eq_to_hom_trans, eq_to_hom_refl, + category.id_comp], + congr, + end, + uniq' := λ c f h, + begin + simp only [skyscraper_presheaf_of_mem_closure₀_from, presheaf.germ, category.assoc], + erw [colimit.ι_desc], + specialize h (op ⟨⊤, trivial⟩), + erw [←h], + simp only [skyscraper_presheaf_cocone_of_mem_closure₀_ι_app, category.assoc, + eq_to_hom_trans_assoc, eq_to_hom_refl, category.id_comp], + exact (category.id_comp _).symm, + end } + +noncomputable def skyscraper_stalk_of_mem_closure₀ {y : X} (h : y ∈ (closure ({p₀} : set X))) : + (skyscraper_presheaf p₀ S ts).stalk y ≅ S := +colimit.iso_colimit_cocone ⟨_, (skyscraper_presheaf_cocone_of_mem_closure₀_is_colimit p₀ S ts h)⟩ -/- +@[simps] def skyscraper_presheaf_cocone_of_not_mem_closure₀ + {y : X} (h : y ∉ (closure ({p₀} : set X))) : + cocone ((open_nhds.inclusion y).op ⋙ skyscraper_presheaf p₀ S ts) := +{ X := star, + ι := + { app := λ U, ts.from _, + naturality' := λ U V inc, ts.hom_ext _ _ } } -dite (p₀ ∈ U) - (λ h, (begin - have := x (𝟙 _) _, +noncomputable def skyscraper_presheaf_of_not_mem_closure₀_from + {y : X} (h : y ∉ (closure ({p₀} : set X))) : + star ⟶ (skyscraper_presheaf p₀ S ts).stalk y := +eq_to_hom (skyscraper_presheaf_obj_of_not_mem S ts $ + (mem_nhds_of_not_mem_closure_singleton p₀ h).some_spec).symm ≫ + (skyscraper_presheaf p₀ S ts).germ (⟨y, (mem_nhds_of_not_mem_closure_singleton p₀ h).some.2⟩ : + (mem_nhds_of_not_mem_closure_singleton p₀ h).some.1) + +noncomputable lemma skyscraper_presheaf_cocone_of_not_mem_closure₀_is_colimit + {y : X} (h : y ∉ (closure ({p₀} : set X))) : + is_colimit (skyscraper_presheaf_cocone_of_not_mem_closure₀ p₀ S ts h) := +{ desc := λ c, ((eq_to_hom ((skyscraper_presheaf_obj_of_not_mem _ _ + (mem_nhds_of_not_mem_closure_singleton p₀ h).some_spec).symm)) ≫ + presheaf.germ (skyscraper_presheaf p₀ S ts) + ⟨y, (mem_nhds_of_not_mem_closure_singleton p₀ h).some.2⟩ ≫ colimit.desc _ _ : star ⟶ _), + fac' := λ c U, + begin + simp only [presheaf.germ, skyscraper_presheaf_cocone_of_not_mem_closure₀_ι_app], + erw [colimit.ι_desc], + dsimp, + by_cases h' : p₀ ∈ (open_nhds.inclusion y).obj (unop U), + { have eq1 : ts.from (ite (p₀ ∈ (open_nhds.inclusion y).obj (unop U)) S star) = + eq_to_hom (if_pos h') ≫ ts.from _ := ts.hom_ext _ _, + rw [eq1, category.assoc, eq_eq_to_hom_comp], + have := c.ι.naturality (hom_of_le (le_top : U.unop ≤ ⊤)).op, + dsimp at this, + have h'' : p₀ ∈ (open_nhds.inclusion y).obj ⊤ := trivial, + split_ifs at this, + rw [category.comp_id, category.id_comp, eq_to_hom_trans, eq_eq_to_hom_comp] at this, + rw this, + clear this, + have := c.ι.naturality + (hom_of_le (le_top : (mem_nhds_of_not_mem_closure_singleton p₀ h).some ≤ ⊤)).op, dsimp at this, - end : c ⟶ S) ≫ (eq_to_hom (skyscraper_presheaf_obj_of_mem S ts h).symm)) - (λ h, ts.from c ≫ (eq_to_hom (skyscraper_presheaf_obj_of_not_mem S ts h).symm)) --/ + have h''' : p₀ ∉ (mem_nhds_of_not_mem_closure_singleton p₀ h).some.1 := + (mem_nhds_of_not_mem_closure_singleton p₀ h).some_spec, + split_ifs at this, + have eq2 : ts.from (ite (p₀ ∈ (open_nhds.inclusion y).obj ⊤) S star) = + eq_to_hom (if_pos h'') ≫ ts.from _ := ts.hom_ext _ _, + rw [category.comp_id, eq2] at this, + simp only [←this, ←category.assoc], + convert eq_whisker _ _, + { ext, refl, }, + { rw [eq_comp_eq_to_hom], + exact ts.hom_ext _ _ } }, + { have eq1 : ts.from (ite (p₀ ∈ (open_nhds.inclusion y).obj (unop U)) S star) = + eq_to_hom (if_neg h') ≫ 𝟙 star := ts.hom_ext _ _, + have eq2 : ts.from (ite (p₀ ∈ (open_nhds.inclusion y).obj ⊤) S star) = + eq_to_hom (if_pos trivial) ≫ ts.from _ := ts.hom_ext _ _, + have eq3 : ts.from (ite (p₀ ∈ (open_nhds.inclusion y).obj + (mem_nhds_of_not_mem_closure_singleton p₀ h).some) S star) = + eq_to_hom (if_neg (mem_nhds_of_not_mem_closure_singleton p₀ h).some_spec) ≫ 𝟙 star := + ts.hom_ext _ _, + rw [eq1, category.comp_id, ←category.assoc, eq_to_hom_trans], + have := c.ι.naturality (hom_of_le (inf_le_left : + (mem_nhds_of_not_mem_closure_singleton p₀ h).some ⊓ unop U ≤ + (mem_nhds_of_not_mem_closure_singleton p₀ h).some)).op, + dsimp at this, + rw [dite_comp, category.comp_id] at this, + simp_rw [category.id_comp, eq_to_hom_trans, eq3, category.comp_id, eq_to_hom_trans] at this, + + + generalize_proofs h₁ h₂ h₃, + have eq_coe : c.ι.app (op ⟨↑(h₁.some), h₃⟩) = + eq_to_hom (by { congr, ext, refl }) ≫ c.ι.app (op h₁.some) ≫ + eq_to_hom (by { congr, ext, refl }), + { symmetry, + rw [eq_eq_to_hom_comp], + have e := c.ι.naturality (eq_to_hom (by {ext, refl}) : + (⟨↑h₁.some, h₃⟩ : open_nhds y) ⟶ h₁.some).op, + dsimp at e, + split_ifs at e with temp, + { exfalso, exact (mem_nhds_of_not_mem_closure_singleton p₀ h).some_spec temp }, + rw [show ts.from (ite (p₀ ∈ (open_nhds.inclusion y).obj h₁.some) S star) = + eq_to_hom (if_neg temp) ≫ 𝟙 star, from ts.hom_ext _ _, category.comp_id, + category.comp_id, eq_to_hom_trans] at e, + rw [←e, category.assoc, eq_eq_to_hom_comp, ←category.assoc, eq_to_hom_trans, eq_to_hom_refl, + eq_to_hom_refl, category.comp_id, category.id_comp], }, + rw [eq_coe, ←this, ←category.assoc, eq_to_hom_trans, eq_eq_to_hom_comp, dite_comp], + have h₅ : p₀ ∉ (open_nhds.inclusion y).obj h₁.some := h₁.some_spec, + have h₆ : p₀ ∉ (open_nhds.inclusion y).obj (h₁.some ⊓ unop U) := λ ⟨r, _⟩, h₅ r, + clear this, + split_ifs, + { exfalso, exact h₅ (by assumption) }, + { exfalso, exact h₆ (by assumption) }, + { exfalso, exact h₅ (by assumption) }, + rw [eq_comp_eq_to_hom, eq_eq_to_hom_comp, ←category.assoc, ←category.assoc, eq_to_hom_trans, + eq_to_hom_refl, category.comp_id], + have := c.ι.naturality (hom_of_le (inf_le_right : + (mem_nhds_of_not_mem_closure_singleton p₀ h).some ⊓ unop U ≤ unop U)).op, + dsimp at this, + rw [dite_comp, category.comp_id] at this, + rw [←this, comp_dite], + clear this, + split_ifs, + rw [eq1, category.comp_id, eq_to_hom_trans, ←category.assoc, eq_to_hom_trans, eq_to_hom_refl, + category.id_comp] } + end, + uniq' := λ c f H, + begin + erw [colimit.ι_desc, ←H], + simp only [skyscraper_presheaf_cocone_of_not_mem_closure₀_ι_app], + dsimp, + have : ts.from (ite (p₀ ∈ (mem_nhds_of_not_mem_closure_singleton p₀ h).some.1) S star) = + eq_to_hom (if_neg (mem_nhds_of_not_mem_closure_singleton p₀ h).some_spec) ≫ 𝟙 star := + ts.hom_ext _ _, + erw [this, category.comp_id, ←category.assoc, eq_to_hom_trans, eq_to_hom_refl, category.id_comp], + end } + +noncomputable def skyscraper_presheaf_stalk_of_not_mem_closure₀ + {y : X} (h : y ∉ (closure ({p₀} : set X))) : + (skyscraper_presheaf p₀ S ts).stalk y ≅ star := +colimit.iso_colimit_cocone ⟨_, (skyscraper_presheaf_cocone_of_not_mem_closure₀_is_colimit _ S _ h)⟩ + end From 5d79d135752f9f526c243caee604987b832471fc Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Thu, 28 Jul 2022 17:04:07 +0100 Subject: [PATCH 34/84] add doc --- src/algebraic_geometry/skyscraper.lean | 56 ++++++++++++++++++++++++++ 1 file changed, 56 insertions(+) diff --git a/src/algebraic_geometry/skyscraper.lean b/src/algebraic_geometry/skyscraper.lean index 3399eb5da0adf..c798d990bedec 100644 --- a/src/algebraic_geometry/skyscraper.lean +++ b/src/algebraic_geometry/skyscraper.lean @@ -1,7 +1,35 @@ +/- +Copyright (c) 2022 Jujian Zhang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jujian Zhang +-/ import algebraic_geometry.sheafed_space import topology.sheaves.sheaf_condition.unique_gluing import topology.sheaves.stalks +/-! +# Skyscraper (pre)sheaves + +A skyscraper (pre)sheaf is a (pre)sheaf supported at a single point: if `p₀ ∈ X` is a specified +point, then the skyscraper (pre)sheaf `𝓕` with value `A` is defined by `U ↦ A` if `p₀ ∈ U` and +`U ↦ *` if `p₀ ∉ A` where `*` is some terminal object. + +## Main definitions + +* `skyscraper_presheaf`: the skyscraper presheaf, to define a skyscraper presheaf at point `p₀` with + value `A`, use `skypscraper_presheaf p₀ A t` where `t` is a proof that `*` is a terminal object. +* `skyscraper_sheaf`: the skyscraper presheaf satisfies the sheaf condition. + +## Main statements + +* `skyscraper_stalk_of_mem_closure₀`: if `y ∈ closure {p₀}` then the stalk of + `skyscraper_presheaf p₀ S t` at `y` is `S`. +* `skyscraper_stalk_of_not_mem_closure₀`: if `y ∉ closure {p₀}` then the stalk of + `skyscraper_presheaf p₀ S t` at `y` is `*`. + +-/ + + section open topological_space @@ -15,6 +43,11 @@ variables {X : Top.{u}} (p₀ : X) {C : Type v} [category.{w} C] (S : C) variables {star : C} (ts : is_terminal star) variable [Π (U : opens X), decidable (p₀ ∈ U)] +/- +A skyscraper presheaf is a presheaf supported at a single point: if `p₀ ∈ X` is a specified +point, then the skyscraper presheaf `𝓕` with value `A` is defined by `U ↦ A` if `p₀ ∈ U` and +`U ↦ *` if `p₀ ∉ A` where `*` is some terminal object. +-/ @[simps] def skyscraper_presheaf : presheaf C X := { obj := λ U, ite (p₀ ∈ unop U) S star, @@ -54,6 +87,11 @@ if_neg h end +/- +A skyscraper sheaf is a sheaf supported at a single point: if `p₀ ∈ X` is a specified +point, then the skyscraper sheaf `𝓕` with value `A` is defined by `U ↦ A` if `p₀ ∈ U` and +`U ↦ *` if `p₀ ∉ A` where `*` is some terminal object. +-/ def skyscraper_sheaf : sheaf C X := ⟨skyscraper_presheaf p₀ S ts, λ c U s hs x hx, ⟨dite (p₀ ∈ U) @@ -154,6 +192,9 @@ begin exact ⟨⟨⟨sᶜ, is_closed.is_open_compl⟩, hy⟩, λ r, r hp₀⟩, end +/-- +The cocone at `S` for the salk functor of `skyscraper_presheaf p₀ S t` when `y ∈ closure {p₀}` +-/ @[simps] def skyscraper_presheaf_cocone_of_mem_closure₀ {y : X} (h : y ∈ (closure ({p₀} : set X))) : cocone ((open_nhds.inclusion y).op ⋙ skyscraper_presheaf p₀ S ts) := { X := S, @@ -179,6 +220,9 @@ end exact ts.hom_ext _ _, }, end } } +/-- +The canonical map `S ⟶ (skyscraper_presheaf p₀ S t).stalk y` when `y ∈ closure {p₀}` +-/ noncomputable def skyscraper_presheaf_of_mem_closure₀_from {y : X} (h : y ∈ (closure ({p₀} : set X))) : S ⟶ (skyscraper_presheaf p₀ S ts).stalk y := @@ -215,10 +259,16 @@ noncomputable lemma skyscraper_presheaf_cocone_of_mem_closure₀_is_colimit exact (category.id_comp _).symm, end } +/-- +If `y ∈ closure {p₀}`, then the stalk of `skyscraper_presheaf p₀ S t` at `y` is `S` +-/ noncomputable def skyscraper_stalk_of_mem_closure₀ {y : X} (h : y ∈ (closure ({p₀} : set X))) : (skyscraper_presheaf p₀ S ts).stalk y ≅ S := colimit.iso_colimit_cocone ⟨_, (skyscraper_presheaf_cocone_of_mem_closure₀_is_colimit p₀ S ts h)⟩ +/-- +The cocone at `*` for the salk functor of `skyscraper_presheaf p₀ S t` when `y ∉ closure {p₀}` +-/ @[simps] def skyscraper_presheaf_cocone_of_not_mem_closure₀ {y : X} (h : y ∉ (closure ({p₀} : set X))) : cocone ((open_nhds.inclusion y).op ⋙ skyscraper_presheaf p₀ S ts) := @@ -227,6 +277,9 @@ colimit.iso_colimit_cocone ⟨_, (skyscraper_presheaf_cocone_of_mem_closure₀_i { app := λ U, ts.from _, naturality' := λ U V inc, ts.hom_ext _ _ } } +/-- +The canonical map `* ⟶ (skyscraper_presheaf p₀ S t).stalk y` when `y ∉ closure {p₀}` +-/ noncomputable def skyscraper_presheaf_of_not_mem_closure₀_from {y : X} (h : y ∉ (closure ({p₀} : set X))) : star ⟶ (skyscraper_presheaf p₀ S ts).stalk y := @@ -336,6 +389,9 @@ noncomputable lemma skyscraper_presheaf_cocone_of_not_mem_closure₀_is_colimit erw [this, category.comp_id, ←category.assoc, eq_to_hom_trans, eq_to_hom_refl, category.id_comp], end } +/-- +If `y ∉ closure {p₀}`, then the stalk of `skyscraper_presheaf p₀ S t` at `y` is `*` +-/ noncomputable def skyscraper_presheaf_stalk_of_not_mem_closure₀ {y : X} (h : y ∉ (closure ({p₀} : set X))) : (skyscraper_presheaf p₀ S ts).stalk y ≅ star := From eb70d55c1280f8e6f4ca72796ede1530b9094575 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Thu, 28 Jul 2022 18:28:07 +0100 Subject: [PATCH 35/84] add doc --- src/algebraic_geometry/skyscraper.lean | 3 +++ 1 file changed, 3 insertions(+) diff --git a/src/algebraic_geometry/skyscraper.lean b/src/algebraic_geometry/skyscraper.lean index c798d990bedec..79232baf1bb04 100644 --- a/src/algebraic_geometry/skyscraper.lean +++ b/src/algebraic_geometry/skyscraper.lean @@ -147,6 +147,9 @@ end section +-- In this section, we calculate the stalks for skyscraper presheaves. We need to restrict universe +-- level. + open topological_space open category_theory category_theory.limits open Top From b3fd04e1a896d224376bbda4f62d2640e9d4a463 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Sun, 31 Jul 2022 20:03:33 +0100 Subject: [PATCH 36/84] add an alternative formualtion --- src/algebraic_geometry/skyscraper.lean | 181 +++++++++++++++++++++++++ 1 file changed, 181 insertions(+) diff --git a/src/algebraic_geometry/skyscraper.lean b/src/algebraic_geometry/skyscraper.lean index 79232baf1bb04..673050b7aca50 100644 --- a/src/algebraic_geometry/skyscraper.lean +++ b/src/algebraic_geometry/skyscraper.lean @@ -401,3 +401,184 @@ noncomputable def skyscraper_presheaf_stalk_of_not_mem_closure₀ colimit.iso_colimit_cocone ⟨_, (skyscraper_presheaf_cocone_of_not_mem_closure₀_is_colimit _ S _ h)⟩ end + +section + +open topological_space +open category_theory category_theory.limits +open Top +open opposite +open_locale classical + +universes u v w + +variables {X : Top.{u}} (p₀ : X) {C : Type v} [category.{w} C] (S : C) +variables {star : C} (ts : is_terminal star) + +def single_point_space : Top := ⟨({p₀} : set X), infer_instance⟩ + +instance : inhabited (single_point_space p₀) := ⟨⟨p₀, rfl⟩⟩ +instance : unique (single_point_space p₀) := unique.subtype_eq p₀ + +@[simps] +noncomputable def single_point_presheaf : presheaf C (single_point_space p₀) := +{ obj := λ U, if U.unop ≠ ⊥ then S else star, + map := λ U V inc, if h : V.unop ≠ ⊥ + then eq_to_hom (if_pos $ λ r, h $ le_bot_iff.mp $ r ▸ le_of_hom inc.unop) ≫ + 𝟙 S ≫ eq_to_hom (if_pos h).symm + else ts.from _ ≫ 𝟙 star ≫ eq_to_hom (if_neg h).symm, + map_id' := λ U, + begin + split_ifs, + { rw [category.id_comp, eq_to_hom_trans, eq_to_hom_refl] }, + { rw [category.id_comp, eq_comp_eq_to_hom, category.id_comp], + exact ts.hom_ext _ _, }, + end, + map_comp' := λ U V W inc1 inc2, + begin + rw comp_dite, + by_cases hW : W.unop ≠ ⊥, + { have hV : V.unop ≠ ⊥, + { intros r, + refine hW (eq_bot_iff.mpr (r ▸ le_of_hom inc2.unop)), }, + have hU : U.unop ≠ ⊥, + { intros r, + refine hW (eq_bot_iff.mpr (r ▸ le_of_hom (inc1 ≫ inc2).unop)) }, + split_ifs, + simp only [category.id_comp, eq_to_hom_trans], }, + { split_ifs; + rw [category.id_comp, eq_comp_eq_to_hom, category.assoc, category.assoc, eq_to_hom_trans, + eq_to_hom_refl, category.comp_id]; + exact ts.hom_ext _ _, }, + end } + +-- def single_point_shef : sheaf C (single_point_space p₀) := +-- ⟨single_point_presheaf p₀ S, +-- λ c U s hs x hx, +-- ⟨dite ((⟨p₀, rfl⟩ : single_point_space p₀) ∈ U) +-- (λ h, begin +-- specialize hs _ h, +-- sorry +-- end) +-- (λ h, begin +-- dsimp, +-- have := s.downward_closed, +-- end), _, _⟩⟩ + +@[simps] +def single_point_inclusion : single_point_space p₀ ⟶ X := +{ to_fun := λ p, p.1, + continuous_to_fun := by continuity } + +@[simps] noncomputable def skyscraper_presheaf_to_pushforward : + skyscraper_presheaf p₀ S ts ⟶ (single_point_inclusion p₀) _* (single_point_presheaf p₀ S ts) := +{ app := λ U, if h : p₀ ∈ U.unop + then eq_to_hom (skyscraper_presheaf_obj_of_mem _ _ h) ≫ 𝟙 S ≫ eq_to_hom + begin + dsimp, + rw if_pos, + erw opens.map_obj _ _ U.unop.2, + change _ ≠ ⊥, + rw opens.ne_bot_iff_nonempty, + refine ⟨⟨p₀, rfl⟩, _⟩, + erw set.mem_preimage, + exact h, + end + else ts.from _ ≫ eq_to_hom + begin + dsimp, + rw if_neg, + push_neg, + rw ←opens.not_nonempty_iff_eq_bot, + rintros ⟨⟨x, hx₁⟩, hx₂⟩, + rw set.mem_singleton_iff at hx₁, + subst hx₁, + erw set.mem_preimage at hx₂, + exact h hx₂, + end, + naturality' := λ U V inc, + begin + by_cases hV : p₀ ∈ V.unop, + { have hU : p₀ ∈ U.unop := le_of_hom inc.unop hV, + have hV' : ¬(opens.map (single_point_inclusion p₀)).obj (unop V) = ⊥, + { change _ ≠ ⊥, + rw opens.ne_bot_iff_nonempty, + refine ⟨⟨p₀, rfl⟩, _⟩, + erw set.mem_preimage, + exact hV, }, + rw [comp_dite, dite_comp], + split_ifs, + { exfalso, exact hV' h, }, + { dsimp, + split_ifs, + simp only [eq_to_hom_trans, category.id_comp], }, }, + { split_ifs; + rw [←category.assoc, eq_comp_eq_to_hom]; + exact ts.hom_ext _ _ }, + end } + +@[simps] noncomputable def pushforward_to_skyscraper_presheaf : + (single_point_inclusion p₀) _* (single_point_presheaf p₀ S ts) ⟶ + skyscraper_presheaf p₀ S ts := +{ app := λ U, if h : p₀ ∈ unop U + then eq_to_hom + begin + dsimp, + rw if_pos, + erw opens.map_obj _ _ U.unop.2, + change _ ≠ ⊥, + rw opens.ne_bot_iff_nonempty, + refine ⟨⟨p₀, rfl⟩, _⟩, + erw set.mem_preimage, + exact h, + end ≫ 𝟙 S ≫ eq_to_hom (skyscraper_presheaf_obj_of_mem _ _ h).symm + else ts.from _ ≫ eq_to_hom (skyscraper_presheaf_obj_of_not_mem _ _ h).symm, + naturality' := λ U V inc, + begin + rw [comp_dite, dite_comp], + by_cases hV : p₀ ∈ V.unop, + { have hU : p₀ ∈ U.unop := le_of_hom inc.unop hV, + have hV' : ¬(opens.map (single_point_inclusion p₀)).obj (unop V) = ⊥, + { change _ ≠ ⊥, + rw opens.ne_bot_iff_nonempty, + refine ⟨⟨p₀, rfl⟩, _⟩, + erw set.mem_preimage, + exact hV, }, + split_ifs, + { exfalso, exact hV' h, }, + { dsimp, + split_ifs, + simp only [eq_to_hom_trans, category.id_comp], }, }, + { split_ifs; + rw [←category.assoc, eq_comp_eq_to_hom]; + exact ts.hom_ext _ _ }, + end } + +noncomputable example : + skyscraper_presheaf p₀ S ts ≅ (single_point_inclusion p₀) _* (single_point_presheaf p₀ S ts) := +{ hom := skyscraper_presheaf_to_pushforward p₀ S ts, + inv := pushforward_to_skyscraper_presheaf p₀ S ts, + hom_inv_id' := + begin + ext U, + dsimp, + split_ifs, + { rw [category.id_comp, eq_to_hom_trans, category.id_comp, eq_to_hom_trans, eq_to_hom_trans, + eq_to_hom_refl], }, + { rw [←category.assoc, eq_comp_eq_to_hom], + exact ts.hom_ext _ _, }, + end, + inv_hom_id' := + begin + ext U, + dsimp, + by_cases hU : p₀ ∈ U.unop, + { split_ifs; + rw [category.id_comp, eq_to_hom_trans, category.id_comp, eq_to_hom_trans, eq_to_hom_trans, + eq_to_hom_refl], }, + { split_ifs; + rw [←category.assoc, eq_comp_eq_to_hom]; + exact ts.hom_ext _ _, }, + end } + +end From 987b81f38ca9e955c71813b8bc9d620225e002f5 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Sun, 31 Jul 2022 20:05:35 +0100 Subject: [PATCH 37/84] fix --- src/algebraic_geometry/skyscraper.lean | 5 +++-- 1 file changed, 3 insertions(+), 2 deletions(-) diff --git a/src/algebraic_geometry/skyscraper.lean b/src/algebraic_geometry/skyscraper.lean index 673050b7aca50..69f7ba1b303a4 100644 --- a/src/algebraic_geometry/skyscraper.lean +++ b/src/algebraic_geometry/skyscraper.lean @@ -547,8 +547,9 @@ def single_point_inclusion : single_point_space p₀ ⟶ X := split_ifs, { exfalso, exact hV' h, }, { dsimp, - split_ifs, - simp only [eq_to_hom_trans, category.id_comp], }, }, + split_ifs; + rw [category.id_comp, eq_to_hom_trans, category.id_comp, eq_to_hom_trans, eq_to_hom_trans, + category.id_comp, eq_to_hom_trans, eq_to_hom_trans, eq_to_hom_trans], }, }, { split_ifs; rw [←category.assoc, eq_comp_eq_to_hom]; exact ts.hom_ext _ _ }, From 942d9e904541bf402d8314b360934371d24db8e7 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Sun, 31 Jul 2022 20:24:59 +0100 Subject: [PATCH 38/84] use specializes notation --- src/algebraic_geometry/skyscraper.lean | 69 ++++++-------------------- 1 file changed, 15 insertions(+), 54 deletions(-) diff --git a/src/algebraic_geometry/skyscraper.lean b/src/algebraic_geometry/skyscraper.lean index 69f7ba1b303a4..d378e064a6ffe 100644 --- a/src/algebraic_geometry/skyscraper.lean +++ b/src/algebraic_geometry/skyscraper.lean @@ -162,43 +162,19 @@ variables {star : C} (ts : is_terminal star) variable [Π (U : opens X), decidable (p₀ ∈ U)] variable [has_colimits C] -lemma mem_nhds_of_closure_singleton {y : X} (U : open_nhds y) (h : y ∈ (closure ({p₀} : set X))) : - p₀ ∈ U.1 := -begin - have := U.2, - contrapose this, - change p₀ ∈ U.1.1ᶜ at this, - change y ∈ U.1.1ᶜ, - have h1 : {p₀} ⊆ U.1.1ᶜ := set.singleton_subset_iff.mpr this, - rw ←is_closed.closure_subset_iff at h1, - { exact h1 h }, - rw is_closed_compl_iff, - exact U.1.2 -end - -lemma mem_nhds_of_not_mem_closure_singleton {y : X} (h : y ∉ (closure ({p₀} : set X))) : +lemma mem_nhds_of_not_mem_closure_singleton {y : X} (h : ¬p₀ ⤳ y) : ∃ (U : open_nhds y), p₀ ∉ U.1 := begin - have : ∃ (s : set X), is_closed s ∧ p₀ ∈ s ∧ y ∉ s, - { contrapose h, - push_neg at h ⊢, - rw mem_closure_iff, - intros t ht hy, - specialize h tᶜ (is_closed_compl_iff.mpr ht), - refine ⟨p₀, _⟩, - contrapose! h, - rw [set.mem_inter_iff, not_and_distrib, set.mem_singleton_iff, eq_self_iff_true, - not_true, or_false] at h, - exact ⟨h, λ r, r hy⟩ }, - rcases this with ⟨s, hs, hp₀, hy⟩, - resetI, - exact ⟨⟨⟨sᶜ, is_closed.is_open_compl⟩, hy⟩, λ r, r hp₀⟩, + have := (not_iff_not.mpr specializes_iff_forall_open).mp h, + push_neg at this, + rcases this with ⟨s, o, h₁, h₂⟩, + exact ⟨⟨⟨s, o⟩, h₁⟩, h₂⟩, end /-- The cocone at `S` for the salk functor of `skyscraper_presheaf p₀ S t` when `y ∈ closure {p₀}` -/ -@[simps] def skyscraper_presheaf_cocone_of_mem_closure₀ {y : X} (h : y ∈ (closure ({p₀} : set X))) : +@[simps] def skyscraper_presheaf_cocone_of_mem_closure₀ {y : X} (h : p₀ ⤳ y) : cocone ((open_nhds.inclusion y).op ⋙ skyscraper_presheaf p₀ S ts) := { X := S, ι := @@ -206,7 +182,7 @@ The cocone at `S` for the salk functor of `skyscraper_presheaf p₀ S t` when `y begin dsimp, rw if_pos, - exact mem_nhds_of_closure_singleton _ _ h, + exact h.mem_open U.unop.1.2 U.unop.2, end ≫ 𝟙 S, naturality' := λ U V inc, begin @@ -226,14 +202,12 @@ The cocone at `S` for the salk functor of `skyscraper_presheaf p₀ S t` when `y /-- The canonical map `S ⟶ (skyscraper_presheaf p₀ S t).stalk y` when `y ∈ closure {p₀}` -/ -noncomputable def skyscraper_presheaf_of_mem_closure₀_from - {y : X} (h : y ∈ (closure ({p₀} : set X))) : +noncomputable def skyscraper_presheaf_of_mem_closure₀_from {y : X} (h : p₀ ⤳ y) : S ⟶ (skyscraper_presheaf p₀ S ts).stalk y := eq_to_hom (skyscraper_presheaf_obj_of_mem S ts (by tauto : p₀ ∈ ⊤)).symm ≫ (skyscraper_presheaf p₀ S ts).germ (⟨y, trivial⟩ : (⊤ : opens X)) -noncomputable lemma skyscraper_presheaf_cocone_of_mem_closure₀_is_colimit - {y : X} (h : y ∈ (closure ({p₀} : set X))) : +noncomputable lemma skyscraper_presheaf_cocone_of_mem_closure₀_is_colimit {y : X} (h : p₀ ⤳ y) : is_colimit (skyscraper_presheaf_cocone_of_mem_closure₀ p₀ S ts h) := { desc := λ c, (skyscraper_presheaf_of_mem_closure₀_from p₀ S ts h ≫ colimit.desc _ _ : S ⟶ c.X), fac' := λ c U, @@ -243,7 +217,7 @@ noncomputable lemma skyscraper_presheaf_cocone_of_mem_closure₀_is_colimit category.assoc, colimit.ι_desc, eq_to_hom_trans_assoc], have := c.ι.naturality (hom_of_le $ (le_top : unop U ≤ _)).op, dsimp at this, - have h' : p₀ ∈ (open_nhds.inclusion y).obj (unop U) := mem_nhds_of_closure_singleton _ _ h, + have h' : p₀ ∈ (open_nhds.inclusion y).obj (unop U) := h.mem_open U.unop.1.2 U.unop.2, have h'' : p₀ ∈ (open_nhds.inclusion y).obj ⊤ := trivial, split_ifs at this, rw [category.comp_id, category.id_comp, eq_to_hom_trans, eq_eq_to_hom_comp] at this, @@ -265,7 +239,7 @@ noncomputable lemma skyscraper_presheaf_cocone_of_mem_closure₀_is_colimit /-- If `y ∈ closure {p₀}`, then the stalk of `skyscraper_presheaf p₀ S t` at `y` is `S` -/ -noncomputable def skyscraper_stalk_of_mem_closure₀ {y : X} (h : y ∈ (closure ({p₀} : set X))) : +noncomputable def skyscraper_stalk_of_mem_closure₀ {y : X} (h : p₀ ⤳ y) : (skyscraper_presheaf p₀ S ts).stalk y ≅ S := colimit.iso_colimit_cocone ⟨_, (skyscraper_presheaf_cocone_of_mem_closure₀_is_colimit p₀ S ts h)⟩ @@ -273,7 +247,7 @@ colimit.iso_colimit_cocone ⟨_, (skyscraper_presheaf_cocone_of_mem_closure₀_i The cocone at `*` for the salk functor of `skyscraper_presheaf p₀ S t` when `y ∉ closure {p₀}` -/ @[simps] def skyscraper_presheaf_cocone_of_not_mem_closure₀ - {y : X} (h : y ∉ (closure ({p₀} : set X))) : + {y : X} (h : ¬p₀ ⤳ y) : cocone ((open_nhds.inclusion y).op ⋙ skyscraper_presheaf p₀ S ts) := { X := star, ι := @@ -284,7 +258,7 @@ The cocone at `*` for the salk functor of `skyscraper_presheaf p₀ S t` when `y The canonical map `* ⟶ (skyscraper_presheaf p₀ S t).stalk y` when `y ∉ closure {p₀}` -/ noncomputable def skyscraper_presheaf_of_not_mem_closure₀_from - {y : X} (h : y ∉ (closure ({p₀} : set X))) : + {y : X} (h : ¬p₀ ⤳ y) : star ⟶ (skyscraper_presheaf p₀ S ts).stalk y := eq_to_hom (skyscraper_presheaf_obj_of_not_mem S ts $ (mem_nhds_of_not_mem_closure_singleton p₀ h).some_spec).symm ≫ @@ -292,7 +266,7 @@ eq_to_hom (skyscraper_presheaf_obj_of_not_mem S ts $ (mem_nhds_of_not_mem_closure_singleton p₀ h).some.1) noncomputable lemma skyscraper_presheaf_cocone_of_not_mem_closure₀_is_colimit - {y : X} (h : y ∉ (closure ({p₀} : set X))) : + {y : X} (h : ¬p₀ ⤳ y) : is_colimit (skyscraper_presheaf_cocone_of_not_mem_closure₀ p₀ S ts h) := { desc := λ c, ((eq_to_hom ((skyscraper_presheaf_obj_of_not_mem _ _ (mem_nhds_of_not_mem_closure_singleton p₀ h).some_spec).symm)) ≫ @@ -396,7 +370,7 @@ noncomputable lemma skyscraper_presheaf_cocone_of_not_mem_closure₀_is_colimit If `y ∉ closure {p₀}`, then the stalk of `skyscraper_presheaf p₀ S t` at `y` is `*` -/ noncomputable def skyscraper_presheaf_stalk_of_not_mem_closure₀ - {y : X} (h : y ∉ (closure ({p₀} : set X))) : + {y : X} (h : ¬p₀ ⤳ y) : (skyscraper_presheaf p₀ S ts).stalk y ≅ star := colimit.iso_colimit_cocone ⟨_, (skyscraper_presheaf_cocone_of_not_mem_closure₀_is_colimit _ S _ h)⟩ @@ -452,19 +426,6 @@ noncomputable def single_point_presheaf : presheaf C (single_point_space p₀) : exact ts.hom_ext _ _, }, end } --- def single_point_shef : sheaf C (single_point_space p₀) := --- ⟨single_point_presheaf p₀ S, --- λ c U s hs x hx, --- ⟨dite ((⟨p₀, rfl⟩ : single_point_space p₀) ∈ U) --- (λ h, begin --- specialize hs _ h, --- sorry --- end) --- (λ h, begin --- dsimp, --- have := s.downward_closed, --- end), _, _⟩⟩ - @[simps] def single_point_inclusion : single_point_space p₀ ⟶ X := { to_fun := λ p, p.1, From 0a3f9d07d7a24719c5e3ac5b1107417a21ca19b2 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Mon, 1 Aug 2022 16:50:32 +0100 Subject: [PATCH 39/84] add documentation --- src/algebraic_geometry/skyscraper.lean | 33 +++++++++++++++++++------- 1 file changed, 24 insertions(+), 9 deletions(-) diff --git a/src/algebraic_geometry/skyscraper.lean b/src/algebraic_geometry/skyscraper.lean index d378e064a6ffe..75aba61767c1d 100644 --- a/src/algebraic_geometry/skyscraper.lean +++ b/src/algebraic_geometry/skyscraper.lean @@ -109,10 +109,7 @@ def skyscraper_sheaf : sheaf C X := split_ifs, rw [category.id_comp, eq_to_hom_trans, category.assoc, eq_to_hom_trans], generalize_proofs h₅, - have := hx _ _ h₂.some_spec.1 h rfl, - work_on_goal 2 { exact h₁.some ⊓ V }, - work_on_goal 2 { exact hom_of_le inf_le_left }, - work_on_goal 2 { exact hom_of_le inf_le_right }, + have := hx (hom_of_le inf_le_left) (hom_of_le inf_le_right) h₂.some_spec.1 h rfl, dsimp at this, have hV' : p₀ ∈ h₁.some ⊓ V := ⟨h₂.some_spec.2, hV⟩, split_ifs at this, @@ -141,8 +138,6 @@ def skyscraper_sheaf : sheaf C X := exact (ts.hom_ext _ _), } end⟩⟩ -example : true := trivial - end section @@ -378,6 +373,11 @@ end section +/-! +Skyscraper sheaf alternatively can be defined as the pushforward of a sheaf on a single point space +`{p₀}` defined by `∅ ↦ *` and `{p₀} ↦ S` under the inclusion `{p₀} ↪ X`. +-/ + open topological_space open category_theory category_theory.limits open Top @@ -389,11 +389,15 @@ universes u v w variables {X : Top.{u}} (p₀ : X) {C : Type v} [category.{w} C] (S : C) variables {star : C} (ts : is_terminal star) +/--The topological space with on a point `{p₀}`. Hence the only open sets are ∅ and ⊤.-/ def single_point_space : Top := ⟨({p₀} : set X), infer_instance⟩ instance : inhabited (single_point_space p₀) := ⟨⟨p₀, rfl⟩⟩ instance : unique (single_point_space p₀) := unique.subtype_eq p₀ +/-- +The presheaf on a single point space `{p₀}` defined by `∅ ↦ *` and `{p₀} ↦ S` +-/ @[simps] noncomputable def single_point_presheaf : presheaf C (single_point_space p₀) := { obj := λ U, if U.unop ≠ ⊥ then S else star, @@ -426,11 +430,16 @@ noncomputable def single_point_presheaf : presheaf C (single_point_space p₀) : exact ts.hom_ext _ _, }, end } -@[simps] -def single_point_inclusion : single_point_space p₀ ⟶ X := +/-- +The trivial inclusion `{p₀} ↪ X`. +-/ +@[simps] def single_point_inclusion : single_point_space p₀ ⟶ X := { to_fun := λ p, p.1, continuous_to_fun := by continuity } +/-- +The morphism from skyscraper presheaf to pushforward sheaf +-/ @[simps] noncomputable def skyscraper_presheaf_to_pushforward : skyscraper_presheaf p₀ S ts ⟶ (single_point_inclusion p₀) _* (single_point_presheaf p₀ S ts) := { app := λ U, if h : p₀ ∈ U.unop @@ -478,6 +487,9 @@ def single_point_inclusion : single_point_space p₀ ⟶ X := exact ts.hom_ext _ _ }, end } +/-- +The morphism from pushforward sheaf to skyscraper presheaf +-/ @[simps] noncomputable def pushforward_to_skyscraper_presheaf : (single_point_inclusion p₀) _* (single_point_presheaf p₀ S ts) ⟶ skyscraper_presheaf p₀ S ts := @@ -516,7 +528,10 @@ def single_point_inclusion : single_point_space p₀ ⟶ X := exact ts.hom_ext _ _ }, end } -noncomputable example : +/-- +Skyscraper presheaf is isomorphic to pushforward of sheaf on single point. +-/ +noncomputable def skyscraper_presheaf_as_pushforward : skyscraper_presheaf p₀ S ts ≅ (single_point_inclusion p₀) _* (single_point_presheaf p₀ S ts) := { hom := skyscraper_presheaf_to_pushforward p₀ S ts, inv := pushforward_to_skyscraper_presheaf p₀ S ts, From b16682a6215403869b19c50d7e9cd2eb938ad06f Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Tue, 2 Aug 2022 14:29:45 +0100 Subject: [PATCH 40/84] adjoint: stalk functor and skyscraper functor --- src/algebraic_geometry/skyscraper.lean | 233 ++++++++++++++++++++++++- 1 file changed, 228 insertions(+), 5 deletions(-) diff --git a/src/algebraic_geometry/skyscraper.lean b/src/algebraic_geometry/skyscraper.lean index 75aba61767c1d..c30c5ac45f1f1 100644 --- a/src/algebraic_geometry/skyscraper.lean +++ b/src/algebraic_geometry/skyscraper.lean @@ -112,11 +112,11 @@ def skyscraper_sheaf : sheaf C X := have := hx (hom_of_le inf_le_left) (hom_of_le inf_le_right) h₂.some_spec.1 h rfl, dsimp at this, have hV' : p₀ ∈ h₁.some ⊓ V := ⟨h₂.some_spec.2, hV⟩, - split_ifs at this, + split_ifs at this; rw [category.id_comp, eq_to_hom_trans, eq_to_hom_trans, eq_comp_eq_to_hom, - category.assoc, eq_to_hom_trans] at this, - generalize_proofs h₆ at this, - rw [this, eq_comp_eq_to_hom], }, + category.assoc, eq_to_hom_trans] at this; + generalize_proofs h₆ at this; + rw [this, category.assoc, eq_to_hom_trans, eq_to_hom_refl, category.comp_id] }, { dsimp, split_ifs, rw [←category.assoc, eq_comp_eq_to_hom], @@ -234,7 +234,8 @@ noncomputable lemma skyscraper_presheaf_cocone_of_mem_closure₀_is_colimit {y : /-- If `y ∈ closure {p₀}`, then the stalk of `skyscraper_presheaf p₀ S t` at `y` is `S` -/ -noncomputable def skyscraper_stalk_of_mem_closure₀ {y : X} (h : p₀ ⤳ y) : +@[reducible] +noncomputable def skyscraper_presheaf_stalk_of_mem_closure₀ {y : X} (h : p₀ ⤳ y) : (skyscraper_presheaf p₀ S ts).stalk y ≅ S := colimit.iso_colimit_cocone ⟨_, (skyscraper_presheaf_cocone_of_mem_closure₀_is_colimit p₀ S ts h)⟩ @@ -364,6 +365,7 @@ noncomputable lemma skyscraper_presheaf_cocone_of_not_mem_closure₀_is_colimit /-- If `y ∉ closure {p₀}`, then the stalk of `skyscraper_presheaf p₀ S t` at `y` is `*` -/ +@[reducible] noncomputable def skyscraper_presheaf_stalk_of_not_mem_closure₀ {y : X} (h : ¬p₀ ⤳ y) : (skyscraper_presheaf p₀ S ts).stalk y ≅ star := @@ -559,3 +561,224 @@ noncomputable def skyscraper_presheaf_as_pushforward : end } end + +section adjoints + + +open topological_space +open category_theory category_theory.limits +open Top +open opposite + +universes u v + +variables {X : Top.{u}} (p₀ : X) {C : Type v} [category.{u} C] +variables {star : C} (ts : is_terminal star) +variable [Π (U : opens X), decidable (p₀ ∈ U)] + +@[simps] +def skyscraper_presheaf_functor : C ⥤ presheaf C X := +{ obj := λ S, skyscraper_presheaf p₀ S ts, + map := λ x y f, + { app := λ U, if h : p₀ ∈ U.unop + then eq_to_hom (skyscraper_presheaf_obj_of_mem _ _ h) ≫ f ≫ + eq_to_hom (skyscraper_presheaf_obj_of_mem _ _ h).symm + else ts.from _ ≫ eq_to_hom (skyscraper_presheaf_obj_of_not_mem _ _ h).symm, + naturality' := λ U V inc, + begin + dsimp, + simp_rw [category.id_comp, eq_to_hom_trans], + by_cases hV : p₀ ∈ V.unop, + { have hU : p₀ ∈ U.unop := le_of_hom inc.unop hV, + split_ifs, + rw [←category.assoc, eq_to_hom_trans, category.assoc, category.assoc, eq_to_hom_trans], + refl, }, + { split_ifs; + rw [←category.assoc, eq_comp_eq_to_hom]; + exact ts.hom_ext _ _ } + end }, + map_id' := λ c, + begin + ext U, + dsimp, + split_ifs, + { simp, }, + { rw [eq_comp_eq_to_hom], + exact ts.hom_ext _ _ }, + end, + map_comp' := λ x y z f g, + begin + ext U, + dsimp, + split_ifs, + { simp }, + { rw [eq_comp_eq_to_hom], + exact ts.hom_ext _ _ }, + end } + + +variable [has_colimits C] + +@[simps] +noncomputable def from_stalk_to_to_skyscraper_presheaf {𝓕 : presheaf C X} {c : C} (f : 𝓕.stalk p₀ ⟶ c) : + 𝓕 ⟶ skyscraper_presheaf p₀ c ts := +{ app := λ U, if h : p₀ ∈ U.unop + then 𝓕.germ ⟨p₀, h⟩ ≫ f ≫ eq_to_hom (skyscraper_presheaf_obj_of_mem _ _ h).symm + else ts.from _ ≫ eq_to_hom (skyscraper_presheaf_obj_of_not_mem _ _ h).symm, + naturality' := λ U V inc, + begin + dsimp, + by_cases hV : p₀ ∈ V.unop, + { have hU : p₀ ∈ U.unop := le_of_hom inc.unop hV, + split_ifs, + erw [←category.assoc, 𝓕.germ_res inc.unop, category.id_comp, eq_to_hom_trans, + category.assoc, category.assoc, eq_to_hom_trans], + refl, }, + { split_ifs, + rw [←category.assoc, eq_comp_eq_to_hom, category.assoc, category.assoc, eq_to_hom_trans, + eq_to_hom_refl, category.comp_id], + exact ts.hom_ext _ _ }, + end } + +@[reducible] +noncomputable def to_skyscraper_presheaf_to_from_stalk {𝓕 : presheaf C X} {c : C} (f : 𝓕 ⟶ skyscraper_presheaf p₀ c ts) : + 𝓕.stalk p₀ ⟶ c := +let CC : cocone ((open_nhds.inclusion p₀).op ⋙ 𝓕) := +{ X := c, + ι := + { app := λ U, f.app (op U.unop.1) ≫ eq_to_hom + begin + dsimp, + rw if_pos, + exact U.unop.2, + end, + naturality' := λ U V inc, + begin + dsimp, + rw [category.comp_id, ←category.assoc, eq_comp_eq_to_hom, category.assoc, eq_to_hom_trans], + generalize_proofs h, + erw f.naturality, + dsimp, + have hV : p₀ ∈ (open_nhds.inclusion p₀).obj V.unop := V.unop.2, + split_ifs, + rw [category.id_comp, eq_to_hom_trans, eq_comp_eq_to_hom, category.assoc, eq_to_hom_trans, + eq_to_hom_refl, category.comp_id], + refl, + end} } in +colimit.desc _ CC + +lemma from_stalk_to_to_skyscraper_presheaf_to_skyscraper_presheaf_to_from_stalk + {𝓕 : presheaf C X} {c : C} (f : 𝓕.stalk p₀ ⟶ c) : +to_skyscraper_presheaf_to_from_stalk p₀ ts (from_stalk_to_to_skyscraper_presheaf p₀ ts f) = f := +begin + ext U, + dsimp, + simp only [subtype.val_eq_coe, unop_op, colimit.ι_desc, from_stalk_to_to_skyscraper_presheaf_app], + split_ifs, + { rw [category.assoc, category.assoc, eq_to_hom_trans, eq_to_hom_refl, category.comp_id], + congr, + dunfold presheaf.germ, + dsimp, + congr, + convert op_unop _, + ext, + refl, }, + { exfalso, exact h U.unop.2, } +end + +lemma to_skyscraper_presheaf_to_from_stalk_from_stalk_to_to_skyscraper_presheaf + {𝓕 : presheaf C X} {c : C} (f : 𝓕 ⟶ skyscraper_presheaf p₀ c ts) : +from_stalk_to_to_skyscraper_presheaf p₀ ts (to_skyscraper_presheaf_to_from_stalk p₀ ts f) = f := +begin + ext U, + dsimp, + split_ifs, + { rw [to_skyscraper_presheaf_to_from_stalk], + dsimp, + dunfold presheaf.germ, + rw [←category.assoc, colimit.ι_desc], + dsimp, + rw [category.assoc, eq_to_hom_trans, eq_to_hom_refl, category.comp_id], }, + { rw [eq_comp_eq_to_hom], + exact ts.hom_ext _ _ , } +end + +@[simps] +noncomputable def stalk_skyscraper_presheaf_adj_unit : + 𝟭 (presheaf C X) ⟶ presheaf.stalk_functor C p₀ ⋙ skyscraper_presheaf_functor p₀ ts := +{ app := λ 𝓕, + { app := λ U, if h : p₀ ∈ U.unop + then 𝓕.germ ⟨p₀, h⟩ ≫ eq_to_hom (if_pos h).symm + else ts.from _ ≫ eq_to_hom (if_neg h).symm, + naturality' := λ U V inc, + begin + dsimp, + by_cases hV : p₀ ∈ V.unop, + { have hU : p₀ ∈ U.unop := le_of_hom inc.unop hV, + split_ifs, + erw [←category.assoc, 𝓕.germ_res inc.unop, category.id_comp, eq_to_hom_trans, category.assoc, eq_to_hom_trans], + congr, }, + { split_ifs, + rw [←category.assoc, eq_comp_eq_to_hom], + exact ts.hom_ext _ _, }, + end }, + naturality' := λ 𝓕 𝓖 f, + begin + ext U, + dsimp, + split_ifs, + { rw [←category.assoc, eq_comp_eq_to_hom], + simp only [category.assoc, eq_to_hom_trans, eq_to_hom_refl, category.comp_id], + rw [←category.assoc _ _ ((presheaf.stalk_functor C p₀).map f), eq_to_hom_trans, eq_to_hom_refl, + category.id_comp], + erw [colimit.ι_map], + congr, }, + { rw [←category.assoc, eq_comp_eq_to_hom], + exact ts.hom_ext _ _ }, + end } + +@[simps] +noncomputable def stalk_skyscraper_presheaf_adj_counit : + skyscraper_presheaf_functor p₀ ts ⋙ presheaf.stalk_functor C p₀ ⟶ 𝟭 C := +{ app := λ c, (skyscraper_presheaf_stalk_of_mem_closure₀ p₀ c ts (specializes_rfl : p₀ ⤳ p₀)).hom, + naturality' := λ x y f, + begin + ext U, + dsimp, + simp only [colimit.iso_colimit_cocone_ι_hom_assoc, skyscraper_presheaf_cocone_of_mem_closure₀_ι_app, category.assoc], + erw [category.id_comp, ←category.assoc, colimit.ι_map], + dsimp, + split_ifs, + { rw [category.assoc, skyscraper_presheaf_stalk_of_mem_closure₀, colimit.iso_colimit_cocone_ι_hom], + dsimp, + simpa, }, + { exfalso, exact h U.unop.2 } + end } + +noncomputable example : presheaf.stalk_functor C p₀ ⊣ skyscraper_presheaf_functor p₀ ts := +{ hom_equiv := λ 𝓕 c, ⟨from_stalk_to_to_skyscraper_presheaf p₀ ts, + to_skyscraper_presheaf_to_from_stalk p₀ ts, + from_stalk_to_to_skyscraper_presheaf_to_skyscraper_presheaf_to_from_stalk p₀ ts, + to_skyscraper_presheaf_to_from_stalk_from_stalk_to_to_skyscraper_presheaf p₀ ts⟩, + unit := stalk_skyscraper_presheaf_adj_unit p₀ ts, + counit := stalk_skyscraper_presheaf_adj_counit p₀ ts, + hom_equiv_unit' := λ 𝓕 𝓖 f, + begin + ext U, + dsimp, + split_ifs, + { simp, }, + { rw [eq_comp_eq_to_hom], + exact ts.hom_ext _ _ }, + end, + hom_equiv_counit' := λ 𝓕 c g, + begin + ext U, + dsimp, + erw [colimit.ι_desc, ←category.assoc, colimit.ι_map, category.assoc, colimit.iso_colimit_cocone_ι_hom], + dsimp, + rw [category.comp_id], + refl, + end } + +end adjoints From 01d6a97a57b20dd4b8cd76e58b9a7de5e9bae2b1 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Tue, 2 Aug 2022 21:54:37 +0100 Subject: [PATCH 41/84] save for now --- src/algebraic_geometry/skyscraper.lean | 35 ++++++++++++++++++- .../preadditive/injective.lean | 23 ++++++++++++ 2 files changed, 57 insertions(+), 1 deletion(-) diff --git a/src/algebraic_geometry/skyscraper.lean b/src/algebraic_geometry/skyscraper.lean index c30c5ac45f1f1..27d9f74c3f808 100644 --- a/src/algebraic_geometry/skyscraper.lean +++ b/src/algebraic_geometry/skyscraper.lean @@ -6,6 +6,7 @@ Authors: Jujian Zhang import algebraic_geometry.sheafed_space import topology.sheaves.sheaf_condition.unique_gluing import topology.sheaves.stalks +import category_theory.preadditive.injective /-! # Skyscraper (pre)sheaves @@ -616,6 +617,22 @@ def skyscraper_presheaf_functor : C ⥤ presheaf C X := exact ts.hom_ext _ _ }, end } +example : true := trivial + +@[simps] +def skyscraper_sheaf_functor : C ⥤ sheaf C X := +{ obj := λ S, skyscraper_sheaf p₀ S ts, + map := λ x y f, ⟨(skyscraper_presheaf_functor p₀ ts).map f⟩, + map_id' := λ c, + begin + ext1, + exact (skyscraper_presheaf_functor p₀ ts).map_id c, + end, + map_comp' := λ x y z f g, + begin + ext1, + exact (skyscraper_presheaf_functor p₀ ts).map_comp f g, + end } variable [has_colimits C] @@ -755,7 +772,7 @@ noncomputable def stalk_skyscraper_presheaf_adj_counit : { exfalso, exact h U.unop.2 } end } -noncomputable example : presheaf.stalk_functor C p₀ ⊣ skyscraper_presheaf_functor p₀ ts := +noncomputable def stalk_skyscraper_presheaf_adj : presheaf.stalk_functor C p₀ ⊣ skyscraper_presheaf_functor p₀ ts := { hom_equiv := λ 𝓕 c, ⟨from_stalk_to_to_skyscraper_presheaf p₀ ts, to_skyscraper_presheaf_to_from_stalk p₀ ts, from_stalk_to_to_skyscraper_presheaf_to_skyscraper_presheaf_to_from_stalk p₀ ts, @@ -781,4 +798,20 @@ noncomputable example : presheaf.stalk_functor C p₀ ⊣ skyscraper_presheaf_fu refl, end } +example : true := trivial + +example : sheaf.forget C X ⋙ presheaf.stalk_functor _ p₀ ⊣ skyscraper_sheaf_functor p₀ ts := +{ hom_equiv := λ 𝓕, _, + unit := _, + counit := _, + hom_equiv_unit' := _, + hom_equiv_counit' := _ } + +section + +-- lemma skyscraper_presheaf_injective (S : C) [injective S] : injective (skyscraper_presheaf p₀ S ts) := +-- injective.injective_of_adjoint (stalk_skyscraper_presheaf_adj p₀ ts) + +end + end adjoints diff --git a/src/category_theory/preadditive/injective.lean b/src/category_theory/preadditive/injective.lean index 69c7841196cbc..0d7a32bcbcc80 100644 --- a/src/category_theory/preadditive/injective.lean +++ b/src/category_theory/preadditive/injective.lean @@ -6,6 +6,7 @@ Authors: Jujian Zhang, Kevin Buzzard import algebra.homology.exact import category_theory.types +import category_theory.functor.epi_mono import category_theory.preadditive.projective import category_theory.limits.shapes.biproducts @@ -174,6 +175,28 @@ begin exact functor.preserves_epimorphisms.iso_iff (coyoneda.obj_op_op _) end +section adjunction + +open category_theory.functor + +universes v₁ v₂ u₁ u₂ + +variables {𝓐 : Type u₁} {𝓑 : Type u₂} [category.{v₁ u₁} 𝓐] [category.{v₂ u₂} 𝓑] +variables {L : 𝓐 ⥤ 𝓑} {R : 𝓑 ⥤ 𝓐} (adj : L ⊣ R) [preserves_monomorphisms L] + +include adj +def injective_of_adjoint {J : 𝓑} [injective J] : injective $ R.obj J := +{ factors := λ A A' g f im, + begin + resetI, + haveI : mono (L.map f) := functor.map_mono L _, + refine ⟨adj.hom_equiv _ _ (factor_thru ((adj.hom_equiv A J).symm g) (L.map f)), _⟩, + apply_fun (adj.hom_equiv _ _).symm using equiv.injective, + simp, + end } + +end adjunction + section enough_injectives variable [enough_injectives C] From 6e6a07709e3c5bb42c9d861ee21a1368be0a0b0b Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Wed, 3 Aug 2022 00:18:26 +0100 Subject: [PATCH 42/84] skyscraper sheaf in Ab is injective when S is inj --- src/algebraic_geometry/skyscraper.lean | 56 ++++++++++++++++++++------ src/topology/sheaves/sheaf.lean | 34 ++++++++++++++++ src/topology/sheaves/stalks.lean | 24 +++++++++++ 3 files changed, 102 insertions(+), 12 deletions(-) diff --git a/src/algebraic_geometry/skyscraper.lean b/src/algebraic_geometry/skyscraper.lean index 27d9f74c3f808..4d4cb4c8e5d1c 100644 --- a/src/algebraic_geometry/skyscraper.lean +++ b/src/algebraic_geometry/skyscraper.lean @@ -7,6 +7,7 @@ import algebraic_geometry.sheafed_space import topology.sheaves.sheaf_condition.unique_gluing import topology.sheaves.stalks import category_theory.preadditive.injective +import algebra.category.Group.abelian /-! # Skyscraper (pre)sheaves @@ -565,7 +566,6 @@ end section adjoints - open topological_space open category_theory category_theory.limits open Top @@ -800,18 +800,50 @@ noncomputable def stalk_skyscraper_presheaf_adj : presheaf.stalk_functor C p₀ example : true := trivial -example : sheaf.forget C X ⋙ presheaf.stalk_functor _ p₀ ⊣ skyscraper_sheaf_functor p₀ ts := -{ hom_equiv := λ 𝓕, _, - unit := _, - counit := _, - hom_equiv_unit' := _, - hom_equiv_counit' := _ } +noncomputable def stalk_skyscraper_sheaf_adj : sheaf.forget C X ⋙ presheaf.stalk_functor _ p₀ ⊣ skyscraper_sheaf_functor p₀ ts := +{ hom_equiv := λ 𝓕 c, + ⟨λ f, ⟨from_stalk_to_to_skyscraper_presheaf p₀ ts f⟩, + λ g, to_skyscraper_presheaf_to_from_stalk p₀ ts g.1, + λ f, from_stalk_to_to_skyscraper_presheaf_to_skyscraper_presheaf_to_from_stalk + p₀ ts f, + begin + intros g, + ext1, + exact to_skyscraper_presheaf_to_from_stalk_from_stalk_to_to_skyscraper_presheaf p₀ ts g.1, + end⟩, + unit := + { app := λ 𝓕, ⟨(stalk_skyscraper_presheaf_adj_unit p₀ ts).app 𝓕.1⟩, + naturality' := λ 𝓐 𝓑 ⟨f⟩, + begin + ext1, + dsimp, + exact (stalk_skyscraper_presheaf_adj_unit p₀ ts).naturality f, + end }, + counit := stalk_skyscraper_presheaf_adj_counit p₀ ts, + hom_equiv_unit' := + begin + intros 𝓐 c f, + ext1, + exact (stalk_skyscraper_presheaf_adj p₀ ts).hom_equiv_unit, + end, + hom_equiv_counit' := λ 𝓐 c f, (stalk_skyscraper_presheaf_adj p₀ ts).hom_equiv_counit } -section +end adjoints --- lemma skyscraper_presheaf_injective (S : C) [injective S] : injective (skyscraper_presheaf p₀ S ts) := --- injective.injective_of_adjoint (stalk_skyscraper_presheaf_adj p₀ ts) +section injective -end +open_locale zero_object +open topological_space +open category_theory category_theory.limits +open Top +open opposite -end adjoints +universe u +variables {X : Top.{u}} (p₀ : X) [Π (U : opens X), decidable (p₀ ∈ U)] + +lemma skyscraper_presheaf_in_Ab_injective (S : Ab.{u}) [injective S] : + injective (skyscraper_sheaf p₀ S (is_zero.is_terminal (is_zero_zero _) : is_terminal (0 : Ab))) := +injective.injective_of_adjoint + (stalk_skyscraper_sheaf_adj p₀ (is_zero.is_terminal (is_zero_zero _) : is_terminal (0 : Ab))) + +end injective diff --git a/src/topology/sheaves/sheaf.lean b/src/topology/sheaves/sheaf.lean index 794e94d5c2aa8..ab487f6660718 100644 --- a/src/topology/sheaves/sheaf.lean +++ b/src/topology/sheaves/sheaf.lean @@ -8,6 +8,7 @@ import category_theory.full_subcategory import category_theory.limits.unit import category_theory.sites.sheaf import category_theory.sites.spaces +import category_theory.sites.sheafification /-! # Sheaves @@ -138,6 +139,39 @@ lemma id_app (F : sheaf C X) (t) : (𝟙 F : F ⟶ F).1.app t = 𝟙 _ := rfl lemma comp_app {F G H : sheaf C X} (f : F ⟶ G) (g : G ⟶ H) (t) : (f ≫ g).1.app t = f.1.app t ≫ g.1.app t := rfl +instance presheaf_mono_of_mono {D : Type u} [category.{w u} D] [concrete_category.{w} D] + [preserves_limits (category_theory.forget D)] + [Π (U : opens X), + preserves_colimits_of_shape ((opens.grothendieck_topology X).cover U)ᵒᵖ (category_theory.forget D)] + [reflects_isomorphisms (category_theory.forget D)] + [has_colimits D] + [∀ (P : (opens X)ᵒᵖ ⥤ D) (U : opens X) (S : (opens.grothendieck_topology X).cover U), has_multiequalizer (S.index P)] + {F G : sheaf D X} (f : F ⟶ G) [mono f] : mono f.1 := +{ right_cancellation := λ P g h eq₀, + begin + set P_plus : sheaf D X := (presheaf_to_Sheaf (opens.grothendieck_topology X) D).obj P, + set g_plus : P_plus ⟶ F := ⟨grothendieck_topology.sheafify_lift (opens.grothendieck_topology X) g F.2⟩ with g_plus_def, + set h_plus : P_plus ⟶ F := ⟨grothendieck_topology.sheafify_lift (opens.grothendieck_topology X) h F.2⟩ with h_plus_def, + have eq₁ : g_plus ≫ f = h_plus ≫ f, + { ext1, + dsimp, + apply grothendieck_topology.sheafify_hom_ext (opens.grothendieck_topology X), + { exact G.2 }, + { simp [eq₀] }, }, + have eq₂ : g_plus = h_plus, + { rwa cancel_mono at eq₁, }, + rw [g_plus_def, h_plus_def, Sheaf.hom.ext_iff] at eq₂, + dsimp at eq₂, + have := grothendieck_topology.to_sheafify_sheafify_lift (opens.grothendieck_topology X) g F.2, + rw ←this, + clear this, + have := grothendieck_topology.to_sheafify_sheafify_lift (opens.grothendieck_topology X) h F.2, + rw ←this, + clear this, + rw eq₂, + end } + + end sheaf end Top diff --git a/src/topology/sheaves/stalks.lean b/src/topology/sheaves/stalks.lean index 5ff3e104a9df6..2522088d9c1b7 100644 --- a/src/topology/sheaves/stalks.lean +++ b/src/topology/sheaves/stalks.lean @@ -9,9 +9,13 @@ import topology.sheaves.sheaf_condition.unique_gluing import category_theory.limits.types import category_theory.limits.preserves.filtered import category_theory.limits.final +import category_theory.adjunction.evaluation import topology.sober import tactic.elementwise import algebra.category.Ring +import algebra.category.Group.colimits +import algebra.category.Group.limits +import algebra.category.Group.filtered_colimits /-! # Stalks @@ -422,6 +426,26 @@ lemma app_injective_iff_stalk_functor_map_injective {F : sheaf C X} ⟨λ h U, app_injective_of_stalk_functor_map_injective f U (λ x, h x.1), stalk_functor_map_injective_of_app_injective f⟩ +/-- +This works for any concrete category where monomorphism = injective functions. +-/ +instance (x : X) : functor.preserves_monomorphisms (sheaf.forget Ab.{v} X ⋙ stalk_functor Ab.{v} x) := +begin + refine ⟨λ 𝓐 𝓑 f im, _⟩, + apply concrete_category.mono_of_injective, + dsimp, + haveI : preserves_limits (forget Ab.{v}) := AddCommGroup.forget_preserves_limits.{v v}, + haveI : mono f.1 := @sheaf.presheaf_mono_of_mono X Ab.{v} _ _ _ _ _ _ _ _ _ f im, + -- this is in another branch + have miffi : ∀ {X Y : Ab.{v}} (f : X ⟶ Y), mono f ↔ function.injective f := sorry, + have := (nat_trans.mono_iff_app_mono _ f.1).mp _, + simp_rw [miffi] at this, + refine (app_injective_iff_stalk_functor_map_injective f.1).mpr _ x, + rintros U, + apply this, + apply_instance, +end + /-- For surjectivity, we are given an arbitrary section `t` and need to find a preimage for it. We claim that it suffices to find preimages *locally*. That is, for each `x : U` we construct a neighborhood `V ≤ U` and a section `s : F.obj (op V))` such that `f.app (op V) s` and `t` From 117fe37e32ec876984f1fccd8e77cfe8e65d8ba5 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Thu, 4 Aug 2022 17:55:02 +0100 Subject: [PATCH 43/84] Apply suggestions from code review Co-authored-by: Markus Himmel --- .../Group/equivalence_Group_AddGroup.lean | 37 +++++-------------- 1 file changed, 9 insertions(+), 28 deletions(-) diff --git a/src/algebra/category/Group/equivalence_Group_AddGroup.lean b/src/algebra/category/Group/equivalence_Group_AddGroup.lean index fb2c2b69e974f..843f3c35354e8 100644 --- a/src/algebra/category/Group/equivalence_Group_AddGroup.lean +++ b/src/algebra/category/Group/equivalence_Group_AddGroup.lean @@ -22,13 +22,8 @@ namespace Group The functor `Group ⥤ AddGroup` by sending `X ↦ additive X` and `f ↦ f`. -/ @[simps] def to_AddGroup : Group ⥤ AddGroup := -{ obj := λ X, ⟨additive X⟩, - map := λ X Y f, - { to_fun := λ x, f x, - map_zero' := by { erw [map_one], refl }, - map_add' := λ x y, by { erw [map_mul], refl } }, - map_id' := λ X, by { ext, refl }, - map_comp' := λ X Y Z f g, by { ext, refl } } +{ obj := λ X, AddGroup.of (additive X), + map := λ X Y, monoid_hom.to_additive } end Group @@ -84,27 +79,13 @@ end AddCommGroup The equivalence of categories between `Group` and `AddGroup` -/ @[simps] def Group_AddGroup_equivalence : Group ≌ AddGroup := -{ functor := Group.to_AddGroup, - inverse := AddGroup.to_Group, - unit_iso := - { hom := - { app := λ X, (𝟙 X), - naturality' := λ _ _ _, by { ext, refl } }, - inv := - { app := λ X, (𝟙 X), - naturality' := λ _ _ _, by { ext, refl } }, - hom_inv_id' := rfl, - inv_hom_id' := rfl }, - counit_iso := - { hom := - { app := λ X, (𝟙 X), - naturality' := λ _ _ _, by { ext, refl } }, - inv := - { app := λ X, (𝟙 X), - naturality' := λ _ _ _, by { ext, refl } }, - hom_inv_id' := rfl, - inv_hom_id' := rfl }, - functor_unit_iso_comp' := λ X, rfl } +equivalence.mk Group.to_AddGroup AddGroup.to_Group + (nat_iso.of_components + (λ X, mul_equiv.to_Group_iso (mul_equiv.multiplicative_additive X)) + (λ X Y f, rfl)) + (nat_iso.of_components + (λ X, add_equiv.to_AddGroup_iso (add_equiv.additive_multiplicative X)) + (λ X Y f, rfl)) /-- The equivalence of categories between `CommGroup` and `AddCommGroup`. From 92d917fc7065a8fb20b53e5cae7a74247f656b91 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Thu, 4 Aug 2022 17:58:49 +0100 Subject: [PATCH 44/84] apply suggestion --- .../Group/equivalence_Group_AddGroup.lean | 30 +++++-------------- src/algebra/hom/equiv.lean | 13 ++++++++ 2 files changed, 21 insertions(+), 22 deletions(-) diff --git a/src/algebra/category/Group/equivalence_Group_AddGroup.lean b/src/algebra/category/Group/equivalence_Group_AddGroup.lean index 843f3c35354e8..b1c3921f24e9f 100644 --- a/src/algebra/category/Group/equivalence_Group_AddGroup.lean +++ b/src/algebra/category/Group/equivalence_Group_AddGroup.lean @@ -90,25 +90,11 @@ equivalence.mk Group.to_AddGroup AddGroup.to_Group /-- The equivalence of categories between `CommGroup` and `AddCommGroup`. -/ -def CommGroup_AddCommGroup_equivalence : CommGroup ≌ AddCommGroup := -{ functor := CommGroup.to_AddCommGroup, - inverse := AddCommGroup.to_CommGroup, - unit_iso := - { hom := - { app := λ X, (𝟙 X), - naturality' := λ _ _ _, by { ext, refl } }, - inv := - { app := λ X, (𝟙 X), - naturality' := λ _ _ _, by { ext, refl } }, - hom_inv_id' := rfl, - inv_hom_id' := rfl }, - counit_iso := - { hom := - { app := λ X, (𝟙 X), - naturality' := λ _ _ _, by { ext, refl } }, - inv := - { app := λ X, (𝟙 X), - naturality' := λ _ _ _, by { ext, refl } }, - hom_inv_id' := rfl, - inv_hom_id' := rfl }, - functor_unit_iso_comp' := λ X, rfl } +@[simps] def CommGroup_AddCommGroup_equivalence : CommGroup ≌ AddCommGroup := +equivalence.mk CommGroup.to_AddCommGroup AddCommGroup.to_CommGroup + (nat_iso.of_components + (λ X, mul_equiv.to_Group_iso (mul_equiv.multiplicative_additive X)) + (λ X Y f, rfl)) + (nat_iso.of_components + (λ X, add_equiv.to_AddGroup_iso (add_equiv.additive_multiplicative X)) + (λ X Y f, rfl)) diff --git a/src/algebra/hom/equiv.lean b/src/algebra/hom/equiv.lean index afaa61705f728..4fcb7c96a05d4 100644 --- a/src/algebra/hom/equiv.lean +++ b/src/algebra/hom/equiv.lean @@ -729,3 +729,16 @@ def mul_equiv.to_additive'' [add_zero_class G] [mul_one_class H] : add_equiv.to_multiplicative''.symm end type_tags + +section +variables (G) (H) + +/-- `additive (multiplicative G)` is just `G`. -/ +def add_equiv.additive_multiplicative [add_zero_class G] : additive (multiplicative G) ≃+ G := +mul_equiv.to_additive'' (mul_equiv.refl (multiplicative G)) + +/-- `multiplicative (additive H)` is just `H`. -/ +def mul_equiv.multiplicative_additive [mul_one_class H] : multiplicative (additive H) ≃* H := +add_equiv.to_multiplicative'' (add_equiv.refl (additive H)) + +end From 5630ac82e96d0586b526239b91d1d08d17b3377d Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Thu, 4 Aug 2022 19:03:53 +0100 Subject: [PATCH 45/84] fix --- src/algebra/category/Group/equivalence_Group_AddGroup.lean | 5 +++-- 1 file changed, 3 insertions(+), 2 deletions(-) diff --git a/src/algebra/category/Group/equivalence_Group_AddGroup.lean b/src/algebra/category/Group/equivalence_Group_AddGroup.lean index b1c3921f24e9f..114cf558fb852 100644 --- a/src/algebra/category/Group/equivalence_Group_AddGroup.lean +++ b/src/algebra/category/Group/equivalence_Group_AddGroup.lean @@ -15,6 +15,7 @@ This file contains two equivalence: sending `X : CommGroup` to `additive X` and `Y : AddCommGroup` to `multiplicative Y`. -/ +open category_theory namespace Group @@ -93,8 +94,8 @@ The equivalence of categories between `CommGroup` and `AddCommGroup`. @[simps] def CommGroup_AddCommGroup_equivalence : CommGroup ≌ AddCommGroup := equivalence.mk CommGroup.to_AddCommGroup AddCommGroup.to_CommGroup (nat_iso.of_components - (λ X, mul_equiv.to_Group_iso (mul_equiv.multiplicative_additive X)) + (λ X, mul_equiv.to_CommGroup_iso (mul_equiv.multiplicative_additive X)) (λ X Y f, rfl)) (nat_iso.of_components - (λ X, add_equiv.to_AddGroup_iso (add_equiv.additive_multiplicative X)) + (λ X, add_equiv.to_AddCommGroup_iso (add_equiv.additive_multiplicative X)) (λ X Y f, rfl)) From aa6b2d46f9c95afb178efce64dca74ff8c4a2f0b Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Thu, 4 Aug 2022 20:46:55 +0100 Subject: [PATCH 46/84] fix --- src/algebra/category/Group/epi_mono.lean | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/src/algebra/category/Group/epi_mono.lean b/src/algebra/category/Group/epi_mono.lean index 92db1b23e3f8e..a4d2559ef7576 100644 --- a/src/algebra/category/Group/epi_mono.lean +++ b/src/algebra/category/Group/epi_mono.lean @@ -6,6 +6,7 @@ Authors: Jujian Zhang import algebra.category.Group.basic import category_theory.epi_mono import group_theory.quotient_group +import algebra.category.Group.equivalence_Group_AddGroup /-! # Monomorphisms and epimorphisms in `Group` @@ -333,7 +334,7 @@ variables {A B : CommGroup.{u}} (f : A ⟶ B) @[to_additive AddCommGroup.ker_eq_bot_of_mono] lemma ker_eq_bot_of_mono [mono f] : f.ker = ⊥ := monoid_hom.ker_eq_bot_of_cancel $ λ u v, - (@cancel_mono _ _ _ _ _ f _ (show Group.of f.ker ⟶ A, from u) _).1 + (@cancel_mono _ _ _ _ _ f _ (show CommGroup.of f.ker ⟶ A, from u) _).1 @[to_additive AddCommGroup.mono_iff_ker_eq_bot] lemma mono_iff_ker_eq_bot : mono f ↔ f.ker = ⊥ := From 40e544e8d189007268eb24ae27df820a13bb27e9 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Sun, 7 Aug 2022 20:00:34 +0100 Subject: [PATCH 47/84] apply change --- .../Group/equivalence_Group_AddGroup.lean | 27 +++++-------------- 1 file changed, 6 insertions(+), 21 deletions(-) diff --git a/src/algebra/category/Group/equivalence_Group_AddGroup.lean b/src/algebra/category/Group/equivalence_Group_AddGroup.lean index 114cf558fb852..3c9073c5a3864 100644 --- a/src/algebra/category/Group/equivalence_Group_AddGroup.lean +++ b/src/algebra/category/Group/equivalence_Group_AddGroup.lean @@ -34,13 +34,8 @@ namespace CommGroup The functor `CommGroup ⥤ AddCommGroup` by sending `X ↦ additive X` and `f ↦ f`. -/ @[simps] def to_AddCommGroup : CommGroup ⥤ AddCommGroup := -{ obj := λ X, ⟨additive X⟩, - map := λ X Y f, - { to_fun := λ x, f x, - map_zero' := by { erw [map_one], refl }, - map_add' := λ x y, by { erw [map_mul], refl } }, - map_id' := λ X, by { ext, refl }, - map_comp' := λ X Y Z f g, by { ext, refl } } +{ obj := λ X, AddCommGroup.of (additive X), + map := λ X Y, monoid_hom.to_additive } end CommGroup @@ -50,13 +45,8 @@ namespace AddGroup The functor `AddGroup ⥤ Group` by sending `X ↦ multiplicative Y` and `f ↦ f`. -/ @[simps] def to_Group : AddGroup ⥤ Group := -{ obj := λ X, ⟨multiplicative X⟩, - map := λ X Y f, - { to_fun := λ x, f x, - map_one' := by { erw [map_zero], refl }, - map_mul' := λ x y, by { erw [map_add], refl } }, - map_id' := λ X, by { ext, refl }, - map_comp' := λ X Y Z f g, by { ext, refl } } +{ obj := λ X, Group.of (multiplicative X), + map := λ X Y, add_monoid_hom.to_multiplicative } end AddGroup @@ -66,13 +56,8 @@ namespace AddCommGroup The functor `AddCommGroup ⥤ CommGroup` by sending `X ↦ multiplicative Y` and `f ↦ f`. -/ @[simps] def to_CommGroup : AddCommGroup ⥤ CommGroup := -{ obj := λ X, ⟨multiplicative X⟩, - map := λ X Y f, - { to_fun := λ x, f x, - map_one' := by { erw [map_zero], refl }, - map_mul' := λ x y, by { erw [map_add], refl } }, - map_id' := λ X, by { ext, refl }, - map_comp' := λ X Y Z f g, by { ext, refl } } +{ obj := λ X, CommGroup.of (multiplicative X), + map := λ X Y, add_monoid_hom.to_multiplicative } end AddCommGroup From 31d54dec70aac5d39450cc072d28c558c77c8152 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Sun, 7 Aug 2022 20:00:50 +0100 Subject: [PATCH 48/84] Update src/algebra/category/Group/epi_mono.lean Co-authored-by: Markus Himmel --- src/algebra/category/Group/epi_mono.lean | 3 +-- 1 file changed, 1 insertion(+), 2 deletions(-) diff --git a/src/algebra/category/Group/epi_mono.lean b/src/algebra/category/Group/epi_mono.lean index a4d2559ef7576..d356e0191546f 100644 --- a/src/algebra/category/Group/epi_mono.lean +++ b/src/algebra/category/Group/epi_mono.lean @@ -321,8 +321,7 @@ begin rwa Group.epi_iff_surjective at i1, end -lemma epi_iff_range_eq_top : - epi f ↔ f.range = ⊤ := +lemma epi_iff_range_eq_top : epi f ↔ f.range = ⊤ := iff.trans (epi_iff_surjective _) (add_subgroup.eq_top_iff' f.range).symm end AddGroup From e305717a1e6ae79e37692166f01b8295bd819747 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Sun, 7 Aug 2022 20:01:21 +0100 Subject: [PATCH 49/84] reorder import --- src/algebra/category/Group/epi_mono.lean | 3 +-- 1 file changed, 1 insertion(+), 2 deletions(-) diff --git a/src/algebra/category/Group/epi_mono.lean b/src/algebra/category/Group/epi_mono.lean index a4d2559ef7576..e9c709fafbfbe 100644 --- a/src/algebra/category/Group/epi_mono.lean +++ b/src/algebra/category/Group/epi_mono.lean @@ -3,10 +3,9 @@ Copyright (c) 2022 Jujian Zhang. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jujian Zhang -/ -import algebra.category.Group.basic +import algebra.category.Group.equivalence_Group_AddGroup import category_theory.epi_mono import group_theory.quotient_group -import algebra.category.Group.equivalence_Group_AddGroup /-! # Monomorphisms and epimorphisms in `Group` From 319101da9f7383273a8ff01e268fe88f03f7c834 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Mon, 8 Aug 2022 15:41:16 +0100 Subject: [PATCH 50/84] initial --- src/topology/sheaves/epi_mono.lean | 95 ++++++++++++++++++++++++++++++ 1 file changed, 95 insertions(+) create mode 100644 src/topology/sheaves/epi_mono.lean diff --git a/src/topology/sheaves/epi_mono.lean b/src/topology/sheaves/epi_mono.lean new file mode 100644 index 0000000000000..0927ac092073f --- /dev/null +++ b/src/topology/sheaves/epi_mono.lean @@ -0,0 +1,95 @@ +/- +Copyright (c) 2022 Jujian Zhang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jujian Zhang +-/ +import topology.sheaves.sheaf +import category_theory.sites.sheafification + +/-! +# Epimorphisms and monomorphisms in category of (pre)sheaves + +Let `X` be a topological space and `F, G` be sheaves on `X`. Then `f : F ⟶ G` is monic iff it is +monic as presheaf morphisms; `f` is epic if `f` is epic as presheaf morphism, but in this case, the +converse is **not** necessarily true. + +-/ + + +universes v u + +open category_theory category_theory.limits +open topological_space +open Top + + +variables {X : Top.{v}} {D : Type u} [category.{v u} D] + +section sheaf + +variables {F G : sheaf D X} (f : F ⟶ G) + +instance mono_of_presheaf_mono [mono f.1] : mono f := +{ right_cancellation := λ P g h eq₀, Sheaf.hom.ext _ _ $ (cancel_mono f.1).mp $ + (Sheaf.hom.ext_iff _ _).mp eq₀ } + +instance sheaf.forget_reflects_mono : (sheaf.forget D X).reflects_monomorphisms := +{ reflects := λ F G f m, @@mono_of_presheaf_mono _ f m } + +instance epi_of_presheaf_epi [epi f.1] : epi f := +{ left_cancellation := λ P g h eq₀, Sheaf.hom.ext _ _ $ (cancel_epi f.1).mp $ + (Sheaf.hom.ext_iff _ _).mp eq₀ } + +instance sheaf_forget_reflects_epi : (sheaf.forget D X).reflects_epimorphisms := +{ reflects := λ F G f m, @@epi_of_presheaf_epi _ f m } + +section to_presheaf + +/-! +Given two sheaves `F, G` on `X` in a cocomplete concrete category `D` where "sheafification" is +possible, then a monic sheaf morphism `f : F ⟶ G` is also a monic presheaf morphism. +-/ + +-- assumptions on target category so that "sheafification" is possible +variables [concrete_category.{v} D] [preserves_limits (forget D)] [has_colimits D] +variables [Π (P : presheaf D X) (U : opens X) (S : (opens.grothendieck_topology X).cover U), + has_multiequalizer (S.index P)] +variables [Π (U : opens X), + preserves_colimits_of_shape ((opens.grothendieck_topology X).cover U)ᵒᵖ (forget D)] +variables [reflects_isomorphisms (forget D)] + +instance presheaf_mono_of_mono [mono f] : mono f.1 := +{ right_cancellation := λ P g h eq₀, + begin + set P_plus : sheaf D X := (presheaf_to_Sheaf (opens.grothendieck_topology X) D).obj P, + set g_plus : P_plus ⟶ F := + ⟨grothendieck_topology.sheafify_lift (opens.grothendieck_topology X) g F.2⟩ with g_plus_def, + set h_plus : P_plus ⟶ F := + ⟨grothendieck_topology.sheafify_lift (opens.grothendieck_topology X) h F.2⟩ with h_plus_def, + have eq₁ : g_plus ≫ f = h_plus ≫ f := Sheaf.hom.ext _ _ (grothendieck_topology.sheafify_hom_ext + _ _ _ G.2 (by simp [*])), + have eq₂ : g_plus = h_plus, + { rwa cancel_mono at eq₁, }, + rw [g_plus_def, h_plus_def, Sheaf.hom.ext_iff] at eq₂, + dsimp at eq₂, + rw ←grothendieck_topology.to_sheafify_sheafify_lift (opens.grothendieck_topology X) g F.2, + rw ←grothendieck_topology.to_sheafify_sheafify_lift (opens.grothendieck_topology X) h F.2, + rw eq₂, + end } + +instance sheaf.forget_preserves_mono : (sheaf.forget D X).preserves_monomorphisms := +{ preserves := λ F G f m, @presheaf_mono_of_mono X D _ F G f _ _ _ _ _ _ m } + +lemma sheaf.mono_iff_presheaf_mono : mono f ↔ mono f.1 := +⟨λ m, @@presheaf_mono_of_mono _ f _ _ _ _ _ _ m, λ m, @@mono_of_presheaf_mono _ f m⟩ + +/- +This is not always true +``` +instance presheaf_epi_of_epi [epi f] : epi f.1 := not true +``` +-/ + +end to_presheaf + +end sheaf From 8f981a6d2b739282caba99481ecbcc2c3f724e5d Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Mon, 8 Aug 2022 16:00:47 +0100 Subject: [PATCH 51/84] fix loop --- src/topology/sheaves/epi_mono.lean | 6 ++++-- 1 file changed, 4 insertions(+), 2 deletions(-) diff --git a/src/topology/sheaves/epi_mono.lean b/src/topology/sheaves/epi_mono.lean index 0927ac092073f..67a4fac9e66d1 100644 --- a/src/topology/sheaves/epi_mono.lean +++ b/src/topology/sheaves/epi_mono.lean @@ -29,7 +29,8 @@ section sheaf variables {F G : sheaf D X} (f : F ⟶ G) -instance mono_of_presheaf_mono [mono f.1] : mono f := +-- Not an instance to prevent class search forming a loop +lemma mono_of_presheaf_mono [mono f.1] : mono f := { right_cancellation := λ P g h eq₀, Sheaf.hom.ext _ _ $ (cancel_mono f.1).mp $ (Sheaf.hom.ext_iff _ _).mp eq₀ } @@ -58,7 +59,8 @@ variables [Π (U : opens X), preserves_colimits_of_shape ((opens.grothendieck_topology X).cover U)ᵒᵖ (forget D)] variables [reflects_isomorphisms (forget D)] -instance presheaf_mono_of_mono [mono f] : mono f.1 := +-- Not an instance to prevent class search forming a loop +lemma presheaf_mono_of_mono [mono f] : mono f.1 := { right_cancellation := λ P g h eq₀, begin set P_plus : sheaf D X := (presheaf_to_Sheaf (opens.grothendieck_topology X) D).obj P, From 47bfa2322079495bb07ded794832e84027b69c2b Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Mon, 8 Aug 2022 16:04:05 +0100 Subject: [PATCH 52/84] minimize diff --- src/topology/sheaves/sheaf.lean | 34 --------------------------------- 1 file changed, 34 deletions(-) diff --git a/src/topology/sheaves/sheaf.lean b/src/topology/sheaves/sheaf.lean index ab487f6660718..794e94d5c2aa8 100644 --- a/src/topology/sheaves/sheaf.lean +++ b/src/topology/sheaves/sheaf.lean @@ -8,7 +8,6 @@ import category_theory.full_subcategory import category_theory.limits.unit import category_theory.sites.sheaf import category_theory.sites.spaces -import category_theory.sites.sheafification /-! # Sheaves @@ -139,39 +138,6 @@ lemma id_app (F : sheaf C X) (t) : (𝟙 F : F ⟶ F).1.app t = 𝟙 _ := rfl lemma comp_app {F G H : sheaf C X} (f : F ⟶ G) (g : G ⟶ H) (t) : (f ≫ g).1.app t = f.1.app t ≫ g.1.app t := rfl -instance presheaf_mono_of_mono {D : Type u} [category.{w u} D] [concrete_category.{w} D] - [preserves_limits (category_theory.forget D)] - [Π (U : opens X), - preserves_colimits_of_shape ((opens.grothendieck_topology X).cover U)ᵒᵖ (category_theory.forget D)] - [reflects_isomorphisms (category_theory.forget D)] - [has_colimits D] - [∀ (P : (opens X)ᵒᵖ ⥤ D) (U : opens X) (S : (opens.grothendieck_topology X).cover U), has_multiequalizer (S.index P)] - {F G : sheaf D X} (f : F ⟶ G) [mono f] : mono f.1 := -{ right_cancellation := λ P g h eq₀, - begin - set P_plus : sheaf D X := (presheaf_to_Sheaf (opens.grothendieck_topology X) D).obj P, - set g_plus : P_plus ⟶ F := ⟨grothendieck_topology.sheafify_lift (opens.grothendieck_topology X) g F.2⟩ with g_plus_def, - set h_plus : P_plus ⟶ F := ⟨grothendieck_topology.sheafify_lift (opens.grothendieck_topology X) h F.2⟩ with h_plus_def, - have eq₁ : g_plus ≫ f = h_plus ≫ f, - { ext1, - dsimp, - apply grothendieck_topology.sheafify_hom_ext (opens.grothendieck_topology X), - { exact G.2 }, - { simp [eq₀] }, }, - have eq₂ : g_plus = h_plus, - { rwa cancel_mono at eq₁, }, - rw [g_plus_def, h_plus_def, Sheaf.hom.ext_iff] at eq₂, - dsimp at eq₂, - have := grothendieck_topology.to_sheafify_sheafify_lift (opens.grothendieck_topology X) g F.2, - rw ←this, - clear this, - have := grothendieck_topology.to_sheafify_sheafify_lift (opens.grothendieck_topology X) h F.2, - rw ←this, - clear this, - rw eq₂, - end } - - end sheaf end Top From c83f8733ec70bdf4aa6a682f15bd537932e331f7 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Mon, 8 Aug 2022 16:13:50 +0100 Subject: [PATCH 53/84] sorry free --- src/topology/sheaves/stalks.lean | 11 ++++++----- 1 file changed, 6 insertions(+), 5 deletions(-) diff --git a/src/topology/sheaves/stalks.lean b/src/topology/sheaves/stalks.lean index 69b56ff09ea4d..ea39bf3cf7738 100644 --- a/src/topology/sheaves/stalks.lean +++ b/src/topology/sheaves/stalks.lean @@ -16,6 +16,8 @@ import algebra.category.Ring import algebra.category.Group.colimits import algebra.category.Group.limits import algebra.category.Group.filtered_colimits +import algebra.category.Group.epi_mono +import topology.sheaves.epi_mono /-! # Stalks @@ -428,17 +430,16 @@ lemma app_injective_iff_stalk_functor_map_injective {F : sheaf C X} /-- This works for any concrete category where monomorphism = injective functions. -/ -instance (x : X) : functor.preserves_monomorphisms (sheaf.forget Ab.{v} X ⋙ stalk_functor Ab.{v} x) := +instance (x : X) : + functor.preserves_monomorphisms (sheaf.forget Ab.{v} X ⋙ stalk_functor Ab.{v} x) := begin refine ⟨λ 𝓐 𝓑 f im, _⟩, apply concrete_category.mono_of_injective, dsimp, haveI : preserves_limits (forget Ab.{v}) := AddCommGroup.forget_preserves_limits.{v v}, - haveI : mono f.1 := @sheaf.presheaf_mono_of_mono X Ab.{v} _ _ _ _ _ _ _ _ _ f im, - -- this is in another branch - have miffi : ∀ {X Y : Ab.{v}} (f : X ⟶ Y), mono f ↔ function.injective f := sorry, + haveI : mono f.1 := @@presheaf_mono_of_mono _ _ _ _ _ _ _ _ im, have := (nat_trans.mono_iff_app_mono _ f.1).mp _, - simp_rw [miffi] at this, + simp_rw [AddCommGroup.mono_iff_injective] at this, refine (app_injective_iff_stalk_functor_map_injective f.1).mpr _ x, rintros U, apply this, From a3b397be47059d9ea6ff80d272a5b2df78a192fd Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Mon, 8 Aug 2022 16:17:20 +0100 Subject: [PATCH 54/84] delete long lines --- src/algebraic_geometry/skyscraper.lean | 32 +++++++++++++++----------- 1 file changed, 19 insertions(+), 13 deletions(-) diff --git a/src/algebraic_geometry/skyscraper.lean b/src/algebraic_geometry/skyscraper.lean index 4d4cb4c8e5d1c..5f072ab36c2a0 100644 --- a/src/algebraic_geometry/skyscraper.lean +++ b/src/algebraic_geometry/skyscraper.lean @@ -361,7 +361,7 @@ noncomputable lemma skyscraper_presheaf_cocone_of_not_mem_closure₀_is_colimit have : ts.from (ite (p₀ ∈ (mem_nhds_of_not_mem_closure_singleton p₀ h).some.1) S star) = eq_to_hom (if_neg (mem_nhds_of_not_mem_closure_singleton p₀ h).some_spec) ≫ 𝟙 star := ts.hom_ext _ _, - erw [this, category.comp_id, ←category.assoc, eq_to_hom_trans, eq_to_hom_refl, category.id_comp], + erw [this, category.comp_id, ←category.assoc, eq_to_hom_trans, eq_to_hom_refl, category.id_comp] end } /-- @@ -637,8 +637,8 @@ def skyscraper_sheaf_functor : C ⥤ sheaf C X := variable [has_colimits C] @[simps] -noncomputable def from_stalk_to_to_skyscraper_presheaf {𝓕 : presheaf C X} {c : C} (f : 𝓕.stalk p₀ ⟶ c) : - 𝓕 ⟶ skyscraper_presheaf p₀ c ts := +noncomputable def from_stalk_to_to_skyscraper_presheaf {𝓕 : presheaf C X} {c : C} + (f : 𝓕.stalk p₀ ⟶ c) : 𝓕 ⟶ skyscraper_presheaf p₀ c ts := { app := λ U, if h : p₀ ∈ U.unop then 𝓕.germ ⟨p₀, h⟩ ≫ f ≫ eq_to_hom (skyscraper_presheaf_obj_of_mem _ _ h).symm else ts.from _ ≫ eq_to_hom (skyscraper_presheaf_obj_of_not_mem _ _ h).symm, @@ -658,8 +658,8 @@ noncomputable def from_stalk_to_to_skyscraper_presheaf {𝓕 : presheaf C X} {c end } @[reducible] -noncomputable def to_skyscraper_presheaf_to_from_stalk {𝓕 : presheaf C X} {c : C} (f : 𝓕 ⟶ skyscraper_presheaf p₀ c ts) : - 𝓕.stalk p₀ ⟶ c := +noncomputable def to_skyscraper_presheaf_to_from_stalk {𝓕 : presheaf C X} {c : C} + (f : 𝓕 ⟶ skyscraper_presheaf p₀ c ts) : 𝓕.stalk p₀ ⟶ c := let CC : cocone ((open_nhds.inclusion p₀).op ⋙ 𝓕) := { X := c, ι := @@ -733,7 +733,8 @@ noncomputable def stalk_skyscraper_presheaf_adj_unit : by_cases hV : p₀ ∈ V.unop, { have hU : p₀ ∈ U.unop := le_of_hom inc.unop hV, split_ifs, - erw [←category.assoc, 𝓕.germ_res inc.unop, category.id_comp, eq_to_hom_trans, category.assoc, eq_to_hom_trans], + erw [←category.assoc, 𝓕.germ_res inc.unop, category.id_comp, eq_to_hom_trans, + category.assoc, eq_to_hom_trans], congr, }, { split_ifs, rw [←category.assoc, eq_comp_eq_to_hom], @@ -746,8 +747,8 @@ noncomputable def stalk_skyscraper_presheaf_adj_unit : split_ifs, { rw [←category.assoc, eq_comp_eq_to_hom], simp only [category.assoc, eq_to_hom_trans, eq_to_hom_refl, category.comp_id], - rw [←category.assoc _ _ ((presheaf.stalk_functor C p₀).map f), eq_to_hom_trans, eq_to_hom_refl, - category.id_comp], + rw [←category.assoc _ _ ((presheaf.stalk_functor C p₀).map f), eq_to_hom_trans, + eq_to_hom_refl, category.id_comp], erw [colimit.ι_map], congr, }, { rw [←category.assoc, eq_comp_eq_to_hom], @@ -762,17 +763,20 @@ noncomputable def stalk_skyscraper_presheaf_adj_counit : begin ext U, dsimp, - simp only [colimit.iso_colimit_cocone_ι_hom_assoc, skyscraper_presheaf_cocone_of_mem_closure₀_ι_app, category.assoc], + simp only [colimit.iso_colimit_cocone_ι_hom_assoc, + skyscraper_presheaf_cocone_of_mem_closure₀_ι_app, category.assoc], erw [category.id_comp, ←category.assoc, colimit.ι_map], dsimp, split_ifs, - { rw [category.assoc, skyscraper_presheaf_stalk_of_mem_closure₀, colimit.iso_colimit_cocone_ι_hom], + { rw [category.assoc, skyscraper_presheaf_stalk_of_mem_closure₀, + colimit.iso_colimit_cocone_ι_hom], dsimp, simpa, }, { exfalso, exact h U.unop.2 } end } -noncomputable def stalk_skyscraper_presheaf_adj : presheaf.stalk_functor C p₀ ⊣ skyscraper_presheaf_functor p₀ ts := +noncomputable def stalk_skyscraper_presheaf_adj : + presheaf.stalk_functor C p₀ ⊣ skyscraper_presheaf_functor p₀ ts := { hom_equiv := λ 𝓕 c, ⟨from_stalk_to_to_skyscraper_presheaf p₀ ts, to_skyscraper_presheaf_to_from_stalk p₀ ts, from_stalk_to_to_skyscraper_presheaf_to_skyscraper_presheaf_to_from_stalk p₀ ts, @@ -792,7 +796,8 @@ noncomputable def stalk_skyscraper_presheaf_adj : presheaf.stalk_functor C p₀ begin ext U, dsimp, - erw [colimit.ι_desc, ←category.assoc, colimit.ι_map, category.assoc, colimit.iso_colimit_cocone_ι_hom], + erw [colimit.ι_desc, ←category.assoc, colimit.ι_map, category.assoc, + colimit.iso_colimit_cocone_ι_hom], dsimp, rw [category.comp_id], refl, @@ -800,7 +805,8 @@ noncomputable def stalk_skyscraper_presheaf_adj : presheaf.stalk_functor C p₀ example : true := trivial -noncomputable def stalk_skyscraper_sheaf_adj : sheaf.forget C X ⋙ presheaf.stalk_functor _ p₀ ⊣ skyscraper_sheaf_functor p₀ ts := +noncomputable def stalk_skyscraper_sheaf_adj : + sheaf.forget C X ⋙ presheaf.stalk_functor _ p₀ ⊣ skyscraper_sheaf_functor p₀ ts := { hom_equiv := λ 𝓕 c, ⟨λ f, ⟨from_stalk_to_to_skyscraper_presheaf p₀ ts f⟩, λ g, to_skyscraper_presheaf_to_from_stalk p₀ ts g.1, From 64b245e5ce5195682e32900cce6fa4f0db174941 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Mon, 8 Aug 2022 16:20:28 +0100 Subject: [PATCH 55/84] initial --- src/algebraic_geometry/skyscraper.lean | 376 +++++++++++++++++++++++++ 1 file changed, 376 insertions(+) create mode 100644 src/algebraic_geometry/skyscraper.lean diff --git a/src/algebraic_geometry/skyscraper.lean b/src/algebraic_geometry/skyscraper.lean new file mode 100644 index 0000000000000..00baba2071639 --- /dev/null +++ b/src/algebraic_geometry/skyscraper.lean @@ -0,0 +1,376 @@ +/- +Copyright (c) 2022 Jujian Zhang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jujian Zhang +-/ +import algebraic_geometry.sheafed_space +import topology.sheaves.sheaf_condition.unique_gluing +import topology.sheaves.stalks +import category_theory.preadditive.injective +import algebra.category.Group.abelian + +/-! +# Skyscraper (pre)sheaves + +A skyscraper (pre)sheaf is a (pre)sheaf supported at a single point: if `p₀ ∈ X` is a specified +point, then the skyscraper (pre)sheaf `𝓕` with value `A` is defined by `U ↦ A` if `p₀ ∈ U` and +`U ↦ *` if `p₀ ∉ A` where `*` is some terminal object. + +## Main definitions + +* `skyscraper_presheaf`: the skyscraper presheaf, to define a skyscraper presheaf at point `p₀` with + value `A`, use `skypscraper_presheaf p₀ A t` where `t` is a proof that `*` is a terminal object. +* `skyscraper_sheaf`: the skyscraper presheaf satisfies the sheaf condition. + +## Main statements + +* `skyscraper_stalk_of_mem_closure₀`: if `y ∈ closure {p₀}` then the stalk of + `skyscraper_presheaf p₀ S t` at `y` is `S`. +* `skyscraper_stalk_of_not_mem_closure₀`: if `y ∉ closure {p₀}` then the stalk of + `skyscraper_presheaf p₀ S t` at `y` is `*`. + +-/ + + +section + +open topological_space +open category_theory category_theory.limits +open Top +open opposite + +universes u v w + +variables {X : Top.{u}} (p₀ : X) {C : Type v} [category.{w} C] (S : C) +variables {star : C} (ts : is_terminal star) +variable [Π (U : opens X), decidable (p₀ ∈ U)] + +/- +A skyscraper presheaf is a presheaf supported at a single point: if `p₀ ∈ X` is a specified +point, then the skyscraper presheaf `𝓕` with value `A` is defined by `U ↦ A` if `p₀ ∈ U` and +`U ↦ *` if `p₀ ∉ A` where `*` is some terminal object. +-/ +@[simps] +def skyscraper_presheaf : presheaf C X := +{ obj := λ U, ite (p₀ ∈ unop U) S star, + map := λ U V i, dite (p₀ ∈ unop V) + (λ h, eq_to_hom (if_pos (le_of_hom i.unop h)) ≫ 𝟙 S ≫ eq_to_hom (if_pos h).symm) + (λ h, ts.from _ ≫ eq_to_hom (if_neg h).symm), + map_id' := λ U, + begin + split_ifs, + { simp, }, + { rw eq_comp_eq_to_hom, + exact ts.hom_ext _ _, }, + end, + map_comp' := λ U V W iVU iWV, + begin + by_cases hW : p₀ ∈ unop W, + { have hV : p₀ ∈ unop V := le_of_hom iWV.unop hW, + have hU : p₀ ∈ unop U := le_of_hom iVU.unop hV, + split_ifs, + simp }, + { split_ifs; + rw eq_comp_eq_to_hom; + exact ts.hom_ext _ _, } + end } + +section + +variables {p₀} + +lemma skyscraper_presheaf_obj_of_mem {U : opens X} (h : p₀ ∈ U) : + (skyscraper_presheaf p₀ S ts).obj (op U) = S := +if_pos h + +lemma skyscraper_presheaf_obj_of_not_mem {U : opens X} (h : p₀ ∉ U) : + (skyscraper_presheaf p₀ S ts).obj (op U) = star := +if_neg h + +end + +/- +A skyscraper sheaf is a sheaf supported at a single point: if `p₀ ∈ X` is a specified +point, then the skyscraper sheaf `𝓕` with value `A` is defined by `U ↦ A` if `p₀ ∈ U` and +`U ↦ *` if `p₀ ∉ A` where `*` is some terminal object. +-/ +def skyscraper_sheaf : sheaf C X := +⟨skyscraper_presheaf p₀ S ts, λ c U s hs x hx, + ⟨dite (p₀ ∈ U) + (λ h, x (hs p₀ h).some_spec.some (hs p₀ h).some_spec.some_spec.1 ≫ + eq_to_hom ((skyscraper_presheaf_obj_of_mem S ts (hs p₀ h).some_spec.some_spec.2).trans + (skyscraper_presheaf_obj_of_mem S ts h).symm)) + (λ h, ts.from c ≫ (eq_to_hom (skyscraper_presheaf_obj_of_not_mem S ts h).symm)), + λ V inc h, + begin + by_cases hV : p₀ ∈ V, + { have hU : p₀ ∈ U := le_of_hom inc hV, + split_ifs, + generalize_proofs h₁ h₂ h₃ h₄, + dsimp, + split_ifs, + rw [category.id_comp, eq_to_hom_trans, category.assoc, eq_to_hom_trans], + generalize_proofs h₅, + have := hx (hom_of_le inf_le_left) (hom_of_le inf_le_right) h₂.some_spec.1 h rfl, + dsimp at this, + have hV' : p₀ ∈ h₁.some ⊓ V := ⟨h₂.some_spec.2, hV⟩, + split_ifs at this; + rw [category.id_comp, eq_to_hom_trans, eq_to_hom_trans, eq_comp_eq_to_hom, + category.assoc, eq_to_hom_trans] at this; + generalize_proofs h₆ at this; + rw [this, category.assoc, eq_to_hom_trans, eq_to_hom_refl, category.comp_id] }, + { dsimp, + split_ifs, + rw [←category.assoc, eq_comp_eq_to_hom], + exact ts.hom_ext _ _, } + end, + λ y (hy : x.is_amalgamation y), + begin + split_ifs, + { generalize_proofs h₁ h₂ h₃ h₄, + have := hy h₂.some h₂.some_spec.1, + dsimp at this, + split_ifs at this with H, + { rw [category.id_comp, eq_to_hom_trans, eq_comp_eq_to_hom] at this, + rw [this, eq_comp_eq_to_hom, category.assoc, eq_to_hom_trans, eq_to_hom_refl, + category.comp_id], }, + { exfalso, + exact H h₂.some_spec.2, }, }, + { rw [←eq_comp_eq_to_hom], + exact (ts.hom_ext _ _), } + end⟩⟩ + +end + +section + +-- In this section, we calculate the stalks for skyscraper presheaves. We need to restrict universe +-- level. + +open topological_space +open category_theory category_theory.limits +open Top +open opposite + +universes u v + +variables {X : Top.{u}} (p₀ : X) {C : Type v} [category.{u} C] (S : C) +variables {star : C} (ts : is_terminal star) +variable [Π (U : opens X), decidable (p₀ ∈ U)] +variable [has_colimits C] + +lemma mem_nhds_of_not_mem_closure_singleton {y : X} (h : ¬p₀ ⤳ y) : + ∃ (U : open_nhds y), p₀ ∉ U.1 := +begin + have := (not_iff_not.mpr specializes_iff_forall_open).mp h, + push_neg at this, + rcases this with ⟨s, o, h₁, h₂⟩, + exact ⟨⟨⟨s, o⟩, h₁⟩, h₂⟩, +end + +/-- +The cocone at `S` for the salk functor of `skyscraper_presheaf p₀ S t` when `y ∈ closure {p₀}` +-/ +@[simps] def skyscraper_presheaf_cocone_of_mem_closure₀ {y : X} (h : p₀ ⤳ y) : + cocone ((open_nhds.inclusion y).op ⋙ skyscraper_presheaf p₀ S ts) := +{ X := S, + ι := + { app := λ U, eq_to_hom + begin + dsimp, + rw if_pos, + exact h.mem_open U.unop.1.2 U.unop.2, + end ≫ 𝟙 S, + naturality' := λ U V inc, + begin + simp only [functor.op_obj, unop_op, functor.comp_map, functor.op_map, skyscraper_presheaf_map, + category.id_comp, eq_to_hom_trans, functor.const_obj_map, category.assoc], + by_cases hV : p₀ ∈ (open_nhds.inclusion y).obj V.unop, + { have hU : p₀ ∈ unop ((open_nhds.inclusion y).op.obj U) := le_of_hom inc.unop hV, + split_ifs, + erw [category.comp_id, category.comp_id, category.comp_id, eq_to_hom_trans], + refl }, + { split_ifs with hU; + erw [category.comp_id, category.comp_id, category.comp_id, eq_comp_eq_to_hom, + eq_comp_eq_to_hom]; + exact ts.hom_ext _ _, }, + end } } + +/-- +The canonical map `S ⟶ (skyscraper_presheaf p₀ S t).stalk y` when `y ∈ closure {p₀}` +-/ +noncomputable def skyscraper_presheaf_of_mem_closure₀_from {y : X} (h : p₀ ⤳ y) : + S ⟶ (skyscraper_presheaf p₀ S ts).stalk y := +eq_to_hom (skyscraper_presheaf_obj_of_mem S ts (by tauto : p₀ ∈ ⊤)).symm ≫ + (skyscraper_presheaf p₀ S ts).germ (⟨y, trivial⟩ : (⊤ : opens X)) + +noncomputable lemma skyscraper_presheaf_cocone_of_mem_closure₀_is_colimit {y : X} (h : p₀ ⤳ y) : + is_colimit (skyscraper_presheaf_cocone_of_mem_closure₀ p₀ S ts h) := +{ desc := λ c, (skyscraper_presheaf_of_mem_closure₀_from p₀ S ts h ≫ colimit.desc _ _ : S ⟶ c.X), + fac' := λ c U, + begin + dsimp, + simp only [skyscraper_presheaf_of_mem_closure₀_from, presheaf.germ, category.comp_id, + category.assoc, colimit.ι_desc, eq_to_hom_trans_assoc], + have := c.ι.naturality (hom_of_le $ (le_top : unop U ≤ _)).op, + dsimp at this, + have h' : p₀ ∈ (open_nhds.inclusion y).obj (unop U) := h.mem_open U.unop.1.2 U.unop.2, + have h'' : p₀ ∈ (open_nhds.inclusion y).obj ⊤ := trivial, + split_ifs at this, + rw [category.comp_id, category.id_comp, eq_to_hom_trans, eq_eq_to_hom_comp] at this, + rw [this, eq_eq_to_hom_comp, ←category.assoc, eq_to_hom_trans, eq_to_hom_refl, + category.id_comp], + congr, + end, + uniq' := λ c f h, + begin + simp only [skyscraper_presheaf_of_mem_closure₀_from, presheaf.germ, category.assoc], + erw [colimit.ι_desc], + specialize h (op ⟨⊤, trivial⟩), + erw [←h], + simp only [skyscraper_presheaf_cocone_of_mem_closure₀_ι_app, category.assoc, + eq_to_hom_trans_assoc, eq_to_hom_refl, category.id_comp], + exact (category.id_comp _).symm, + end } + +/-- +If `y ∈ closure {p₀}`, then the stalk of `skyscraper_presheaf p₀ S t` at `y` is `S` +-/ +@[reducible] +noncomputable def skyscraper_presheaf_stalk_of_mem_closure₀ {y : X} (h : p₀ ⤳ y) : + (skyscraper_presheaf p₀ S ts).stalk y ≅ S := +colimit.iso_colimit_cocone ⟨_, (skyscraper_presheaf_cocone_of_mem_closure₀_is_colimit p₀ S ts h)⟩ + +/-- +The cocone at `*` for the salk functor of `skyscraper_presheaf p₀ S t` when `y ∉ closure {p₀}` +-/ +@[simps] def skyscraper_presheaf_cocone_of_not_mem_closure₀ + {y : X} (h : ¬p₀ ⤳ y) : + cocone ((open_nhds.inclusion y).op ⋙ skyscraper_presheaf p₀ S ts) := +{ X := star, + ι := + { app := λ U, ts.from _, + naturality' := λ U V inc, ts.hom_ext _ _ } } + +/-- +The canonical map `* ⟶ (skyscraper_presheaf p₀ S t).stalk y` when `y ∉ closure {p₀}` +-/ +noncomputable def skyscraper_presheaf_of_not_mem_closure₀_from + {y : X} (h : ¬p₀ ⤳ y) : + star ⟶ (skyscraper_presheaf p₀ S ts).stalk y := +eq_to_hom (skyscraper_presheaf_obj_of_not_mem S ts $ + (mem_nhds_of_not_mem_closure_singleton p₀ h).some_spec).symm ≫ + (skyscraper_presheaf p₀ S ts).germ (⟨y, (mem_nhds_of_not_mem_closure_singleton p₀ h).some.2⟩ : + (mem_nhds_of_not_mem_closure_singleton p₀ h).some.1) + +noncomputable lemma skyscraper_presheaf_cocone_of_not_mem_closure₀_is_colimit + {y : X} (h : ¬p₀ ⤳ y) : + is_colimit (skyscraper_presheaf_cocone_of_not_mem_closure₀ p₀ S ts h) := +{ desc := λ c, ((eq_to_hom ((skyscraper_presheaf_obj_of_not_mem _ _ + (mem_nhds_of_not_mem_closure_singleton p₀ h).some_spec).symm)) ≫ + presheaf.germ (skyscraper_presheaf p₀ S ts) + ⟨y, (mem_nhds_of_not_mem_closure_singleton p₀ h).some.2⟩ ≫ colimit.desc _ _ : star ⟶ _), + fac' := λ c U, + begin + simp only [presheaf.germ, skyscraper_presheaf_cocone_of_not_mem_closure₀_ι_app], + erw [colimit.ι_desc], + dsimp, + by_cases h' : p₀ ∈ (open_nhds.inclusion y).obj (unop U), + { have eq1 : ts.from (ite (p₀ ∈ (open_nhds.inclusion y).obj (unop U)) S star) = + eq_to_hom (if_pos h') ≫ ts.from _ := ts.hom_ext _ _, + rw [eq1, category.assoc, eq_eq_to_hom_comp], + have := c.ι.naturality (hom_of_le (le_top : U.unop ≤ ⊤)).op, + dsimp at this, + have h'' : p₀ ∈ (open_nhds.inclusion y).obj ⊤ := trivial, + split_ifs at this, + rw [category.comp_id, category.id_comp, eq_to_hom_trans, eq_eq_to_hom_comp] at this, + rw this, + clear this, + have := c.ι.naturality + (hom_of_le (le_top : (mem_nhds_of_not_mem_closure_singleton p₀ h).some ≤ ⊤)).op, + dsimp at this, + have h''' : p₀ ∉ (mem_nhds_of_not_mem_closure_singleton p₀ h).some.1 := + (mem_nhds_of_not_mem_closure_singleton p₀ h).some_spec, + split_ifs at this, + have eq2 : ts.from (ite (p₀ ∈ (open_nhds.inclusion y).obj ⊤) S star) = + eq_to_hom (if_pos h'') ≫ ts.from _ := ts.hom_ext _ _, + rw [category.comp_id, eq2] at this, + simp only [←this, ←category.assoc], + convert eq_whisker _ _, + { ext, refl, }, + { rw [eq_comp_eq_to_hom], + exact ts.hom_ext _ _ } }, + { have eq1 : ts.from (ite (p₀ ∈ (open_nhds.inclusion y).obj (unop U)) S star) = + eq_to_hom (if_neg h') ≫ 𝟙 star := ts.hom_ext _ _, + have eq2 : ts.from (ite (p₀ ∈ (open_nhds.inclusion y).obj ⊤) S star) = + eq_to_hom (if_pos trivial) ≫ ts.from _ := ts.hom_ext _ _, + have eq3 : ts.from (ite (p₀ ∈ (open_nhds.inclusion y).obj + (mem_nhds_of_not_mem_closure_singleton p₀ h).some) S star) = + eq_to_hom (if_neg (mem_nhds_of_not_mem_closure_singleton p₀ h).some_spec) ≫ 𝟙 star := + ts.hom_ext _ _, + rw [eq1, category.comp_id, ←category.assoc, eq_to_hom_trans], + have := c.ι.naturality (hom_of_le (inf_le_left : + (mem_nhds_of_not_mem_closure_singleton p₀ h).some ⊓ unop U ≤ + (mem_nhds_of_not_mem_closure_singleton p₀ h).some)).op, + dsimp at this, + rw [dite_comp, category.comp_id] at this, + simp_rw [category.id_comp, eq_to_hom_trans, eq3, category.comp_id, eq_to_hom_trans] at this, + + + generalize_proofs h₁ h₂ h₃, + have eq_coe : c.ι.app (op ⟨↑(h₁.some), h₃⟩) = + eq_to_hom (by { congr, ext, refl }) ≫ c.ι.app (op h₁.some) ≫ + eq_to_hom (by { congr, ext, refl }), + { symmetry, + rw [eq_eq_to_hom_comp], + have e := c.ι.naturality (eq_to_hom (by {ext, refl}) : + (⟨↑h₁.some, h₃⟩ : open_nhds y) ⟶ h₁.some).op, + dsimp at e, + split_ifs at e with temp, + { exfalso, exact (mem_nhds_of_not_mem_closure_singleton p₀ h).some_spec temp }, + rw [show ts.from (ite (p₀ ∈ (open_nhds.inclusion y).obj h₁.some) S star) = + eq_to_hom (if_neg temp) ≫ 𝟙 star, from ts.hom_ext _ _, category.comp_id, + category.comp_id, eq_to_hom_trans] at e, + rw [←e, category.assoc, eq_eq_to_hom_comp, ←category.assoc, eq_to_hom_trans, eq_to_hom_refl, + eq_to_hom_refl, category.comp_id, category.id_comp], }, + rw [eq_coe, ←this, ←category.assoc, eq_to_hom_trans, eq_eq_to_hom_comp, dite_comp], + have h₅ : p₀ ∉ (open_nhds.inclusion y).obj h₁.some := h₁.some_spec, + have h₆ : p₀ ∉ (open_nhds.inclusion y).obj (h₁.some ⊓ unop U) := λ ⟨r, _⟩, h₅ r, + clear this, + split_ifs, + { exfalso, exact h₅ (by assumption) }, + { exfalso, exact h₆ (by assumption) }, + { exfalso, exact h₅ (by assumption) }, + rw [eq_comp_eq_to_hom, eq_eq_to_hom_comp, ←category.assoc, ←category.assoc, eq_to_hom_trans, + eq_to_hom_refl, category.comp_id], + have := c.ι.naturality (hom_of_le (inf_le_right : + (mem_nhds_of_not_mem_closure_singleton p₀ h).some ⊓ unop U ≤ unop U)).op, + dsimp at this, + rw [dite_comp, category.comp_id] at this, + rw [←this, comp_dite], + clear this, + split_ifs, + rw [eq1, category.comp_id, eq_to_hom_trans, ←category.assoc, eq_to_hom_trans, eq_to_hom_refl, + category.id_comp] } + end, + uniq' := λ c f H, + begin + erw [colimit.ι_desc, ←H], + simp only [skyscraper_presheaf_cocone_of_not_mem_closure₀_ι_app], + dsimp, + have : ts.from (ite (p₀ ∈ (mem_nhds_of_not_mem_closure_singleton p₀ h).some.1) S star) = + eq_to_hom (if_neg (mem_nhds_of_not_mem_closure_singleton p₀ h).some_spec) ≫ 𝟙 star := + ts.hom_ext _ _, + erw [this, category.comp_id, ←category.assoc, eq_to_hom_trans, eq_to_hom_refl, category.id_comp] + end } + +/-- +If `y ∉ closure {p₀}`, then the stalk of `skyscraper_presheaf p₀ S t` at `y` is `*` +-/ +@[reducible] +noncomputable def skyscraper_presheaf_stalk_of_not_mem_closure₀ + {y : X} (h : ¬p₀ ⤳ y) : + (skyscraper_presheaf p₀ S ts).stalk y ≅ star := +colimit.iso_colimit_cocone ⟨_, (skyscraper_presheaf_cocone_of_not_mem_closure₀_is_colimit _ S _ h)⟩ + +end From 95a366342543eb6231005edfd43cfaf4bdec30ce Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Mon, 8 Aug 2022 16:21:26 +0100 Subject: [PATCH 56/84] move file --- src/{algebraic_geometry => topology/sheaves}/skyscraper.lean | 0 1 file changed, 0 insertions(+), 0 deletions(-) rename src/{algebraic_geometry => topology/sheaves}/skyscraper.lean (100%) diff --git a/src/algebraic_geometry/skyscraper.lean b/src/topology/sheaves/skyscraper.lean similarity index 100% rename from src/algebraic_geometry/skyscraper.lean rename to src/topology/sheaves/skyscraper.lean From 93f4f21958a8320d68fe547a23824c878ee0d484 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Mon, 8 Aug 2022 17:40:26 +0100 Subject: [PATCH 57/84] need cache --- .../concrete_category/basic.lean | 29 +++++++++++++++++++ 1 file changed, 29 insertions(+) diff --git a/src/category_theory/concrete_category/basic.lean b/src/category_theory/concrete_category/basic.lean index 29e9eb2079c0b..214947584dfad 100644 --- a/src/category_theory/concrete_category/basic.lean +++ b/src/category_theory/concrete_category/basic.lean @@ -5,6 +5,7 @@ Authors: Scott Morrison, Johannes Hölzl, Reid Barton, Sean Leather, Yury Kudrya -/ import category_theory.types import category_theory.functor.epi_mono +import category_theory.limits.constructions.epi_mono /-! # Concrete categories @@ -38,6 +39,8 @@ universes w v v' u namespace category_theory +open category_theory.limits + /-- A concrete category is a category `C` with a fixed faithful functor `forget : C ⥤ Type`. @@ -129,11 +132,37 @@ lemma concrete_category.mono_of_injective {X Y : C} (f : X ⟶ Y) (i : function. mono f := (forget C).mono_of_mono_map ((mono_iff_injective f).2 i) +lemma concrete_category.injective_of_mono_of_preserves_pullback {X Y : C} (f : X ⟶ Y) [mono f] + [preserves_limits_of_shape walking_cospan (forget C)] : function.injective f := +begin + haveI := category_theory.preserves_monomorphisms_of_preserves_limits_of_shape (forget C), + haveI : mono ((forget C).map f) := infer_instance, + exact (mono_iff_injective ((forget C).map f)).mp infer_instance, +end + +lemma concrete_category.mono_iff_injective_of_preserves_pullback {X Y : C} (f : X ⟶ Y) [mono f] + [preserves_limits_of_shape walking_cospan (forget C)] : mono f ↔ function.injective f := +⟨λ m, concrete_category.injective_of_mono_of_preserves_pullback _, + concrete_category.mono_of_injective f⟩ + /-- In any concrete category, surjective morphisms are epimorphisms. -/ lemma concrete_category.epi_of_surjective {X Y : C} (f : X ⟶ Y) (s : function.surjective f) : epi f := (forget C).epi_of_epi_map ((epi_iff_surjective f).2 s) +lemma concrete_category.surjective_of_epi_of_preserves_pushout {X Y : C} (f : X ⟶ Y) [epi f] + [preserves_colimits_of_shape walking_span (forget C)] : function.surjective f := +begin + haveI := category_theory.preserves_epimorphisms_of_preserves_colimits_of_shape (forget C), + haveI : epi ((forget C).map f) := infer_instance, + exact (epi_iff_surjective ((forget C).map f)).mp infer_instance +end + +lemma concrete_category.epi_iff_surjective_of_preserves_pushout {X Y : C} (f : X ⟶ Y) + [preserves_colimits_of_shape walking_span (forget C)] : epi f ↔ function.surjective f := +⟨λ m, @@concrete_category.surjective_of_epi_of_preserves_pushout _ _ _ m _, + concrete_category.epi_of_surjective f⟩ + @[simp] lemma concrete_category.has_coe_to_fun_Type {X Y : Type u} (f : X ⟶ Y) : coe_fn f = f := rfl From 3c964fd0976b81ddd6726b2bbe0aedac80a8566f Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Wed, 10 Aug 2022 19:56:41 +0100 Subject: [PATCH 58/84] need cache --- src/category_theory/eq_to_hom.lean | 10 ++++++++++ 1 file changed, 10 insertions(+) diff --git a/src/category_theory/eq_to_hom.lean b/src/category_theory/eq_to_hom.lean index 19e66aee7315d..17d1c67b40c34 100644 --- a/src/category_theory/eq_to_hom.lean +++ b/src/category_theory/eq_to_hom.lean @@ -45,6 +45,16 @@ def eq_to_hom {X Y : C} (p : X = Y) : X ⟶ Y := by rw p; exact 𝟙 _ eq_to_hom p ≫ eq_to_hom q = eq_to_hom (p.trans q) := by { cases p, cases q, simp, } +lemma eq_comp_eq_to_hom {X Y Y' : C} (p : Y = Y') (f : X ⟶ Y) (g : X ⟶ Y') : + f ≫ eq_to_hom p = g ↔ f = g ≫ eq_to_hom p.symm := +{ mp := λ h, h ▸ by simp, + mpr := λ h, by simp [eq_whisker h (eq_to_hom p)] } + +lemma eq_eq_to_hom_comp {X X' Y : C} (p : X = X') (f : X ⟶ Y) (g : X' ⟶ Y) : + eq_to_hom p ≫ g = f ↔ g = eq_to_hom p.symm ≫ f := +{ mp := λ h, h ▸ by simp, + mpr := λ h, h ▸ by simp [whisker_eq _ h] } + /-- If we (perhaps unintentionally) perform equational rewriting on the source object of a morphism, From 085be8fdcf0b982077607dd74dd0658d7f0b3bed Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Wed, 10 Aug 2022 23:01:26 +0100 Subject: [PATCH 59/84] make it better --- src/topology/sheaves/skyscraper.lean | 151 +++++++++++++-------------- 1 file changed, 70 insertions(+), 81 deletions(-) diff --git a/src/topology/sheaves/skyscraper.lean b/src/topology/sheaves/skyscraper.lean index 00baba2071639..f83f64c183c77 100644 --- a/src/topology/sheaves/skyscraper.lean +++ b/src/topology/sheaves/skyscraper.lean @@ -31,6 +31,7 @@ point, then the skyscraper (pre)sheaf `𝓕` with value `A` is defined by `U ↦ -/ +noncomputable theory section @@ -42,26 +43,24 @@ open opposite universes u v w variables {X : Top.{u}} (p₀ : X) {C : Type v} [category.{w} C] (S : C) -variables {star : C} (ts : is_terminal star) -variable [Π (U : opens X), decidable (p₀ ∈ U)] +variables [has_terminal C] [Π (U : opens X), decidable (p₀ ∈ U)] /- A skyscraper presheaf is a presheaf supported at a single point: if `p₀ ∈ X` is a specified point, then the skyscraper presheaf `𝓕` with value `A` is defined by `U ↦ A` if `p₀ ∈ U` and `U ↦ *` if `p₀ ∉ A` where `*` is some terminal object. -/ -@[simps] -def skyscraper_presheaf : presheaf C X := -{ obj := λ U, ite (p₀ ∈ unop U) S star, +@[simps] def skyscraper_presheaf : presheaf C X := +{ obj := λ U, ite (p₀ ∈ unop U) S (terminal C), map := λ U V i, dite (p₀ ∈ unop V) (λ h, eq_to_hom (if_pos (le_of_hom i.unop h)) ≫ 𝟙 S ≫ eq_to_hom (if_pos h).symm) - (λ h, ts.from _ ≫ eq_to_hom (if_neg h).symm), + (λ h, terminal.from _ ≫ eq_to_hom (if_neg h).symm), map_id' := λ U, begin split_ifs, { simp, }, { rw eq_comp_eq_to_hom, - exact ts.hom_ext _ _, }, + exact terminal_is_terminal.hom_ext _ _, }, end, map_comp' := λ U V W iVU iWV, begin @@ -72,7 +71,7 @@ def skyscraper_presheaf : presheaf C X := simp }, { split_ifs; rw eq_comp_eq_to_hom; - exact ts.hom_ext _ _, } + exact terminal_is_terminal.hom_ext _ _, } end } section @@ -80,11 +79,11 @@ section variables {p₀} lemma skyscraper_presheaf_obj_of_mem {U : opens X} (h : p₀ ∈ U) : - (skyscraper_presheaf p₀ S ts).obj (op U) = S := + (skyscraper_presheaf p₀ S).obj (op U) = S := if_pos h lemma skyscraper_presheaf_obj_of_not_mem {U : opens X} (h : p₀ ∉ U) : - (skyscraper_presheaf p₀ S ts).obj (op U) = star := + (skyscraper_presheaf p₀ S).obj (op U) = (terminal C) := if_neg h end @@ -95,12 +94,12 @@ point, then the skyscraper sheaf `𝓕` with value `A` is defined by `U ↦ A` i `U ↦ *` if `p₀ ∉ A` where `*` is some terminal object. -/ def skyscraper_sheaf : sheaf C X := -⟨skyscraper_presheaf p₀ S ts, λ c U s hs x hx, +⟨skyscraper_presheaf p₀ S, λ c U s hs x hx, ⟨dite (p₀ ∈ U) (λ h, x (hs p₀ h).some_spec.some (hs p₀ h).some_spec.some_spec.1 ≫ - eq_to_hom ((skyscraper_presheaf_obj_of_mem S ts (hs p₀ h).some_spec.some_spec.2).trans - (skyscraper_presheaf_obj_of_mem S ts h).symm)) - (λ h, ts.from c ≫ (eq_to_hom (skyscraper_presheaf_obj_of_not_mem S ts h).symm)), + eq_to_hom ((skyscraper_presheaf_obj_of_mem S (hs p₀ h).some_spec.some_spec.2).trans + (skyscraper_presheaf_obj_of_mem S h).symm)) + (λ h, terminal.from c ≫ (eq_to_hom (skyscraper_presheaf_obj_of_not_mem S h).symm)), λ V inc h, begin by_cases hV : p₀ ∈ V, @@ -122,7 +121,7 @@ def skyscraper_sheaf : sheaf C X := { dsimp, split_ifs, rw [←category.assoc, eq_comp_eq_to_hom], - exact ts.hom_ext _ _, } + exact terminal_is_terminal.hom_ext _ _, } end, λ y (hy : x.is_amalgamation y), begin @@ -137,7 +136,7 @@ def skyscraper_sheaf : sheaf C X := { exfalso, exact H h₂.some_spec.2, }, }, { rw [←eq_comp_eq_to_hom], - exact (ts.hom_ext _ _), } + exact terminal_is_terminal.hom_ext _ _, } end⟩⟩ end @@ -155,9 +154,8 @@ open opposite universes u v variables {X : Top.{u}} (p₀ : X) {C : Type v} [category.{u} C] (S : C) -variables {star : C} (ts : is_terminal star) -variable [Π (U : opens X), decidable (p₀ ∈ U)] -variable [has_colimits C] +variables [has_terminal C] +variables [Π (U : opens X), decidable (p₀ ∈ U)] [has_colimits C] lemma mem_nhds_of_not_mem_closure_singleton {y : X} (h : ¬p₀ ⤳ y) : ∃ (U : open_nhds y), p₀ ∉ U.1 := @@ -172,7 +170,7 @@ end The cocone at `S` for the salk functor of `skyscraper_presheaf p₀ S t` when `y ∈ closure {p₀}` -/ @[simps] def skyscraper_presheaf_cocone_of_mem_closure₀ {y : X} (h : p₀ ⤳ y) : - cocone ((open_nhds.inclusion y).op ⋙ skyscraper_presheaf p₀ S ts) := + cocone ((open_nhds.inclusion y).op ⋙ skyscraper_presheaf p₀ S) := { X := S, ι := { app := λ U, eq_to_hom @@ -193,20 +191,20 @@ The cocone at `S` for the salk functor of `skyscraper_presheaf p₀ S t` when `y { split_ifs with hU; erw [category.comp_id, category.comp_id, category.comp_id, eq_comp_eq_to_hom, eq_comp_eq_to_hom]; - exact ts.hom_ext _ _, }, + exact terminal_is_terminal.hom_ext _ _, }, end } } /-- The canonical map `S ⟶ (skyscraper_presheaf p₀ S t).stalk y` when `y ∈ closure {p₀}` -/ noncomputable def skyscraper_presheaf_of_mem_closure₀_from {y : X} (h : p₀ ⤳ y) : - S ⟶ (skyscraper_presheaf p₀ S ts).stalk y := -eq_to_hom (skyscraper_presheaf_obj_of_mem S ts (by tauto : p₀ ∈ ⊤)).symm ≫ - (skyscraper_presheaf p₀ S ts).germ (⟨y, trivial⟩ : (⊤ : opens X)) + S ⟶ (skyscraper_presheaf p₀ S).stalk y := +eq_to_hom (skyscraper_presheaf_obj_of_mem S (by tauto : p₀ ∈ ⊤)).symm ≫ + (skyscraper_presheaf p₀ S).germ (⟨y, trivial⟩ : (⊤ : opens X)) noncomputable lemma skyscraper_presheaf_cocone_of_mem_closure₀_is_colimit {y : X} (h : p₀ ⤳ y) : - is_colimit (skyscraper_presheaf_cocone_of_mem_closure₀ p₀ S ts h) := -{ desc := λ c, (skyscraper_presheaf_of_mem_closure₀_from p₀ S ts h ≫ colimit.desc _ _ : S ⟶ c.X), + is_colimit (skyscraper_presheaf_cocone_of_mem_closure₀ p₀ S h) := +{ desc := λ c, (skyscraper_presheaf_of_mem_closure₀_from p₀ S h ≫ colimit.desc _ _ : S ⟶ c.X), fac' := λ c U, begin dsimp, @@ -238,46 +236,45 @@ If `y ∈ closure {p₀}`, then the stalk of `skyscraper_presheaf p₀ S t` at ` -/ @[reducible] noncomputable def skyscraper_presheaf_stalk_of_mem_closure₀ {y : X} (h : p₀ ⤳ y) : - (skyscraper_presheaf p₀ S ts).stalk y ≅ S := -colimit.iso_colimit_cocone ⟨_, (skyscraper_presheaf_cocone_of_mem_closure₀_is_colimit p₀ S ts h)⟩ + (skyscraper_presheaf p₀ S).stalk y ≅ S := +colimit.iso_colimit_cocone ⟨_, (skyscraper_presheaf_cocone_of_mem_closure₀_is_colimit p₀ S h)⟩ /-- The cocone at `*` for the salk functor of `skyscraper_presheaf p₀ S t` when `y ∉ closure {p₀}` -/ -@[simps] def skyscraper_presheaf_cocone_of_not_mem_closure₀ - {y : X} (h : ¬p₀ ⤳ y) : - cocone ((open_nhds.inclusion y).op ⋙ skyscraper_presheaf p₀ S ts) := -{ X := star, +@[simps] def skyscraper_presheaf_cocone_of_not_mem_closure₀ {y : X} (h : ¬p₀ ⤳ y) : + cocone ((open_nhds.inclusion y).op ⋙ skyscraper_presheaf p₀ S) := +{ X := terminal C, ι := - { app := λ U, ts.from _, - naturality' := λ U V inc, ts.hom_ext _ _ } } + { app := λ U, terminal.from _, + naturality' := λ U V inc, terminal_is_terminal.hom_ext _ _ } } /-- The canonical map `* ⟶ (skyscraper_presheaf p₀ S t).stalk y` when `y ∉ closure {p₀}` -/ noncomputable def skyscraper_presheaf_of_not_mem_closure₀_from {y : X} (h : ¬p₀ ⤳ y) : - star ⟶ (skyscraper_presheaf p₀ S ts).stalk y := -eq_to_hom (skyscraper_presheaf_obj_of_not_mem S ts $ + terminal C ⟶ (skyscraper_presheaf p₀ S).stalk y := +eq_to_hom (skyscraper_presheaf_obj_of_not_mem S $ (mem_nhds_of_not_mem_closure_singleton p₀ h).some_spec).symm ≫ - (skyscraper_presheaf p₀ S ts).germ (⟨y, (mem_nhds_of_not_mem_closure_singleton p₀ h).some.2⟩ : + (skyscraper_presheaf p₀ S).germ (⟨y, (mem_nhds_of_not_mem_closure_singleton p₀ h).some.2⟩ : (mem_nhds_of_not_mem_closure_singleton p₀ h).some.1) noncomputable lemma skyscraper_presheaf_cocone_of_not_mem_closure₀_is_colimit {y : X} (h : ¬p₀ ⤳ y) : - is_colimit (skyscraper_presheaf_cocone_of_not_mem_closure₀ p₀ S ts h) := -{ desc := λ c, ((eq_to_hom ((skyscraper_presheaf_obj_of_not_mem _ _ + is_colimit (skyscraper_presheaf_cocone_of_not_mem_closure₀ p₀ S h) := +{ desc := λ c, (eq_to_hom ((skyscraper_presheaf_obj_of_not_mem _ (mem_nhds_of_not_mem_closure_singleton p₀ h).some_spec).symm)) ≫ - presheaf.germ (skyscraper_presheaf p₀ S ts) - ⟨y, (mem_nhds_of_not_mem_closure_singleton p₀ h).some.2⟩ ≫ colimit.desc _ _ : star ⟶ _), + presheaf.germ (skyscraper_presheaf p₀ S) + ⟨y, (mem_nhds_of_not_mem_closure_singleton p₀ h).some.2⟩ ≫ colimit.desc _ _, fac' := λ c U, begin simp only [presheaf.germ, skyscraper_presheaf_cocone_of_not_mem_closure₀_ι_app], erw [colimit.ι_desc], dsimp, by_cases h' : p₀ ∈ (open_nhds.inclusion y).obj (unop U), - { have eq1 : ts.from (ite (p₀ ∈ (open_nhds.inclusion y).obj (unop U)) S star) = - eq_to_hom (if_pos h') ≫ ts.from _ := ts.hom_ext _ _, + { have eq1 : terminal.from (ite (p₀ ∈ (open_nhds.inclusion y).obj (unop U)) S (terminal C)) = + eq_to_hom (if_pos h') ≫ terminal.from _ := terminal_is_terminal.hom_ext _ _, rw [eq1, category.assoc, eq_eq_to_hom_comp], have := c.ι.naturality (hom_of_le (le_top : U.unop ≤ ⊤)).op, dsimp at this, @@ -292,22 +289,23 @@ noncomputable lemma skyscraper_presheaf_cocone_of_not_mem_closure₀_is_colimit have h''' : p₀ ∉ (mem_nhds_of_not_mem_closure_singleton p₀ h).some.1 := (mem_nhds_of_not_mem_closure_singleton p₀ h).some_spec, split_ifs at this, - have eq2 : ts.from (ite (p₀ ∈ (open_nhds.inclusion y).obj ⊤) S star) = - eq_to_hom (if_pos h'') ≫ ts.from _ := ts.hom_ext _ _, + have eq2 : terminal.from (ite (p₀ ∈ (open_nhds.inclusion y).obj ⊤) S (terminal C)) = + eq_to_hom (if_pos h'') ≫ terminal.from _ := terminal_is_terminal.hom_ext _ _, rw [category.comp_id, eq2] at this, simp only [←this, ←category.assoc], convert eq_whisker _ _, { ext, refl, }, { rw [eq_comp_eq_to_hom], - exact ts.hom_ext _ _ } }, - { have eq1 : ts.from (ite (p₀ ∈ (open_nhds.inclusion y).obj (unop U)) S star) = - eq_to_hom (if_neg h') ≫ 𝟙 star := ts.hom_ext _ _, - have eq2 : ts.from (ite (p₀ ∈ (open_nhds.inclusion y).obj ⊤) S star) = - eq_to_hom (if_pos trivial) ≫ ts.from _ := ts.hom_ext _ _, - have eq3 : ts.from (ite (p₀ ∈ (open_nhds.inclusion y).obj - (mem_nhds_of_not_mem_closure_singleton p₀ h).some) S star) = - eq_to_hom (if_neg (mem_nhds_of_not_mem_closure_singleton p₀ h).some_spec) ≫ 𝟙 star := - ts.hom_ext _ _, + exact terminal_is_terminal.hom_ext _ _ } }, + { + have eq1 : terminal.from (ite (p₀ ∈ (open_nhds.inclusion y).obj (unop U)) S (terminal C)) = + eq_to_hom (if_neg h') ≫ 𝟙 (terminal C) := terminal_is_terminal.hom_ext _ _, + have eq2 : terminal.from (ite (p₀ ∈ (open_nhds.inclusion y).obj ⊤) S (terminal C)) = + eq_to_hom (if_pos trivial) ≫ terminal.from _ := terminal_is_terminal.hom_ext _ _, + have eq3 : terminal.from (ite (p₀ ∈ (open_nhds.inclusion y).obj + (mem_nhds_of_not_mem_closure_singleton p₀ h).some) S (terminal C)) = + eq_to_hom (if_neg (mem_nhds_of_not_mem_closure_singleton p₀ h).some_spec) ≫ 𝟙 (terminal C):= + terminal_is_terminal.hom_ext _ _, rw [eq1, category.comp_id, ←category.assoc, eq_to_hom_trans], have := c.ι.naturality (hom_of_le (inf_le_left : (mem_nhds_of_not_mem_closure_singleton p₀ h).some ⊓ unop U ≤ @@ -315,53 +313,45 @@ noncomputable lemma skyscraper_presheaf_cocone_of_not_mem_closure₀_is_colimit dsimp at this, rw [dite_comp, category.comp_id] at this, simp_rw [category.id_comp, eq_to_hom_trans, eq3, category.comp_id, eq_to_hom_trans] at this, - - generalize_proofs h₁ h₂ h₃, - have eq_coe : c.ι.app (op ⟨↑(h₁.some), h₃⟩) = + have eq_coe : c.ι.app (op ⟨h₁.some.1, h₃⟩) = eq_to_hom (by { congr, ext, refl }) ≫ c.ι.app (op h₁.some) ≫ eq_to_hom (by { congr, ext, refl }), { symmetry, rw [eq_eq_to_hom_comp], have e := c.ι.naturality (eq_to_hom (by {ext, refl}) : - (⟨↑h₁.some, h₃⟩ : open_nhds y) ⟶ h₁.some).op, + (⟨h₁.some.1, h₃⟩ : open_nhds y) ⟶ h₁.some).op, dsimp at e, split_ifs at e with temp, { exfalso, exact (mem_nhds_of_not_mem_closure_singleton p₀ h).some_spec temp }, - rw [show ts.from (ite (p₀ ∈ (open_nhds.inclusion y).obj h₁.some) S star) = - eq_to_hom (if_neg temp) ≫ 𝟙 star, from ts.hom_ext _ _, category.comp_id, - category.comp_id, eq_to_hom_trans] at e, + rw [show terminal.from (ite (p₀ ∈ (open_nhds.inclusion y).obj h₁.some) S (terminal C)) = + eq_to_hom (if_neg temp) ≫ 𝟙 (terminal C), from terminal_is_terminal.hom_ext _ _, + category.comp_id, category.comp_id, eq_to_hom_trans] at e, rw [←e, category.assoc, eq_eq_to_hom_comp, ←category.assoc, eq_to_hom_trans, eq_to_hom_refl, eq_to_hom_refl, category.comp_id, category.id_comp], }, - rw [eq_coe, ←this, ←category.assoc, eq_to_hom_trans, eq_eq_to_hom_comp, dite_comp], - have h₅ : p₀ ∉ (open_nhds.inclusion y).obj h₁.some := h₁.some_spec, - have h₆ : p₀ ∉ (open_nhds.inclusion y).obj (h₁.some ⊓ unop U) := λ ⟨r, _⟩, h₅ r, + simp only [eq_coe], + rw [←this, ←category.assoc, eq_to_hom_trans, eq_to_hom_refl, category.comp_id, comp_dite], clear this, + generalize_proofs, + have h₄ : p₀ ∉ (open_nhds.inclusion y).obj (h₁.some ⊓ unop U) := λ r, h' r.2, split_ifs, - { exfalso, exact h₅ (by assumption) }, - { exfalso, exact h₆ (by assumption) }, - { exfalso, exact h₅ (by assumption) }, - rw [eq_comp_eq_to_hom, eq_eq_to_hom_comp, ←category.assoc, ←category.assoc, eq_to_hom_trans, - eq_to_hom_refl, category.comp_id], have := c.ι.naturality (hom_of_le (inf_le_right : (mem_nhds_of_not_mem_closure_singleton p₀ h).some ⊓ unop U ≤ unop U)).op, dsimp at this, rw [dite_comp, category.comp_id] at this, - rw [←this, comp_dite], - clear this, - split_ifs, - rw [eq1, category.comp_id, eq_to_hom_trans, ←category.assoc, eq_to_hom_trans, eq_to_hom_refl, - category.id_comp] } + split_ifs at this, + rw [←this, eq1, category.comp_id, eq_to_hom_trans, ←category.assoc, eq_to_hom_trans], } end, uniq' := λ c f H, begin erw [colimit.ι_desc, ←H], simp only [skyscraper_presheaf_cocone_of_not_mem_closure₀_ι_app], dsimp, - have : ts.from (ite (p₀ ∈ (mem_nhds_of_not_mem_closure_singleton p₀ h).some.1) S star) = - eq_to_hom (if_neg (mem_nhds_of_not_mem_closure_singleton p₀ h).some_spec) ≫ 𝟙 star := - ts.hom_ext _ _, - erw [this, category.comp_id, ←category.assoc, eq_to_hom_trans, eq_to_hom_refl, category.id_comp] + have : terminal.from + (ite (p₀ ∈ (mem_nhds_of_not_mem_closure_singleton p₀ h).some.1) S (terminal C)) = + eq_to_hom (if_neg (mem_nhds_of_not_mem_closure_singleton p₀ h).some_spec) := + terminal_is_terminal.hom_ext _ _, + erw [this,←category.assoc, eq_to_hom_trans, eq_to_hom_refl, category.id_comp] end } /-- @@ -369,8 +359,7 @@ If `y ∉ closure {p₀}`, then the stalk of `skyscraper_presheaf p₀ S t` at ` -/ @[reducible] noncomputable def skyscraper_presheaf_stalk_of_not_mem_closure₀ - {y : X} (h : ¬p₀ ⤳ y) : - (skyscraper_presheaf p₀ S ts).stalk y ≅ star := -colimit.iso_colimit_cocone ⟨_, (skyscraper_presheaf_cocone_of_not_mem_closure₀_is_colimit _ S _ h)⟩ + {y : X} (h : ¬p₀ ⤳ y) : (skyscraper_presheaf p₀ S).stalk y ≅ terminal C := +colimit.iso_colimit_cocone ⟨_, (skyscraper_presheaf_cocone_of_not_mem_closure₀_is_colimit _ S h)⟩ end From adfff44c6b3f41ff9192b2b088243a4540862228 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Mon, 15 Aug 2022 22:19:11 +0100 Subject: [PATCH 60/84] fix linter error --- src/topology/sheaves/skyscraper.lean | 114 ++++++++++++++------------- 1 file changed, 58 insertions(+), 56 deletions(-) diff --git a/src/topology/sheaves/skyscraper.lean b/src/topology/sheaves/skyscraper.lean index f83f64c183c77..3ba84e1f912b4 100644 --- a/src/topology/sheaves/skyscraper.lean +++ b/src/topology/sheaves/skyscraper.lean @@ -18,16 +18,16 @@ point, then the skyscraper (pre)sheaf `𝓕` with value `A` is defined by `U ↦ ## Main definitions -* `skyscraper_presheaf`: the skyscraper presheaf, to define a skyscraper presheaf at point `p₀` with - value `A`, use `skypscraper_presheaf p₀ A t` where `t` is a proof that `*` is a terminal object. +* `skyscraper_presheaf`: `skyscraper_presheaf p₀ A` is the skyscraper presheaf at point `p₀` with + value `A`. * `skyscraper_sheaf`: the skyscraper presheaf satisfies the sheaf condition. ## Main statements * `skyscraper_stalk_of_mem_closure₀`: if `y ∈ closure {p₀}` then the stalk of - `skyscraper_presheaf p₀ S t` at `y` is `S`. + `skyscraper_presheaf p₀ S` at `y` is `S`. * `skyscraper_stalk_of_not_mem_closure₀`: if `y ∉ closure {p₀}` then the stalk of - `skyscraper_presheaf p₀ S t` at `y` is `*`. + `skyscraper_presheaf p₀ S` at `y` is `*` the terminal object. -/ @@ -45,7 +45,7 @@ universes u v w variables {X : Top.{u}} (p₀ : X) {C : Type v} [category.{w} C] (S : C) variables [has_terminal C] [Π (U : opens X), decidable (p₀ ∈ U)] -/- +/-- A skyscraper presheaf is a presheaf supported at a single point: if `p₀ ∈ X` is a specified point, then the skyscraper presheaf `𝓕` with value `A` is defined by `U ↦ A` if `p₀ ∈ U` and `U ↦ *` if `p₀ ∉ A` where `*` is some terminal object. @@ -88,7 +88,7 @@ if_neg h end -/- +/-- A skyscraper sheaf is a sheaf supported at a single point: if `p₀ ∈ X` is a specified point, then the skyscraper sheaf `𝓕` with value `A` is defined by `U ↦ A` if `p₀ ∈ U` and `U ↦ *` if `p₀ ∉ A` where `*` is some terminal object. @@ -153,9 +153,7 @@ open opposite universes u v -variables {X : Top.{u}} (p₀ : X) {C : Type v} [category.{u} C] (S : C) -variables [has_terminal C] -variables [Π (U : opens X), decidable (p₀ ∈ U)] [has_colimits C] +variables {X : Top.{u}} (p₀ : X) {C : Type v} [category.{u} C] (S : C) [has_terminal C] lemma mem_nhds_of_not_mem_closure_singleton {y : X} (h : ¬p₀ ⤳ y) : ∃ (U : open_nhds y), p₀ ∉ U.1 := @@ -166,8 +164,9 @@ begin exact ⟨⟨⟨s, o⟩, h₁⟩, h₂⟩, end +variable [Π (U : opens X), decidable (p₀ ∈ U)] /-- -The cocone at `S` for the salk functor of `skyscraper_presheaf p₀ S t` when `y ∈ closure {p₀}` +The cocone at `S` for the salk functor of `skyscraper_presheaf p₀ S` when `y ∈ closure {p₀}` -/ @[simps] def skyscraper_presheaf_cocone_of_mem_closure₀ {y : X} (h : p₀ ⤳ y) : cocone ((open_nhds.inclusion y).op ⋙ skyscraper_presheaf p₀ S) := @@ -197,18 +196,22 @@ The cocone at `S` for the salk functor of `skyscraper_presheaf p₀ S t` when `y /-- The canonical map `S ⟶ (skyscraper_presheaf p₀ S t).stalk y` when `y ∈ closure {p₀}` -/ -noncomputable def skyscraper_presheaf_of_mem_closure₀_from {y : X} (h : p₀ ⤳ y) : +noncomputable def skyscraper_presheaf_from [has_colimits C] {y : X} : S ⟶ (skyscraper_presheaf p₀ S).stalk y := eq_to_hom (skyscraper_presheaf_obj_of_mem S (by tauto : p₀ ∈ ⊤)).symm ≫ (skyscraper_presheaf p₀ S).germ (⟨y, trivial⟩ : (⊤ : opens X)) -noncomputable lemma skyscraper_presheaf_cocone_of_mem_closure₀_is_colimit {y : X} (h : p₀ ⤳ y) : - is_colimit (skyscraper_presheaf_cocone_of_mem_closure₀ p₀ S h) := -{ desc := λ c, (skyscraper_presheaf_of_mem_closure₀_from p₀ S h ≫ colimit.desc _ _ : S ⟶ c.X), +/-- +The cocone at `S` for the salk functor of `skyscraper_presheaf p₀ S` when `y ∈ closure {p₀}` is a +colimit +-/ +noncomputable def skyscraper_presheaf_cocone_of_mem_closure₀_is_colimit [has_colimits C] + {y : X} (h : p₀ ⤳ y) : is_colimit (skyscraper_presheaf_cocone_of_mem_closure₀ p₀ S h) := +{ desc := λ c, (skyscraper_presheaf_from p₀ S ≫ colimit.desc _ _ : S ⟶ c.X), fac' := λ c U, begin dsimp, - simp only [skyscraper_presheaf_of_mem_closure₀_from, presheaf.germ, category.comp_id, + simp only [skyscraper_presheaf_from, presheaf.germ, category.comp_id, category.assoc, colimit.ι_desc, eq_to_hom_trans_assoc], have := c.ι.naturality (hom_of_le $ (le_top : unop U ≤ _)).op, dsimp at this, @@ -222,7 +225,7 @@ noncomputable lemma skyscraper_presheaf_cocone_of_mem_closure₀_is_colimit {y : end, uniq' := λ c f h, begin - simp only [skyscraper_presheaf_of_mem_closure₀_from, presheaf.germ, category.assoc], + simp only [skyscraper_presheaf_from, presheaf.germ, category.assoc], erw [colimit.ι_desc], specialize h (op ⟨⊤, trivial⟩), erw [←h], @@ -232,17 +235,17 @@ noncomputable lemma skyscraper_presheaf_cocone_of_mem_closure₀_is_colimit {y : end } /-- -If `y ∈ closure {p₀}`, then the stalk of `skyscraper_presheaf p₀ S t` at `y` is `S` +If `y ∈ closure {p₀}`, then the stalk of `skyscraper_presheaf p₀ S` at `y` is `S` -/ @[reducible] -noncomputable def skyscraper_presheaf_stalk_of_mem_closure₀ {y : X} (h : p₀ ⤳ y) : - (skyscraper_presheaf p₀ S).stalk y ≅ S := +noncomputable def skyscraper_presheaf_stalk_of_mem_closure₀ [has_colimits C] + {y : X} (h : p₀ ⤳ y) : (skyscraper_presheaf p₀ S).stalk y ≅ S := colimit.iso_colimit_cocone ⟨_, (skyscraper_presheaf_cocone_of_mem_closure₀_is_colimit p₀ S h)⟩ /-- -The cocone at `*` for the salk functor of `skyscraper_presheaf p₀ S t` when `y ∉ closure {p₀}` +The cocone at `*` for the salk functor of `skyscraper_presheaf p₀ S` when `y ∉ closure {p₀}` -/ -@[simps] def skyscraper_presheaf_cocone_of_not_mem_closure₀ {y : X} (h : ¬p₀ ⤳ y) : +@[simps] def skyscraper_presheaf_cocone {y : X} : cocone ((open_nhds.inclusion y).op ⋙ skyscraper_presheaf p₀ S) := { X := terminal C, ι := @@ -250,26 +253,28 @@ The cocone at `*` for the salk functor of `skyscraper_presheaf p₀ S t` when `y naturality' := λ U V inc, terminal_is_terminal.hom_ext _ _ } } /-- -The canonical map `* ⟶ (skyscraper_presheaf p₀ S t).stalk y` when `y ∉ closure {p₀}` +The canonical map `* ⟶ (skyscraper_presheaf p₀ S).stalk y` when `y ∉ closure {p₀}` -/ -noncomputable def skyscraper_presheaf_of_not_mem_closure₀_from - {y : X} (h : ¬p₀ ⤳ y) : - terminal C ⟶ (skyscraper_presheaf p₀ S).stalk y := +noncomputable def skyscraper_presheaf_of_not_mem_closure₀_from [has_colimits C] + {y : X} (h : ¬p₀ ⤳ y) : terminal C ⟶ (skyscraper_presheaf p₀ S).stalk y := eq_to_hom (skyscraper_presheaf_obj_of_not_mem S $ (mem_nhds_of_not_mem_closure_singleton p₀ h).some_spec).symm ≫ (skyscraper_presheaf p₀ S).germ (⟨y, (mem_nhds_of_not_mem_closure_singleton p₀ h).some.2⟩ : (mem_nhds_of_not_mem_closure_singleton p₀ h).some.1) -noncomputable lemma skyscraper_presheaf_cocone_of_not_mem_closure₀_is_colimit - {y : X} (h : ¬p₀ ⤳ y) : - is_colimit (skyscraper_presheaf_cocone_of_not_mem_closure₀ p₀ S h) := +/-- +The cocone at `*` for the salk functor of `skyscraper_presheaf p₀ S` when `y ∉ closure {p₀}` is a +colimit +-/ +noncomputable def skyscraper_presheaf_cocone_of_not_mem_closure₀_is_colimit [has_colimits C] + {y : X} (h : ¬p₀ ⤳ y) : is_colimit (skyscraper_presheaf_cocone p₀ S) := { desc := λ c, (eq_to_hom ((skyscraper_presheaf_obj_of_not_mem _ (mem_nhds_of_not_mem_closure_singleton p₀ h).some_spec).symm)) ≫ presheaf.germ (skyscraper_presheaf p₀ S) ⟨y, (mem_nhds_of_not_mem_closure_singleton p₀ h).some.2⟩ ≫ colimit.desc _ _, fac' := λ c U, begin - simp only [presheaf.germ, skyscraper_presheaf_cocone_of_not_mem_closure₀_ι_app], + simp only [presheaf.germ, skyscraper_presheaf_cocone_ι_app], erw [colimit.ι_desc], dsimp, by_cases h' : p₀ ∈ (open_nhds.inclusion y).obj (unop U), @@ -281,7 +286,7 @@ noncomputable lemma skyscraper_presheaf_cocone_of_not_mem_closure₀_is_colimit have h'' : p₀ ∈ (open_nhds.inclusion y).obj ⊤ := trivial, split_ifs at this, rw [category.comp_id, category.id_comp, eq_to_hom_trans, eq_eq_to_hom_comp] at this, - rw this, + rw [this, ←category.assoc _ _ (c.ι.app (op ⊤)), eq_to_hom_trans], clear this, have := c.ι.naturality (hom_of_le (le_top : (mem_nhds_of_not_mem_closure_singleton p₀ h).some ≤ ⊤)).op, @@ -291,74 +296,71 @@ noncomputable lemma skyscraper_presheaf_cocone_of_not_mem_closure₀_is_colimit split_ifs at this, have eq2 : terminal.from (ite (p₀ ∈ (open_nhds.inclusion y).obj ⊤) S (terminal C)) = eq_to_hom (if_pos h'') ≫ terminal.from _ := terminal_is_terminal.hom_ext _ _, - rw [category.comp_id, eq2] at this, + rw [category.comp_id, eq2, category.assoc, category.assoc] at this, simp only [←this, ←category.assoc], convert eq_whisker _ _, { ext, refl, }, { rw [eq_comp_eq_to_hom], exact terminal_is_terminal.hom_ext _ _ } }, - { - have eq1 : terminal.from (ite (p₀ ∈ (open_nhds.inclusion y).obj (unop U)) S (terminal C)) = - eq_to_hom (if_neg h') ≫ 𝟙 (terminal C) := terminal_is_terminal.hom_ext _ _, + { have eq1 : terminal.from (ite (p₀ ∈ (open_nhds.inclusion y).obj (unop U)) S (terminal C)) = + eq_to_hom (if_neg h') := terminal_is_terminal.hom_ext _ _, have eq2 : terminal.from (ite (p₀ ∈ (open_nhds.inclusion y).obj ⊤) S (terminal C)) = eq_to_hom (if_pos trivial) ≫ terminal.from _ := terminal_is_terminal.hom_ext _ _, have eq3 : terminal.from (ite (p₀ ∈ (open_nhds.inclusion y).obj (mem_nhds_of_not_mem_closure_singleton p₀ h).some) S (terminal C)) = - eq_to_hom (if_neg (mem_nhds_of_not_mem_closure_singleton p₀ h).some_spec) ≫ 𝟙 (terminal C):= + eq_to_hom (if_neg (mem_nhds_of_not_mem_closure_singleton p₀ h).some_spec):= terminal_is_terminal.hom_ext _ _, - rw [eq1, category.comp_id, ←category.assoc, eq_to_hom_trans], + rw [eq1, ←category.assoc, eq_to_hom_trans], have := c.ι.naturality (hom_of_le (inf_le_left : (mem_nhds_of_not_mem_closure_singleton p₀ h).some ⊓ unop U ≤ (mem_nhds_of_not_mem_closure_singleton p₀ h).some)).op, dsimp at this, rw [dite_comp, category.comp_id] at this, - simp_rw [category.id_comp, eq_to_hom_trans, eq3, category.comp_id, eq_to_hom_trans] at this, - generalize_proofs h₁ h₂ h₃, - have eq_coe : c.ι.app (op ⟨h₁.some.1, h₃⟩) = + simp_rw [category.id_comp, eq_to_hom_trans, eq3, eq_to_hom_trans] at this, + generalize_proofs h₁ h₂ h₃ h₄ at this, + have h' : p₀ ∉ (open_nhds.inclusion y).obj (h₁.some ⊓ unop U) := λ r, h₁.some_spec r.1, + split_ifs at this, + { exfalso, exact h₁.some_spec (by assumption), }, + have eq_coe : c.ι.app (op ⟨h₁.some.1, h₂⟩) = eq_to_hom (by { congr, ext, refl }) ≫ c.ι.app (op h₁.some) ≫ eq_to_hom (by { congr, ext, refl }), { symmetry, rw [eq_eq_to_hom_comp], have e := c.ι.naturality (eq_to_hom (by {ext, refl}) : - (⟨h₁.some.1, h₃⟩ : open_nhds y) ⟶ h₁.some).op, + (⟨h₁.some.1, h₂⟩ : open_nhds y) ⟶ h₁.some).op, dsimp at e, split_ifs at e with temp, - { exfalso, exact (mem_nhds_of_not_mem_closure_singleton p₀ h).some_spec temp }, - rw [show terminal.from (ite (p₀ ∈ (open_nhds.inclusion y).obj h₁.some) S (terminal C)) = - eq_to_hom (if_neg temp) ≫ 𝟙 (terminal C), from terminal_is_terminal.hom_ext _ _, - category.comp_id, category.comp_id, eq_to_hom_trans] at e, - rw [←e, category.assoc, eq_eq_to_hom_comp, ←category.assoc, eq_to_hom_trans, eq_to_hom_refl, - eq_to_hom_refl, category.comp_id, category.id_comp], }, - simp only [eq_coe], - rw [←this, ←category.assoc, eq_to_hom_trans, eq_to_hom_refl, category.comp_id, comp_dite], + rw [category.comp_id, show terminal.from (ite (p₀ ∈ (open_nhds.inclusion y).obj h₁.some) + S (terminal C)) = eq_to_hom (if_neg h₁.some_spec), from terminal_is_terminal.hom_ext _ _, + eq_to_hom_trans, eq_to_hom_refl, category.id_comp] at e, + rw [←e, eq_to_hom_refl, category.comp_id, eq_to_hom_refl, category.id_comp], }, + erw [eq_coe, ←this, ←category.assoc, eq_to_hom_trans, eq_to_hom_refl, category.comp_id, + ←category.assoc, eq_to_hom_trans], clear this, - generalize_proofs, - have h₄ : p₀ ∉ (open_nhds.inclusion y).obj (h₁.some ⊓ unop U) := λ r, h' r.2, - split_ifs, have := c.ι.naturality (hom_of_le (inf_le_right : (mem_nhds_of_not_mem_closure_singleton p₀ h).some ⊓ unop U ≤ unop U)).op, dsimp at this, rw [dite_comp, category.comp_id] at this, split_ifs at this, - rw [←this, eq1, category.comp_id, eq_to_hom_trans, ←category.assoc, eq_to_hom_trans], } + rw [←this, eq1, eq_to_hom_trans], } end, uniq' := λ c f H, begin erw [colimit.ι_desc, ←H], - simp only [skyscraper_presheaf_cocone_of_not_mem_closure₀_ι_app], + simp only [skyscraper_presheaf_cocone_ι_app], dsimp, have : terminal.from (ite (p₀ ∈ (mem_nhds_of_not_mem_closure_singleton p₀ h).some.1) S (terminal C)) = eq_to_hom (if_neg (mem_nhds_of_not_mem_closure_singleton p₀ h).some_spec) := terminal_is_terminal.hom_ext _ _, - erw [this,←category.assoc, eq_to_hom_trans, eq_to_hom_refl, category.id_comp] + erw [this, ←category.assoc, eq_to_hom_trans, eq_to_hom_refl, category.id_comp] end } /-- -If `y ∉ closure {p₀}`, then the stalk of `skyscraper_presheaf p₀ S t` at `y` is `*` +If `y ∉ closure {p₀}`, then the stalk of `skyscraper_presheaf p₀ S` at `y` is `*` -/ @[reducible] -noncomputable def skyscraper_presheaf_stalk_of_not_mem_closure₀ +noncomputable def skyscraper_presheaf_stalk_of_not_mem_closure₀ [has_colimits C] {y : X} (h : ¬p₀ ⤳ y) : (skyscraper_presheaf p₀ S).stalk y ≅ terminal C := colimit.iso_colimit_cocone ⟨_, (skyscraper_presheaf_cocone_of_not_mem_closure₀_is_colimit _ S h)⟩ From fb707bc788d56456ef7298b4811750b95fd2758a Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Mon, 15 Aug 2022 22:33:52 +0100 Subject: [PATCH 61/84] fix --- src/topology/sheaves/skyscraper.lean | 29 +++++++++------------------- 1 file changed, 9 insertions(+), 20 deletions(-) diff --git a/src/topology/sheaves/skyscraper.lean b/src/topology/sheaves/skyscraper.lean index 3ba84e1f912b4..40c3962546044 100644 --- a/src/topology/sheaves/skyscraper.lean +++ b/src/topology/sheaves/skyscraper.lean @@ -59,8 +59,7 @@ point, then the skyscraper presheaf `𝓕` with value `A` is defined by `U ↦ A begin split_ifs, { simp, }, - { rw eq_comp_eq_to_hom, - exact terminal_is_terminal.hom_ext _ _, }, + { simpa only [eq_comp_eq_to_hom] using terminal_is_terminal.hom_ext _ _, }, end, map_comp' := λ U V W iVU iWV, begin @@ -70,8 +69,7 @@ point, then the skyscraper presheaf `𝓕` with value `A` is defined by `U ↦ A split_ifs, simp }, { split_ifs; - rw eq_comp_eq_to_hom; - exact terminal_is_terminal.hom_ext _ _, } + simpa only [eq_comp_eq_to_hom] using terminal_is_terminal.hom_ext _ _, } end } section @@ -79,12 +77,10 @@ section variables {p₀} lemma skyscraper_presheaf_obj_of_mem {U : opens X} (h : p₀ ∈ U) : - (skyscraper_presheaf p₀ S).obj (op U) = S := -if_pos h + (skyscraper_presheaf p₀ S).obj (op U) = S := if_pos h lemma skyscraper_presheaf_obj_of_not_mem {U : opens X} (h : p₀ ∉ U) : - (skyscraper_presheaf p₀ S).obj (op U) = (terminal C) := -if_neg h + (skyscraper_presheaf p₀ S).obj (op U) = terminal C := if_neg h end @@ -143,8 +139,8 @@ end section --- In this section, we calculate the stalks for skyscraper presheaves. We need to restrict universe --- level. +-- In this section, we calculate the stalks for skyscraper presheaves. +-- We need to restrict universe level. open topological_space open category_theory category_theory.limits @@ -172,12 +168,7 @@ The cocone at `S` for the salk functor of `skyscraper_presheaf p₀ S` when `y cocone ((open_nhds.inclusion y).op ⋙ skyscraper_presheaf p₀ S) := { X := S, ι := - { app := λ U, eq_to_hom - begin - dsimp, - rw if_pos, - exact h.mem_open U.unop.1.2 U.unop.2, - end ≫ 𝟙 S, + { app := λ U, eq_to_hom $ if_pos (h.mem_open U.unop.1.2 U.unop.2), naturality' := λ U V inc, begin simp only [functor.op_obj, unop_op, functor.comp_map, functor.op_map, skyscraper_presheaf_map, @@ -185,11 +176,10 @@ The cocone at `S` for the salk functor of `skyscraper_presheaf p₀ S` when `y by_cases hV : p₀ ∈ (open_nhds.inclusion y).obj V.unop, { have hU : p₀ ∈ unop ((open_nhds.inclusion y).op.obj U) := le_of_hom inc.unop hV, split_ifs, - erw [category.comp_id, category.comp_id, category.comp_id, eq_to_hom_trans], + erw [eq_to_hom_trans, category.comp_id], refl }, { split_ifs with hU; - erw [category.comp_id, category.comp_id, category.comp_id, eq_comp_eq_to_hom, - eq_comp_eq_to_hom]; + erw [category.comp_id, category.assoc, eq_to_hom_trans, eq_comp_eq_to_hom, eq_to_hom_trans]; exact terminal_is_terminal.hom_ext _ _, }, end } } @@ -231,7 +221,6 @@ noncomputable def skyscraper_presheaf_cocone_of_mem_closure₀_is_colimit [has_c erw [←h], simp only [skyscraper_presheaf_cocone_of_mem_closure₀_ι_app, category.assoc, eq_to_hom_trans_assoc, eq_to_hom_refl, category.id_comp], - exact (category.id_comp _).symm, end } /-- From aa6a325b9f827dd248e5ed6331ad2b5e6f83cf68 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Tue, 16 Aug 2022 02:37:02 +0100 Subject: [PATCH 62/84] move fiel --- src/algebraic_geometry/skyscraper.lean | 855 ------------------------- 1 file changed, 855 deletions(-) delete mode 100644 src/algebraic_geometry/skyscraper.lean diff --git a/src/algebraic_geometry/skyscraper.lean b/src/algebraic_geometry/skyscraper.lean deleted file mode 100644 index 5f072ab36c2a0..0000000000000 --- a/src/algebraic_geometry/skyscraper.lean +++ /dev/null @@ -1,855 +0,0 @@ -/- -Copyright (c) 2022 Jujian Zhang. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Jujian Zhang --/ -import algebraic_geometry.sheafed_space -import topology.sheaves.sheaf_condition.unique_gluing -import topology.sheaves.stalks -import category_theory.preadditive.injective -import algebra.category.Group.abelian - -/-! -# Skyscraper (pre)sheaves - -A skyscraper (pre)sheaf is a (pre)sheaf supported at a single point: if `p₀ ∈ X` is a specified -point, then the skyscraper (pre)sheaf `𝓕` with value `A` is defined by `U ↦ A` if `p₀ ∈ U` and -`U ↦ *` if `p₀ ∉ A` where `*` is some terminal object. - -## Main definitions - -* `skyscraper_presheaf`: the skyscraper presheaf, to define a skyscraper presheaf at point `p₀` with - value `A`, use `skypscraper_presheaf p₀ A t` where `t` is a proof that `*` is a terminal object. -* `skyscraper_sheaf`: the skyscraper presheaf satisfies the sheaf condition. - -## Main statements - -* `skyscraper_stalk_of_mem_closure₀`: if `y ∈ closure {p₀}` then the stalk of - `skyscraper_presheaf p₀ S t` at `y` is `S`. -* `skyscraper_stalk_of_not_mem_closure₀`: if `y ∉ closure {p₀}` then the stalk of - `skyscraper_presheaf p₀ S t` at `y` is `*`. - --/ - - -section - -open topological_space -open category_theory category_theory.limits -open Top -open opposite - -universes u v w - -variables {X : Top.{u}} (p₀ : X) {C : Type v} [category.{w} C] (S : C) -variables {star : C} (ts : is_terminal star) -variable [Π (U : opens X), decidable (p₀ ∈ U)] - -/- -A skyscraper presheaf is a presheaf supported at a single point: if `p₀ ∈ X` is a specified -point, then the skyscraper presheaf `𝓕` with value `A` is defined by `U ↦ A` if `p₀ ∈ U` and -`U ↦ *` if `p₀ ∉ A` where `*` is some terminal object. --/ -@[simps] -def skyscraper_presheaf : presheaf C X := -{ obj := λ U, ite (p₀ ∈ unop U) S star, - map := λ U V i, dite (p₀ ∈ unop V) - (λ h, eq_to_hom (if_pos (le_of_hom i.unop h)) ≫ 𝟙 S ≫ eq_to_hom (if_pos h).symm) - (λ h, ts.from _ ≫ eq_to_hom (if_neg h).symm), - map_id' := λ U, - begin - split_ifs, - { simp, }, - { rw eq_comp_eq_to_hom, - exact ts.hom_ext _ _, }, - end, - map_comp' := λ U V W iVU iWV, - begin - by_cases hW : p₀ ∈ unop W, - { have hV : p₀ ∈ unop V := le_of_hom iWV.unop hW, - have hU : p₀ ∈ unop U := le_of_hom iVU.unop hV, - split_ifs, - simp }, - { split_ifs; - rw eq_comp_eq_to_hom; - exact ts.hom_ext _ _, } - end } - -section - -variables {p₀} - -lemma skyscraper_presheaf_obj_of_mem {U : opens X} (h : p₀ ∈ U) : - (skyscraper_presheaf p₀ S ts).obj (op U) = S := -if_pos h - -lemma skyscraper_presheaf_obj_of_not_mem {U : opens X} (h : p₀ ∉ U) : - (skyscraper_presheaf p₀ S ts).obj (op U) = star := -if_neg h - -end - -/- -A skyscraper sheaf is a sheaf supported at a single point: if `p₀ ∈ X` is a specified -point, then the skyscraper sheaf `𝓕` with value `A` is defined by `U ↦ A` if `p₀ ∈ U` and -`U ↦ *` if `p₀ ∉ A` where `*` is some terminal object. --/ -def skyscraper_sheaf : sheaf C X := -⟨skyscraper_presheaf p₀ S ts, λ c U s hs x hx, - ⟨dite (p₀ ∈ U) - (λ h, x (hs p₀ h).some_spec.some (hs p₀ h).some_spec.some_spec.1 ≫ - eq_to_hom ((skyscraper_presheaf_obj_of_mem S ts (hs p₀ h).some_spec.some_spec.2).trans - (skyscraper_presheaf_obj_of_mem S ts h).symm)) - (λ h, ts.from c ≫ (eq_to_hom (skyscraper_presheaf_obj_of_not_mem S ts h).symm)), - λ V inc h, - begin - by_cases hV : p₀ ∈ V, - { have hU : p₀ ∈ U := le_of_hom inc hV, - split_ifs, - generalize_proofs h₁ h₂ h₃ h₄, - dsimp, - split_ifs, - rw [category.id_comp, eq_to_hom_trans, category.assoc, eq_to_hom_trans], - generalize_proofs h₅, - have := hx (hom_of_le inf_le_left) (hom_of_le inf_le_right) h₂.some_spec.1 h rfl, - dsimp at this, - have hV' : p₀ ∈ h₁.some ⊓ V := ⟨h₂.some_spec.2, hV⟩, - split_ifs at this; - rw [category.id_comp, eq_to_hom_trans, eq_to_hom_trans, eq_comp_eq_to_hom, - category.assoc, eq_to_hom_trans] at this; - generalize_proofs h₆ at this; - rw [this, category.assoc, eq_to_hom_trans, eq_to_hom_refl, category.comp_id] }, - { dsimp, - split_ifs, - rw [←category.assoc, eq_comp_eq_to_hom], - exact ts.hom_ext _ _, } - end, - λ y (hy : x.is_amalgamation y), - begin - split_ifs, - { generalize_proofs h₁ h₂ h₃ h₄, - have := hy h₂.some h₂.some_spec.1, - dsimp at this, - split_ifs at this with H, - { rw [category.id_comp, eq_to_hom_trans, eq_comp_eq_to_hom] at this, - rw [this, eq_comp_eq_to_hom, category.assoc, eq_to_hom_trans, eq_to_hom_refl, - category.comp_id], }, - { exfalso, - exact H h₂.some_spec.2, }, }, - { rw [←eq_comp_eq_to_hom], - exact (ts.hom_ext _ _), } - end⟩⟩ - -end - -section - --- In this section, we calculate the stalks for skyscraper presheaves. We need to restrict universe --- level. - -open topological_space -open category_theory category_theory.limits -open Top -open opposite - -universes u v - -variables {X : Top.{u}} (p₀ : X) {C : Type v} [category.{u} C] (S : C) -variables {star : C} (ts : is_terminal star) -variable [Π (U : opens X), decidable (p₀ ∈ U)] -variable [has_colimits C] - -lemma mem_nhds_of_not_mem_closure_singleton {y : X} (h : ¬p₀ ⤳ y) : - ∃ (U : open_nhds y), p₀ ∉ U.1 := -begin - have := (not_iff_not.mpr specializes_iff_forall_open).mp h, - push_neg at this, - rcases this with ⟨s, o, h₁, h₂⟩, - exact ⟨⟨⟨s, o⟩, h₁⟩, h₂⟩, -end - -/-- -The cocone at `S` for the salk functor of `skyscraper_presheaf p₀ S t` when `y ∈ closure {p₀}` --/ -@[simps] def skyscraper_presheaf_cocone_of_mem_closure₀ {y : X} (h : p₀ ⤳ y) : - cocone ((open_nhds.inclusion y).op ⋙ skyscraper_presheaf p₀ S ts) := -{ X := S, - ι := - { app := λ U, eq_to_hom - begin - dsimp, - rw if_pos, - exact h.mem_open U.unop.1.2 U.unop.2, - end ≫ 𝟙 S, - naturality' := λ U V inc, - begin - simp only [functor.op_obj, unop_op, functor.comp_map, functor.op_map, skyscraper_presheaf_map, - category.id_comp, eq_to_hom_trans, functor.const_obj_map, category.assoc], - by_cases hV : p₀ ∈ (open_nhds.inclusion y).obj V.unop, - { have hU : p₀ ∈ unop ((open_nhds.inclusion y).op.obj U) := le_of_hom inc.unop hV, - split_ifs, - erw [category.comp_id, category.comp_id, category.comp_id, eq_to_hom_trans], - refl }, - { split_ifs with hU; - erw [category.comp_id, category.comp_id, category.comp_id, eq_comp_eq_to_hom, - eq_comp_eq_to_hom]; - exact ts.hom_ext _ _, }, - end } } - -/-- -The canonical map `S ⟶ (skyscraper_presheaf p₀ S t).stalk y` when `y ∈ closure {p₀}` --/ -noncomputable def skyscraper_presheaf_of_mem_closure₀_from {y : X} (h : p₀ ⤳ y) : - S ⟶ (skyscraper_presheaf p₀ S ts).stalk y := -eq_to_hom (skyscraper_presheaf_obj_of_mem S ts (by tauto : p₀ ∈ ⊤)).symm ≫ - (skyscraper_presheaf p₀ S ts).germ (⟨y, trivial⟩ : (⊤ : opens X)) - -noncomputable lemma skyscraper_presheaf_cocone_of_mem_closure₀_is_colimit {y : X} (h : p₀ ⤳ y) : - is_colimit (skyscraper_presheaf_cocone_of_mem_closure₀ p₀ S ts h) := -{ desc := λ c, (skyscraper_presheaf_of_mem_closure₀_from p₀ S ts h ≫ colimit.desc _ _ : S ⟶ c.X), - fac' := λ c U, - begin - dsimp, - simp only [skyscraper_presheaf_of_mem_closure₀_from, presheaf.germ, category.comp_id, - category.assoc, colimit.ι_desc, eq_to_hom_trans_assoc], - have := c.ι.naturality (hom_of_le $ (le_top : unop U ≤ _)).op, - dsimp at this, - have h' : p₀ ∈ (open_nhds.inclusion y).obj (unop U) := h.mem_open U.unop.1.2 U.unop.2, - have h'' : p₀ ∈ (open_nhds.inclusion y).obj ⊤ := trivial, - split_ifs at this, - rw [category.comp_id, category.id_comp, eq_to_hom_trans, eq_eq_to_hom_comp] at this, - rw [this, eq_eq_to_hom_comp, ←category.assoc, eq_to_hom_trans, eq_to_hom_refl, - category.id_comp], - congr, - end, - uniq' := λ c f h, - begin - simp only [skyscraper_presheaf_of_mem_closure₀_from, presheaf.germ, category.assoc], - erw [colimit.ι_desc], - specialize h (op ⟨⊤, trivial⟩), - erw [←h], - simp only [skyscraper_presheaf_cocone_of_mem_closure₀_ι_app, category.assoc, - eq_to_hom_trans_assoc, eq_to_hom_refl, category.id_comp], - exact (category.id_comp _).symm, - end } - -/-- -If `y ∈ closure {p₀}`, then the stalk of `skyscraper_presheaf p₀ S t` at `y` is `S` --/ -@[reducible] -noncomputable def skyscraper_presheaf_stalk_of_mem_closure₀ {y : X} (h : p₀ ⤳ y) : - (skyscraper_presheaf p₀ S ts).stalk y ≅ S := -colimit.iso_colimit_cocone ⟨_, (skyscraper_presheaf_cocone_of_mem_closure₀_is_colimit p₀ S ts h)⟩ - -/-- -The cocone at `*` for the salk functor of `skyscraper_presheaf p₀ S t` when `y ∉ closure {p₀}` --/ -@[simps] def skyscraper_presheaf_cocone_of_not_mem_closure₀ - {y : X} (h : ¬p₀ ⤳ y) : - cocone ((open_nhds.inclusion y).op ⋙ skyscraper_presheaf p₀ S ts) := -{ X := star, - ι := - { app := λ U, ts.from _, - naturality' := λ U V inc, ts.hom_ext _ _ } } - -/-- -The canonical map `* ⟶ (skyscraper_presheaf p₀ S t).stalk y` when `y ∉ closure {p₀}` --/ -noncomputable def skyscraper_presheaf_of_not_mem_closure₀_from - {y : X} (h : ¬p₀ ⤳ y) : - star ⟶ (skyscraper_presheaf p₀ S ts).stalk y := -eq_to_hom (skyscraper_presheaf_obj_of_not_mem S ts $ - (mem_nhds_of_not_mem_closure_singleton p₀ h).some_spec).symm ≫ - (skyscraper_presheaf p₀ S ts).germ (⟨y, (mem_nhds_of_not_mem_closure_singleton p₀ h).some.2⟩ : - (mem_nhds_of_not_mem_closure_singleton p₀ h).some.1) - -noncomputable lemma skyscraper_presheaf_cocone_of_not_mem_closure₀_is_colimit - {y : X} (h : ¬p₀ ⤳ y) : - is_colimit (skyscraper_presheaf_cocone_of_not_mem_closure₀ p₀ S ts h) := -{ desc := λ c, ((eq_to_hom ((skyscraper_presheaf_obj_of_not_mem _ _ - (mem_nhds_of_not_mem_closure_singleton p₀ h).some_spec).symm)) ≫ - presheaf.germ (skyscraper_presheaf p₀ S ts) - ⟨y, (mem_nhds_of_not_mem_closure_singleton p₀ h).some.2⟩ ≫ colimit.desc _ _ : star ⟶ _), - fac' := λ c U, - begin - simp only [presheaf.germ, skyscraper_presheaf_cocone_of_not_mem_closure₀_ι_app], - erw [colimit.ι_desc], - dsimp, - by_cases h' : p₀ ∈ (open_nhds.inclusion y).obj (unop U), - { have eq1 : ts.from (ite (p₀ ∈ (open_nhds.inclusion y).obj (unop U)) S star) = - eq_to_hom (if_pos h') ≫ ts.from _ := ts.hom_ext _ _, - rw [eq1, category.assoc, eq_eq_to_hom_comp], - have := c.ι.naturality (hom_of_le (le_top : U.unop ≤ ⊤)).op, - dsimp at this, - have h'' : p₀ ∈ (open_nhds.inclusion y).obj ⊤ := trivial, - split_ifs at this, - rw [category.comp_id, category.id_comp, eq_to_hom_trans, eq_eq_to_hom_comp] at this, - rw this, - clear this, - have := c.ι.naturality - (hom_of_le (le_top : (mem_nhds_of_not_mem_closure_singleton p₀ h).some ≤ ⊤)).op, - dsimp at this, - have h''' : p₀ ∉ (mem_nhds_of_not_mem_closure_singleton p₀ h).some.1 := - (mem_nhds_of_not_mem_closure_singleton p₀ h).some_spec, - split_ifs at this, - have eq2 : ts.from (ite (p₀ ∈ (open_nhds.inclusion y).obj ⊤) S star) = - eq_to_hom (if_pos h'') ≫ ts.from _ := ts.hom_ext _ _, - rw [category.comp_id, eq2] at this, - simp only [←this, ←category.assoc], - convert eq_whisker _ _, - { ext, refl, }, - { rw [eq_comp_eq_to_hom], - exact ts.hom_ext _ _ } }, - { have eq1 : ts.from (ite (p₀ ∈ (open_nhds.inclusion y).obj (unop U)) S star) = - eq_to_hom (if_neg h') ≫ 𝟙 star := ts.hom_ext _ _, - have eq2 : ts.from (ite (p₀ ∈ (open_nhds.inclusion y).obj ⊤) S star) = - eq_to_hom (if_pos trivial) ≫ ts.from _ := ts.hom_ext _ _, - have eq3 : ts.from (ite (p₀ ∈ (open_nhds.inclusion y).obj - (mem_nhds_of_not_mem_closure_singleton p₀ h).some) S star) = - eq_to_hom (if_neg (mem_nhds_of_not_mem_closure_singleton p₀ h).some_spec) ≫ 𝟙 star := - ts.hom_ext _ _, - rw [eq1, category.comp_id, ←category.assoc, eq_to_hom_trans], - have := c.ι.naturality (hom_of_le (inf_le_left : - (mem_nhds_of_not_mem_closure_singleton p₀ h).some ⊓ unop U ≤ - (mem_nhds_of_not_mem_closure_singleton p₀ h).some)).op, - dsimp at this, - rw [dite_comp, category.comp_id] at this, - simp_rw [category.id_comp, eq_to_hom_trans, eq3, category.comp_id, eq_to_hom_trans] at this, - - - generalize_proofs h₁ h₂ h₃, - have eq_coe : c.ι.app (op ⟨↑(h₁.some), h₃⟩) = - eq_to_hom (by { congr, ext, refl }) ≫ c.ι.app (op h₁.some) ≫ - eq_to_hom (by { congr, ext, refl }), - { symmetry, - rw [eq_eq_to_hom_comp], - have e := c.ι.naturality (eq_to_hom (by {ext, refl}) : - (⟨↑h₁.some, h₃⟩ : open_nhds y) ⟶ h₁.some).op, - dsimp at e, - split_ifs at e with temp, - { exfalso, exact (mem_nhds_of_not_mem_closure_singleton p₀ h).some_spec temp }, - rw [show ts.from (ite (p₀ ∈ (open_nhds.inclusion y).obj h₁.some) S star) = - eq_to_hom (if_neg temp) ≫ 𝟙 star, from ts.hom_ext _ _, category.comp_id, - category.comp_id, eq_to_hom_trans] at e, - rw [←e, category.assoc, eq_eq_to_hom_comp, ←category.assoc, eq_to_hom_trans, eq_to_hom_refl, - eq_to_hom_refl, category.comp_id, category.id_comp], }, - rw [eq_coe, ←this, ←category.assoc, eq_to_hom_trans, eq_eq_to_hom_comp, dite_comp], - have h₅ : p₀ ∉ (open_nhds.inclusion y).obj h₁.some := h₁.some_spec, - have h₆ : p₀ ∉ (open_nhds.inclusion y).obj (h₁.some ⊓ unop U) := λ ⟨r, _⟩, h₅ r, - clear this, - split_ifs, - { exfalso, exact h₅ (by assumption) }, - { exfalso, exact h₆ (by assumption) }, - { exfalso, exact h₅ (by assumption) }, - rw [eq_comp_eq_to_hom, eq_eq_to_hom_comp, ←category.assoc, ←category.assoc, eq_to_hom_trans, - eq_to_hom_refl, category.comp_id], - have := c.ι.naturality (hom_of_le (inf_le_right : - (mem_nhds_of_not_mem_closure_singleton p₀ h).some ⊓ unop U ≤ unop U)).op, - dsimp at this, - rw [dite_comp, category.comp_id] at this, - rw [←this, comp_dite], - clear this, - split_ifs, - rw [eq1, category.comp_id, eq_to_hom_trans, ←category.assoc, eq_to_hom_trans, eq_to_hom_refl, - category.id_comp] } - end, - uniq' := λ c f H, - begin - erw [colimit.ι_desc, ←H], - simp only [skyscraper_presheaf_cocone_of_not_mem_closure₀_ι_app], - dsimp, - have : ts.from (ite (p₀ ∈ (mem_nhds_of_not_mem_closure_singleton p₀ h).some.1) S star) = - eq_to_hom (if_neg (mem_nhds_of_not_mem_closure_singleton p₀ h).some_spec) ≫ 𝟙 star := - ts.hom_ext _ _, - erw [this, category.comp_id, ←category.assoc, eq_to_hom_trans, eq_to_hom_refl, category.id_comp] - end } - -/-- -If `y ∉ closure {p₀}`, then the stalk of `skyscraper_presheaf p₀ S t` at `y` is `*` --/ -@[reducible] -noncomputable def skyscraper_presheaf_stalk_of_not_mem_closure₀ - {y : X} (h : ¬p₀ ⤳ y) : - (skyscraper_presheaf p₀ S ts).stalk y ≅ star := -colimit.iso_colimit_cocone ⟨_, (skyscraper_presheaf_cocone_of_not_mem_closure₀_is_colimit _ S _ h)⟩ - -end - -section - -/-! -Skyscraper sheaf alternatively can be defined as the pushforward of a sheaf on a single point space -`{p₀}` defined by `∅ ↦ *` and `{p₀} ↦ S` under the inclusion `{p₀} ↪ X`. --/ - -open topological_space -open category_theory category_theory.limits -open Top -open opposite -open_locale classical - -universes u v w - -variables {X : Top.{u}} (p₀ : X) {C : Type v} [category.{w} C] (S : C) -variables {star : C} (ts : is_terminal star) - -/--The topological space with on a point `{p₀}`. Hence the only open sets are ∅ and ⊤.-/ -def single_point_space : Top := ⟨({p₀} : set X), infer_instance⟩ - -instance : inhabited (single_point_space p₀) := ⟨⟨p₀, rfl⟩⟩ -instance : unique (single_point_space p₀) := unique.subtype_eq p₀ - -/-- -The presheaf on a single point space `{p₀}` defined by `∅ ↦ *` and `{p₀} ↦ S` --/ -@[simps] -noncomputable def single_point_presheaf : presheaf C (single_point_space p₀) := -{ obj := λ U, if U.unop ≠ ⊥ then S else star, - map := λ U V inc, if h : V.unop ≠ ⊥ - then eq_to_hom (if_pos $ λ r, h $ le_bot_iff.mp $ r ▸ le_of_hom inc.unop) ≫ - 𝟙 S ≫ eq_to_hom (if_pos h).symm - else ts.from _ ≫ 𝟙 star ≫ eq_to_hom (if_neg h).symm, - map_id' := λ U, - begin - split_ifs, - { rw [category.id_comp, eq_to_hom_trans, eq_to_hom_refl] }, - { rw [category.id_comp, eq_comp_eq_to_hom, category.id_comp], - exact ts.hom_ext _ _, }, - end, - map_comp' := λ U V W inc1 inc2, - begin - rw comp_dite, - by_cases hW : W.unop ≠ ⊥, - { have hV : V.unop ≠ ⊥, - { intros r, - refine hW (eq_bot_iff.mpr (r ▸ le_of_hom inc2.unop)), }, - have hU : U.unop ≠ ⊥, - { intros r, - refine hW (eq_bot_iff.mpr (r ▸ le_of_hom (inc1 ≫ inc2).unop)) }, - split_ifs, - simp only [category.id_comp, eq_to_hom_trans], }, - { split_ifs; - rw [category.id_comp, eq_comp_eq_to_hom, category.assoc, category.assoc, eq_to_hom_trans, - eq_to_hom_refl, category.comp_id]; - exact ts.hom_ext _ _, }, - end } - -/-- -The trivial inclusion `{p₀} ↪ X`. --/ -@[simps] def single_point_inclusion : single_point_space p₀ ⟶ X := -{ to_fun := λ p, p.1, - continuous_to_fun := by continuity } - -/-- -The morphism from skyscraper presheaf to pushforward sheaf --/ -@[simps] noncomputable def skyscraper_presheaf_to_pushforward : - skyscraper_presheaf p₀ S ts ⟶ (single_point_inclusion p₀) _* (single_point_presheaf p₀ S ts) := -{ app := λ U, if h : p₀ ∈ U.unop - then eq_to_hom (skyscraper_presheaf_obj_of_mem _ _ h) ≫ 𝟙 S ≫ eq_to_hom - begin - dsimp, - rw if_pos, - erw opens.map_obj _ _ U.unop.2, - change _ ≠ ⊥, - rw opens.ne_bot_iff_nonempty, - refine ⟨⟨p₀, rfl⟩, _⟩, - erw set.mem_preimage, - exact h, - end - else ts.from _ ≫ eq_to_hom - begin - dsimp, - rw if_neg, - push_neg, - rw ←opens.not_nonempty_iff_eq_bot, - rintros ⟨⟨x, hx₁⟩, hx₂⟩, - rw set.mem_singleton_iff at hx₁, - subst hx₁, - erw set.mem_preimage at hx₂, - exact h hx₂, - end, - naturality' := λ U V inc, - begin - by_cases hV : p₀ ∈ V.unop, - { have hU : p₀ ∈ U.unop := le_of_hom inc.unop hV, - have hV' : ¬(opens.map (single_point_inclusion p₀)).obj (unop V) = ⊥, - { change _ ≠ ⊥, - rw opens.ne_bot_iff_nonempty, - refine ⟨⟨p₀, rfl⟩, _⟩, - erw set.mem_preimage, - exact hV, }, - rw [comp_dite, dite_comp], - split_ifs, - { exfalso, exact hV' h, }, - { dsimp, - split_ifs, - simp only [eq_to_hom_trans, category.id_comp], }, }, - { split_ifs; - rw [←category.assoc, eq_comp_eq_to_hom]; - exact ts.hom_ext _ _ }, - end } - -/-- -The morphism from pushforward sheaf to skyscraper presheaf --/ -@[simps] noncomputable def pushforward_to_skyscraper_presheaf : - (single_point_inclusion p₀) _* (single_point_presheaf p₀ S ts) ⟶ - skyscraper_presheaf p₀ S ts := -{ app := λ U, if h : p₀ ∈ unop U - then eq_to_hom - begin - dsimp, - rw if_pos, - erw opens.map_obj _ _ U.unop.2, - change _ ≠ ⊥, - rw opens.ne_bot_iff_nonempty, - refine ⟨⟨p₀, rfl⟩, _⟩, - erw set.mem_preimage, - exact h, - end ≫ 𝟙 S ≫ eq_to_hom (skyscraper_presheaf_obj_of_mem _ _ h).symm - else ts.from _ ≫ eq_to_hom (skyscraper_presheaf_obj_of_not_mem _ _ h).symm, - naturality' := λ U V inc, - begin - rw [comp_dite, dite_comp], - by_cases hV : p₀ ∈ V.unop, - { have hU : p₀ ∈ U.unop := le_of_hom inc.unop hV, - have hV' : ¬(opens.map (single_point_inclusion p₀)).obj (unop V) = ⊥, - { change _ ≠ ⊥, - rw opens.ne_bot_iff_nonempty, - refine ⟨⟨p₀, rfl⟩, _⟩, - erw set.mem_preimage, - exact hV, }, - split_ifs, - { exfalso, exact hV' h, }, - { dsimp, - split_ifs; - rw [category.id_comp, eq_to_hom_trans, category.id_comp, eq_to_hom_trans, eq_to_hom_trans, - category.id_comp, eq_to_hom_trans, eq_to_hom_trans, eq_to_hom_trans], }, }, - { split_ifs; - rw [←category.assoc, eq_comp_eq_to_hom]; - exact ts.hom_ext _ _ }, - end } - -/-- -Skyscraper presheaf is isomorphic to pushforward of sheaf on single point. --/ -noncomputable def skyscraper_presheaf_as_pushforward : - skyscraper_presheaf p₀ S ts ≅ (single_point_inclusion p₀) _* (single_point_presheaf p₀ S ts) := -{ hom := skyscraper_presheaf_to_pushforward p₀ S ts, - inv := pushforward_to_skyscraper_presheaf p₀ S ts, - hom_inv_id' := - begin - ext U, - dsimp, - split_ifs, - { rw [category.id_comp, eq_to_hom_trans, category.id_comp, eq_to_hom_trans, eq_to_hom_trans, - eq_to_hom_refl], }, - { rw [←category.assoc, eq_comp_eq_to_hom], - exact ts.hom_ext _ _, }, - end, - inv_hom_id' := - begin - ext U, - dsimp, - by_cases hU : p₀ ∈ U.unop, - { split_ifs; - rw [category.id_comp, eq_to_hom_trans, category.id_comp, eq_to_hom_trans, eq_to_hom_trans, - eq_to_hom_refl], }, - { split_ifs; - rw [←category.assoc, eq_comp_eq_to_hom]; - exact ts.hom_ext _ _, }, - end } - -end - -section adjoints - -open topological_space -open category_theory category_theory.limits -open Top -open opposite - -universes u v - -variables {X : Top.{u}} (p₀ : X) {C : Type v} [category.{u} C] -variables {star : C} (ts : is_terminal star) -variable [Π (U : opens X), decidable (p₀ ∈ U)] - -@[simps] -def skyscraper_presheaf_functor : C ⥤ presheaf C X := -{ obj := λ S, skyscraper_presheaf p₀ S ts, - map := λ x y f, - { app := λ U, if h : p₀ ∈ U.unop - then eq_to_hom (skyscraper_presheaf_obj_of_mem _ _ h) ≫ f ≫ - eq_to_hom (skyscraper_presheaf_obj_of_mem _ _ h).symm - else ts.from _ ≫ eq_to_hom (skyscraper_presheaf_obj_of_not_mem _ _ h).symm, - naturality' := λ U V inc, - begin - dsimp, - simp_rw [category.id_comp, eq_to_hom_trans], - by_cases hV : p₀ ∈ V.unop, - { have hU : p₀ ∈ U.unop := le_of_hom inc.unop hV, - split_ifs, - rw [←category.assoc, eq_to_hom_trans, category.assoc, category.assoc, eq_to_hom_trans], - refl, }, - { split_ifs; - rw [←category.assoc, eq_comp_eq_to_hom]; - exact ts.hom_ext _ _ } - end }, - map_id' := λ c, - begin - ext U, - dsimp, - split_ifs, - { simp, }, - { rw [eq_comp_eq_to_hom], - exact ts.hom_ext _ _ }, - end, - map_comp' := λ x y z f g, - begin - ext U, - dsimp, - split_ifs, - { simp }, - { rw [eq_comp_eq_to_hom], - exact ts.hom_ext _ _ }, - end } - -example : true := trivial - -@[simps] -def skyscraper_sheaf_functor : C ⥤ sheaf C X := -{ obj := λ S, skyscraper_sheaf p₀ S ts, - map := λ x y f, ⟨(skyscraper_presheaf_functor p₀ ts).map f⟩, - map_id' := λ c, - begin - ext1, - exact (skyscraper_presheaf_functor p₀ ts).map_id c, - end, - map_comp' := λ x y z f g, - begin - ext1, - exact (skyscraper_presheaf_functor p₀ ts).map_comp f g, - end } - -variable [has_colimits C] - -@[simps] -noncomputable def from_stalk_to_to_skyscraper_presheaf {𝓕 : presheaf C X} {c : C} - (f : 𝓕.stalk p₀ ⟶ c) : 𝓕 ⟶ skyscraper_presheaf p₀ c ts := -{ app := λ U, if h : p₀ ∈ U.unop - then 𝓕.germ ⟨p₀, h⟩ ≫ f ≫ eq_to_hom (skyscraper_presheaf_obj_of_mem _ _ h).symm - else ts.from _ ≫ eq_to_hom (skyscraper_presheaf_obj_of_not_mem _ _ h).symm, - naturality' := λ U V inc, - begin - dsimp, - by_cases hV : p₀ ∈ V.unop, - { have hU : p₀ ∈ U.unop := le_of_hom inc.unop hV, - split_ifs, - erw [←category.assoc, 𝓕.germ_res inc.unop, category.id_comp, eq_to_hom_trans, - category.assoc, category.assoc, eq_to_hom_trans], - refl, }, - { split_ifs, - rw [←category.assoc, eq_comp_eq_to_hom, category.assoc, category.assoc, eq_to_hom_trans, - eq_to_hom_refl, category.comp_id], - exact ts.hom_ext _ _ }, - end } - -@[reducible] -noncomputable def to_skyscraper_presheaf_to_from_stalk {𝓕 : presheaf C X} {c : C} - (f : 𝓕 ⟶ skyscraper_presheaf p₀ c ts) : 𝓕.stalk p₀ ⟶ c := -let CC : cocone ((open_nhds.inclusion p₀).op ⋙ 𝓕) := -{ X := c, - ι := - { app := λ U, f.app (op U.unop.1) ≫ eq_to_hom - begin - dsimp, - rw if_pos, - exact U.unop.2, - end, - naturality' := λ U V inc, - begin - dsimp, - rw [category.comp_id, ←category.assoc, eq_comp_eq_to_hom, category.assoc, eq_to_hom_trans], - generalize_proofs h, - erw f.naturality, - dsimp, - have hV : p₀ ∈ (open_nhds.inclusion p₀).obj V.unop := V.unop.2, - split_ifs, - rw [category.id_comp, eq_to_hom_trans, eq_comp_eq_to_hom, category.assoc, eq_to_hom_trans, - eq_to_hom_refl, category.comp_id], - refl, - end} } in -colimit.desc _ CC - -lemma from_stalk_to_to_skyscraper_presheaf_to_skyscraper_presheaf_to_from_stalk - {𝓕 : presheaf C X} {c : C} (f : 𝓕.stalk p₀ ⟶ c) : -to_skyscraper_presheaf_to_from_stalk p₀ ts (from_stalk_to_to_skyscraper_presheaf p₀ ts f) = f := -begin - ext U, - dsimp, - simp only [subtype.val_eq_coe, unop_op, colimit.ι_desc, from_stalk_to_to_skyscraper_presheaf_app], - split_ifs, - { rw [category.assoc, category.assoc, eq_to_hom_trans, eq_to_hom_refl, category.comp_id], - congr, - dunfold presheaf.germ, - dsimp, - congr, - convert op_unop _, - ext, - refl, }, - { exfalso, exact h U.unop.2, } -end - -lemma to_skyscraper_presheaf_to_from_stalk_from_stalk_to_to_skyscraper_presheaf - {𝓕 : presheaf C X} {c : C} (f : 𝓕 ⟶ skyscraper_presheaf p₀ c ts) : -from_stalk_to_to_skyscraper_presheaf p₀ ts (to_skyscraper_presheaf_to_from_stalk p₀ ts f) = f := -begin - ext U, - dsimp, - split_ifs, - { rw [to_skyscraper_presheaf_to_from_stalk], - dsimp, - dunfold presheaf.germ, - rw [←category.assoc, colimit.ι_desc], - dsimp, - rw [category.assoc, eq_to_hom_trans, eq_to_hom_refl, category.comp_id], }, - { rw [eq_comp_eq_to_hom], - exact ts.hom_ext _ _ , } -end - -@[simps] -noncomputable def stalk_skyscraper_presheaf_adj_unit : - 𝟭 (presheaf C X) ⟶ presheaf.stalk_functor C p₀ ⋙ skyscraper_presheaf_functor p₀ ts := -{ app := λ 𝓕, - { app := λ U, if h : p₀ ∈ U.unop - then 𝓕.germ ⟨p₀, h⟩ ≫ eq_to_hom (if_pos h).symm - else ts.from _ ≫ eq_to_hom (if_neg h).symm, - naturality' := λ U V inc, - begin - dsimp, - by_cases hV : p₀ ∈ V.unop, - { have hU : p₀ ∈ U.unop := le_of_hom inc.unop hV, - split_ifs, - erw [←category.assoc, 𝓕.germ_res inc.unop, category.id_comp, eq_to_hom_trans, - category.assoc, eq_to_hom_trans], - congr, }, - { split_ifs, - rw [←category.assoc, eq_comp_eq_to_hom], - exact ts.hom_ext _ _, }, - end }, - naturality' := λ 𝓕 𝓖 f, - begin - ext U, - dsimp, - split_ifs, - { rw [←category.assoc, eq_comp_eq_to_hom], - simp only [category.assoc, eq_to_hom_trans, eq_to_hom_refl, category.comp_id], - rw [←category.assoc _ _ ((presheaf.stalk_functor C p₀).map f), eq_to_hom_trans, - eq_to_hom_refl, category.id_comp], - erw [colimit.ι_map], - congr, }, - { rw [←category.assoc, eq_comp_eq_to_hom], - exact ts.hom_ext _ _ }, - end } - -@[simps] -noncomputable def stalk_skyscraper_presheaf_adj_counit : - skyscraper_presheaf_functor p₀ ts ⋙ presheaf.stalk_functor C p₀ ⟶ 𝟭 C := -{ app := λ c, (skyscraper_presheaf_stalk_of_mem_closure₀ p₀ c ts (specializes_rfl : p₀ ⤳ p₀)).hom, - naturality' := λ x y f, - begin - ext U, - dsimp, - simp only [colimit.iso_colimit_cocone_ι_hom_assoc, - skyscraper_presheaf_cocone_of_mem_closure₀_ι_app, category.assoc], - erw [category.id_comp, ←category.assoc, colimit.ι_map], - dsimp, - split_ifs, - { rw [category.assoc, skyscraper_presheaf_stalk_of_mem_closure₀, - colimit.iso_colimit_cocone_ι_hom], - dsimp, - simpa, }, - { exfalso, exact h U.unop.2 } - end } - -noncomputable def stalk_skyscraper_presheaf_adj : - presheaf.stalk_functor C p₀ ⊣ skyscraper_presheaf_functor p₀ ts := -{ hom_equiv := λ 𝓕 c, ⟨from_stalk_to_to_skyscraper_presheaf p₀ ts, - to_skyscraper_presheaf_to_from_stalk p₀ ts, - from_stalk_to_to_skyscraper_presheaf_to_skyscraper_presheaf_to_from_stalk p₀ ts, - to_skyscraper_presheaf_to_from_stalk_from_stalk_to_to_skyscraper_presheaf p₀ ts⟩, - unit := stalk_skyscraper_presheaf_adj_unit p₀ ts, - counit := stalk_skyscraper_presheaf_adj_counit p₀ ts, - hom_equiv_unit' := λ 𝓕 𝓖 f, - begin - ext U, - dsimp, - split_ifs, - { simp, }, - { rw [eq_comp_eq_to_hom], - exact ts.hom_ext _ _ }, - end, - hom_equiv_counit' := λ 𝓕 c g, - begin - ext U, - dsimp, - erw [colimit.ι_desc, ←category.assoc, colimit.ι_map, category.assoc, - colimit.iso_colimit_cocone_ι_hom], - dsimp, - rw [category.comp_id], - refl, - end } - -example : true := trivial - -noncomputable def stalk_skyscraper_sheaf_adj : - sheaf.forget C X ⋙ presheaf.stalk_functor _ p₀ ⊣ skyscraper_sheaf_functor p₀ ts := -{ hom_equiv := λ 𝓕 c, - ⟨λ f, ⟨from_stalk_to_to_skyscraper_presheaf p₀ ts f⟩, - λ g, to_skyscraper_presheaf_to_from_stalk p₀ ts g.1, - λ f, from_stalk_to_to_skyscraper_presheaf_to_skyscraper_presheaf_to_from_stalk - p₀ ts f, - begin - intros g, - ext1, - exact to_skyscraper_presheaf_to_from_stalk_from_stalk_to_to_skyscraper_presheaf p₀ ts g.1, - end⟩, - unit := - { app := λ 𝓕, ⟨(stalk_skyscraper_presheaf_adj_unit p₀ ts).app 𝓕.1⟩, - naturality' := λ 𝓐 𝓑 ⟨f⟩, - begin - ext1, - dsimp, - exact (stalk_skyscraper_presheaf_adj_unit p₀ ts).naturality f, - end }, - counit := stalk_skyscraper_presheaf_adj_counit p₀ ts, - hom_equiv_unit' := - begin - intros 𝓐 c f, - ext1, - exact (stalk_skyscraper_presheaf_adj p₀ ts).hom_equiv_unit, - end, - hom_equiv_counit' := λ 𝓐 c f, (stalk_skyscraper_presheaf_adj p₀ ts).hom_equiv_counit } - -end adjoints - -section injective - -open_locale zero_object -open topological_space -open category_theory category_theory.limits -open Top -open opposite - -universe u -variables {X : Top.{u}} (p₀ : X) [Π (U : opens X), decidable (p₀ ∈ U)] - -lemma skyscraper_presheaf_in_Ab_injective (S : Ab.{u}) [injective S] : - injective (skyscraper_sheaf p₀ S (is_zero.is_terminal (is_zero_zero _) : is_terminal (0 : Ab))) := -injective.injective_of_adjoint - (stalk_skyscraper_sheaf_adj p₀ (is_zero.is_terminal (is_zero_zero _) : is_terminal (0 : Ab))) - -end injective From c7c9f5f4ecd1c38231dadf50b1fda45a1d5ea366 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Tue, 16 Aug 2022 02:37:49 +0100 Subject: [PATCH 63/84] move file --- src/topology/sheaves/skyscraper.lean | 855 +++++++++++++++++++++++++++ 1 file changed, 855 insertions(+) create mode 100644 src/topology/sheaves/skyscraper.lean diff --git a/src/topology/sheaves/skyscraper.lean b/src/topology/sheaves/skyscraper.lean new file mode 100644 index 0000000000000..5f072ab36c2a0 --- /dev/null +++ b/src/topology/sheaves/skyscraper.lean @@ -0,0 +1,855 @@ +/- +Copyright (c) 2022 Jujian Zhang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jujian Zhang +-/ +import algebraic_geometry.sheafed_space +import topology.sheaves.sheaf_condition.unique_gluing +import topology.sheaves.stalks +import category_theory.preadditive.injective +import algebra.category.Group.abelian + +/-! +# Skyscraper (pre)sheaves + +A skyscraper (pre)sheaf is a (pre)sheaf supported at a single point: if `p₀ ∈ X` is a specified +point, then the skyscraper (pre)sheaf `𝓕` with value `A` is defined by `U ↦ A` if `p₀ ∈ U` and +`U ↦ *` if `p₀ ∉ A` where `*` is some terminal object. + +## Main definitions + +* `skyscraper_presheaf`: the skyscraper presheaf, to define a skyscraper presheaf at point `p₀` with + value `A`, use `skypscraper_presheaf p₀ A t` where `t` is a proof that `*` is a terminal object. +* `skyscraper_sheaf`: the skyscraper presheaf satisfies the sheaf condition. + +## Main statements + +* `skyscraper_stalk_of_mem_closure₀`: if `y ∈ closure {p₀}` then the stalk of + `skyscraper_presheaf p₀ S t` at `y` is `S`. +* `skyscraper_stalk_of_not_mem_closure₀`: if `y ∉ closure {p₀}` then the stalk of + `skyscraper_presheaf p₀ S t` at `y` is `*`. + +-/ + + +section + +open topological_space +open category_theory category_theory.limits +open Top +open opposite + +universes u v w + +variables {X : Top.{u}} (p₀ : X) {C : Type v} [category.{w} C] (S : C) +variables {star : C} (ts : is_terminal star) +variable [Π (U : opens X), decidable (p₀ ∈ U)] + +/- +A skyscraper presheaf is a presheaf supported at a single point: if `p₀ ∈ X` is a specified +point, then the skyscraper presheaf `𝓕` with value `A` is defined by `U ↦ A` if `p₀ ∈ U` and +`U ↦ *` if `p₀ ∉ A` where `*` is some terminal object. +-/ +@[simps] +def skyscraper_presheaf : presheaf C X := +{ obj := λ U, ite (p₀ ∈ unop U) S star, + map := λ U V i, dite (p₀ ∈ unop V) + (λ h, eq_to_hom (if_pos (le_of_hom i.unop h)) ≫ 𝟙 S ≫ eq_to_hom (if_pos h).symm) + (λ h, ts.from _ ≫ eq_to_hom (if_neg h).symm), + map_id' := λ U, + begin + split_ifs, + { simp, }, + { rw eq_comp_eq_to_hom, + exact ts.hom_ext _ _, }, + end, + map_comp' := λ U V W iVU iWV, + begin + by_cases hW : p₀ ∈ unop W, + { have hV : p₀ ∈ unop V := le_of_hom iWV.unop hW, + have hU : p₀ ∈ unop U := le_of_hom iVU.unop hV, + split_ifs, + simp }, + { split_ifs; + rw eq_comp_eq_to_hom; + exact ts.hom_ext _ _, } + end } + +section + +variables {p₀} + +lemma skyscraper_presheaf_obj_of_mem {U : opens X} (h : p₀ ∈ U) : + (skyscraper_presheaf p₀ S ts).obj (op U) = S := +if_pos h + +lemma skyscraper_presheaf_obj_of_not_mem {U : opens X} (h : p₀ ∉ U) : + (skyscraper_presheaf p₀ S ts).obj (op U) = star := +if_neg h + +end + +/- +A skyscraper sheaf is a sheaf supported at a single point: if `p₀ ∈ X` is a specified +point, then the skyscraper sheaf `𝓕` with value `A` is defined by `U ↦ A` if `p₀ ∈ U` and +`U ↦ *` if `p₀ ∉ A` where `*` is some terminal object. +-/ +def skyscraper_sheaf : sheaf C X := +⟨skyscraper_presheaf p₀ S ts, λ c U s hs x hx, + ⟨dite (p₀ ∈ U) + (λ h, x (hs p₀ h).some_spec.some (hs p₀ h).some_spec.some_spec.1 ≫ + eq_to_hom ((skyscraper_presheaf_obj_of_mem S ts (hs p₀ h).some_spec.some_spec.2).trans + (skyscraper_presheaf_obj_of_mem S ts h).symm)) + (λ h, ts.from c ≫ (eq_to_hom (skyscraper_presheaf_obj_of_not_mem S ts h).symm)), + λ V inc h, + begin + by_cases hV : p₀ ∈ V, + { have hU : p₀ ∈ U := le_of_hom inc hV, + split_ifs, + generalize_proofs h₁ h₂ h₃ h₄, + dsimp, + split_ifs, + rw [category.id_comp, eq_to_hom_trans, category.assoc, eq_to_hom_trans], + generalize_proofs h₅, + have := hx (hom_of_le inf_le_left) (hom_of_le inf_le_right) h₂.some_spec.1 h rfl, + dsimp at this, + have hV' : p₀ ∈ h₁.some ⊓ V := ⟨h₂.some_spec.2, hV⟩, + split_ifs at this; + rw [category.id_comp, eq_to_hom_trans, eq_to_hom_trans, eq_comp_eq_to_hom, + category.assoc, eq_to_hom_trans] at this; + generalize_proofs h₆ at this; + rw [this, category.assoc, eq_to_hom_trans, eq_to_hom_refl, category.comp_id] }, + { dsimp, + split_ifs, + rw [←category.assoc, eq_comp_eq_to_hom], + exact ts.hom_ext _ _, } + end, + λ y (hy : x.is_amalgamation y), + begin + split_ifs, + { generalize_proofs h₁ h₂ h₃ h₄, + have := hy h₂.some h₂.some_spec.1, + dsimp at this, + split_ifs at this with H, + { rw [category.id_comp, eq_to_hom_trans, eq_comp_eq_to_hom] at this, + rw [this, eq_comp_eq_to_hom, category.assoc, eq_to_hom_trans, eq_to_hom_refl, + category.comp_id], }, + { exfalso, + exact H h₂.some_spec.2, }, }, + { rw [←eq_comp_eq_to_hom], + exact (ts.hom_ext _ _), } + end⟩⟩ + +end + +section + +-- In this section, we calculate the stalks for skyscraper presheaves. We need to restrict universe +-- level. + +open topological_space +open category_theory category_theory.limits +open Top +open opposite + +universes u v + +variables {X : Top.{u}} (p₀ : X) {C : Type v} [category.{u} C] (S : C) +variables {star : C} (ts : is_terminal star) +variable [Π (U : opens X), decidable (p₀ ∈ U)] +variable [has_colimits C] + +lemma mem_nhds_of_not_mem_closure_singleton {y : X} (h : ¬p₀ ⤳ y) : + ∃ (U : open_nhds y), p₀ ∉ U.1 := +begin + have := (not_iff_not.mpr specializes_iff_forall_open).mp h, + push_neg at this, + rcases this with ⟨s, o, h₁, h₂⟩, + exact ⟨⟨⟨s, o⟩, h₁⟩, h₂⟩, +end + +/-- +The cocone at `S` for the salk functor of `skyscraper_presheaf p₀ S t` when `y ∈ closure {p₀}` +-/ +@[simps] def skyscraper_presheaf_cocone_of_mem_closure₀ {y : X} (h : p₀ ⤳ y) : + cocone ((open_nhds.inclusion y).op ⋙ skyscraper_presheaf p₀ S ts) := +{ X := S, + ι := + { app := λ U, eq_to_hom + begin + dsimp, + rw if_pos, + exact h.mem_open U.unop.1.2 U.unop.2, + end ≫ 𝟙 S, + naturality' := λ U V inc, + begin + simp only [functor.op_obj, unop_op, functor.comp_map, functor.op_map, skyscraper_presheaf_map, + category.id_comp, eq_to_hom_trans, functor.const_obj_map, category.assoc], + by_cases hV : p₀ ∈ (open_nhds.inclusion y).obj V.unop, + { have hU : p₀ ∈ unop ((open_nhds.inclusion y).op.obj U) := le_of_hom inc.unop hV, + split_ifs, + erw [category.comp_id, category.comp_id, category.comp_id, eq_to_hom_trans], + refl }, + { split_ifs with hU; + erw [category.comp_id, category.comp_id, category.comp_id, eq_comp_eq_to_hom, + eq_comp_eq_to_hom]; + exact ts.hom_ext _ _, }, + end } } + +/-- +The canonical map `S ⟶ (skyscraper_presheaf p₀ S t).stalk y` when `y ∈ closure {p₀}` +-/ +noncomputable def skyscraper_presheaf_of_mem_closure₀_from {y : X} (h : p₀ ⤳ y) : + S ⟶ (skyscraper_presheaf p₀ S ts).stalk y := +eq_to_hom (skyscraper_presheaf_obj_of_mem S ts (by tauto : p₀ ∈ ⊤)).symm ≫ + (skyscraper_presheaf p₀ S ts).germ (⟨y, trivial⟩ : (⊤ : opens X)) + +noncomputable lemma skyscraper_presheaf_cocone_of_mem_closure₀_is_colimit {y : X} (h : p₀ ⤳ y) : + is_colimit (skyscraper_presheaf_cocone_of_mem_closure₀ p₀ S ts h) := +{ desc := λ c, (skyscraper_presheaf_of_mem_closure₀_from p₀ S ts h ≫ colimit.desc _ _ : S ⟶ c.X), + fac' := λ c U, + begin + dsimp, + simp only [skyscraper_presheaf_of_mem_closure₀_from, presheaf.germ, category.comp_id, + category.assoc, colimit.ι_desc, eq_to_hom_trans_assoc], + have := c.ι.naturality (hom_of_le $ (le_top : unop U ≤ _)).op, + dsimp at this, + have h' : p₀ ∈ (open_nhds.inclusion y).obj (unop U) := h.mem_open U.unop.1.2 U.unop.2, + have h'' : p₀ ∈ (open_nhds.inclusion y).obj ⊤ := trivial, + split_ifs at this, + rw [category.comp_id, category.id_comp, eq_to_hom_trans, eq_eq_to_hom_comp] at this, + rw [this, eq_eq_to_hom_comp, ←category.assoc, eq_to_hom_trans, eq_to_hom_refl, + category.id_comp], + congr, + end, + uniq' := λ c f h, + begin + simp only [skyscraper_presheaf_of_mem_closure₀_from, presheaf.germ, category.assoc], + erw [colimit.ι_desc], + specialize h (op ⟨⊤, trivial⟩), + erw [←h], + simp only [skyscraper_presheaf_cocone_of_mem_closure₀_ι_app, category.assoc, + eq_to_hom_trans_assoc, eq_to_hom_refl, category.id_comp], + exact (category.id_comp _).symm, + end } + +/-- +If `y ∈ closure {p₀}`, then the stalk of `skyscraper_presheaf p₀ S t` at `y` is `S` +-/ +@[reducible] +noncomputable def skyscraper_presheaf_stalk_of_mem_closure₀ {y : X} (h : p₀ ⤳ y) : + (skyscraper_presheaf p₀ S ts).stalk y ≅ S := +colimit.iso_colimit_cocone ⟨_, (skyscraper_presheaf_cocone_of_mem_closure₀_is_colimit p₀ S ts h)⟩ + +/-- +The cocone at `*` for the salk functor of `skyscraper_presheaf p₀ S t` when `y ∉ closure {p₀}` +-/ +@[simps] def skyscraper_presheaf_cocone_of_not_mem_closure₀ + {y : X} (h : ¬p₀ ⤳ y) : + cocone ((open_nhds.inclusion y).op ⋙ skyscraper_presheaf p₀ S ts) := +{ X := star, + ι := + { app := λ U, ts.from _, + naturality' := λ U V inc, ts.hom_ext _ _ } } + +/-- +The canonical map `* ⟶ (skyscraper_presheaf p₀ S t).stalk y` when `y ∉ closure {p₀}` +-/ +noncomputable def skyscraper_presheaf_of_not_mem_closure₀_from + {y : X} (h : ¬p₀ ⤳ y) : + star ⟶ (skyscraper_presheaf p₀ S ts).stalk y := +eq_to_hom (skyscraper_presheaf_obj_of_not_mem S ts $ + (mem_nhds_of_not_mem_closure_singleton p₀ h).some_spec).symm ≫ + (skyscraper_presheaf p₀ S ts).germ (⟨y, (mem_nhds_of_not_mem_closure_singleton p₀ h).some.2⟩ : + (mem_nhds_of_not_mem_closure_singleton p₀ h).some.1) + +noncomputable lemma skyscraper_presheaf_cocone_of_not_mem_closure₀_is_colimit + {y : X} (h : ¬p₀ ⤳ y) : + is_colimit (skyscraper_presheaf_cocone_of_not_mem_closure₀ p₀ S ts h) := +{ desc := λ c, ((eq_to_hom ((skyscraper_presheaf_obj_of_not_mem _ _ + (mem_nhds_of_not_mem_closure_singleton p₀ h).some_spec).symm)) ≫ + presheaf.germ (skyscraper_presheaf p₀ S ts) + ⟨y, (mem_nhds_of_not_mem_closure_singleton p₀ h).some.2⟩ ≫ colimit.desc _ _ : star ⟶ _), + fac' := λ c U, + begin + simp only [presheaf.germ, skyscraper_presheaf_cocone_of_not_mem_closure₀_ι_app], + erw [colimit.ι_desc], + dsimp, + by_cases h' : p₀ ∈ (open_nhds.inclusion y).obj (unop U), + { have eq1 : ts.from (ite (p₀ ∈ (open_nhds.inclusion y).obj (unop U)) S star) = + eq_to_hom (if_pos h') ≫ ts.from _ := ts.hom_ext _ _, + rw [eq1, category.assoc, eq_eq_to_hom_comp], + have := c.ι.naturality (hom_of_le (le_top : U.unop ≤ ⊤)).op, + dsimp at this, + have h'' : p₀ ∈ (open_nhds.inclusion y).obj ⊤ := trivial, + split_ifs at this, + rw [category.comp_id, category.id_comp, eq_to_hom_trans, eq_eq_to_hom_comp] at this, + rw this, + clear this, + have := c.ι.naturality + (hom_of_le (le_top : (mem_nhds_of_not_mem_closure_singleton p₀ h).some ≤ ⊤)).op, + dsimp at this, + have h''' : p₀ ∉ (mem_nhds_of_not_mem_closure_singleton p₀ h).some.1 := + (mem_nhds_of_not_mem_closure_singleton p₀ h).some_spec, + split_ifs at this, + have eq2 : ts.from (ite (p₀ ∈ (open_nhds.inclusion y).obj ⊤) S star) = + eq_to_hom (if_pos h'') ≫ ts.from _ := ts.hom_ext _ _, + rw [category.comp_id, eq2] at this, + simp only [←this, ←category.assoc], + convert eq_whisker _ _, + { ext, refl, }, + { rw [eq_comp_eq_to_hom], + exact ts.hom_ext _ _ } }, + { have eq1 : ts.from (ite (p₀ ∈ (open_nhds.inclusion y).obj (unop U)) S star) = + eq_to_hom (if_neg h') ≫ 𝟙 star := ts.hom_ext _ _, + have eq2 : ts.from (ite (p₀ ∈ (open_nhds.inclusion y).obj ⊤) S star) = + eq_to_hom (if_pos trivial) ≫ ts.from _ := ts.hom_ext _ _, + have eq3 : ts.from (ite (p₀ ∈ (open_nhds.inclusion y).obj + (mem_nhds_of_not_mem_closure_singleton p₀ h).some) S star) = + eq_to_hom (if_neg (mem_nhds_of_not_mem_closure_singleton p₀ h).some_spec) ≫ 𝟙 star := + ts.hom_ext _ _, + rw [eq1, category.comp_id, ←category.assoc, eq_to_hom_trans], + have := c.ι.naturality (hom_of_le (inf_le_left : + (mem_nhds_of_not_mem_closure_singleton p₀ h).some ⊓ unop U ≤ + (mem_nhds_of_not_mem_closure_singleton p₀ h).some)).op, + dsimp at this, + rw [dite_comp, category.comp_id] at this, + simp_rw [category.id_comp, eq_to_hom_trans, eq3, category.comp_id, eq_to_hom_trans] at this, + + + generalize_proofs h₁ h₂ h₃, + have eq_coe : c.ι.app (op ⟨↑(h₁.some), h₃⟩) = + eq_to_hom (by { congr, ext, refl }) ≫ c.ι.app (op h₁.some) ≫ + eq_to_hom (by { congr, ext, refl }), + { symmetry, + rw [eq_eq_to_hom_comp], + have e := c.ι.naturality (eq_to_hom (by {ext, refl}) : + (⟨↑h₁.some, h₃⟩ : open_nhds y) ⟶ h₁.some).op, + dsimp at e, + split_ifs at e with temp, + { exfalso, exact (mem_nhds_of_not_mem_closure_singleton p₀ h).some_spec temp }, + rw [show ts.from (ite (p₀ ∈ (open_nhds.inclusion y).obj h₁.some) S star) = + eq_to_hom (if_neg temp) ≫ 𝟙 star, from ts.hom_ext _ _, category.comp_id, + category.comp_id, eq_to_hom_trans] at e, + rw [←e, category.assoc, eq_eq_to_hom_comp, ←category.assoc, eq_to_hom_trans, eq_to_hom_refl, + eq_to_hom_refl, category.comp_id, category.id_comp], }, + rw [eq_coe, ←this, ←category.assoc, eq_to_hom_trans, eq_eq_to_hom_comp, dite_comp], + have h₅ : p₀ ∉ (open_nhds.inclusion y).obj h₁.some := h₁.some_spec, + have h₆ : p₀ ∉ (open_nhds.inclusion y).obj (h₁.some ⊓ unop U) := λ ⟨r, _⟩, h₅ r, + clear this, + split_ifs, + { exfalso, exact h₅ (by assumption) }, + { exfalso, exact h₆ (by assumption) }, + { exfalso, exact h₅ (by assumption) }, + rw [eq_comp_eq_to_hom, eq_eq_to_hom_comp, ←category.assoc, ←category.assoc, eq_to_hom_trans, + eq_to_hom_refl, category.comp_id], + have := c.ι.naturality (hom_of_le (inf_le_right : + (mem_nhds_of_not_mem_closure_singleton p₀ h).some ⊓ unop U ≤ unop U)).op, + dsimp at this, + rw [dite_comp, category.comp_id] at this, + rw [←this, comp_dite], + clear this, + split_ifs, + rw [eq1, category.comp_id, eq_to_hom_trans, ←category.assoc, eq_to_hom_trans, eq_to_hom_refl, + category.id_comp] } + end, + uniq' := λ c f H, + begin + erw [colimit.ι_desc, ←H], + simp only [skyscraper_presheaf_cocone_of_not_mem_closure₀_ι_app], + dsimp, + have : ts.from (ite (p₀ ∈ (mem_nhds_of_not_mem_closure_singleton p₀ h).some.1) S star) = + eq_to_hom (if_neg (mem_nhds_of_not_mem_closure_singleton p₀ h).some_spec) ≫ 𝟙 star := + ts.hom_ext _ _, + erw [this, category.comp_id, ←category.assoc, eq_to_hom_trans, eq_to_hom_refl, category.id_comp] + end } + +/-- +If `y ∉ closure {p₀}`, then the stalk of `skyscraper_presheaf p₀ S t` at `y` is `*` +-/ +@[reducible] +noncomputable def skyscraper_presheaf_stalk_of_not_mem_closure₀ + {y : X} (h : ¬p₀ ⤳ y) : + (skyscraper_presheaf p₀ S ts).stalk y ≅ star := +colimit.iso_colimit_cocone ⟨_, (skyscraper_presheaf_cocone_of_not_mem_closure₀_is_colimit _ S _ h)⟩ + +end + +section + +/-! +Skyscraper sheaf alternatively can be defined as the pushforward of a sheaf on a single point space +`{p₀}` defined by `∅ ↦ *` and `{p₀} ↦ S` under the inclusion `{p₀} ↪ X`. +-/ + +open topological_space +open category_theory category_theory.limits +open Top +open opposite +open_locale classical + +universes u v w + +variables {X : Top.{u}} (p₀ : X) {C : Type v} [category.{w} C] (S : C) +variables {star : C} (ts : is_terminal star) + +/--The topological space with on a point `{p₀}`. Hence the only open sets are ∅ and ⊤.-/ +def single_point_space : Top := ⟨({p₀} : set X), infer_instance⟩ + +instance : inhabited (single_point_space p₀) := ⟨⟨p₀, rfl⟩⟩ +instance : unique (single_point_space p₀) := unique.subtype_eq p₀ + +/-- +The presheaf on a single point space `{p₀}` defined by `∅ ↦ *` and `{p₀} ↦ S` +-/ +@[simps] +noncomputable def single_point_presheaf : presheaf C (single_point_space p₀) := +{ obj := λ U, if U.unop ≠ ⊥ then S else star, + map := λ U V inc, if h : V.unop ≠ ⊥ + then eq_to_hom (if_pos $ λ r, h $ le_bot_iff.mp $ r ▸ le_of_hom inc.unop) ≫ + 𝟙 S ≫ eq_to_hom (if_pos h).symm + else ts.from _ ≫ 𝟙 star ≫ eq_to_hom (if_neg h).symm, + map_id' := λ U, + begin + split_ifs, + { rw [category.id_comp, eq_to_hom_trans, eq_to_hom_refl] }, + { rw [category.id_comp, eq_comp_eq_to_hom, category.id_comp], + exact ts.hom_ext _ _, }, + end, + map_comp' := λ U V W inc1 inc2, + begin + rw comp_dite, + by_cases hW : W.unop ≠ ⊥, + { have hV : V.unop ≠ ⊥, + { intros r, + refine hW (eq_bot_iff.mpr (r ▸ le_of_hom inc2.unop)), }, + have hU : U.unop ≠ ⊥, + { intros r, + refine hW (eq_bot_iff.mpr (r ▸ le_of_hom (inc1 ≫ inc2).unop)) }, + split_ifs, + simp only [category.id_comp, eq_to_hom_trans], }, + { split_ifs; + rw [category.id_comp, eq_comp_eq_to_hom, category.assoc, category.assoc, eq_to_hom_trans, + eq_to_hom_refl, category.comp_id]; + exact ts.hom_ext _ _, }, + end } + +/-- +The trivial inclusion `{p₀} ↪ X`. +-/ +@[simps] def single_point_inclusion : single_point_space p₀ ⟶ X := +{ to_fun := λ p, p.1, + continuous_to_fun := by continuity } + +/-- +The morphism from skyscraper presheaf to pushforward sheaf +-/ +@[simps] noncomputable def skyscraper_presheaf_to_pushforward : + skyscraper_presheaf p₀ S ts ⟶ (single_point_inclusion p₀) _* (single_point_presheaf p₀ S ts) := +{ app := λ U, if h : p₀ ∈ U.unop + then eq_to_hom (skyscraper_presheaf_obj_of_mem _ _ h) ≫ 𝟙 S ≫ eq_to_hom + begin + dsimp, + rw if_pos, + erw opens.map_obj _ _ U.unop.2, + change _ ≠ ⊥, + rw opens.ne_bot_iff_nonempty, + refine ⟨⟨p₀, rfl⟩, _⟩, + erw set.mem_preimage, + exact h, + end + else ts.from _ ≫ eq_to_hom + begin + dsimp, + rw if_neg, + push_neg, + rw ←opens.not_nonempty_iff_eq_bot, + rintros ⟨⟨x, hx₁⟩, hx₂⟩, + rw set.mem_singleton_iff at hx₁, + subst hx₁, + erw set.mem_preimage at hx₂, + exact h hx₂, + end, + naturality' := λ U V inc, + begin + by_cases hV : p₀ ∈ V.unop, + { have hU : p₀ ∈ U.unop := le_of_hom inc.unop hV, + have hV' : ¬(opens.map (single_point_inclusion p₀)).obj (unop V) = ⊥, + { change _ ≠ ⊥, + rw opens.ne_bot_iff_nonempty, + refine ⟨⟨p₀, rfl⟩, _⟩, + erw set.mem_preimage, + exact hV, }, + rw [comp_dite, dite_comp], + split_ifs, + { exfalso, exact hV' h, }, + { dsimp, + split_ifs, + simp only [eq_to_hom_trans, category.id_comp], }, }, + { split_ifs; + rw [←category.assoc, eq_comp_eq_to_hom]; + exact ts.hom_ext _ _ }, + end } + +/-- +The morphism from pushforward sheaf to skyscraper presheaf +-/ +@[simps] noncomputable def pushforward_to_skyscraper_presheaf : + (single_point_inclusion p₀) _* (single_point_presheaf p₀ S ts) ⟶ + skyscraper_presheaf p₀ S ts := +{ app := λ U, if h : p₀ ∈ unop U + then eq_to_hom + begin + dsimp, + rw if_pos, + erw opens.map_obj _ _ U.unop.2, + change _ ≠ ⊥, + rw opens.ne_bot_iff_nonempty, + refine ⟨⟨p₀, rfl⟩, _⟩, + erw set.mem_preimage, + exact h, + end ≫ 𝟙 S ≫ eq_to_hom (skyscraper_presheaf_obj_of_mem _ _ h).symm + else ts.from _ ≫ eq_to_hom (skyscraper_presheaf_obj_of_not_mem _ _ h).symm, + naturality' := λ U V inc, + begin + rw [comp_dite, dite_comp], + by_cases hV : p₀ ∈ V.unop, + { have hU : p₀ ∈ U.unop := le_of_hom inc.unop hV, + have hV' : ¬(opens.map (single_point_inclusion p₀)).obj (unop V) = ⊥, + { change _ ≠ ⊥, + rw opens.ne_bot_iff_nonempty, + refine ⟨⟨p₀, rfl⟩, _⟩, + erw set.mem_preimage, + exact hV, }, + split_ifs, + { exfalso, exact hV' h, }, + { dsimp, + split_ifs; + rw [category.id_comp, eq_to_hom_trans, category.id_comp, eq_to_hom_trans, eq_to_hom_trans, + category.id_comp, eq_to_hom_trans, eq_to_hom_trans, eq_to_hom_trans], }, }, + { split_ifs; + rw [←category.assoc, eq_comp_eq_to_hom]; + exact ts.hom_ext _ _ }, + end } + +/-- +Skyscraper presheaf is isomorphic to pushforward of sheaf on single point. +-/ +noncomputable def skyscraper_presheaf_as_pushforward : + skyscraper_presheaf p₀ S ts ≅ (single_point_inclusion p₀) _* (single_point_presheaf p₀ S ts) := +{ hom := skyscraper_presheaf_to_pushforward p₀ S ts, + inv := pushforward_to_skyscraper_presheaf p₀ S ts, + hom_inv_id' := + begin + ext U, + dsimp, + split_ifs, + { rw [category.id_comp, eq_to_hom_trans, category.id_comp, eq_to_hom_trans, eq_to_hom_trans, + eq_to_hom_refl], }, + { rw [←category.assoc, eq_comp_eq_to_hom], + exact ts.hom_ext _ _, }, + end, + inv_hom_id' := + begin + ext U, + dsimp, + by_cases hU : p₀ ∈ U.unop, + { split_ifs; + rw [category.id_comp, eq_to_hom_trans, category.id_comp, eq_to_hom_trans, eq_to_hom_trans, + eq_to_hom_refl], }, + { split_ifs; + rw [←category.assoc, eq_comp_eq_to_hom]; + exact ts.hom_ext _ _, }, + end } + +end + +section adjoints + +open topological_space +open category_theory category_theory.limits +open Top +open opposite + +universes u v + +variables {X : Top.{u}} (p₀ : X) {C : Type v} [category.{u} C] +variables {star : C} (ts : is_terminal star) +variable [Π (U : opens X), decidable (p₀ ∈ U)] + +@[simps] +def skyscraper_presheaf_functor : C ⥤ presheaf C X := +{ obj := λ S, skyscraper_presheaf p₀ S ts, + map := λ x y f, + { app := λ U, if h : p₀ ∈ U.unop + then eq_to_hom (skyscraper_presheaf_obj_of_mem _ _ h) ≫ f ≫ + eq_to_hom (skyscraper_presheaf_obj_of_mem _ _ h).symm + else ts.from _ ≫ eq_to_hom (skyscraper_presheaf_obj_of_not_mem _ _ h).symm, + naturality' := λ U V inc, + begin + dsimp, + simp_rw [category.id_comp, eq_to_hom_trans], + by_cases hV : p₀ ∈ V.unop, + { have hU : p₀ ∈ U.unop := le_of_hom inc.unop hV, + split_ifs, + rw [←category.assoc, eq_to_hom_trans, category.assoc, category.assoc, eq_to_hom_trans], + refl, }, + { split_ifs; + rw [←category.assoc, eq_comp_eq_to_hom]; + exact ts.hom_ext _ _ } + end }, + map_id' := λ c, + begin + ext U, + dsimp, + split_ifs, + { simp, }, + { rw [eq_comp_eq_to_hom], + exact ts.hom_ext _ _ }, + end, + map_comp' := λ x y z f g, + begin + ext U, + dsimp, + split_ifs, + { simp }, + { rw [eq_comp_eq_to_hom], + exact ts.hom_ext _ _ }, + end } + +example : true := trivial + +@[simps] +def skyscraper_sheaf_functor : C ⥤ sheaf C X := +{ obj := λ S, skyscraper_sheaf p₀ S ts, + map := λ x y f, ⟨(skyscraper_presheaf_functor p₀ ts).map f⟩, + map_id' := λ c, + begin + ext1, + exact (skyscraper_presheaf_functor p₀ ts).map_id c, + end, + map_comp' := λ x y z f g, + begin + ext1, + exact (skyscraper_presheaf_functor p₀ ts).map_comp f g, + end } + +variable [has_colimits C] + +@[simps] +noncomputable def from_stalk_to_to_skyscraper_presheaf {𝓕 : presheaf C X} {c : C} + (f : 𝓕.stalk p₀ ⟶ c) : 𝓕 ⟶ skyscraper_presheaf p₀ c ts := +{ app := λ U, if h : p₀ ∈ U.unop + then 𝓕.germ ⟨p₀, h⟩ ≫ f ≫ eq_to_hom (skyscraper_presheaf_obj_of_mem _ _ h).symm + else ts.from _ ≫ eq_to_hom (skyscraper_presheaf_obj_of_not_mem _ _ h).symm, + naturality' := λ U V inc, + begin + dsimp, + by_cases hV : p₀ ∈ V.unop, + { have hU : p₀ ∈ U.unop := le_of_hom inc.unop hV, + split_ifs, + erw [←category.assoc, 𝓕.germ_res inc.unop, category.id_comp, eq_to_hom_trans, + category.assoc, category.assoc, eq_to_hom_trans], + refl, }, + { split_ifs, + rw [←category.assoc, eq_comp_eq_to_hom, category.assoc, category.assoc, eq_to_hom_trans, + eq_to_hom_refl, category.comp_id], + exact ts.hom_ext _ _ }, + end } + +@[reducible] +noncomputable def to_skyscraper_presheaf_to_from_stalk {𝓕 : presheaf C X} {c : C} + (f : 𝓕 ⟶ skyscraper_presheaf p₀ c ts) : 𝓕.stalk p₀ ⟶ c := +let CC : cocone ((open_nhds.inclusion p₀).op ⋙ 𝓕) := +{ X := c, + ι := + { app := λ U, f.app (op U.unop.1) ≫ eq_to_hom + begin + dsimp, + rw if_pos, + exact U.unop.2, + end, + naturality' := λ U V inc, + begin + dsimp, + rw [category.comp_id, ←category.assoc, eq_comp_eq_to_hom, category.assoc, eq_to_hom_trans], + generalize_proofs h, + erw f.naturality, + dsimp, + have hV : p₀ ∈ (open_nhds.inclusion p₀).obj V.unop := V.unop.2, + split_ifs, + rw [category.id_comp, eq_to_hom_trans, eq_comp_eq_to_hom, category.assoc, eq_to_hom_trans, + eq_to_hom_refl, category.comp_id], + refl, + end} } in +colimit.desc _ CC + +lemma from_stalk_to_to_skyscraper_presheaf_to_skyscraper_presheaf_to_from_stalk + {𝓕 : presheaf C X} {c : C} (f : 𝓕.stalk p₀ ⟶ c) : +to_skyscraper_presheaf_to_from_stalk p₀ ts (from_stalk_to_to_skyscraper_presheaf p₀ ts f) = f := +begin + ext U, + dsimp, + simp only [subtype.val_eq_coe, unop_op, colimit.ι_desc, from_stalk_to_to_skyscraper_presheaf_app], + split_ifs, + { rw [category.assoc, category.assoc, eq_to_hom_trans, eq_to_hom_refl, category.comp_id], + congr, + dunfold presheaf.germ, + dsimp, + congr, + convert op_unop _, + ext, + refl, }, + { exfalso, exact h U.unop.2, } +end + +lemma to_skyscraper_presheaf_to_from_stalk_from_stalk_to_to_skyscraper_presheaf + {𝓕 : presheaf C X} {c : C} (f : 𝓕 ⟶ skyscraper_presheaf p₀ c ts) : +from_stalk_to_to_skyscraper_presheaf p₀ ts (to_skyscraper_presheaf_to_from_stalk p₀ ts f) = f := +begin + ext U, + dsimp, + split_ifs, + { rw [to_skyscraper_presheaf_to_from_stalk], + dsimp, + dunfold presheaf.germ, + rw [←category.assoc, colimit.ι_desc], + dsimp, + rw [category.assoc, eq_to_hom_trans, eq_to_hom_refl, category.comp_id], }, + { rw [eq_comp_eq_to_hom], + exact ts.hom_ext _ _ , } +end + +@[simps] +noncomputable def stalk_skyscraper_presheaf_adj_unit : + 𝟭 (presheaf C X) ⟶ presheaf.stalk_functor C p₀ ⋙ skyscraper_presheaf_functor p₀ ts := +{ app := λ 𝓕, + { app := λ U, if h : p₀ ∈ U.unop + then 𝓕.germ ⟨p₀, h⟩ ≫ eq_to_hom (if_pos h).symm + else ts.from _ ≫ eq_to_hom (if_neg h).symm, + naturality' := λ U V inc, + begin + dsimp, + by_cases hV : p₀ ∈ V.unop, + { have hU : p₀ ∈ U.unop := le_of_hom inc.unop hV, + split_ifs, + erw [←category.assoc, 𝓕.germ_res inc.unop, category.id_comp, eq_to_hom_trans, + category.assoc, eq_to_hom_trans], + congr, }, + { split_ifs, + rw [←category.assoc, eq_comp_eq_to_hom], + exact ts.hom_ext _ _, }, + end }, + naturality' := λ 𝓕 𝓖 f, + begin + ext U, + dsimp, + split_ifs, + { rw [←category.assoc, eq_comp_eq_to_hom], + simp only [category.assoc, eq_to_hom_trans, eq_to_hom_refl, category.comp_id], + rw [←category.assoc _ _ ((presheaf.stalk_functor C p₀).map f), eq_to_hom_trans, + eq_to_hom_refl, category.id_comp], + erw [colimit.ι_map], + congr, }, + { rw [←category.assoc, eq_comp_eq_to_hom], + exact ts.hom_ext _ _ }, + end } + +@[simps] +noncomputable def stalk_skyscraper_presheaf_adj_counit : + skyscraper_presheaf_functor p₀ ts ⋙ presheaf.stalk_functor C p₀ ⟶ 𝟭 C := +{ app := λ c, (skyscraper_presheaf_stalk_of_mem_closure₀ p₀ c ts (specializes_rfl : p₀ ⤳ p₀)).hom, + naturality' := λ x y f, + begin + ext U, + dsimp, + simp only [colimit.iso_colimit_cocone_ι_hom_assoc, + skyscraper_presheaf_cocone_of_mem_closure₀_ι_app, category.assoc], + erw [category.id_comp, ←category.assoc, colimit.ι_map], + dsimp, + split_ifs, + { rw [category.assoc, skyscraper_presheaf_stalk_of_mem_closure₀, + colimit.iso_colimit_cocone_ι_hom], + dsimp, + simpa, }, + { exfalso, exact h U.unop.2 } + end } + +noncomputable def stalk_skyscraper_presheaf_adj : + presheaf.stalk_functor C p₀ ⊣ skyscraper_presheaf_functor p₀ ts := +{ hom_equiv := λ 𝓕 c, ⟨from_stalk_to_to_skyscraper_presheaf p₀ ts, + to_skyscraper_presheaf_to_from_stalk p₀ ts, + from_stalk_to_to_skyscraper_presheaf_to_skyscraper_presheaf_to_from_stalk p₀ ts, + to_skyscraper_presheaf_to_from_stalk_from_stalk_to_to_skyscraper_presheaf p₀ ts⟩, + unit := stalk_skyscraper_presheaf_adj_unit p₀ ts, + counit := stalk_skyscraper_presheaf_adj_counit p₀ ts, + hom_equiv_unit' := λ 𝓕 𝓖 f, + begin + ext U, + dsimp, + split_ifs, + { simp, }, + { rw [eq_comp_eq_to_hom], + exact ts.hom_ext _ _ }, + end, + hom_equiv_counit' := λ 𝓕 c g, + begin + ext U, + dsimp, + erw [colimit.ι_desc, ←category.assoc, colimit.ι_map, category.assoc, + colimit.iso_colimit_cocone_ι_hom], + dsimp, + rw [category.comp_id], + refl, + end } + +example : true := trivial + +noncomputable def stalk_skyscraper_sheaf_adj : + sheaf.forget C X ⋙ presheaf.stalk_functor _ p₀ ⊣ skyscraper_sheaf_functor p₀ ts := +{ hom_equiv := λ 𝓕 c, + ⟨λ f, ⟨from_stalk_to_to_skyscraper_presheaf p₀ ts f⟩, + λ g, to_skyscraper_presheaf_to_from_stalk p₀ ts g.1, + λ f, from_stalk_to_to_skyscraper_presheaf_to_skyscraper_presheaf_to_from_stalk + p₀ ts f, + begin + intros g, + ext1, + exact to_skyscraper_presheaf_to_from_stalk_from_stalk_to_to_skyscraper_presheaf p₀ ts g.1, + end⟩, + unit := + { app := λ 𝓕, ⟨(stalk_skyscraper_presheaf_adj_unit p₀ ts).app 𝓕.1⟩, + naturality' := λ 𝓐 𝓑 ⟨f⟩, + begin + ext1, + dsimp, + exact (stalk_skyscraper_presheaf_adj_unit p₀ ts).naturality f, + end }, + counit := stalk_skyscraper_presheaf_adj_counit p₀ ts, + hom_equiv_unit' := + begin + intros 𝓐 c f, + ext1, + exact (stalk_skyscraper_presheaf_adj p₀ ts).hom_equiv_unit, + end, + hom_equiv_counit' := λ 𝓐 c f, (stalk_skyscraper_presheaf_adj p₀ ts).hom_equiv_counit } + +end adjoints + +section injective + +open_locale zero_object +open topological_space +open category_theory category_theory.limits +open Top +open opposite + +universe u +variables {X : Top.{u}} (p₀ : X) [Π (U : opens X), decidable (p₀ ∈ U)] + +lemma skyscraper_presheaf_in_Ab_injective (S : Ab.{u}) [injective S] : + injective (skyscraper_sheaf p₀ S (is_zero.is_terminal (is_zero_zero _) : is_terminal (0 : Ab))) := +injective.injective_of_adjoint + (stalk_skyscraper_sheaf_adj p₀ (is_zero.is_terminal (is_zero_zero _) : is_terminal (0 : Ab))) + +end injective From 8c38e63cf30605b4006b883a8e3e68708693c490 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Tue, 16 Aug 2022 02:38:41 +0100 Subject: [PATCH 64/84] delete unnecessary --- src/topology/sheaves/epi_mono.lean | 97 ------------------------------ 1 file changed, 97 deletions(-) delete mode 100644 src/topology/sheaves/epi_mono.lean diff --git a/src/topology/sheaves/epi_mono.lean b/src/topology/sheaves/epi_mono.lean deleted file mode 100644 index 67a4fac9e66d1..0000000000000 --- a/src/topology/sheaves/epi_mono.lean +++ /dev/null @@ -1,97 +0,0 @@ -/- -Copyright (c) 2022 Jujian Zhang. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Jujian Zhang --/ -import topology.sheaves.sheaf -import category_theory.sites.sheafification - -/-! -# Epimorphisms and monomorphisms in category of (pre)sheaves - -Let `X` be a topological space and `F, G` be sheaves on `X`. Then `f : F ⟶ G` is monic iff it is -monic as presheaf morphisms; `f` is epic if `f` is epic as presheaf morphism, but in this case, the -converse is **not** necessarily true. - --/ - - -universes v u - -open category_theory category_theory.limits -open topological_space -open Top - - -variables {X : Top.{v}} {D : Type u} [category.{v u} D] - -section sheaf - -variables {F G : sheaf D X} (f : F ⟶ G) - --- Not an instance to prevent class search forming a loop -lemma mono_of_presheaf_mono [mono f.1] : mono f := -{ right_cancellation := λ P g h eq₀, Sheaf.hom.ext _ _ $ (cancel_mono f.1).mp $ - (Sheaf.hom.ext_iff _ _).mp eq₀ } - -instance sheaf.forget_reflects_mono : (sheaf.forget D X).reflects_monomorphisms := -{ reflects := λ F G f m, @@mono_of_presheaf_mono _ f m } - -instance epi_of_presheaf_epi [epi f.1] : epi f := -{ left_cancellation := λ P g h eq₀, Sheaf.hom.ext _ _ $ (cancel_epi f.1).mp $ - (Sheaf.hom.ext_iff _ _).mp eq₀ } - -instance sheaf_forget_reflects_epi : (sheaf.forget D X).reflects_epimorphisms := -{ reflects := λ F G f m, @@epi_of_presheaf_epi _ f m } - -section to_presheaf - -/-! -Given two sheaves `F, G` on `X` in a cocomplete concrete category `D` where "sheafification" is -possible, then a monic sheaf morphism `f : F ⟶ G` is also a monic presheaf morphism. --/ - --- assumptions on target category so that "sheafification" is possible -variables [concrete_category.{v} D] [preserves_limits (forget D)] [has_colimits D] -variables [Π (P : presheaf D X) (U : opens X) (S : (opens.grothendieck_topology X).cover U), - has_multiequalizer (S.index P)] -variables [Π (U : opens X), - preserves_colimits_of_shape ((opens.grothendieck_topology X).cover U)ᵒᵖ (forget D)] -variables [reflects_isomorphisms (forget D)] - --- Not an instance to prevent class search forming a loop -lemma presheaf_mono_of_mono [mono f] : mono f.1 := -{ right_cancellation := λ P g h eq₀, - begin - set P_plus : sheaf D X := (presheaf_to_Sheaf (opens.grothendieck_topology X) D).obj P, - set g_plus : P_plus ⟶ F := - ⟨grothendieck_topology.sheafify_lift (opens.grothendieck_topology X) g F.2⟩ with g_plus_def, - set h_plus : P_plus ⟶ F := - ⟨grothendieck_topology.sheafify_lift (opens.grothendieck_topology X) h F.2⟩ with h_plus_def, - have eq₁ : g_plus ≫ f = h_plus ≫ f := Sheaf.hom.ext _ _ (grothendieck_topology.sheafify_hom_ext - _ _ _ G.2 (by simp [*])), - have eq₂ : g_plus = h_plus, - { rwa cancel_mono at eq₁, }, - rw [g_plus_def, h_plus_def, Sheaf.hom.ext_iff] at eq₂, - dsimp at eq₂, - rw ←grothendieck_topology.to_sheafify_sheafify_lift (opens.grothendieck_topology X) g F.2, - rw ←grothendieck_topology.to_sheafify_sheafify_lift (opens.grothendieck_topology X) h F.2, - rw eq₂, - end } - -instance sheaf.forget_preserves_mono : (sheaf.forget D X).preserves_monomorphisms := -{ preserves := λ F G f m, @presheaf_mono_of_mono X D _ F G f _ _ _ _ _ _ m } - -lemma sheaf.mono_iff_presheaf_mono : mono f ↔ mono f.1 := -⟨λ m, @@presheaf_mono_of_mono _ f _ _ _ _ _ _ m, λ m, @@mono_of_presheaf_mono _ f m⟩ - -/- -This is not always true -``` -instance presheaf_epi_of_epi [epi f] : epi f.1 := not true -``` --/ - -end to_presheaf - -end sheaf From dc5b5d0d852aef8095092b7bc0c53fec1471e582 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Tue, 16 Aug 2022 02:50:44 +0100 Subject: [PATCH 65/84] need cache --- src/topology/sheaves/skyscraper.lean | 43 ++++++++++++++-------------- 1 file changed, 21 insertions(+), 22 deletions(-) diff --git a/src/topology/sheaves/skyscraper.lean b/src/topology/sheaves/skyscraper.lean index d8620992bb1b1..b21b62c3a9251 100644 --- a/src/topology/sheaves/skyscraper.lean +++ b/src/topology/sheaves/skyscraper.lean @@ -370,8 +370,7 @@ open_locale classical universes u v w -variables {X : Top.{u}} (p₀ : X) {C : Type v} [category.{w} C] (S : C) -variables {star : C} (ts : is_terminal star) +variables {X : Top.{u}} (p₀ : X) {C : Type v} [category.{w} C] (S : C) [has_terminal C] /--The topological space with on a point `{p₀}`. Hence the only open sets are ∅ and ⊤.-/ def single_point_space : Top := ⟨({p₀} : set X), infer_instance⟩ @@ -384,17 +383,17 @@ The presheaf on a single point space `{p₀}` defined by `∅ ↦ *` and `{p₀} -/ @[simps] noncomputable def single_point_presheaf : presheaf C (single_point_space p₀) := -{ obj := λ U, if U.unop ≠ ⊥ then S else star, +{ obj := λ U, if U.unop ≠ ⊥ then S else terminal C, map := λ U V inc, if h : V.unop ≠ ⊥ then eq_to_hom (if_pos $ λ r, h $ le_bot_iff.mp $ r ▸ le_of_hom inc.unop) ≫ 𝟙 S ≫ eq_to_hom (if_pos h).symm - else ts.from _ ≫ 𝟙 star ≫ eq_to_hom (if_neg h).symm, + else terminal.from _ ≫ eq_to_hom (if_neg h).symm, map_id' := λ U, begin split_ifs, { rw [category.id_comp, eq_to_hom_trans, eq_to_hom_refl] }, - { rw [category.id_comp, eq_comp_eq_to_hom, category.id_comp], - exact ts.hom_ext _ _, }, + { rw [eq_comp_eq_to_hom, category.id_comp], + exact terminal_is_terminal.hom_ext _ _, }, end, map_comp' := λ U V W inc1 inc2, begin @@ -409,9 +408,9 @@ noncomputable def single_point_presheaf : presheaf C (single_point_space p₀) : split_ifs, simp only [category.id_comp, eq_to_hom_trans], }, { split_ifs; - rw [category.id_comp, eq_comp_eq_to_hom, category.assoc, category.assoc, eq_to_hom_trans, + rw [eq_comp_eq_to_hom, category.assoc, category.assoc, eq_to_hom_trans, eq_to_hom_refl, category.comp_id]; - exact ts.hom_ext _ _, }, + exact terminal_is_terminal.hom_ext _ _, }, end } /-- @@ -425,9 +424,9 @@ The trivial inclusion `{p₀} ↪ X`. The morphism from skyscraper presheaf to pushforward sheaf -/ @[simps] noncomputable def skyscraper_presheaf_to_pushforward : - skyscraper_presheaf p₀ S ts ⟶ (single_point_inclusion p₀) _* (single_point_presheaf p₀ S ts) := + skyscraper_presheaf p₀ S ⟶ (single_point_inclusion p₀) _* (single_point_presheaf p₀ S) := { app := λ U, if h : p₀ ∈ U.unop - then eq_to_hom (skyscraper_presheaf_obj_of_mem _ _ h) ≫ 𝟙 S ≫ eq_to_hom + then eq_to_hom (skyscraper_presheaf_obj_of_mem _ h) ≫ eq_to_hom begin dsimp, rw if_pos, @@ -438,7 +437,7 @@ The morphism from skyscraper presheaf to pushforward sheaf erw set.mem_preimage, exact h, end - else ts.from _ ≫ eq_to_hom + else terminal.from _ ≫ eq_to_hom begin dsimp, rw if_neg, @@ -468,15 +467,15 @@ The morphism from skyscraper presheaf to pushforward sheaf simp only [eq_to_hom_trans, category.id_comp], }, }, { split_ifs; rw [←category.assoc, eq_comp_eq_to_hom]; - exact ts.hom_ext _ _ }, + exact terminal_is_terminal.hom_ext _ _ }, end } /-- The morphism from pushforward sheaf to skyscraper presheaf -/ @[simps] noncomputable def pushforward_to_skyscraper_presheaf : - (single_point_inclusion p₀) _* (single_point_presheaf p₀ S ts) ⟶ - skyscraper_presheaf p₀ S ts := + (single_point_inclusion p₀) _* (single_point_presheaf p₀ S) ⟶ + skyscraper_presheaf p₀ S := { app := λ U, if h : p₀ ∈ unop U then eq_to_hom begin @@ -488,8 +487,8 @@ The morphism from pushforward sheaf to skyscraper presheaf refine ⟨⟨p₀, rfl⟩, _⟩, erw set.mem_preimage, exact h, - end ≫ 𝟙 S ≫ eq_to_hom (skyscraper_presheaf_obj_of_mem _ _ h).symm - else ts.from _ ≫ eq_to_hom (skyscraper_presheaf_obj_of_not_mem _ _ h).symm, + end ≫ eq_to_hom (skyscraper_presheaf_obj_of_mem _ h).symm + else terminal.from _ ≫ eq_to_hom (skyscraper_presheaf_obj_of_not_mem _ h).symm, naturality' := λ U V inc, begin rw [comp_dite, dite_comp], @@ -505,26 +504,26 @@ The morphism from pushforward sheaf to skyscraper presheaf { exfalso, exact hV' h, }, { dsimp, split_ifs; - rw [category.id_comp, eq_to_hom_trans, category.id_comp, eq_to_hom_trans, eq_to_hom_trans, + rw [eq_to_hom_trans, category.id_comp, eq_to_hom_trans, eq_to_hom_trans, category.id_comp, eq_to_hom_trans, eq_to_hom_trans, eq_to_hom_trans], }, }, { split_ifs; rw [←category.assoc, eq_comp_eq_to_hom]; - exact ts.hom_ext _ _ }, + exact terminal_is_terminal.hom_ext _ _ }, end } /-- Skyscraper presheaf is isomorphic to pushforward of sheaf on single point. -/ noncomputable def skyscraper_presheaf_as_pushforward : - skyscraper_presheaf p₀ S ts ≅ (single_point_inclusion p₀) _* (single_point_presheaf p₀ S ts) := -{ hom := skyscraper_presheaf_to_pushforward p₀ S ts, - inv := pushforward_to_skyscraper_presheaf p₀ S ts, + skyscraper_presheaf p₀ S ≅ (single_point_inclusion p₀) _* (single_point_presheaf p₀ S) := +{ hom := skyscraper_presheaf_to_pushforward p₀ S, + inv := pushforward_to_skyscraper_presheaf p₀ S, hom_inv_id' := begin ext U, dsimp, split_ifs, - { rw [category.id_comp, eq_to_hom_trans, category.id_comp, eq_to_hom_trans, eq_to_hom_trans, + { rw [eq_to_hom_trans, category.id_comp, eq_to_hom_trans, eq_to_hom_trans, eq_to_hom_refl], }, { rw [←category.assoc, eq_comp_eq_to_hom], exact ts.hom_ext _ _, }, From 8b13323d116c8c4af33462aa28570d820b12b4ea Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Tue, 16 Aug 2022 03:49:46 +0100 Subject: [PATCH 66/84] Merge remote-tracking branch 'origin/master' into jjaassoonn/sky --- src/topology/sheaves/stalks.lean | 1 - 1 file changed, 1 deletion(-) diff --git a/src/topology/sheaves/stalks.lean b/src/topology/sheaves/stalks.lean index ea39bf3cf7738..2712297985015 100644 --- a/src/topology/sheaves/stalks.lean +++ b/src/topology/sheaves/stalks.lean @@ -17,7 +17,6 @@ import algebra.category.Group.colimits import algebra.category.Group.limits import algebra.category.Group.filtered_colimits import algebra.category.Group.epi_mono -import topology.sheaves.epi_mono /-! # Stalks From cf4bbe8d9eff947a23e5f4b469a0b6230f1974c2 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Tue, 16 Aug 2022 04:12:58 +0100 Subject: [PATCH 67/84] fix --- src/topology/sheaves/stalks.lean | 14 ++++++++------ 1 file changed, 8 insertions(+), 6 deletions(-) diff --git a/src/topology/sheaves/stalks.lean b/src/topology/sheaves/stalks.lean index 2712297985015..b7fffdfb09adc 100644 --- a/src/topology/sheaves/stalks.lean +++ b/src/topology/sheaves/stalks.lean @@ -17,6 +17,7 @@ import algebra.category.Group.colimits import algebra.category.Group.limits import algebra.category.Group.filtered_colimits import algebra.category.Group.epi_mono +import category_theory.sites.sheafification /-! # Stalks @@ -430,19 +431,20 @@ lemma app_injective_iff_stalk_functor_map_injective {F : sheaf C X} This works for any concrete category where monomorphism = injective functions. -/ instance (x : X) : - functor.preserves_monomorphisms (sheaf.forget Ab.{v} X ⋙ stalk_functor Ab.{v} x) := + functor.preserves_monomorphisms (sheaf.forget C X ⋙ stalk_functor C x) := begin refine ⟨λ 𝓐 𝓑 f im, _⟩, apply concrete_category.mono_of_injective, dsimp, - haveI : preserves_limits (forget Ab.{v}) := AddCommGroup.forget_preserves_limits.{v v}, - haveI : mono f.1 := @@presheaf_mono_of_mono _ _ _ _ _ _ _ _ im, - have := (nat_trans.mono_iff_app_mono _ f.1).mp _, - simp_rw [AddCommGroup.mono_iff_injective] at this, + haveI : preserves_limits_of_shape walking_cospan (forget C) := infer_instance, + haveI im2 : mono f.1 := @@category_theory.presheaf_mono_of_mono _ _ _ _ _ _ _ _ _ _ _ im, + have im3 := (nat_trans.mono_iff_app_mono _ f.1).mp im2, + have : ∀ (c : (opens X)ᵒᵖ), function.injective (f.1.app c) := + λ c, (@@concrete_category.mono_iff_injective_of_preserves_pullback _ _ (f.1.app c) (im3 c) _).mp + (im3 c), refine (app_injective_iff_stalk_functor_map_injective f.1).mpr _ x, rintros U, apply this, - apply_instance, end /-- For surjectivity, we are given an arbitrary section `t` and need to find a preimage for it. From 081ff8409a7c0317d1d3e198f67d473d2d7ba954 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Tue, 16 Aug 2022 07:53:24 +0100 Subject: [PATCH 68/84] generalize --- src/topology/sheaves/skyscraper.lean | 126 +++++++++++++-------------- src/topology/sheaves/stalks.lean | 11 ++- 2 files changed, 70 insertions(+), 67 deletions(-) diff --git a/src/topology/sheaves/skyscraper.lean b/src/topology/sheaves/skyscraper.lean index b21b62c3a9251..8ab8a81288914 100644 --- a/src/topology/sheaves/skyscraper.lean +++ b/src/topology/sheaves/skyscraper.lean @@ -6,6 +6,7 @@ Authors: Jujian Zhang import algebraic_geometry.sheafed_space import topology.sheaves.sheaf_condition.unique_gluing import topology.sheaves.stalks +import category_theory.sites.sheafification import category_theory.preadditive.injective import algebra.category.Group.abelian @@ -523,10 +524,9 @@ noncomputable def skyscraper_presheaf_as_pushforward : ext U, dsimp, split_ifs, - { rw [eq_to_hom_trans, category.id_comp, eq_to_hom_trans, eq_to_hom_trans, - eq_to_hom_refl], }, + { rw [eq_to_hom_trans, eq_to_hom_trans, eq_to_hom_trans, eq_to_hom_refl], }, { rw [←category.assoc, eq_comp_eq_to_hom], - exact ts.hom_ext _ _, }, + exact terminal_is_terminal.hom_ext _ _, }, end, inv_hom_id' := begin @@ -534,11 +534,10 @@ noncomputable def skyscraper_presheaf_as_pushforward : dsimp, by_cases hU : p₀ ∈ U.unop, { split_ifs; - rw [category.id_comp, eq_to_hom_trans, category.id_comp, eq_to_hom_trans, eq_to_hom_trans, - eq_to_hom_refl], }, + rw [eq_to_hom_trans, eq_to_hom_trans, eq_to_hom_trans, eq_to_hom_refl], }, { split_ifs; rw [←category.assoc, eq_comp_eq_to_hom]; - exact ts.hom_ext _ _, }, + exact terminal_is_terminal.hom_ext _ _, }, end } end @@ -553,17 +552,17 @@ open opposite universes u v variables {X : Top.{u}} (p₀ : X) {C : Type v} [category.{u} C] -variables {star : C} (ts : is_terminal star) +variables [has_terminal C] -- {star : C} (ts : is_terminal star) variable [Π (U : opens X), decidable (p₀ ∈ U)] @[simps] def skyscraper_presheaf_functor : C ⥤ presheaf C X := -{ obj := λ S, skyscraper_presheaf p₀ S ts, +{ obj := λ S, skyscraper_presheaf p₀ S, map := λ x y f, { app := λ U, if h : p₀ ∈ U.unop - then eq_to_hom (skyscraper_presheaf_obj_of_mem _ _ h) ≫ f ≫ - eq_to_hom (skyscraper_presheaf_obj_of_mem _ _ h).symm - else ts.from _ ≫ eq_to_hom (skyscraper_presheaf_obj_of_not_mem _ _ h).symm, + then eq_to_hom (skyscraper_presheaf_obj_of_mem _ h) ≫ f ≫ + eq_to_hom (skyscraper_presheaf_obj_of_mem _ h).symm + else terminal.from _ ≫ eq_to_hom (skyscraper_presheaf_obj_of_not_mem _ h).symm, naturality' := λ U V inc, begin dsimp, @@ -575,7 +574,7 @@ def skyscraper_presheaf_functor : C ⥤ presheaf C X := refl, }, { split_ifs; rw [←category.assoc, eq_comp_eq_to_hom]; - exact ts.hom_ext _ _ } + exact terminal_is_terminal.hom_ext _ _ } end }, map_id' := λ c, begin @@ -584,7 +583,7 @@ def skyscraper_presheaf_functor : C ⥤ presheaf C X := split_ifs, { simp, }, { rw [eq_comp_eq_to_hom], - exact ts.hom_ext _ _ }, + exact terminal_is_terminal.hom_ext _ _ }, end, map_comp' := λ x y z f g, begin @@ -593,34 +592,32 @@ def skyscraper_presheaf_functor : C ⥤ presheaf C X := split_ifs, { simp }, { rw [eq_comp_eq_to_hom], - exact ts.hom_ext _ _ }, + exact terminal_is_terminal.hom_ext _ _ }, end } -example : true := trivial - @[simps] def skyscraper_sheaf_functor : C ⥤ sheaf C X := -{ obj := λ S, skyscraper_sheaf p₀ S ts, - map := λ x y f, ⟨(skyscraper_presheaf_functor p₀ ts).map f⟩, +{ obj := λ S, skyscraper_sheaf p₀ S, + map := λ x y f, ⟨(skyscraper_presheaf_functor p₀).map f⟩, map_id' := λ c, begin ext1, - exact (skyscraper_presheaf_functor p₀ ts).map_id c, + exact (skyscraper_presheaf_functor p₀).map_id c, end, map_comp' := λ x y z f g, begin ext1, - exact (skyscraper_presheaf_functor p₀ ts).map_comp f g, + exact (skyscraper_presheaf_functor p₀).map_comp f g, end } variable [has_colimits C] @[simps] noncomputable def from_stalk_to_to_skyscraper_presheaf {𝓕 : presheaf C X} {c : C} - (f : 𝓕.stalk p₀ ⟶ c) : 𝓕 ⟶ skyscraper_presheaf p₀ c ts := + (f : 𝓕.stalk p₀ ⟶ c) : 𝓕 ⟶ skyscraper_presheaf p₀ c := { app := λ U, if h : p₀ ∈ U.unop - then 𝓕.germ ⟨p₀, h⟩ ≫ f ≫ eq_to_hom (skyscraper_presheaf_obj_of_mem _ _ h).symm - else ts.from _ ≫ eq_to_hom (skyscraper_presheaf_obj_of_not_mem _ _ h).symm, + then 𝓕.germ ⟨p₀, h⟩ ≫ f ≫ eq_to_hom (skyscraper_presheaf_obj_of_mem _ h).symm + else terminal.from _ ≫ eq_to_hom (skyscraper_presheaf_obj_of_not_mem _ h).symm, naturality' := λ U V inc, begin dsimp, @@ -633,12 +630,12 @@ noncomputable def from_stalk_to_to_skyscraper_presheaf {𝓕 : presheaf C X} {c { split_ifs, rw [←category.assoc, eq_comp_eq_to_hom, category.assoc, category.assoc, eq_to_hom_trans, eq_to_hom_refl, category.comp_id], - exact ts.hom_ext _ _ }, + exact terminal_is_terminal.hom_ext _ _ }, end } @[reducible] noncomputable def to_skyscraper_presheaf_to_from_stalk {𝓕 : presheaf C X} {c : C} - (f : 𝓕 ⟶ skyscraper_presheaf p₀ c ts) : 𝓕.stalk p₀ ⟶ c := + (f : 𝓕 ⟶ skyscraper_presheaf p₀ c) : 𝓕.stalk p₀ ⟶ c := let CC : cocone ((open_nhds.inclusion p₀).op ⋙ 𝓕) := { X := c, ι := @@ -665,7 +662,7 @@ colimit.desc _ CC lemma from_stalk_to_to_skyscraper_presheaf_to_skyscraper_presheaf_to_from_stalk {𝓕 : presheaf C X} {c : C} (f : 𝓕.stalk p₀ ⟶ c) : -to_skyscraper_presheaf_to_from_stalk p₀ ts (from_stalk_to_to_skyscraper_presheaf p₀ ts f) = f := +to_skyscraper_presheaf_to_from_stalk p₀ (from_stalk_to_to_skyscraper_presheaf p₀ f) = f := begin ext U, dsimp, @@ -683,8 +680,8 @@ begin end lemma to_skyscraper_presheaf_to_from_stalk_from_stalk_to_to_skyscraper_presheaf - {𝓕 : presheaf C X} {c : C} (f : 𝓕 ⟶ skyscraper_presheaf p₀ c ts) : -from_stalk_to_to_skyscraper_presheaf p₀ ts (to_skyscraper_presheaf_to_from_stalk p₀ ts f) = f := + {𝓕 : presheaf C X} {c : C} (f : 𝓕 ⟶ skyscraper_presheaf p₀ c) : +from_stalk_to_to_skyscraper_presheaf p₀ (to_skyscraper_presheaf_to_from_stalk p₀ f) = f := begin ext U, dsimp, @@ -696,16 +693,16 @@ begin dsimp, rw [category.assoc, eq_to_hom_trans, eq_to_hom_refl, category.comp_id], }, { rw [eq_comp_eq_to_hom], - exact ts.hom_ext _ _ , } + exact terminal_is_terminal.hom_ext _ _ , } end @[simps] noncomputable def stalk_skyscraper_presheaf_adj_unit : - 𝟭 (presheaf C X) ⟶ presheaf.stalk_functor C p₀ ⋙ skyscraper_presheaf_functor p₀ ts := + 𝟭 (presheaf C X) ⟶ presheaf.stalk_functor C p₀ ⋙ skyscraper_presheaf_functor p₀ := { app := λ 𝓕, { app := λ U, if h : p₀ ∈ U.unop then 𝓕.germ ⟨p₀, h⟩ ≫ eq_to_hom (if_pos h).symm - else ts.from _ ≫ eq_to_hom (if_neg h).symm, + else terminal.from _ ≫ eq_to_hom (if_neg h).symm, naturality' := λ U V inc, begin dsimp, @@ -717,7 +714,7 @@ noncomputable def stalk_skyscraper_presheaf_adj_unit : congr, }, { split_ifs, rw [←category.assoc, eq_comp_eq_to_hom], - exact ts.hom_ext _ _, }, + exact terminal_is_terminal.hom_ext _ _, }, end }, naturality' := λ 𝓕 𝓖 f, begin @@ -731,20 +728,20 @@ noncomputable def stalk_skyscraper_presheaf_adj_unit : erw [colimit.ι_map], congr, }, { rw [←category.assoc, eq_comp_eq_to_hom], - exact ts.hom_ext _ _ }, + exact terminal_is_terminal.hom_ext _ _ }, end } @[simps] noncomputable def stalk_skyscraper_presheaf_adj_counit : - skyscraper_presheaf_functor p₀ ts ⋙ presheaf.stalk_functor C p₀ ⟶ 𝟭 C := -{ app := λ c, (skyscraper_presheaf_stalk_of_mem_closure₀ p₀ c ts (specializes_rfl : p₀ ⤳ p₀)).hom, + skyscraper_presheaf_functor p₀ ⋙ presheaf.stalk_functor C p₀ ⟶ 𝟭 C := +{ app := λ c, (skyscraper_presheaf_stalk_of_mem_closure₀ p₀ c (specializes_rfl : p₀ ⤳ p₀)).hom, naturality' := λ x y f, begin ext U, dsimp, simp only [colimit.iso_colimit_cocone_ι_hom_assoc, skyscraper_presheaf_cocone_of_mem_closure₀_ι_app, category.assoc], - erw [category.id_comp, ←category.assoc, colimit.ι_map], + erw [←category.assoc, colimit.ι_map], dsimp, split_ifs, { rw [category.assoc, skyscraper_presheaf_stalk_of_mem_closure₀, @@ -755,13 +752,13 @@ noncomputable def stalk_skyscraper_presheaf_adj_counit : end } noncomputable def stalk_skyscraper_presheaf_adj : - presheaf.stalk_functor C p₀ ⊣ skyscraper_presheaf_functor p₀ ts := -{ hom_equiv := λ 𝓕 c, ⟨from_stalk_to_to_skyscraper_presheaf p₀ ts, - to_skyscraper_presheaf_to_from_stalk p₀ ts, - from_stalk_to_to_skyscraper_presheaf_to_skyscraper_presheaf_to_from_stalk p₀ ts, - to_skyscraper_presheaf_to_from_stalk_from_stalk_to_to_skyscraper_presheaf p₀ ts⟩, - unit := stalk_skyscraper_presheaf_adj_unit p₀ ts, - counit := stalk_skyscraper_presheaf_adj_counit p₀ ts, + presheaf.stalk_functor C p₀ ⊣ skyscraper_presheaf_functor p₀ := +{ hom_equiv := λ 𝓕 c, ⟨from_stalk_to_to_skyscraper_presheaf p₀, + to_skyscraper_presheaf_to_from_stalk p₀, + from_stalk_to_to_skyscraper_presheaf_to_skyscraper_presheaf_to_from_stalk p₀, + to_skyscraper_presheaf_to_from_stalk_from_stalk_to_to_skyscraper_presheaf p₀⟩, + unit := stalk_skyscraper_presheaf_adj_unit p₀, + counit := stalk_skyscraper_presheaf_adj_counit p₀, hom_equiv_unit' := λ 𝓕 𝓖 f, begin ext U, @@ -769,7 +766,7 @@ noncomputable def stalk_skyscraper_presheaf_adj : split_ifs, { simp, }, { rw [eq_comp_eq_to_hom], - exact ts.hom_ext _ _ }, + exact terminal_is_terminal.hom_ext _ _ }, end, hom_equiv_counit' := λ 𝓕 c g, begin @@ -778,57 +775,60 @@ noncomputable def stalk_skyscraper_presheaf_adj : erw [colimit.ι_desc, ←category.assoc, colimit.ι_map, category.assoc, colimit.iso_colimit_cocone_ι_hom], dsimp, - rw [category.comp_id], + rw [eq_comp_eq_to_hom, category.assoc, eq_to_hom_trans, eq_to_hom_refl, + category.comp_id], refl, end } -example : true := trivial - +variable (C) noncomputable def stalk_skyscraper_sheaf_adj : - sheaf.forget C X ⋙ presheaf.stalk_functor _ p₀ ⊣ skyscraper_sheaf_functor p₀ ts := + sheaf.forget C X ⋙ presheaf.stalk_functor _ p₀ ⊣ skyscraper_sheaf_functor p₀ := { hom_equiv := λ 𝓕 c, - ⟨λ f, ⟨from_stalk_to_to_skyscraper_presheaf p₀ ts f⟩, - λ g, to_skyscraper_presheaf_to_from_stalk p₀ ts g.1, + ⟨λ f, ⟨from_stalk_to_to_skyscraper_presheaf p₀ f⟩, + λ g, to_skyscraper_presheaf_to_from_stalk p₀ g.1, λ f, from_stalk_to_to_skyscraper_presheaf_to_skyscraper_presheaf_to_from_stalk - p₀ ts f, + p₀ f, begin intros g, ext1, - exact to_skyscraper_presheaf_to_from_stalk_from_stalk_to_to_skyscraper_presheaf p₀ ts g.1, + exact to_skyscraper_presheaf_to_from_stalk_from_stalk_to_to_skyscraper_presheaf p₀ g.1, end⟩, unit := - { app := λ 𝓕, ⟨(stalk_skyscraper_presheaf_adj_unit p₀ ts).app 𝓕.1⟩, + { app := λ 𝓕, ⟨(stalk_skyscraper_presheaf_adj_unit p₀).app 𝓕.1⟩, naturality' := λ 𝓐 𝓑 ⟨f⟩, begin ext1, dsimp, - exact (stalk_skyscraper_presheaf_adj_unit p₀ ts).naturality f, + exact (stalk_skyscraper_presheaf_adj_unit p₀).naturality f, end }, - counit := stalk_skyscraper_presheaf_adj_counit p₀ ts, + counit := stalk_skyscraper_presheaf_adj_counit p₀, hom_equiv_unit' := begin intros 𝓐 c f, ext1, - exact (stalk_skyscraper_presheaf_adj p₀ ts).hom_equiv_unit, + exact (stalk_skyscraper_presheaf_adj p₀).hom_equiv_unit, end, - hom_equiv_counit' := λ 𝓐 c f, (stalk_skyscraper_presheaf_adj p₀ ts).hom_equiv_counit } + hom_equiv_counit' := λ 𝓐 c f, (stalk_skyscraper_presheaf_adj p₀).hom_equiv_counit } end adjoints section injective -open_locale zero_object +-- need to restrict universe level again + open topological_space open category_theory category_theory.limits open Top open opposite -universe u +universes u variables {X : Top.{u}} (p₀ : X) [Π (U : opens X), decidable (p₀ ∈ U)] +variables {C : Type u} [category.{u} C] [concrete_category.{u} C] [has_colimits C] +variables [has_limits C] [preserves_limits (forget C)] [reflects_isomorphisms (forget C)] +variables [preserves_filtered_colimits (forget C)] -lemma skyscraper_presheaf_in_Ab_injective (S : Ab.{u}) [injective S] : - injective (skyscraper_sheaf p₀ S (is_zero.is_terminal (is_zero_zero _) : is_terminal (0 : Ab))) := -injective.injective_of_adjoint - (stalk_skyscraper_sheaf_adj p₀ (is_zero.is_terminal (is_zero_zero _) : is_terminal (0 : Ab))) +lemma skyscraper_presheaf_in_Ab_injective (S : C) [injective S] : + injective (skyscraper_sheaf p₀ S) := +injective.injective_of_adjoint (stalk_skyscraper_sheaf_adj p₀ C) S end injective diff --git a/src/topology/sheaves/stalks.lean b/src/topology/sheaves/stalks.lean index b7fffdfb09adc..478ebfbb6b8df 100644 --- a/src/topology/sheaves/stalks.lean +++ b/src/topology/sheaves/stalks.lean @@ -427,10 +427,11 @@ lemma app_injective_iff_stalk_functor_map_injective {F : sheaf C X} ⟨λ h U, app_injective_of_stalk_functor_map_injective f U (λ x, h x.1), stalk_functor_map_injective_of_app_injective f⟩ -/-- -This works for any concrete category where monomorphism = injective functions. --/ -instance (x : X) : +section + +variable (C) + +instance stalk_functor_preserves_mono (x : X) : functor.preserves_monomorphisms (sheaf.forget C X ⋙ stalk_functor C x) := begin refine ⟨λ 𝓐 𝓑 f im, _⟩, @@ -447,6 +448,8 @@ begin apply this, end +end + /-- For surjectivity, we are given an arbitrary section `t` and need to find a preimage for it. We claim that it suffices to find preimages *locally*. That is, for each `x : U` we construct a neighborhood `V ≤ U` and a section `s : F.obj (op V))` such that `f.app (op V) s` and `t` From 1cf1bbb5eeb4f8029eff60b5b97cff91a15f65e6 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Tue, 16 Aug 2022 07:54:19 +0100 Subject: [PATCH 69/84] minimize import --- src/topology/sheaves/skyscraper.lean | 1 - src/topology/sheaves/stalks.lean | 5 ----- 2 files changed, 6 deletions(-) diff --git a/src/topology/sheaves/skyscraper.lean b/src/topology/sheaves/skyscraper.lean index 8ab8a81288914..477775b5498b9 100644 --- a/src/topology/sheaves/skyscraper.lean +++ b/src/topology/sheaves/skyscraper.lean @@ -8,7 +8,6 @@ import topology.sheaves.sheaf_condition.unique_gluing import topology.sheaves.stalks import category_theory.sites.sheafification import category_theory.preadditive.injective -import algebra.category.Group.abelian /-! # Skyscraper (pre)sheaves diff --git a/src/topology/sheaves/stalks.lean b/src/topology/sheaves/stalks.lean index 478ebfbb6b8df..b72dea32f2121 100644 --- a/src/topology/sheaves/stalks.lean +++ b/src/topology/sheaves/stalks.lean @@ -12,11 +12,6 @@ import category_theory.limits.final import category_theory.adjunction.evaluation import topology.sober import tactic.elementwise -import algebra.category.Ring -import algebra.category.Group.colimits -import algebra.category.Group.limits -import algebra.category.Group.filtered_colimits -import algebra.category.Group.epi_mono import category_theory.sites.sheafification /-! From 67ad63477e66a8d804312d3506aa59a2144789aa Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Tue, 16 Aug 2022 07:55:40 +0100 Subject: [PATCH 70/84] fix --- src/topology/sheaves/stalks.lean | 1 + 1 file changed, 1 insertion(+) diff --git a/src/topology/sheaves/stalks.lean b/src/topology/sheaves/stalks.lean index b72dea32f2121..5e048dfdb43be 100644 --- a/src/topology/sheaves/stalks.lean +++ b/src/topology/sheaves/stalks.lean @@ -12,6 +12,7 @@ import category_theory.limits.final import category_theory.adjunction.evaluation import topology.sober import tactic.elementwise +import algebra.category.Ring import category_theory.sites.sheafification /-! From 54292536801184f13f6b171c9a66e3a9dd5e36fb Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Wed, 17 Aug 2022 06:32:49 +0100 Subject: [PATCH 71/84] need cache --- .../concrete_category/basic.lean | 4 +- src/topology/sheaves/godement.lean | 86 +++++++++++++++++++ src/topology/sheaves/skyscraper.lean | 2 +- 3 files changed, 89 insertions(+), 3 deletions(-) create mode 100644 src/topology/sheaves/godement.lean diff --git a/src/category_theory/concrete_category/basic.lean b/src/category_theory/concrete_category/basic.lean index 214947584dfad..44d1596471c36 100644 --- a/src/category_theory/concrete_category/basic.lean +++ b/src/category_theory/concrete_category/basic.lean @@ -140,9 +140,9 @@ begin exact (mono_iff_injective ((forget C).map f)).mp infer_instance, end -lemma concrete_category.mono_iff_injective_of_preserves_pullback {X Y : C} (f : X ⟶ Y) [mono f] +lemma concrete_category.mono_iff_injective_of_preserves_pullback {X Y : C} (f : X ⟶ Y) [preserves_limits_of_shape walking_cospan (forget C)] : mono f ↔ function.injective f := -⟨λ m, concrete_category.injective_of_mono_of_preserves_pullback _, +⟨λ m, by resetI; exact concrete_category.injective_of_mono_of_preserves_pullback _, concrete_category.mono_of_injective f⟩ /-- In any concrete category, surjective morphisms are epimorphisms. -/ diff --git a/src/topology/sheaves/godement.lean b/src/topology/sheaves/godement.lean new file mode 100644 index 0000000000000..64966d1064050 --- /dev/null +++ b/src/topology/sheaves/godement.lean @@ -0,0 +1,86 @@ +import topology.sheaves.sheaf +import topology.sheaves.limits +import topology.sheaves.skyscraper +import topology.sheaves.stalks +import category_theory.preadditive.injective + +noncomputable theory + +section presheaf + +open Top +open topological_space +open category_theory +open category_theory.limits + +universes u v + +variables {X : Top.{u}} {C : Type v} [category.{u} C] +variables [has_limits C] [has_terminal C] [has_colimits C] +variables [Π (x : X) (U : opens X), decidable (x ∈ U)] +variable (𝓕 : presheaf C X) + +def godement : presheaf C X := +∏ (λ x, skyscraper_presheaf x (𝓕.stalk x) : X → presheaf C X) + +@[simps] def godement_app (U : (opens X)ᵒᵖ) : + (godement 𝓕).obj U ≅ ∏ (λ x, (skyscraper_presheaf x (𝓕.stalk x)).obj U) := +limit_obj_iso_limit_comp_evaluation _ _ ≪≫ +{ hom := lim_map + { app := λ _, 𝟙 _, + naturality' := by { rintros ⟨x⟩ ⟨y⟩ ⟨⟨(rfl : x = y)⟩⟩, refl } }, + inv := lim_map + { app := λ _, 𝟙 _, + naturality' := by { rintros ⟨x⟩ ⟨y⟩ ⟨⟨(rfl : x = y)⟩⟩, refl } }, + hom_inv_id' := + begin + dsimp, + ext, + erw [category.assoc, lim_map_π, ←category.assoc, lim_map_π, category.id_comp, category.comp_id, + category.comp_id], + end, + inv_hom_id' := + begin + dsimp, + ext, + erw [category.assoc, lim_map_π, ←category.assoc, lim_map_π, category.comp_id, category.id_comp, + category.comp_id], + end } + +@[reducible] +def to_godement_aux (U : (opens X)ᵒᵖ) : + 𝓕.obj U ⟶ ∏ λ (x : X), ite (x ∈ opposite.unop U) (𝓕.stalk x) (⊤_ C) := +pi.lift (λ x, if m : x ∈ U.unop + then 𝓕.germ ⟨x, m⟩ ≫ eq_to_hom ((skyscraper_presheaf_obj_of_mem _ m).symm.trans (by congr) : + 𝓕.stalk (⟨x, m⟩ : U.unop) = (skyscraper_presheaf x (𝓕.stalk x)).obj U) + else terminal.from _ ≫ eq_to_hom (skyscraper_presheaf_obj_of_not_mem _ m).symm) + +@[simps] def to_godement : 𝓕 ⟶ godement 𝓕 := +{ app := λ U, to_godement_aux 𝓕 U ≫ (godement_app 𝓕 U).inv, + naturality' := + begin + intros U V inc, + ext ⟨x⟩, + dsimp, + simp only [category.assoc, limit_obj_iso_limit_comp_evaluation_inv_π_app, lim_map_π, + category.comp_id, nat_trans.naturality], + simp only [←category.assoc _ _ ((skyscraper_presheaf _ _).map inc), + limit_obj_iso_limit_comp_evaluation_inv_π_app, lim_map_π, category.comp_id], + simp only [limit.lift_π, fan.mk_π_app, skyscraper_presheaf_map, category.id_comp, + eq_to_hom_trans, comp_dite], + dsimp, + split_ifs with hV, + { have hU : x ∈ U.unop := (le_of_hom inc.unop) hV, + split_ifs, + erw [category.assoc, eq_to_hom_trans, ←category.assoc, eq_comp_eq_to_hom, + category.assoc, eq_to_hom_trans, eq_to_hom_refl, category.comp_id, 𝓕.germ_res inc.unop], + refl }, + { rw [←category.assoc, eq_comp_eq_to_hom, category.assoc, category.assoc, eq_to_hom_trans, + eq_to_hom_refl, category.comp_id], + exact terminal_is_terminal.hom_ext _ _ }, + end } + +lemma godement_is_sheaf (h : 𝓕.is_sheaf) : (godement 𝓕).is_sheaf := +_ + +end presheaf diff --git a/src/topology/sheaves/skyscraper.lean b/src/topology/sheaves/skyscraper.lean index 477775b5498b9..be349be4c2e63 100644 --- a/src/topology/sheaves/skyscraper.lean +++ b/src/topology/sheaves/skyscraper.lean @@ -826,7 +826,7 @@ variables {C : Type u} [category.{u} C] [concrete_category.{u} C] [has_colimits variables [has_limits C] [preserves_limits (forget C)] [reflects_isomorphisms (forget C)] variables [preserves_filtered_colimits (forget C)] -lemma skyscraper_presheaf_in_Ab_injective (S : C) [injective S] : +lemma skyscraper_sheaf_injective (S : C) [injective S] : injective (skyscraper_sheaf p₀ S) := injective.injective_of_adjoint (stalk_skyscraper_sheaf_adj p₀ C) S From 0dbcde21e2ddd81ad05ecdc419fbcdafbe31d914 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Wed, 17 Aug 2022 07:15:55 +0100 Subject: [PATCH 72/84] fix and add lemma about mono in sheaves --- src/topology/sheaves/godement.lean | 37 ++++++++++++++++++++++-------- src/topology/sheaves/stalks.lean | 26 +++++++++++++++++---- 2 files changed, 49 insertions(+), 14 deletions(-) diff --git a/src/topology/sheaves/godement.lean b/src/topology/sheaves/godement.lean index 64966d1064050..9a2824fc34031 100644 --- a/src/topology/sheaves/godement.lean +++ b/src/topology/sheaves/godement.lean @@ -15,16 +15,16 @@ open category_theory.limits universes u v -variables {X : Top.{u}} {C : Type v} [category.{u} C] +variables {X : Top.{u}} {C : Type u} [category.{u} C] variables [has_limits C] [has_terminal C] [has_colimits C] variables [Π (x : X) (U : opens X), decidable (x ∈ U)] -variable (𝓕 : presheaf C X) +variables (𝓕 : presheaf C X) (𝓖 : sheaf C X) -def godement : presheaf C X := +def godement_presheaf : presheaf C X := ∏ (λ x, skyscraper_presheaf x (𝓕.stalk x) : X → presheaf C X) -@[simps] def godement_app (U : (opens X)ᵒᵖ) : - (godement 𝓕).obj U ≅ ∏ (λ x, (skyscraper_presheaf x (𝓕.stalk x)).obj U) := +@[simps] def godement_presheaf_app (U : (opens X)ᵒᵖ) : + (godement_presheaf 𝓕).obj U ≅ ∏ (λ x, (skyscraper_presheaf x (𝓕.stalk x)).obj U) := limit_obj_iso_limit_comp_evaluation _ _ ≪≫ { hom := lim_map { app := λ _, 𝟙 _, @@ -48,15 +48,15 @@ limit_obj_iso_limit_comp_evaluation _ _ ≪≫ end } @[reducible] -def to_godement_aux (U : (opens X)ᵒᵖ) : +def to_godement_presheaf_aux (U : (opens X)ᵒᵖ) : 𝓕.obj U ⟶ ∏ λ (x : X), ite (x ∈ opposite.unop U) (𝓕.stalk x) (⊤_ C) := pi.lift (λ x, if m : x ∈ U.unop then 𝓕.germ ⟨x, m⟩ ≫ eq_to_hom ((skyscraper_presheaf_obj_of_mem _ m).symm.trans (by congr) : 𝓕.stalk (⟨x, m⟩ : U.unop) = (skyscraper_presheaf x (𝓕.stalk x)).obj U) else terminal.from _ ≫ eq_to_hom (skyscraper_presheaf_obj_of_not_mem _ m).symm) -@[simps] def to_godement : 𝓕 ⟶ godement 𝓕 := -{ app := λ U, to_godement_aux 𝓕 U ≫ (godement_app 𝓕 U).inv, +@[simps] def to_godement_presheaf : 𝓕 ⟶ godement_presheaf 𝓕 := +{ app := λ U, to_godement_presheaf_aux 𝓕 U ≫ (godement_presheaf_app 𝓕 U).inv, naturality' := begin intros U V inc, @@ -80,7 +80,24 @@ pi.lift (λ x, if m : x ∈ U.unop exact terminal_is_terminal.hom_ext _ _ }, end } -lemma godement_is_sheaf (h : 𝓕.is_sheaf) : (godement 𝓕).is_sheaf := -_ +lemma godement_presheaf_is_sheaf (h : 𝓕.is_sheaf) : (godement_presheaf 𝓕).is_sheaf := +limit_is_sheaf _ $ λ ⟨x⟩, (skyscraper_sheaf x _).2 + +def godement_sheaf : sheaf C X := +⟨godement_presheaf 𝓖.1, godement_presheaf_is_sheaf _ 𝓖.2⟩ + +def godement_sheaf_app (U : (opens X)ᵒᵖ) : + (godement_sheaf 𝓖).1.obj U ≅ ∏ (λ x, (skyscraper_presheaf x (𝓖.presheaf.stalk x)).obj U) := +godement_presheaf_app 𝓖.1 U + +def to_godement_sheaf : 𝓖 ⟶ godement_sheaf 𝓖 := +⟨to_godement_presheaf 𝓖.1⟩ + +variables [concrete_category.{u} C] [preserves_limits (forget C)] +variables [Π (U : opens X), preserves_colimits_of_shape + ((opens.grothendieck_topology X).cover U)ᵒᵖ (forget C)] +variables [reflects_isomorphisms (forget C)] + +instance : mono $ to_godement_sheaf 𝓖 := sorry end presheaf diff --git a/src/topology/sheaves/stalks.lean b/src/topology/sheaves/stalks.lean index 5e048dfdb43be..b51fe0bf60de5 100644 --- a/src/topology/sheaves/stalks.lean +++ b/src/topology/sheaves/stalks.lean @@ -432,18 +432,36 @@ instance stalk_functor_preserves_mono (x : X) : begin refine ⟨λ 𝓐 𝓑 f im, _⟩, apply concrete_category.mono_of_injective, - dsimp, haveI : preserves_limits_of_shape walking_cospan (forget C) := infer_instance, haveI im2 : mono f.1 := @@category_theory.presheaf_mono_of_mono _ _ _ _ _ _ _ _ _ _ _ im, have im3 := (nat_trans.mono_iff_app_mono _ f.1).mp im2, - have : ∀ (c : (opens X)ᵒᵖ), function.injective (f.1.app c) := - λ c, (@@concrete_category.mono_iff_injective_of_preserves_pullback _ _ (f.1.app c) (im3 c) _).mp - (im3 c), + have : ∀ (c : (opens X)ᵒᵖ), function.injective (f.1.app c) := λ c, + (@@concrete_category.mono_iff_injective_of_preserves_pullback _ _ (f.1.app c) _).mp (im3 c), refine (app_injective_iff_stalk_functor_map_injective f.1).mpr _ x, rintros U, apply this, end +instance stalk_mono_of_mono (x : X) {F G : sheaf C X} (f : F ⟶ G) [mono f] : + mono $ (stalk_functor C x).map f.1 := +begin + have := functor.map_mono (sheaf.forget C X ⋙ stalk_functor C x) f, + exact this, +end + +instance mono_of_stalk_mono {F G : sheaf C X} (f : F ⟶ G) + [Π (x : X), mono $ (stalk_functor C x).map f.1] : + mono f := +begin + rw [Sheaf.hom.mono_iff_presheaf_mono, nat_trans.mono_iff_app_mono], + intros U, + rw concrete_category.mono_iff_injective_of_preserves_pullback, + refine app_injective_of_stalk_functor_map_injective f.1 U.unop _, + rintros ⟨x, hx⟩, + rw ←concrete_category.mono_iff_injective_of_preserves_pullback, + apply_instance, +end + end /-- For surjectivity, we are given an arbitrary section `t` and need to find a preimage for it. From 729c649668694fe61c89b5c1648610fccda81292 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Wed, 17 Aug 2022 07:23:20 +0100 Subject: [PATCH 73/84] fixx --- src/topology/sheaves/godement.lean | 3 ++- src/topology/sheaves/stalks.lean | 4 ++++ 2 files changed, 6 insertions(+), 1 deletion(-) diff --git a/src/topology/sheaves/godement.lean b/src/topology/sheaves/godement.lean index 9a2824fc34031..4285fc9c53f08 100644 --- a/src/topology/sheaves/godement.lean +++ b/src/topology/sheaves/godement.lean @@ -98,6 +98,7 @@ variables [Π (U : opens X), preserves_colimits_of_shape ((opens.grothendieck_topology X).cover U)ᵒᵖ (forget C)] variables [reflects_isomorphisms (forget C)] -instance : mono $ to_godement_sheaf 𝓖 := sorry +instance : mono $ to_godement_sheaf 𝓖 := +sorry end presheaf diff --git a/src/topology/sheaves/stalks.lean b/src/topology/sheaves/stalks.lean index b51fe0bf60de5..fd147878b77a6 100644 --- a/src/topology/sheaves/stalks.lean +++ b/src/topology/sheaves/stalks.lean @@ -462,6 +462,10 @@ begin apply_instance, end +lemma mono_iff_stalk_mono {F G : sheaf C X} (f : F ⟶ G) : + mono f ↔ ∀ x, mono ((stalk_functor C x).map f.1) := +⟨λ m, by { resetI, apply_instance }, λ m, by { resetI, apply_instance }⟩ + end /-- For surjectivity, we are given an arbitrary section `t` and need to find a preimage for it. From ffe5571f61a0ef5ba50b5462eb11246c29c49663 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Wed, 17 Aug 2022 08:08:02 +0100 Subject: [PATCH 74/84] commit --- src/topology/sheaves/godement.lean | 20 ++++++++++++++++++-- 1 file changed, 18 insertions(+), 2 deletions(-) diff --git a/src/topology/sheaves/godement.lean b/src/topology/sheaves/godement.lean index 4285fc9c53f08..484c932b8a5bd 100644 --- a/src/topology/sheaves/godement.lean +++ b/src/topology/sheaves/godement.lean @@ -80,6 +80,17 @@ pi.lift (λ x, if m : x ∈ U.unop exact terminal_is_terminal.hom_ext _ _ }, end } +lemma godement_presheaf_stalk [decidable_eq X] (x : X) : + (godement_presheaf 𝓕).stalk x ≅ 𝓕.stalk x := +let ccc : colimit_cocone ((open_nhds.inclusion x).op ⋙ 𝓕) := +{ cocone := + { X := (godement_presheaf 𝓕).stalk x, + ι := + { app := λ U, _, + naturality' := _ } }, + is_colimit := _ } in +(colimit.iso_colimit_cocone ccc) + lemma godement_presheaf_is_sheaf (h : 𝓕.is_sheaf) : (godement_presheaf 𝓕).is_sheaf := limit_is_sheaf _ $ λ ⟨x⟩, (skyscraper_sheaf x _).2 @@ -98,7 +109,12 @@ variables [Π (U : opens X), preserves_colimits_of_shape ((opens.grothendieck_topology X).cover U)ᵒᵖ (forget C)] variables [reflects_isomorphisms (forget C)] -instance : mono $ to_godement_sheaf 𝓖 := -sorry +instance [preserves_filtered_colimits (forget C)] : mono $ to_godement_sheaf 𝓖 := +begin + rw presheaf.mono_iff_stalk_mono, + intros x, + change mono ((presheaf.stalk_functor C x).map (to_godement_presheaf 𝓖.1)), + dsimp [presheaf.stalk_functor], +end end presheaf From 6cfe30f7f321b479b735444667b292c1fa2f1098 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Fri, 19 Aug 2022 04:01:22 +0100 Subject: [PATCH 75/84] defined some auxilary isomorphism --- src/topology/sheaves/godement.lean | 107 +++++++++++++++++++++++++++-- 1 file changed, 100 insertions(+), 7 deletions(-) diff --git a/src/topology/sheaves/godement.lean b/src/topology/sheaves/godement.lean index 484c932b8a5bd..0ee2a026ae6de 100644 --- a/src/topology/sheaves/godement.lean +++ b/src/topology/sheaves/godement.lean @@ -23,7 +23,7 @@ variables (𝓕 : presheaf C X) (𝓖 : sheaf C X) def godement_presheaf : presheaf C X := ∏ (λ x, skyscraper_presheaf x (𝓕.stalk x) : X → presheaf C X) -@[simps] def godement_presheaf_app (U : (opens X)ᵒᵖ) : +@[simps] def godement_presheaf_obj (U : (opens X)ᵒᵖ) : (godement_presheaf 𝓕).obj U ≅ ∏ (λ x, (skyscraper_presheaf x (𝓕.stalk x)).obj U) := limit_obj_iso_limit_comp_evaluation _ _ ≪≫ { hom := lim_map @@ -56,7 +56,7 @@ pi.lift (λ x, if m : x ∈ U.unop else terminal.from _ ≫ eq_to_hom (skyscraper_presheaf_obj_of_not_mem _ m).symm) @[simps] def to_godement_presheaf : 𝓕 ⟶ godement_presheaf 𝓕 := -{ app := λ U, to_godement_presheaf_aux 𝓕 U ≫ (godement_presheaf_app 𝓕 U).inv, +{ app := λ U, to_godement_presheaf_aux 𝓕 U ≫ (godement_presheaf_obj 𝓕 U).inv, naturality' := begin intros U V inc, @@ -97,9 +97,9 @@ limit_is_sheaf _ $ λ ⟨x⟩, (skyscraper_sheaf x _).2 def godement_sheaf : sheaf C X := ⟨godement_presheaf 𝓖.1, godement_presheaf_is_sheaf _ 𝓖.2⟩ -def godement_sheaf_app (U : (opens X)ᵒᵖ) : +def godement_sheaf_obj (U : (opens X)ᵒᵖ) : (godement_sheaf 𝓖).1.obj U ≅ ∏ (λ x, (skyscraper_presheaf x (𝓖.presheaf.stalk x)).obj U) := -godement_presheaf_app 𝓖.1 U +godement_presheaf_obj 𝓖.1 U def to_godement_sheaf : 𝓖 ⟶ godement_sheaf 𝓖 := ⟨to_godement_presheaf 𝓖.1⟩ @@ -107,14 +107,107 @@ def to_godement_sheaf : 𝓖 ⟶ godement_sheaf 𝓖 := variables [concrete_category.{u} C] [preserves_limits (forget C)] variables [Π (U : opens X), preserves_colimits_of_shape ((opens.grothendieck_topology X).cover U)ᵒᵖ (forget C)] -variables [reflects_isomorphisms (forget C)] +variables [reflects_isomorphisms (forget C)] [preserves_filtered_colimits (forget C)] -instance [preserves_filtered_colimits (forget C)] : mono $ to_godement_sheaf 𝓖 := +def sheaf_in_Type : sheaf C X ⥤ sheaf (Type u) X := +{ obj := λ F, ⟨F.1 ⋙ forget C, (presheaf.is_sheaf_iff_is_sheaf_comp (forget C) F.1).mp F.2⟩, + map := λ F G f, Sheaf.hom.mk + { app := λ U, (forget C).map (f.1.app U), + naturality' := λ U V inc, by erw [←(forget C).map_comp, ←(forget C).map_comp, f.1.naturality] }, + map_id' := λ F, by { ext, dsimp, rw [id_apply] }, + map_comp' := λ F G H f g, by { ext, dsimp, rw [comp_apply] } } + +example : true := trivial + + +def godement_sheaf_in_Type : sheaf (Type u) X := sheaf_in_Type.obj (godement_sheaf 𝓖) + +def godement_sheaf_in_Type_obj_aux1 (U : (opens X)ᵒᵖ) : + (godement_sheaf_in_Type 𝓖).1.obj U ≅ + (forget C).obj ∏ (λ x, (skyscraper_presheaf x (𝓖.presheaf.stalk x)).obj U) := +{ hom := (forget C).map $ (godement_sheaf_obj 𝓖 U).hom, + inv := (forget C).map $ (godement_sheaf_obj 𝓖 U).inv, + hom_inv_id' := by { rw [←(forget C).map_comp, iso.hom_inv_id, (forget C).map_id], refl, }, + inv_hom_id' := by { rw [←(forget C).map_comp, iso.inv_hom_id, (forget C).map_id], } } + +example : true := trivial + +def godement_sheaf_in_Type_obj_aux2 (U : (opens X)ᵒᵖ) : + (forget C).obj ∏ (λ x, (skyscraper_presheaf x (𝓖.presheaf.stalk x)).obj U) ≅ + ∏ (λ x, (forget C).obj ((skyscraper_presheaf x (𝓖.presheaf.stalk x)).obj U)) := +{ hom := pi.lift $ λ p, (forget C).map $ pi.π _ p, + inv := + begin + refine lim_map _ ≫ (preserves_limit_iso (forget C) _).inv, + refine { app := λ p x, x, naturality' := _ }, + rintros ⟨p⟩ ⟨q⟩ ⟨⟨(eq1 : p = q)⟩⟩, + dsimp, + induction eq1, + ext1, + dsimp, + simp only [discrete.functor_map_id, types_id_apply, id_apply], + end, + hom_inv_id' := + begin + rw [←category.assoc, limit.lift_map, iso.comp_inv_eq, category.id_comp], + refine limit.hom_ext _, + rintros ⟨p⟩, + rw [preserves_limits_iso_hom_π, limit.lift_π], + simpa only [cones.postcompose_obj_π, nat_trans.comp_app, fan.mk_π_app, forget_map_eq_coe], + end, + inv_hom_id' := _ } + +def godement_sheaf_in_Type_obj (U : (opens X)ᵒᵖ) : + (godement_sheaf_in_Type 𝓖).1.obj U ≅ + ∏ (λ x, (forget C).obj $ (skyscraper_presheaf x (𝓖.presheaf.stalk x)).obj U) := +godement_sheaf_in_Type_obj_aux1 𝓖 U ≪≫ +{ hom := _, + inv := _, + hom_inv_id' := _, + inv_hom_id' := _ } + +lemma to_godement_sheaf_app_section (U : (opens X)ᵒᵖ) (s : (sheaf_in_Type.obj 𝓖).1.obj U) (x : U.unop) : + -- (forget C).map (godement_sheaf_obj 𝓖 U).hom ((sheaf_in_Type.map $ to_godement_sheaf 𝓖).1.app U s) = + -- (forget C).obj (𝓖.1.stalk ) := + sorry := +begin + have := (forget C).map (godement_sheaf_obj 𝓖 U).hom, +end + + +lemma to_godement_sheaf_app_injective (U : opens X) : + function.injective $ (forget C).map ((to_godement_sheaf 𝓖).1.app (opposite.op U)) := +λ x y eq1, +begin + change (sheaf_in_Type.obj 𝓖).1.obj (opposite.op U) at x, + change (sheaf_in_Type.obj 𝓖).1.obj (opposite.op U) at y, + change (sheaf_in_Type.map (to_godement_sheaf 𝓖)).1.app (opposite.op U) x = + (sheaf_in_Type.map (to_godement_sheaf 𝓖)).1.app (opposite.op U) y at eq1, + apply presheaf.section_ext, + /- + U : opens ↥X, + x y : (sheaf_in_Type.obj 𝓖).val.obj (opposite.op U), + eq1 : + (sheaf_in_Type.map (to_godement_sheaf 𝓖)).val.app (opposite.op U) x = + (sheaf_in_Type.map (to_godement_sheaf 𝓖)).val.app (opposite.op U) y, + p : ↥U + ⊢ ⇑(𝓖.presheaf.germ p) x = ⇑(𝓖.presheaf.germ p) y + -/ + + intros p, + have := presheaf.germ_ext 𝓖.1 U p.2; + -- ext, + sorry +end + +instance : mono $ to_godement_sheaf 𝓖 := begin rw presheaf.mono_iff_stalk_mono, intros x, change mono ((presheaf.stalk_functor C x).map (to_godement_presheaf 𝓖.1)), - dsimp [presheaf.stalk_functor], + rw concrete_category.mono_iff_injective_of_preserves_pullback, + exact (presheaf.app_injective_iff_stalk_functor_map_injective (to_godement_presheaf 𝓖.1)).mpr + (to_godement_sheaf_app_injective 𝓖) x, end end presheaf From 036fc2f8bc4692972f1ce809ca0067fc47868899 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Fri, 19 Aug 2022 04:18:11 +0100 Subject: [PATCH 76/84] save for now --- src/topology/sheaves/godement.lean | 32 ++++++++++++------------------ 1 file changed, 13 insertions(+), 19 deletions(-) diff --git a/src/topology/sheaves/godement.lean b/src/topology/sheaves/godement.lean index 0ee2a026ae6de..81a6cd53e5a93 100644 --- a/src/topology/sheaves/godement.lean +++ b/src/topology/sheaves/godement.lean @@ -122,17 +122,7 @@ example : true := trivial def godement_sheaf_in_Type : sheaf (Type u) X := sheaf_in_Type.obj (godement_sheaf 𝓖) -def godement_sheaf_in_Type_obj_aux1 (U : (opens X)ᵒᵖ) : - (godement_sheaf_in_Type 𝓖).1.obj U ≅ - (forget C).obj ∏ (λ x, (skyscraper_presheaf x (𝓖.presheaf.stalk x)).obj U) := -{ hom := (forget C).map $ (godement_sheaf_obj 𝓖 U).hom, - inv := (forget C).map $ (godement_sheaf_obj 𝓖 U).inv, - hom_inv_id' := by { rw [←(forget C).map_comp, iso.hom_inv_id, (forget C).map_id], refl, }, - inv_hom_id' := by { rw [←(forget C).map_comp, iso.inv_hom_id, (forget C).map_id], } } - -example : true := trivial - -def godement_sheaf_in_Type_obj_aux2 (U : (opens X)ᵒᵖ) : +def godement_sheaf_in_Type_obj_aux (U : (opens X)ᵒᵖ) : (forget C).obj ∏ (λ x, (skyscraper_presheaf x (𝓖.presheaf.stalk x)).obj U) ≅ ∏ (λ x, (forget C).obj ((skyscraper_presheaf x (𝓖.presheaf.stalk x)).obj U)) := { hom := pi.lift $ λ p, (forget C).map $ pi.π _ p, @@ -155,23 +145,27 @@ def godement_sheaf_in_Type_obj_aux2 (U : (opens X)ᵒᵖ) : rw [preserves_limits_iso_hom_π, limit.lift_π], simpa only [cones.postcompose_obj_π, nat_trans.comp_app, fan.mk_π_app, forget_map_eq_coe], end, - inv_hom_id' := _ } + inv_hom_id' := + begin + ext1 ⟨p⟩, + rw [category.assoc, limit.lift_π, fan.mk_π_app, category.assoc, preserves_limits_iso_inv_π, + lim_map_π, category.id_comp], + refl, + end } def godement_sheaf_in_Type_obj (U : (opens X)ᵒᵖ) : (godement_sheaf_in_Type 𝓖).1.obj U ≅ ∏ (λ x, (forget C).obj $ (skyscraper_presheaf x (𝓖.presheaf.stalk x)).obj U) := -godement_sheaf_in_Type_obj_aux1 𝓖 U ≪≫ -{ hom := _, - inv := _, - hom_inv_id' := _, - inv_hom_id' := _ } +((forget C).map_iso $ godement_sheaf_obj 𝓖 U) ≪≫ godement_sheaf_in_Type_obj_aux 𝓖 U -lemma to_godement_sheaf_app_section (U : (opens X)ᵒᵖ) (s : (sheaf_in_Type.obj 𝓖).1.obj U) (x : U.unop) : +lemma to_godement_sheaf_app_section (U : (opens X)ᵒᵖ) (s : (sheaf_in_Type.obj 𝓖).1.obj U) + (x : U.unop) : -- (forget C).map (godement_sheaf_obj 𝓖 U).hom ((sheaf_in_Type.map $ to_godement_sheaf 𝓖).1.app U s) = -- (forget C).obj (𝓖.1.stalk ) := sorry := begin - have := (forget C).map (godement_sheaf_obj 𝓖 U).hom, + have := (sheaf_in_Type.map (to_godement_sheaf 𝓖)).1.app U s, + have := (godement_sheaf_in_Type_obj 𝓖 U).hom, end From 55d4046aaadeb5fc3d44fa59ba0ece839aca2e15 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Fri, 19 Aug 2022 06:33:32 +0100 Subject: [PATCH 77/84] commit --- src/topology/sheaves/godement.lean | 65 +++++++++++++++++++++++++--- src/topology/sheaves/skyscraper.lean | 57 ++++++++++++++++++++++++ 2 files changed, 115 insertions(+), 7 deletions(-) diff --git a/src/topology/sheaves/godement.lean b/src/topology/sheaves/godement.lean index 81a6cd53e5a93..48de5451bdff2 100644 --- a/src/topology/sheaves/godement.lean +++ b/src/topology/sheaves/godement.lean @@ -158,14 +158,65 @@ def godement_sheaf_in_Type_obj (U : (opens X)ᵒᵖ) : ∏ (λ x, (forget C).obj $ (skyscraper_presheaf x (𝓖.presheaf.stalk x)).obj U) := ((forget C).map_iso $ godement_sheaf_obj 𝓖 U) ≪≫ godement_sheaf_in_Type_obj_aux 𝓖 U -lemma to_godement_sheaf_app_section (U : (opens X)ᵒᵖ) (s : (sheaf_in_Type.obj 𝓖).1.obj U) - (x : U.unop) : - -- (forget C).map (godement_sheaf_obj 𝓖 U).hom ((sheaf_in_Type.map $ to_godement_sheaf 𝓖).1.app U s) = - -- (forget C).obj (𝓖.1.stalk ) := - sorry := +def sheaf_in_Type_skyscraper_sheaf (x : X) (c : C) : + (sheaf_in_Type.obj $ skyscraper_sheaf x c) ≅ + skyscraper_sheaf x ((forget C).obj c) := +{ hom := Sheaf.hom.mk + { app := λ U, eq_to_hom + begin + change (forget C).obj _ = (skyscraper_sheaf x ((forget C).obj c)).1.obj _, + by_cases hU : x ∈ U.unop, + { erw [skyscraper_presheaf_obj_of_mem _ hU, skyscraper_presheaf_obj_of_mem _ hU], }, + { erw [skyscraper_presheaf_obj_of_not_mem _ hU, skyscraper_presheaf_obj_of_not_mem _ hU], + -- need to make "skyscraper_presheaf_with_specified_terminal_object" + }, + end, + naturality' := λ U V inc, + begin + dsimp, + rw [comp_dite, dite_comp], + by_cases hV : x ∈ V.unop, + { have hU : x ∈ U.unop := le_of_hom inc.unop hV, + split_ifs, + -- generalize_proofs _ h1 h2 h3 h4, + change (sheaf_in_Type.obj (skyscraper_sheaf x c)).val.map inc ≫ + (forget C).map (eq_to_hom _) ≫ eq_to_hom _ = + ((forget C).map (eq_to_hom _) ≫ eq_to_hom _) ≫ + (skyscraper_sheaf x ((forget C).obj c)).val.map inc, + rw [eq_to_hom_map, eq_to_hom_map, eq_to_hom_trans, eq_to_hom_trans], + generalize_proofs _ h1 h2, + + sorry, + -- have := (sheaf_in_Type.obj (skyscraper_sheaf x c)).1.map_comp, + -- unfold_coes, + -- erw show (forget C).map (eq_to_hom h1) = _, from rfl, + -- erw [eq_to_hom_map (forget C) h1], + }, + { split_ifs, + rw [←category.assoc, eq_comp_eq_to_hom], + exact terminal_is_terminal.hom_ext _ _, }, + end }, + inv := Sheaf.hom.mk + { app := λ U, if m : x ∈ U.unop + then eq_to_hom _ + else _, + naturality' := _ }, + hom_inv_id' := _, + inv_hom_id' := _ } + +lemma stalk_bundles (U : (opens X)ᵒᵖ) (s : (sheaf_in_Type.obj 𝓖).1.obj U) + (x : U.unop) : presheaf.stalk (sheaf_in_Type.obj 𝓖).1 x := +-- let S : (sheaf_in_Type.obj (skyscraper_sheaf x.val (𝓖.presheaf.stalk x.val))).1.obj U := +-- in +-- presheaf.germ begin - have := (sheaf_in_Type.map (to_godement_sheaf 𝓖)).1.app U s, - have := (godement_sheaf_in_Type_obj 𝓖 U).hom, + have : (sheaf_in_Type.obj (skyscraper_sheaf x.val (𝓖.presheaf.stalk x.val))).1.obj U, + { }, + -- (types.product_iso _).hom (((godement_sheaf_in_Type_obj 𝓖 U).hom ((sheaf_in_Type.map (to_godement_sheaf 𝓖)).1.app U s))) x.1, + dsimp at S, + have := (sheaf_in_Type.obj 𝓖).presheaf.germ x, + -- have := , + have := (types.product_iso _).hom (((godement_sheaf_in_Type_obj 𝓖 U).hom ((sheaf_in_Type.map (to_godement_sheaf 𝓖)).1.app U s))) x.1, end diff --git a/src/topology/sheaves/skyscraper.lean b/src/topology/sheaves/skyscraper.lean index be349be4c2e63..017749e6d9600 100644 --- a/src/topology/sheaves/skyscraper.lean +++ b/src/topology/sheaves/skyscraper.lean @@ -72,6 +72,63 @@ point, then the skyscraper presheaf `𝓕` with value `A` is defined by `U ↦ A simpa only [eq_comp_eq_to_hom] using terminal_is_terminal.hom_ext _ _, } end } +@[simps] def skyscraper_presheaf' {c : C} (it : is_terminal c) : presheaf C X := +{ obj := λ U, ite (p₀ ∈ unop U) S c, + map := λ U V i, dite (p₀ ∈ unop V) + (λ h, eq_to_hom (if_pos (le_of_hom i.unop h)) ≫ 𝟙 S ≫ eq_to_hom (if_pos h).symm) + (λ h, it.from _ ≫ eq_to_hom (if_neg h).symm), + map_id' := λ U, + begin + split_ifs, + { simp, }, + { simpa only [eq_comp_eq_to_hom] using it.hom_ext _ _, }, + end, + map_comp' := λ U V W iVU iWV, + begin + by_cases hW : p₀ ∈ unop W, + { have hV : p₀ ∈ unop V := le_of_hom iWV.unop hW, + have hU : p₀ ∈ unop U := le_of_hom iVU.unop hV, + split_ifs, + simp }, + { split_ifs; + simpa only [eq_comp_eq_to_hom] using it.hom_ext _ _, } + end } + +def skyscraper_presheaf_iso {c : C} (it : is_terminal c) : + (skyscraper_presheaf p₀ S) ≅ (skyscraper_presheaf' p₀ S it) := +nat_iso.of_components + (λ U, iso.mk + (if m : p₀ ∈ U.unop + then eq_to_hom (by simp [if_pos m]) + else it.from _ ≫ eq_to_hom (by simp [if_neg m])) + (if m : p₀ ∈ U.unop + then eq_to_hom (by simp [if_pos m]) + else terminal.from _ ≫ eq_to_hom (by simp [if_neg m])) + begin + dsimp, + split_ifs, + { rw [eq_to_hom_trans, eq_to_hom_refl], }, + { rw [←category.assoc, eq_comp_eq_to_hom], + exact terminal_is_terminal.hom_ext _ _, } + end + begin + dsimp, + split_ifs, + { rw [eq_to_hom_trans, eq_to_hom_refl], }, + { rw [←category.assoc, eq_comp_eq_to_hom], + exact it.hom_ext _ _, } + end) $ λ U V inc, + begin + dsimp, + by_cases hV : p₀ ∈ V.unop, + { have hU : p₀ ∈ U.unop := le_of_hom inc.unop hV, + split_ifs, + simp only [category.id_comp, eq_to_hom_trans, eq_to_hom_refl], }, + { split_ifs; + rw [←category.assoc, eq_comp_eq_to_hom]; + exact it.hom_ext _ _, }, + end + section variables {p₀} From 84c015d5480d4cc8c7b9845d94a7a7d34064d272 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Fri, 19 Aug 2022 07:17:34 +0100 Subject: [PATCH 78/84] save --- src/topology/sheaves/skyscraper.lean | 11 +++++++++++ 1 file changed, 11 insertions(+) diff --git a/src/topology/sheaves/skyscraper.lean b/src/topology/sheaves/skyscraper.lean index 017749e6d9600..9bfe0407f2064 100644 --- a/src/topology/sheaves/skyscraper.lean +++ b/src/topology/sheaves/skyscraper.lean @@ -192,6 +192,17 @@ def skyscraper_sheaf : sheaf C X := exact terminal_is_terminal.hom_ext _ _, } end⟩⟩ +def skyscraper_sheaf' {c : C} (it : is_terminal c) : sheaf C X := +⟨skyscraper_presheaf' p₀ S it, + (presheaf.is_sheaf_iso_iff (skyscraper_presheaf_iso p₀ S it)).mp (skyscraper_sheaf p₀ S).2⟩ + +def skyscraper_sheaf_iso {c : C} (it : is_terminal c) : + skyscraper_sheaf p₀ S ≅ skyscraper_sheaf' p₀ S it := +{ hom := Sheaf.hom.mk (skyscraper_presheaf_iso p₀ S it).hom, + inv := Sheaf.hom.mk (skyscraper_presheaf_iso p₀ S it).inv, + hom_inv_id' := by { ext, dsimp, simp only [iso.hom_inv_id_app] }, + inv_hom_id' := by { ext, dsimp, simp only [iso.inv_hom_id_app] } } + end section From e941031e6efa01b4984919b592102e21e80c372e Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Fri, 19 Aug 2022 07:37:13 +0100 Subject: [PATCH 79/84] more api --- src/topology/sheaves/godement.lean | 104 ++++++++++++++++++----------- 1 file changed, 64 insertions(+), 40 deletions(-) diff --git a/src/topology/sheaves/godement.lean b/src/topology/sheaves/godement.lean index 48de5451bdff2..a72bc02279aea 100644 --- a/src/topology/sheaves/godement.lean +++ b/src/topology/sheaves/godement.lean @@ -80,16 +80,16 @@ pi.lift (λ x, if m : x ∈ U.unop exact terminal_is_terminal.hom_ext _ _ }, end } -lemma godement_presheaf_stalk [decidable_eq X] (x : X) : - (godement_presheaf 𝓕).stalk x ≅ 𝓕.stalk x := -let ccc : colimit_cocone ((open_nhds.inclusion x).op ⋙ 𝓕) := -{ cocone := - { X := (godement_presheaf 𝓕).stalk x, - ι := - { app := λ U, _, - naturality' := _ } }, - is_colimit := _ } in -(colimit.iso_colimit_cocone ccc) +-- lemma godement_presheaf_stalk [decidable_eq X] (x : X) : +-- (godement_presheaf 𝓕).stalk x ≅ 𝓕.stalk x := +-- let ccc : colimit_cocone ((open_nhds.inclusion x).op ⋙ 𝓕) := +-- { cocone := +-- { X := (godement_presheaf 𝓕).stalk x, +-- ι := +-- { app := λ U, _, +-- naturality' := _ } }, +-- is_colimit := _ } in +-- (colimit.iso_colimit_cocone ccc) lemma godement_presheaf_is_sheaf (h : 𝓕.is_sheaf) : (godement_presheaf 𝓕).is_sheaf := limit_is_sheaf _ $ λ ⟨x⟩, (skyscraper_sheaf x _).2 @@ -161,48 +161,72 @@ def godement_sheaf_in_Type_obj (U : (opens X)ᵒᵖ) : def sheaf_in_Type_skyscraper_sheaf (x : X) (c : C) : (sheaf_in_Type.obj $ skyscraper_sheaf x c) ≅ skyscraper_sheaf x ((forget C).obj c) := +let e1 := preserves_limit_iso (forget C) (functor.empty.{0} C) in +have it : is_terminal $ (forget C).obj (terminal C), +from is_terminal.of_iso +begin + rw show functor.empty C ⋙ forget C = functor.empty (Type u), from functor.empty_ext' _ _, + exact terminal_is_terminal, +end e1.symm, +let e : (sheaf_in_Type.obj $ skyscraper_sheaf x c) ≅ skyscraper_sheaf' x ((forget C).obj c) it := { hom := Sheaf.hom.mk { app := λ U, eq_to_hom begin - change (forget C).obj _ = (skyscraper_sheaf x ((forget C).obj c)).1.obj _, + change (forget C).obj _ = (skyscraper_sheaf' x ((forget C).obj c) _).1.obj _, by_cases hU : x ∈ U.unop, - { erw [skyscraper_presheaf_obj_of_mem _ hU, skyscraper_presheaf_obj_of_mem _ hU], }, - { erw [skyscraper_presheaf_obj_of_not_mem _ hU, skyscraper_presheaf_obj_of_not_mem _ hU], - -- need to make "skyscraper_presheaf_with_specified_terminal_object" - }, + { erw [skyscraper_presheaf_obj_of_mem _ hU, skyscraper_sheaf'], + dsimp, split_ifs, refl, }, + { erw [skyscraper_presheaf_obj_of_not_mem _ hU, skyscraper_sheaf'], + dsimp, split_ifs, refl, }, end, naturality' := λ U V inc, begin - dsimp, - rw [comp_dite, dite_comp], + dsimp [skyscraper_sheaf', sheaf_in_Type, skyscraper_sheaf], + change (forget C).map _ ≫ _ = _, by_cases hV : x ∈ V.unop, { have hU : x ∈ U.unop := le_of_hom inc.unop hV, split_ifs, - -- generalize_proofs _ h1 h2 h3 h4, - change (sheaf_in_Type.obj (skyscraper_sheaf x c)).val.map inc ≫ - (forget C).map (eq_to_hom _) ≫ eq_to_hom _ = - ((forget C).map (eq_to_hom _) ≫ eq_to_hom _) ≫ - (skyscraper_sheaf x ((forget C).obj c)).val.map inc, - rw [eq_to_hom_map, eq_to_hom_map, eq_to_hom_trans, eq_to_hom_trans], - generalize_proofs _ h1 h2, - - sorry, - -- have := (sheaf_in_Type.obj (skyscraper_sheaf x c)).1.map_comp, - -- unfold_coes, - -- erw show (forget C).map (eq_to_hom h1) = _, from rfl, - -- erw [eq_to_hom_map (forget C) h1], - }, - { split_ifs, - rw [←category.assoc, eq_comp_eq_to_hom], - exact terminal_is_terminal.hom_ext _ _, }, + simp only [category.id_comp, eq_to_hom_trans, eq_to_hom_refl, eq_to_hom_map], + refl, }, + { split_ifs; + symmetry; + rw [←category.assoc, eq_comp_eq_to_hom]; + exact it.hom_ext _ _, }, end }, inv := Sheaf.hom.mk - { app := λ U, if m : x ∈ U.unop - then eq_to_hom _ - else _, - naturality' := _ }, - hom_inv_id' := _, - inv_hom_id' := _ } + { app := λ U, eq_to_hom + begin + change (skyscraper_sheaf' x ((forget C).obj c) _).1.obj _ = (forget C).obj _, + by_cases hU : x ∈ U.unop, + { erw [skyscraper_presheaf_obj_of_mem _ hU, skyscraper_sheaf'], + dsimp, split_ifs, refl, }, + { erw [skyscraper_presheaf_obj_of_not_mem _ hU, skyscraper_sheaf'], + dsimp, split_ifs, refl, }, + end, + naturality' := λ U V inc, + begin + dsimp [skyscraper_sheaf', sheaf_in_Type, skyscraper_sheaf], + change _ = _ ≫ (forget C).map _, + by_cases hV : x ∈ V.unop, + { have hU : x ∈ U.unop := le_of_hom inc.unop hV, + split_ifs, + simp only [category.id_comp, eq_to_hom_trans, eq_to_hom_refl, eq_to_hom_map], + refl, }, + { split_ifs; + rw [eq_comp_eq_to_hom, eq_comp_eq_to_hom]; + exact it.hom_ext _ _, }, + end }, + hom_inv_id' := + begin + ext : 3, dsimp, + rw [eq_to_hom_trans, eq_to_hom_refl], + end, + inv_hom_id' := + begin + ext : 3, dsimp, + rw [eq_to_hom_trans, eq_to_hom_refl], + end } in +e.trans $ (skyscraper_sheaf_iso _ _ it).symm lemma stalk_bundles (U : (opens X)ᵒᵖ) (s : (sheaf_in_Type.obj 𝓖).1.obj U) (x : U.unop) : presheaf.stalk (sheaf_in_Type.obj 𝓖).1 x := From 98575524c56378e0273ef3052b03522df29f773e Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Sun, 21 Aug 2022 00:42:07 +0100 Subject: [PATCH 80/84] finish godement resolution --- src/topology/sheaves/godement.lean | 121 +++++++++++++++++------------ 1 file changed, 73 insertions(+), 48 deletions(-) diff --git a/src/topology/sheaves/godement.lean b/src/topology/sheaves/godement.lean index a72bc02279aea..5b226ecefb794 100644 --- a/src/topology/sheaves/godement.lean +++ b/src/topology/sheaves/godement.lean @@ -55,6 +55,23 @@ pi.lift (λ x, if m : x ∈ U.unop 𝓕.stalk (⟨x, m⟩ : U.unop) = (skyscraper_presheaf x (𝓕.stalk x)).obj U) else terminal.from _ ≫ eq_to_hom (skyscraper_presheaf_obj_of_not_mem _ m).symm) +@[reducible] +def to_godement_presheaf_aux_comp_π {U : (opens X)ᵒᵖ} (p : U.unop) : + 𝓕.obj U ⟶ 𝓕.stalk p := +to_godement_presheaf_aux 𝓕 U ≫ pi.π _ p ≫ eq_to_hom (if_pos p.2) + +lemma to_godement_presheaf_aux_comp_π_eq {U : (opens X)ᵒᵖ} (p : U.unop) : + to_godement_presheaf_aux_comp_π 𝓕 p = presheaf.germ 𝓕 p := +begin + dunfold to_godement_presheaf_aux_comp_π presheaf.germ to_godement_presheaf_aux, + rw [←category.assoc, limit.lift_π], + simp only [fan.mk_π_app], + split_ifs, + { rw [category.assoc, eq_to_hom_trans, eq_to_hom_refl, category.comp_id], + refl }, + { exfalso, exact h p.2, }, +end + @[simps] def to_godement_presheaf : 𝓕 ⟶ godement_presheaf 𝓕 := { app := λ U, to_godement_presheaf_aux 𝓕 U ≫ (godement_presheaf_obj 𝓕 U).inv, naturality' := @@ -80,17 +97,6 @@ pi.lift (λ x, if m : x ∈ U.unop exact terminal_is_terminal.hom_ext _ _ }, end } --- lemma godement_presheaf_stalk [decidable_eq X] (x : X) : --- (godement_presheaf 𝓕).stalk x ≅ 𝓕.stalk x := --- let ccc : colimit_cocone ((open_nhds.inclusion x).op ⋙ 𝓕) := --- { cocone := --- { X := (godement_presheaf 𝓕).stalk x, --- ι := --- { app := λ U, _, --- naturality' := _ } }, --- is_colimit := _ } in --- (colimit.iso_colimit_cocone ccc) - lemma godement_presheaf_is_sheaf (h : 𝓕.is_sheaf) : (godement_presheaf 𝓕).is_sheaf := limit_is_sheaf _ $ λ ⟨x⟩, (skyscraper_sheaf x _).2 @@ -109,7 +115,7 @@ variables [Π (U : opens X), preserves_colimits_of_shape ((opens.grothendieck_topology X).cover U)ᵒᵖ (forget C)] variables [reflects_isomorphisms (forget C)] [preserves_filtered_colimits (forget C)] -def sheaf_in_Type : sheaf C X ⥤ sheaf (Type u) X := +@[simps] def sheaf_in_Type : sheaf C X ⥤ sheaf (Type u) X := { obj := λ F, ⟨F.1 ⋙ forget C, (presheaf.is_sheaf_iff_is_sheaf_comp (forget C) F.1).mp F.2⟩, map := λ F G f, Sheaf.hom.mk { app := λ U, (forget C).map (f.1.app U), @@ -228,55 +234,74 @@ let e : (sheaf_in_Type.obj $ skyscraper_sheaf x c) ≅ skyscraper_sheaf' x ((for end } in e.trans $ (skyscraper_sheaf_iso _ _ it).symm -lemma stalk_bundles (U : (opens X)ᵒᵖ) (s : (sheaf_in_Type.obj 𝓖).1.obj U) - (x : U.unop) : presheaf.stalk (sheaf_in_Type.obj 𝓖).1 x := --- let S : (sheaf_in_Type.obj (skyscraper_sheaf x.val (𝓖.presheaf.stalk x.val))).1.obj U := --- in --- presheaf.germ +example : true := trivial + +/-- +`x y : 𝓖(U)`, and `p ∈ U` +`` +-/ +def stalk_bundles_eq0 (U : (opens X)ᵒᵖ) (x y : (sheaf_in_Type.obj 𝓖).1.obj U) + (eq1 : (sheaf_in_Type.map (to_godement_sheaf 𝓖)).val.app U x = + (sheaf_in_Type.map (to_godement_sheaf 𝓖)).val.app U y) (p : U.unop) : + (forget C).map (to_godement_presheaf_aux 𝓖.presheaf U) x = + (forget C).map (to_godement_presheaf_aux 𝓖.presheaf U) y := begin - have : (sheaf_in_Type.obj (skyscraper_sheaf x.val (𝓖.presheaf.stalk x.val))).1.obj U, - { }, - -- (types.product_iso _).hom (((godement_sheaf_in_Type_obj 𝓖 U).hom ((sheaf_in_Type.map (to_godement_sheaf 𝓖)).1.app U s))) x.1, - dsimp at S, - have := (sheaf_in_Type.obj 𝓖).presheaf.germ x, - -- have := , - have := (types.product_iso _).hom (((godement_sheaf_in_Type_obj 𝓖 U).hom ((sheaf_in_Type.map (to_godement_sheaf 𝓖)).1.app U s))) x.1, + change (forget C).map ((to_godement_presheaf 𝓖.presheaf).app _) x = + (forget C).map ((to_godement_presheaf 𝓖.presheaf).app _) y at eq1, + dsimp at eq1, + change (forget C).map _ x = (forget C).map _ y at eq1, + have eq2 := congr_arg ((forget C).map (limit_obj_iso_limit_comp_evaluation (discrete.functor _) U).hom) eq1, + dsimp at eq2, + erw [←comp_apply, ←comp_apply, ←category.assoc] at eq2, + simp only [category.assoc, iso.inv_hom_id, category.comp_id] at eq2, + set α : nat_trans (discrete.functor (λ (x : ↥X), ite (x ∈ opposite.unop U) (𝓖.presheaf.stalk x) (⊤_ C))) + (discrete.functor (λ (x : ↥X), skyscraper_presheaf x (𝓖.presheaf.stalk x)) ⋙ + (evaluation (opens ↥X)ᵒᵖ C).obj U) := _, + change (forget C).map (_ ≫ lim_map α) x = (forget C).map (_ ≫ lim_map α) y at eq2, + haveI : is_iso (lim_map α), + { refine is_iso.mk ⟨lim_map { app := λ x, 𝟙 _, naturality' := _ }, _, _⟩, + { rintros ⟨x⟩ ⟨y⟩ ⟨⟨eq0 : x = y⟩⟩, subst eq0, refl}, + { ext1, simp only [category.assoc, lim_map_π, category.comp_id, category.id_comp], }, + { ext1, simp only [category.assoc, lim_map_π, category.comp_id, category.id_comp], }, }, + have eq3 := congr_arg ((forget C).map (inv (lim_map α))) eq2, + change ((forget C).map _ ≫ (forget C).map _) _ = ((forget C).map _ ≫ (forget C).map _) _ at eq3, + simpa only [←(forget C).map_comp, category.assoc, is_iso.hom_inv_id,category.comp_id] using eq3, end +def stalk_bundles_eq (U : (opens X)ᵒᵖ) (x y : (sheaf_in_Type.obj 𝓖).1.obj U) + (eq1 : (sheaf_in_Type.map (to_godement_sheaf 𝓖)).val.app U x = + (sheaf_in_Type.map (to_godement_sheaf 𝓖)).val.app U y) (p : U.unop) : + (forget C).map (𝓖.presheaf.germ p) x = (forget C).map (𝓖.presheaf.germ p) y := -lemma to_godement_sheaf_app_injective (U : opens X) : - function.injective $ (forget C).map ((to_godement_sheaf 𝓖).1.app (opposite.op U)) := -λ x y eq1, begin - change (sheaf_in_Type.obj 𝓖).1.obj (opposite.op U) at x, - change (sheaf_in_Type.obj 𝓖).1.obj (opposite.op U) at y, - change (sheaf_in_Type.map (to_godement_sheaf 𝓖)).1.app (opposite.op U) x = - (sheaf_in_Type.map (to_godement_sheaf 𝓖)).1.app (opposite.op U) y at eq1, - apply presheaf.section_ext, - /- - U : opens ↥X, - x y : (sheaf_in_Type.obj 𝓖).val.obj (opposite.op U), - eq1 : - (sheaf_in_Type.map (to_godement_sheaf 𝓖)).val.app (opposite.op U) x = - (sheaf_in_Type.map (to_godement_sheaf 𝓖)).val.app (opposite.op U) y, - p : ↥U - ⊢ ⇑(𝓖.presheaf.germ p) x = ⇑(𝓖.presheaf.germ p) y - -/ - - intros p, - have := presheaf.germ_ext 𝓖.1 U p.2; - -- ext, - sorry + have eq1' := stalk_bundles_eq0 𝓖 U x y eq1 p, + have eq1'' : (forget C).map (to_godement_presheaf_aux_comp_π 𝓖.presheaf p) x = + (forget C).map (to_godement_presheaf_aux_comp_π 𝓖.presheaf p) y, + { dsimp at eq1' ⊢, + dunfold to_godement_presheaf_aux_comp_π, + simp only [comp_apply, eq1'], }, + rwa to_godement_presheaf_aux_comp_π_eq at eq1'', end +example : true := trivial + +def forget_stalk_iso (U : (opens X)ᵒᵖ) (x : U.unop) : + (forget C).obj (𝓖.presheaf.stalk x) ≅ (sheaf_in_Type.obj 𝓖).presheaf.stalk x := +preserves_colimit_iso _ _ + +lemma to_godement_sheaf_app_injective (U : opens X) : + function.injective $ (forget C).map ((to_godement_sheaf 𝓖).1.app (opposite.op U)) := +λ x y eq1, presheaf.section_ext _ _ _ _ (λ p, stalk_bundles_eq 𝓖 (opposite.op U) x y eq1 p) + instance : mono $ to_godement_sheaf 𝓖 := begin rw presheaf.mono_iff_stalk_mono, intros x, change mono ((presheaf.stalk_functor C x).map (to_godement_presheaf 𝓖.1)), rw concrete_category.mono_iff_injective_of_preserves_pullback, - exact (presheaf.app_injective_iff_stalk_functor_map_injective (to_godement_presheaf 𝓖.1)).mpr - (to_godement_sheaf_app_injective 𝓖) x, + refine (presheaf.app_injective_iff_stalk_functor_map_injective (to_godement_presheaf 𝓖.1)).mpr + (λ U x y eq1, presheaf.section_ext _ _ _ _ (λ p, stalk_bundles_eq 𝓖 (opposite.op U) x y eq1 p)) + x, end end presheaf From ef039e9dae7760a4d6b0ac4ef4c8f440fbad391b Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Sun, 21 Aug 2022 07:29:34 +0100 Subject: [PATCH 81/84] cache --- .../limits/shapes/products.lean | 13 ++ src/topology/sheaves/godement.lean | 207 +++++++----------- 2 files changed, 91 insertions(+), 129 deletions(-) diff --git a/src/category_theory/limits/shapes/products.lean b/src/category_theory/limits/shapes/products.lean index 4bf49c165731f..8ecbc896498c8 100644 --- a/src/category_theory/limits/shapes/products.lean +++ b/src/category_theory/limits/shapes/products.lean @@ -142,6 +142,19 @@ from a family of morphisms between the factors. abbreviation pi.map {f g : β → C} [has_product f] [has_product g] (p : Π b, f b ⟶ g b) : ∏ f ⟶ ∏ g := lim_map (discrete.nat_trans (λ X, p X.as)) + +instance pi.map_mono {f g : β → C} [has_product f] [has_product g] + (p : Π b, f b ⟶ g b) [Π i, mono (p i)] : mono (pi.map p) := +{ right_cancellation := λ c α₁ α₂ eq1, + begin + ext ⟨j⟩, dsimp, + have eq2 := eq_whisker eq1 (pi.π _ j), + rw [category.assoc, lim_map_π, category.assoc, lim_map_π, + ←category.assoc, ←category.assoc] at eq2, + simp only [discrete.nat_trans_app] at eq2, dsimp at eq2, + rwa cancel_mono at eq2, + end } + /-- Construct an isomorphism between categorical products (indexed by the same type) from a family of isomorphisms between the factors. diff --git a/src/topology/sheaves/godement.lean b/src/topology/sheaves/godement.lean index 5b226ecefb794..d038a3082acbc 100644 --- a/src/topology/sheaves/godement.lean +++ b/src/topology/sheaves/godement.lean @@ -123,123 +123,6 @@ variables [reflects_isomorphisms (forget C)] [preserves_filtered_colimits (forge map_id' := λ F, by { ext, dsimp, rw [id_apply] }, map_comp' := λ F G H f g, by { ext, dsimp, rw [comp_apply] } } -example : true := trivial - - -def godement_sheaf_in_Type : sheaf (Type u) X := sheaf_in_Type.obj (godement_sheaf 𝓖) - -def godement_sheaf_in_Type_obj_aux (U : (opens X)ᵒᵖ) : - (forget C).obj ∏ (λ x, (skyscraper_presheaf x (𝓖.presheaf.stalk x)).obj U) ≅ - ∏ (λ x, (forget C).obj ((skyscraper_presheaf x (𝓖.presheaf.stalk x)).obj U)) := -{ hom := pi.lift $ λ p, (forget C).map $ pi.π _ p, - inv := - begin - refine lim_map _ ≫ (preserves_limit_iso (forget C) _).inv, - refine { app := λ p x, x, naturality' := _ }, - rintros ⟨p⟩ ⟨q⟩ ⟨⟨(eq1 : p = q)⟩⟩, - dsimp, - induction eq1, - ext1, - dsimp, - simp only [discrete.functor_map_id, types_id_apply, id_apply], - end, - hom_inv_id' := - begin - rw [←category.assoc, limit.lift_map, iso.comp_inv_eq, category.id_comp], - refine limit.hom_ext _, - rintros ⟨p⟩, - rw [preserves_limits_iso_hom_π, limit.lift_π], - simpa only [cones.postcompose_obj_π, nat_trans.comp_app, fan.mk_π_app, forget_map_eq_coe], - end, - inv_hom_id' := - begin - ext1 ⟨p⟩, - rw [category.assoc, limit.lift_π, fan.mk_π_app, category.assoc, preserves_limits_iso_inv_π, - lim_map_π, category.id_comp], - refl, - end } - -def godement_sheaf_in_Type_obj (U : (opens X)ᵒᵖ) : - (godement_sheaf_in_Type 𝓖).1.obj U ≅ - ∏ (λ x, (forget C).obj $ (skyscraper_presheaf x (𝓖.presheaf.stalk x)).obj U) := -((forget C).map_iso $ godement_sheaf_obj 𝓖 U) ≪≫ godement_sheaf_in_Type_obj_aux 𝓖 U - -def sheaf_in_Type_skyscraper_sheaf (x : X) (c : C) : - (sheaf_in_Type.obj $ skyscraper_sheaf x c) ≅ - skyscraper_sheaf x ((forget C).obj c) := -let e1 := preserves_limit_iso (forget C) (functor.empty.{0} C) in -have it : is_terminal $ (forget C).obj (terminal C), -from is_terminal.of_iso -begin - rw show functor.empty C ⋙ forget C = functor.empty (Type u), from functor.empty_ext' _ _, - exact terminal_is_terminal, -end e1.symm, -let e : (sheaf_in_Type.obj $ skyscraper_sheaf x c) ≅ skyscraper_sheaf' x ((forget C).obj c) it := -{ hom := Sheaf.hom.mk - { app := λ U, eq_to_hom - begin - change (forget C).obj _ = (skyscraper_sheaf' x ((forget C).obj c) _).1.obj _, - by_cases hU : x ∈ U.unop, - { erw [skyscraper_presheaf_obj_of_mem _ hU, skyscraper_sheaf'], - dsimp, split_ifs, refl, }, - { erw [skyscraper_presheaf_obj_of_not_mem _ hU, skyscraper_sheaf'], - dsimp, split_ifs, refl, }, - end, - naturality' := λ U V inc, - begin - dsimp [skyscraper_sheaf', sheaf_in_Type, skyscraper_sheaf], - change (forget C).map _ ≫ _ = _, - by_cases hV : x ∈ V.unop, - { have hU : x ∈ U.unop := le_of_hom inc.unop hV, - split_ifs, - simp only [category.id_comp, eq_to_hom_trans, eq_to_hom_refl, eq_to_hom_map], - refl, }, - { split_ifs; - symmetry; - rw [←category.assoc, eq_comp_eq_to_hom]; - exact it.hom_ext _ _, }, - end }, - inv := Sheaf.hom.mk - { app := λ U, eq_to_hom - begin - change (skyscraper_sheaf' x ((forget C).obj c) _).1.obj _ = (forget C).obj _, - by_cases hU : x ∈ U.unop, - { erw [skyscraper_presheaf_obj_of_mem _ hU, skyscraper_sheaf'], - dsimp, split_ifs, refl, }, - { erw [skyscraper_presheaf_obj_of_not_mem _ hU, skyscraper_sheaf'], - dsimp, split_ifs, refl, }, - end, - naturality' := λ U V inc, - begin - dsimp [skyscraper_sheaf', sheaf_in_Type, skyscraper_sheaf], - change _ = _ ≫ (forget C).map _, - by_cases hV : x ∈ V.unop, - { have hU : x ∈ U.unop := le_of_hom inc.unop hV, - split_ifs, - simp only [category.id_comp, eq_to_hom_trans, eq_to_hom_refl, eq_to_hom_map], - refl, }, - { split_ifs; - rw [eq_comp_eq_to_hom, eq_comp_eq_to_hom]; - exact it.hom_ext _ _, }, - end }, - hom_inv_id' := - begin - ext : 3, dsimp, - rw [eq_to_hom_trans, eq_to_hom_refl], - end, - inv_hom_id' := - begin - ext : 3, dsimp, - rw [eq_to_hom_trans, eq_to_hom_refl], - end } in -e.trans $ (skyscraper_sheaf_iso _ _ it).symm - -example : true := trivial - -/-- -`x y : 𝓖(U)`, and `p ∈ U` -`` --/ def stalk_bundles_eq0 (U : (opens X)ᵒᵖ) (x y : (sheaf_in_Type.obj 𝓖).1.obj U) (eq1 : (sheaf_in_Type.map (to_godement_sheaf 𝓖)).val.app U x = (sheaf_in_Type.map (to_godement_sheaf 𝓖)).val.app U y) (p : U.unop) : @@ -272,7 +155,6 @@ def stalk_bundles_eq (U : (opens X)ᵒᵖ) (x y : (sheaf_in_Type.obj 𝓖).1.obj (eq1 : (sheaf_in_Type.map (to_godement_sheaf 𝓖)).val.app U x = (sheaf_in_Type.map (to_godement_sheaf 𝓖)).val.app U y) (p : U.unop) : (forget C).map (𝓖.presheaf.germ p) x = (forget C).map (𝓖.presheaf.germ p) y := - begin have eq1' := stalk_bundles_eq0 𝓖 U x y eq1 p, have eq1'' : (forget C).map (to_godement_presheaf_aux_comp_π 𝓖.presheaf p) x = @@ -285,23 +167,90 @@ end example : true := trivial -def forget_stalk_iso (U : (opens X)ᵒᵖ) (x : U.unop) : - (forget C).obj (𝓖.presheaf.stalk x) ≅ (sheaf_in_Type.obj 𝓖).presheaf.stalk x := -preserves_colimit_iso _ _ - -lemma to_godement_sheaf_app_injective (U : opens X) : - function.injective $ (forget C).map ((to_godement_sheaf 𝓖).1.app (opposite.op U)) := -λ x y eq1, presheaf.section_ext _ _ _ _ (λ p, stalk_bundles_eq 𝓖 (opposite.op U) x y eq1 p) - instance : mono $ to_godement_sheaf 𝓖 := begin rw presheaf.mono_iff_stalk_mono, intros x, change mono ((presheaf.stalk_functor C x).map (to_godement_presheaf 𝓖.1)), rw concrete_category.mono_iff_injective_of_preserves_pullback, - refine (presheaf.app_injective_iff_stalk_functor_map_injective (to_godement_presheaf 𝓖.1)).mpr - (λ U x y eq1, presheaf.section_ext _ _ _ _ (λ p, stalk_bundles_eq 𝓖 (opposite.op U) x y eq1 p)) - x, + exact (presheaf.app_injective_iff_stalk_functor_map_injective (to_godement_presheaf 𝓖.1)).mpr + (λ U x y H, presheaf.section_ext _ _ _ _ (λ p, stalk_bundles_eq 𝓖 (opposite.op U) x y H p)) x, +end + +section enough_injectives + +variables [enough_injectives C] + +namespace sheaf_enough_inj_aux + +def injective_sheaf : sheaf C X := +⟨∏ (λ x, skyscraper_presheaf x (injective.under $ 𝓕.stalk x) : X → presheaf C X), + limit_is_sheaf _ $ λ ⟨x⟩, (skyscraper_sheaf x _).2⟩ + +def injective_sheaf_iso : + injective_sheaf 𝓖.1 ≅ + ∏ (λ x, skyscraper_sheaf x (injective.under $ 𝓖.presheaf.stalk x)) := +{ hom := Sheaf.hom.mk $ eq_to_hom begin + change limit _ = limit _, congr, apply category_theory.functor.ext, + { rintros ⟨p⟩ ⟨q⟩ ⟨⟨(eq1 : p = q)⟩⟩, subst eq1, + rw [eq_to_hom_refl, category.id_comp, eq_to_hom_refl, category.comp_id], refl, }, + { rintros ⟨p⟩, dsimp, refl, }, + end ≫ (preserves_limit_iso (sheaf.forget C X) _).inv, + inv := Sheaf.hom.mk $ (preserves_limit_iso (sheaf.forget C X) _).hom ≫ eq_to_hom begin + change limit _ = limit _, congr, apply category_theory.functor.ext, + { rintros ⟨p⟩ ⟨q⟩ ⟨⟨(eq1 : p = q)⟩⟩, subst eq1, + rw [eq_to_hom_refl, category.id_comp, eq_to_hom_refl, category.comp_id], refl, }, + { rintros ⟨p⟩, dsimp, refl, }, + end, + hom_inv_id' := + begin + ext ⟨p⟩ U, dsimp, + rw [←category.assoc, category.assoc _ _ ((preserves_limit_iso (sheaf.forget C X) _).hom.app U), + iso.inv_hom_id_app, category.comp_id, category.id_comp, ←nat_trans.comp_app, eq_to_hom_trans, + eq_to_hom_refl], + convert category.id_comp _, + end, + inv_hom_id' := + begin + ext ⟨p⟩ U, dsimp, + rw [←category.assoc, category.assoc _ _ (eq_to_hom _), eq_to_hom_trans, eq_to_hom_refl, + category.comp_id, iso.hom_inv_id], + end } + +local notation `J` := injective_sheaf 𝓖.1 + +instance injective_J : injective J := +injective.of_iso (injective_sheaf_iso 𝓖).symm $ +@@injective.category_theory.limits.pi_obj.injective _ _ _ $ λ p, +(skyscraper_sheaf_injective p _ : injective + (skyscraper_sheaf p (injective.under (𝓖.presheaf.stalk p)))) + +def to_J : 𝓖 ⟶ J := +Sheaf.hom.mk $ to_godement_presheaf _ ≫ + pi.map (λ p, (skyscraper_presheaf_functor p).map $ injective.ι _) + +instance mono_to_J : mono (to_J 𝓖) := +(Sheaf.hom.mono_iff_presheaf_mono _ _ _).mpr +begin + haveI t1 : mono (to_godement_sheaf 𝓖) := infer_instance, + rw Sheaf.hom.mono_iff_presheaf_mono at t1, + change mono (to_godement_presheaf 𝓖.1) at t1, + resetI, + haveI t2 : mono (pi.map (λ p, (skyscraper_presheaf_functor p).map + (injective.ι (presheaf.stalk 𝓖.val p)))), + { sorry }, + apply mono_comp, end +end sheaf_enough_inj_aux + +instance : enough_injectives (sheaf C X) := +{ presentation := λ 𝓖, nonempty.intro + { J := sheaf_enough_inj_aux.injective_sheaf 𝓖.1, + injective := sheaf_enough_inj_aux.injective_J _, + f := sheaf_enough_inj_aux.to_J _, + mono := sheaf_enough_inj_aux.mono_to_J 𝓖 } } + +end enough_injectives + end presheaf From 06689d18af87b97aed5ad209f4dd94ea1a724bba Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Sun, 21 Aug 2022 07:43:06 +0100 Subject: [PATCH 82/84] finish --- src/topology/sheaves/godement.lean | 10 +++++++--- 1 file changed, 7 insertions(+), 3 deletions(-) diff --git a/src/topology/sheaves/godement.lean b/src/topology/sheaves/godement.lean index d038a3082acbc..92b1e4a54cce4 100644 --- a/src/topology/sheaves/godement.lean +++ b/src/topology/sheaves/godement.lean @@ -236,9 +236,13 @@ begin rw Sheaf.hom.mono_iff_presheaf_mono at t1, change mono (to_godement_presheaf 𝓖.1) at t1, resetI, - haveI t2 : mono (pi.map (λ p, (skyscraper_presheaf_functor p).map - (injective.ι (presheaf.stalk 𝓖.val p)))), - { sorry }, + have t21 : ∀ (p : X), mono ((skyscraper_presheaf_functor p).map + (injective.ι (presheaf.stalk 𝓖.val p))), + { exact λ p, @@functor.map_mono _ _ (skyscraper_presheaf_functor p) + (functor.preserves_monomorphisms_of_adjunction (stalk_skyscraper_presheaf_adj p)) + (injective.ι (presheaf.stalk 𝓖.val p)) _, }, + haveI t22 : mono (pi.map (λ p, (skyscraper_presheaf_functor p).map + (injective.ι (presheaf.stalk 𝓖.val p)))) := @pi.map_mono _ _ _ _ _ _ _ _ t21, apply mono_comp, end From 02957538d9ef6ff5272a5b7d22c889135bdaf363 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Sun, 21 Aug 2022 07:59:13 +0100 Subject: [PATCH 83/84] add some more lemma --- .../limits/shapes/products.lean | 19 ++++++++++++++----- 1 file changed, 14 insertions(+), 5 deletions(-) diff --git a/src/category_theory/limits/shapes/products.lean b/src/category_theory/limits/shapes/products.lean index 8ecbc896498c8..4875d848ac28d 100644 --- a/src/category_theory/limits/shapes/products.lean +++ b/src/category_theory/limits/shapes/products.lean @@ -144,15 +144,13 @@ abbreviation pi.map {f g : β → C} [has_product f] [has_product g] lim_map (discrete.nat_trans (λ X, p X.as)) instance pi.map_mono {f g : β → C} [has_product f] [has_product g] - (p : Π b, f b ⟶ g b) [Π i, mono (p i)] : mono (pi.map p) := + (p : Π b, f b ⟶ g b) [Π i, mono (p i)] : mono $ pi.map p := { right_cancellation := λ c α₁ α₂ eq1, begin ext ⟨j⟩, dsimp, have eq2 := eq_whisker eq1 (pi.π _ j), - rw [category.assoc, lim_map_π, category.assoc, lim_map_π, - ←category.assoc, ←category.assoc] at eq2, - simp only [discrete.nat_trans_app] at eq2, dsimp at eq2, - rwa cancel_mono at eq2, + rwa [category.assoc, lim_map_π, category.assoc, lim_map_π, ←category.assoc, ←category.assoc, + discrete.nat_trans_app, cancel_mono] at eq2, end } /-- @@ -169,6 +167,17 @@ from a family of morphisms between the factors. abbreviation sigma.map {f g : β → C} [has_coproduct f] [has_coproduct g] (p : Π b, f b ⟶ g b) : ∐ f ⟶ ∐ g := colim_map (discrete.nat_trans (λ X, p X.as)) + +instance sigma.map_epi {f g : β → C} [has_coproduct f] [has_coproduct g] + (p : Π b, f b ⟶ g b) [Π i, epi (p i)] : epi $ sigma.map p := +{ left_cancellation := λ c α₁ α₂ eq1, + begin + ext ⟨j⟩, dsimp, + have eq2 := whisker_eq (sigma.ι _ j) eq1, + rwa [←category.assoc, ι_colim_map, ←category.assoc, ι_colim_map, discrete.nat_trans_app, + category.assoc, category.assoc, cancel_epi] at eq2, + end } + /-- Construct an isomorphism between categorical coproducts (indexed by the same type) from a family of isomorphisms between the factors. From 28dc545db40e29caa3bf6a1db545083e055f40df Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Sun, 21 Aug 2022 08:19:51 +0100 Subject: [PATCH 84/84] add some docs --- src/topology/sheaves/godement.lean | 35 ++++++++++++++++++++++++++---- 1 file changed, 31 insertions(+), 4 deletions(-) diff --git a/src/topology/sheaves/godement.lean b/src/topology/sheaves/godement.lean index 92b1e4a54cce4..b1430cd6a6080 100644 --- a/src/topology/sheaves/godement.lean +++ b/src/topology/sheaves/godement.lean @@ -1,9 +1,38 @@ +/- +Copyright (c) 2022 Jujian Zhang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jujian Zhang +-/ import topology.sheaves.sheaf import topology.sheaves.limits import topology.sheaves.skyscraper import topology.sheaves.stalks import category_theory.preadditive.injective +/-! +# Godement resolution + +For a presheaf `𝓕 : (opens X)ᵒᵖ ⥤ C`, we can embedded `𝓕` into a sheaf `∏ₓ skyscraper(𝓕ₓ)` where +`x` ranges over `X` and `𝓕 ⟶ ∏ₓ skyscraper(𝓕ₓ)` is mono. + +## Main definition +* `godement_presheaf`: for a presheaf `𝓕`, its Godement presheaf is `∏ₓ skyscraper(𝓕ₓ)` +* `to_godement_presheaf`: the canonical map `𝓕 ⟶ godement_presheaf 𝓕` sending `s : 𝓕(U)` to a + bundle of stalks `x ↦ sₓ`. +* `godement_sheaf`: the Godement presheaf of any presheaf is a sheaf. +* `to_godement_sheaf`: for a sheaf `𝓖`, the canonical sheaf morphism `𝓖 ⟶ godement_sheaf 𝓖.1`. +For a sheaf `𝓖 : sheaf C X` where `C` is concrete and the forgetful functor preserves limits and +filtered colimits +* `mono_to_godement_sheaf`: the canonical map `𝓖 ⟶ godement_sheaf 𝓖` is mono. +If further `C` has enough injectives +* `sheaf_enough_inj_aux.injective_sheaf`: since each `𝓕ₓ` can be embedded into `𝓕ₓ ⟶ I(x)` via a + monomorphism, `𝓖` can be embedded into `∏ₓ skyscraper(I(x))`. +* `sheaf_enough_inj_aux.injective_J`: `∏ₓ skyscraper(I(x))` is injective. +* `sheaf_enough_inj_aux.to_J_mono`: the canonical map `𝓖 ⟶ ∏ₓ skyscraper(I(x))` is mono. +* `sheaf_has_enough_injectives`: the category of sheaves on `X` in `C` has enough injectives. +-/ + + noncomputable theory section presheaf @@ -111,8 +140,6 @@ def to_godement_sheaf : 𝓖 ⟶ godement_sheaf 𝓖 := ⟨to_godement_presheaf 𝓖.1⟩ variables [concrete_category.{u} C] [preserves_limits (forget C)] -variables [Π (U : opens X), preserves_colimits_of_shape - ((opens.grothendieck_topology X).cover U)ᵒᵖ (forget C)] variables [reflects_isomorphisms (forget C)] [preserves_filtered_colimits (forget C)] @[simps] def sheaf_in_Type : sheaf C X ⥤ sheaf (Type u) X := @@ -167,7 +194,7 @@ end example : true := trivial -instance : mono $ to_godement_sheaf 𝓖 := +instance mono_to_godement_sheaf : mono $ to_godement_sheaf 𝓖 := begin rw presheaf.mono_iff_stalk_mono, intros x, @@ -248,7 +275,7 @@ end end sheaf_enough_inj_aux -instance : enough_injectives (sheaf C X) := +instance sheaf_has_enough_injectives : enough_injectives (sheaf C X) := { presentation := λ 𝓖, nonempty.intro { J := sheaf_enough_inj_aux.injective_sheaf 𝓖.1, injective := sheaf_enough_inj_aux.injective_J _,