From 44b3f42373ba4f6daed7c340d262112db2468e97 Mon Sep 17 00:00:00 2001 From: leanprover-community-bot Date: Sat, 1 Jul 2023 04:27:34 +0000 Subject: [PATCH 01/61] chore(scripts): update nolints.txt (#19227) I am happy to remove some nolints for you! --- scripts/style-exceptions.txt | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/scripts/style-exceptions.txt b/scripts/style-exceptions.txt index c8f9e355af435..c04c94940278a 100644 --- a/scripts/style-exceptions.txt +++ b/scripts/style-exceptions.txt @@ -1,4 +1,4 @@ -src/control/basic.lean : line 9 : ERR_MOD : Module docstring missing, or too late +src/control/basic.lean : line 10 : ERR_MOD : Module docstring missing, or too late src/control/monad/cont.lean : line 13 : ERR_MOD : Module docstring missing, or too late src/control/monad/writer.lean : line 11 : ERR_MOD : Module docstring missing, or too late src/control/traversable/derive.lean : line 11 : ERR_MOD : Module docstring missing, or too late From a176cb1219e300e85793d44583dede42377b51af Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sun, 2 Jul 2023 06:51:23 +0000 Subject: [PATCH 02/61] feat(linear_algebra/matrix/schur_complement): invertibility of block matrices (#19156) Co-authored-by: Jeremy Tan Jie Rui Co-authored-by: Jireh Loreaux --- .../preadditive/biproducts.lean | 7 + .../matrix/schur_complement.lean | 153 ++++++++++++++++++ 2 files changed, 160 insertions(+) diff --git a/src/category_theory/preadditive/biproducts.lean b/src/category_theory/preadditive/biproducts.lean index b6f15e50a9d78..e50ebb7b168d3 100644 --- a/src/category_theory/preadditive/biproducts.lean +++ b/src/category_theory/preadditive/biproducts.lean @@ -46,6 +46,13 @@ In (or between) preadditive categories, * A functor preserves a biproduct if and only if it preserves the corresponding product if and only if it preserves the corresponding coproduct. + +There are connections between this material and the special case of the category whose morphisms are +matrices over a ring, in particular the Schur complement (see +`linear_algebra.matrix.schur_complement`). In particular, the declarations +`category_theory.biprod.iso_elim`, `category_theory.biprod.gaussian` +and `matrix.invertible_of_from_blocks₁₁_invertible` are all closely related. + -/ open category_theory diff --git a/src/linear_algebra/matrix/schur_complement.lean b/src/linear_algebra/matrix/schur_complement.lean index 440f1fbe308be..1c354e0f09131 100644 --- a/src/linear_algebra/matrix/schur_complement.lean +++ b/src/linear_algebra/matrix/schur_complement.lean @@ -3,6 +3,7 @@ Copyright (c) 2022 Alexander Bentkamp. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Alexander Bentkamp, Eric Wieser, Jeremy Avigad, Johan Commelin -/ +import data.matrix.invertible import linear_algebra.matrix.nonsingular_inverse import linear_algebra.matrix.pos_def @@ -11,6 +12,11 @@ import linear_algebra.matrix.pos_def This file proves properties of 2×2 block matrices `[A B; C D]` that relate to the Schur complement `D - C⬝A⁻¹⬝B`. +Some of the results here generalize to 2×2 matrices in a category, rather than just a ring. A few +results in this direction can be found in the the file `cateogry_theory.preadditive.biproducts`, +especially the declarations `category_theory.biprod.gaussian` and `category_theory.biprod.iso_elim`. +Compare with `matrix.invertible_of_from_blocks₁₁_invertible`. + ## Main results * `matrix.det_from_blocks₁₁`, `matrix.det_from_blocks₂₂`: determinant of a block matrix in terms of @@ -215,6 +221,153 @@ end end triangular +/-! ### 2×2 block matrices -/ + +section block + +/-! #### General 2×2 block matrices-/ + +/-- A block matrix is invertible if the bottom right corner and the corresponding schur complement +is. -/ +def from_blocks₂₂_invertible + (A : matrix m m α) (B : matrix m n α) (C : matrix n m α) (D : matrix n n α) + [invertible D] [invertible (A - B⬝⅟D⬝C)] : + invertible (from_blocks A B C D) := +begin + -- factor `from_blocks` via `from_blocks_eq_of_invertible₂₂`, and state the inverse we expect + refine invertible.copy' _ _ + (from_blocks + (⅟(A - B⬝⅟D⬝C)) (-(⅟(A - B⬝⅟D⬝C)⬝B⬝⅟D)) + (-(⅟D⬝C⬝⅟(A - B⬝⅟D⬝C))) (⅟D + ⅟D⬝C⬝⅟(A - B⬝⅟D⬝C)⬝B⬝⅟D)) + (from_blocks_eq_of_invertible₂₂ _ _ _ _) _, + { -- the product is invertible because all the factors are + letI : invertible (1 : matrix n n α) := invertible_one, + letI : invertible (1 : matrix m m α) := invertible_one, + refine invertible.matrix_mul _ (from_blocks_zero₁₂_invertible _ _ _), + exact invertible.matrix_mul (from_blocks_zero₂₁_invertible _ _ _) + (from_blocks_zero₂₁_invertible _ _ _) }, + { -- unfold the `invertible` instances to get the raw factors + show _ = from_blocks 1 0 (-(1 ⬝ (⅟ D ⬝ C) ⬝ 1)) 1 + ⬝ (from_blocks (⅟ (A - B ⬝ ⅟ D ⬝ C)) (-(⅟ (A - B ⬝ ⅟ D ⬝ C) ⬝ 0 ⬝ ⅟ D)) 0 (⅟ D) + ⬝ from_blocks 1 (-(1 ⬝ (B ⬝ ⅟ D) ⬝ 1)) 0 1), + -- combine into a single block matrix + simp only [from_blocks_multiply, inv_of_one, matrix.one_mul, matrix.mul_one, matrix.zero_mul, + matrix.mul_zero, add_zero, zero_add, neg_zero, matrix.mul_neg, matrix.neg_mul, neg_neg, + ←matrix.mul_assoc, add_comm], }, +end + +/-- A block matrix is invertible if the top left corner and the corresponding schur complement +is. -/ +def from_blocks₁₁_invertible + (A : matrix m m α) (B : matrix m n α) (C : matrix n m α) (D : matrix n n α) + [invertible A] [invertible (D - C⬝⅟A⬝B)] : + invertible (from_blocks A B C D) := +begin + -- we argue by symmetry + letI := from_blocks₂₂_invertible D C B A, + letI iDCBA + := + submatrix_equiv_invertible (from_blocks D C B A) (equiv.sum_comm _ _) (equiv.sum_comm _ _), + exact iDCBA.copy' _ + (from_blocks + (⅟A + ⅟A⬝B⬝⅟(D - C⬝⅟A⬝B)⬝C⬝⅟A) (-(⅟A⬝B⬝⅟(D - C⬝⅟A⬝B))) + (-(⅟(D - C⬝⅟A⬝B)⬝C⬝⅟A)) (⅟(D - C⬝⅟A⬝B))) + (from_blocks_submatrix_sum_swap_sum_swap _ _ _ _).symm + (from_blocks_submatrix_sum_swap_sum_swap _ _ _ _).symm, +end + +lemma inv_of_from_blocks₂₂_eq + (A : matrix m m α) (B : matrix m n α) (C : matrix n m α) (D : matrix n n α) + [invertible D] [invertible (A - B⬝⅟D⬝C)] [invertible (from_blocks A B C D)] : + ⅟(from_blocks A B C D) = from_blocks + (⅟(A - B⬝⅟D⬝C)) (-(⅟(A - B⬝⅟D⬝C)⬝B⬝⅟D)) + (-(⅟D⬝C⬝⅟(A - B⬝⅟D⬝C))) (⅟D + ⅟D⬝C⬝⅟(A - B⬝⅟D⬝C)⬝B⬝⅟D):= +begin + letI := from_blocks₂₂_invertible A B C D, + convert (rfl : ⅟(from_blocks A B C D) = _), +end + +lemma inv_of_from_blocks₁₁_eq + (A : matrix m m α) (B : matrix m n α) (C : matrix n m α) (D : matrix n n α) + [invertible A] [invertible (D - C⬝⅟A⬝B)] [invertible (from_blocks A B C D)] : + ⅟(from_blocks A B C D) = from_blocks + (⅟A + ⅟A⬝B⬝⅟(D - C⬝⅟A⬝B)⬝C⬝⅟A) (-(⅟A⬝B⬝⅟(D - C⬝⅟A⬝B))) + (-(⅟(D - C⬝⅟A⬝B)⬝C⬝⅟A)) (⅟(D - C⬝⅟A⬝B)) := +begin + letI := from_blocks₁₁_invertible A B C D, + convert (rfl : ⅟(from_blocks A B C D) = _), +end + +/-- If a block matrix is invertible and so is its bottom left element, then so is the corresponding +Schur complement. -/ +def invertible_of_from_blocks₂₂_invertible + (A : matrix m m α) (B : matrix m n α) (C : matrix n m α) (D : matrix n n α) + [invertible D] [invertible (from_blocks A B C D)] : invertible (A - B⬝⅟D⬝C) := +begin + suffices : invertible (from_blocks (A - B ⬝ ⅟ D ⬝ C) 0 0 D), + { exactI (invertible_of_from_blocks_zero₁₂_invertible (A - B ⬝ ⅟ D ⬝ C) 0 D).1 }, + letI : invertible (1 : matrix n n α) := invertible_one, + letI : invertible (1 : matrix m m α) := invertible_one, + letI iDC : invertible (from_blocks 1 0 (⅟ D ⬝ C) 1 : matrix (m ⊕ n) (m ⊕ n) α) := + from_blocks_zero₁₂_invertible _ _ _, + letI iBD : invertible (from_blocks 1 (B ⬝ ⅟ D) 0 1 : matrix(m ⊕ n) (m ⊕ n) α) := + from_blocks_zero₂₁_invertible _ _ _, + letI iBDC := invertible.copy ‹_› _ (from_blocks_eq_of_invertible₂₂ A B C D).symm, + refine (iBD.matrix_mul_left _).symm _, + refine (iDC.matrix_mul_right _).symm iBDC, +end + +/-- If a block matrix is invertible and so is its bottom left element, then so is the corresponding +Schur complement. -/ +def invertible_of_from_blocks₁₁_invertible + (A : matrix m m α) (B : matrix m n α) (C : matrix n m α) (D : matrix n n α) + [invertible A] [invertible (from_blocks A B C D)] : invertible (D - C⬝⅟A⬝B) := +begin + -- another symmetry argument + letI iABCD' := + submatrix_equiv_invertible (from_blocks A B C D) (equiv.sum_comm _ _) (equiv.sum_comm _ _), + letI iDCBA := iABCD'.copy _ (from_blocks_submatrix_sum_swap_sum_swap _ _ _ _).symm, + refine invertible_of_from_blocks₂₂_invertible D C B A, +end + +/-- `matrix.invertible_of_from_blocks₂₂_invertible` and `matrix.from_blocks₂₂_invertible` as an +equivalence. -/ +def invertible_equiv_from_blocks₂₂_invertible + (A : matrix m m α) (B : matrix m n α) (C : matrix n m α) (D : matrix n n α) + [invertible D] : invertible (from_blocks A B C D) ≃ invertible (A - B⬝⅟D⬝C) := +{ to_fun := λ iABCD, by exactI invertible_of_from_blocks₂₂_invertible _ _ _ _, + inv_fun := λ i_schur,by exactI from_blocks₂₂_invertible _ _ _ _, + left_inv := λ iABCD, subsingleton.elim _ _, + right_inv := λ i_schur, subsingleton.elim _ _ } + +/-- `matrix.invertible_of_from_blocks₁₁_invertible` and `matrix.from_blocks₁₁_invertible` as an +equivalence. -/ +def invertible_equiv_from_blocks₁₁_invertible + (A : matrix m m α) (B : matrix m n α) (C : matrix n m α) (D : matrix n n α) + [invertible A] : invertible (from_blocks A B C D) ≃ invertible (D - C⬝⅟A⬝B) := +{ to_fun := λ iABCD, by exactI invertible_of_from_blocks₁₁_invertible _ _ _ _, + inv_fun := λ i_schur,by exactI from_blocks₁₁_invertible _ _ _ _, + left_inv := λ iABCD, subsingleton.elim _ _, + right_inv := λ i_schur, subsingleton.elim _ _ } + +/-- If the bottom-left element of a block matrix is invertible, then the whole matrix is invertible +iff the corresponding schur complement is. -/ +lemma is_unit_from_blocks_iff_of_invertible₂₂ + {A : matrix m m α} {B : matrix m n α} {C : matrix n m α} {D : matrix n n α} [invertible D] : + is_unit (from_blocks A B C D) ↔ is_unit (A - B⬝⅟D⬝C) := +by simp only [← nonempty_invertible_iff_is_unit, + (invertible_equiv_from_blocks₂₂_invertible A B C D).nonempty_congr] + +/-- If the top-right element of a block matrix is invertible, then the whole matrix is invertible +iff the corresponding schur complement is. -/ +lemma is_unit_from_blocks_iff_of_invertible₁₁ + {A : matrix m m α} {B : matrix m n α} {C : matrix n m α} {D : matrix n n α} [invertible A] : + is_unit (from_blocks A B C D) ↔ is_unit (D - C⬝⅟A⬝B) := +by simp only [← nonempty_invertible_iff_is_unit, + (invertible_equiv_from_blocks₁₁_invertible A B C D).nonempty_congr] + +end block + /-! ### Lemmas about `matrix.det` -/ section det From e473c3198bb41f68560cab68a0529c854b618833 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sun, 2 Jul 2023 14:35:26 +0000 Subject: [PATCH 03/61] refactor: redefine `bundle.total_space` (#19221) - Use a custom structure for `bundle.total_space`. - Use `bundle.total_space.mk` instead of `bundle.total_space_mk`. - Use `bundle.total_space.to_prod` instead of `equiv.sigma_equiv_prod`. - Use `bundle.total_space.mk'` (scoped notation) to specify `F`. - Rename `bundle.trivial.proj_snd` to `bundle.total_space.trivial_snd` to allow dot notation. Should we just use `bundle.total_space.snd` since `bundle.trivial` is now reducible? - Add an unused argument to `bundle.total_space`. - Make `bundle.trivial` and `bundle.continuous_linear_map` reducible. - Drop instances that are no longer needed. --- src/data/bundle.lean | 126 +++++----- src/geometry/manifold/cont_mdiff_mfderiv.lean | 47 ++-- src/geometry/manifold/mfderiv.lean | 49 ++-- .../manifold/vector_bundle/basic.lean | 81 ++++--- src/geometry/manifold/vector_bundle/hom.lean | 14 +- .../manifold/vector_bundle/pullback.lean | 2 +- .../vector_bundle/smooth_section.lean | 18 +- .../manifold/vector_bundle/tangent.lean | 18 +- src/topology/covering.lean | 4 +- src/topology/fiber_bundle/basic.lean | 131 +++++----- src/topology/fiber_bundle/constructions.lean | 135 ++++++----- src/topology/fiber_bundle/trivialization.lean | 73 +++--- src/topology/vector_bundle/basic.lean | 223 +++++++++--------- src/topology/vector_bundle/constructions.lean | 16 +- src/topology/vector_bundle/hom.lean | 112 ++++----- 15 files changed, 501 insertions(+), 548 deletions(-) diff --git a/src/data/bundle.lean b/src/data/bundle.lean index 68cfd8bbb915b..b916ce0d12457 100644 --- a/src/data/bundle.lean +++ b/src/data/bundle.lean @@ -15,15 +15,26 @@ should contain all possible results that do not involve any topology. We represent a bundle `E` over a base space `B` as a dependent type `E : B → Type*`. -We provide a type synonym of `Σ x, E x` as `bundle.total_space E`, to be able to endow it with -a topology which is not the disjoint union topology `sigma.topological_space`. In general, the -constructions of fiber bundles we will make will be of this form. +We define `bundle.total_space F E` to be the type of pairs `⟨b, x⟩`, where `b : B` and `x : E x`. +This type is isomorphic to `Σ x, E x` and uses an extra argument `F` for reasons explained below. In +general, the constructions of fiber bundles we will make will be of this form. ## Main Definitions * `bundle.total_space` the total space of a bundle. * `bundle.total_space.proj` the projection from the total space to the base space. -* `bundle.total_space_mk` the constructor for the total space. +* `bundle.total_space.mk` the constructor for the total space. + +## Implementation Notes + +- We use a custom structure for the total space of a bundle instead of using a type synonym for the + canonical disjoint union `Σ x, E x` because the total space usually has a different topology and + Lean 4 `simp` fails to apply lemmas about `Σ x, E x` to elements of the total space. + +- The definition of `bundle.total_space` has an unused argument `F`. The reason is that in some + constructions (e.g., `bundle.continuous_linear_map.vector_bundle`) we need access to the atlas of + trivializations of original fiber bundles to construct the topology on the total space of the new + fiber bundle. ## References - https://en.wikipedia.org/wiki/Bundle_(mathematics) @@ -31,61 +42,65 @@ constructions of fiber bundles we will make will be of this form. namespace bundle -variables {B : Type*} (E : B → Type*) +variables {B F : Type*} (E : B → Type*) /-- -`bundle.total_space E` is the total space of the bundle `Σ x, E x`. -This type synonym is used to avoid conflicts with general sigma types. +`bundle.total_space E` is the total space of the bundle. It consists of pairs +`(proj : B, snd : E proj)`. -/ -def total_space := Σ x, E x +@[ext] +structure total_space (F : Type*) (E : B → Type*) := +(proj : B) +(snd : E proj) instance [inhabited B] [inhabited (E default)] : - inhabited (total_space E) := ⟨⟨default, default⟩⟩ + inhabited (total_space F E) := ⟨⟨default, default⟩⟩ variables {E} /-- `bundle.total_space.proj` is the canonical projection `bundle.total_space E → B` from the total space to the base space. -/ -@[simp, reducible] def total_space.proj : total_space E → B := sigma.fst +add_decl_doc total_space.proj -- this notation won't be used in the pretty-printer localized "notation `π` := @bundle.total_space.proj _" in bundle -/-- Constructor for the total space of a bundle. -/ -@[simp, reducible] def total_space_mk (b : B) (a : E b) : - bundle.total_space E := ⟨b, a⟩ - -lemma total_space.proj_mk {x : B} {y : E x} : (total_space_mk x y).proj = x := -rfl - -lemma sigma_mk_eq_total_space_mk {x : B} {y : E x} : sigma.mk x y = total_space_mk x y := -rfl +-- TODO: try `abbrev` in Lean 4 +localized "notation `total_space.mk'` F:max := @bundle.total_space.mk _ F _" in bundle lemma total_space.mk_cast {x x' : B} (h : x = x') (b : E x) : - total_space_mk x' (cast (congr_arg E h) b) = total_space_mk x b := + total_space.mk' F x' (cast (congr_arg E h) b) = total_space.mk x b := by { subst h, refl } -lemma total_space.eta (z : total_space E) : - total_space_mk z.proj z.2 = z := -sigma.eta z +instance {x : B} : has_coe_t (E x) (total_space F E) := ⟨total_space.mk x⟩ -instance {x : B} : has_coe_t (E x) (total_space E) := ⟨total_space_mk x⟩ +@[simp] lemma total_space.coe_proj (x : B) (v : E x) : (v : total_space F E).proj = x := rfl +@[simp] lemma total_space.coe_snd {x : B} {y : E x} : (y : total_space F E).snd = y := rfl -@[simp] lemma coe_fst (x : B) (v : E x) : (v : total_space E).fst = x := rfl -@[simp] lemma coe_snd {x : B} {y : E x} : (y : total_space E).snd = y := rfl +lemma total_space.coe_eq_mk {x : B} (v : E x) : (v : total_space F E) = total_space.mk x v := rfl -lemma to_total_space_coe {x : B} (v : E x) : (v : total_space E) = total_space_mk x v := rfl +lemma total_space.eta (z : total_space F E) : + total_space.mk z.proj z.2 = z := +by cases z; refl -- notation for the direct sum of two bundles over the same base notation E₁ ` ×ᵇ `:100 E₂ := λ x, E₁ x × E₂ x /-- `bundle.trivial B F` is the trivial bundle over `B` of fiber `F`. -/ -def trivial (B : Type*) (F : Type*) : B → Type* := function.const B F - -instance {F : Type*} [inhabited F] {b : B} : inhabited (bundle.trivial B F b) := ⟨(default : F)⟩ +@[reducible, nolint unused_arguments] +def trivial (B : Type*) (F : Type*) : B → Type* := λ _, F /-- The trivial bundle, unlike other bundles, has a canonical projection on the fiber. -/ -def trivial.proj_snd (B : Type*) (F : Type*) : total_space (bundle.trivial B F) → F := sigma.snd +def total_space.trivial_snd (B : Type*) (F : Type*) : total_space F (bundle.trivial B F) → F := +total_space.snd + +/-- A trivial bundle is equivalent to the product `B × F`. -/ +@[simps { attrs := [`simp, `mfld_simps] }] +def total_space.to_prod (B F : Type*) : total_space F (λ _ : B, F) ≃ B × F := +{ to_fun := λ x, (x.1, x.2), + inv_fun := λ x, ⟨x.1, x.2⟩, + left_inv := λ ⟨_, _⟩, rfl, + right_inv := λ ⟨_, _⟩, rfl } section pullback @@ -93,55 +108,36 @@ variable {B' : Type*} /-- The pullback of a bundle `E` over a base `B` under a map `f : B' → B`, denoted by `pullback f E` or `f *ᵖ E`, is the bundle over `B'` whose fiber over `b'` is `E (f b')`. -/ -@[nolint has_nonempty_instance] def pullback (f : B' → B) (E : B → Type*) := λ x, E (f x) +def pullback (f : B' → B) (E : B → Type*) : B' → Type* := λ x, E (f x) + +notation f ` *ᵖ ` E:max := pullback f E -notation f ` *ᵖ ` E := pullback f E +instance {f : B' → B} {x : B'} [nonempty (E (f x))] : nonempty (f *ᵖ E x) := ‹nonempty (E (f x))› /-- Natural embedding of the total space of `f *ᵖ E` into `B' × total_space E`. -/ @[simp] def pullback_total_space_embedding (f : B' → B) : - total_space (f *ᵖ E) → B' × total_space E := -λ z, (z.proj, total_space_mk (f z.proj) z.2) + total_space F (f *ᵖ E) → B' × total_space F E := +λ z, (z.proj, total_space.mk (f z.proj) z.2) /-- The base map `f : B' → B` lifts to a canonical map on the total spaces. -/ -def pullback.lift (f : B' → B) : total_space (f *ᵖ E) → total_space E := -λ z, total_space_mk (f z.proj) z.2 +@[simps { attrs := [`simp, `mfld_simps] }] +def pullback.lift (f : B' → B) : total_space F (f *ᵖ E) → total_space F E := +λ z, ⟨f z.proj, z.2⟩ -@[simp] lemma pullback.proj_lift (f : B' → B) (x : total_space (f *ᵖ E)) : - (pullback.lift f x).proj = f x.1 := -rfl - -@[simp] lemma pullback.lift_mk (f : B' → B) (x : B') (y : E (f x)) : - pullback.lift f (total_space_mk x y) = total_space_mk (f x) y := -rfl - -lemma pullback_total_space_embedding_snd (f : B' → B) (x : total_space (f *ᵖ E)) : - (pullback_total_space_embedding f x).2 = pullback.lift f x := +@[simp, mfld_simps] lemma pullback.lift_mk (f : B' → B) (x : B') (y : E (f x)) : + pullback.lift f (total_space.mk' F x y) = ⟨f x, y⟩ := rfl end pullback section fiber_structures -variable [∀ x, add_comm_monoid (E x)] - -@[simp] lemma coe_snd_map_apply (x : B) (v w : E x) : - (↑(v + w) : total_space E).snd = (v : total_space E).snd + (w : total_space E).snd := rfl +@[simp] lemma coe_snd_map_apply [∀ x, has_add (E x)] (x : B) (v w : E x) : + (↑(v + w) : total_space F E).snd = (v : total_space F E).snd + (w : total_space F E).snd := rfl -variables (R : Type*) [semiring R] [∀ x, module R (E x)] - -@[simp] lemma coe_snd_map_smul (x : B) (r : R) (v : E x) : - (↑(r • v) : total_space E).snd = r • (v : total_space E).snd := rfl +@[simp] lemma coe_snd_map_smul {R} [∀ x, has_smul R (E x)] (x : B) (r : R) (v : E x) : + (↑(r • v) : total_space F E).snd = r • (v : total_space F E).snd := rfl end fiber_structures -section trivial_instances - -variables {F : Type*} {R : Type*} [semiring R] (b : B) - -instance [add_comm_monoid F] : add_comm_monoid (bundle.trivial B F b) := ‹add_comm_monoid F› -instance [add_comm_group F] : add_comm_group (bundle.trivial B F b) := ‹add_comm_group F› -instance [add_comm_monoid F] [module R F] : module R (bundle.trivial B F b) := ‹module R F› - -end trivial_instances - end bundle diff --git a/src/geometry/manifold/cont_mdiff_mfderiv.lean b/src/geometry/manifold/cont_mdiff_mfderiv.lean index 048675244325c..9efd7835a3bd5 100644 --- a/src/geometry/manifold/cont_mdiff_mfderiv.lean +++ b/src/geometry/manifold/cont_mdiff_mfderiv.lean @@ -248,7 +248,7 @@ space are model spaces in models with corners. The general fact is proved in lemma cont_mdiff_on.continuous_on_tangent_map_within_aux {f : H → H'} {s : set H} (hf : cont_mdiff_on I I' n f s) (hn : 1 ≤ n) (hs : unique_mdiff_on I s) : - continuous_on (tangent_map_within I I' f s) (π (tangent_space I) ⁻¹' s) := + continuous_on (tangent_map_within I I' f s) (π E (tangent_space I) ⁻¹' s) := begin suffices h : continuous_on (λ (p : H × E), (f p.fst, (fderiv_within 𝕜 (written_in_ext_chart_at I I' p.fst f) (I.symm ⁻¹' s ∩ range I) @@ -258,7 +258,7 @@ begin have B := ((tangent_bundle_model_space_homeomorph H' I').symm.continuous.comp_continuous_on h) .comp' A, have : (univ ∩ ⇑(tangent_bundle_model_space_homeomorph H I) ⁻¹' (prod.fst ⁻¹' s)) = - π (tangent_space I) ⁻¹' s, + π E (tangent_space I) ⁻¹' s, by { ext ⟨x, v⟩, simp only with mfld_simps }, rw this at B, apply B.congr, @@ -307,7 +307,8 @@ are model spaces in models with corners. The general fact is proved in lemma cont_mdiff_on.cont_mdiff_on_tangent_map_within_aux {f : H → H'} {s : set H} (hf : cont_mdiff_on I I' n f s) (hmn : m + 1 ≤ n) (hs : unique_mdiff_on I s) : - cont_mdiff_on I.tangent I'.tangent m (tangent_map_within I I' f s) (π (tangent_space I) ⁻¹' s) := + cont_mdiff_on I.tangent I'.tangent m (tangent_map_within I I' f s) + (π E (tangent_space I) ⁻¹' s) := begin have m_le_n : m ≤ n, { apply le_trans _ hmn, @@ -324,13 +325,13 @@ begin rw cont_mdiff_on_iff, refine ⟨hf.continuous_on_tangent_map_within_aux one_le_n hs, λp q, _⟩, have A : range I ×ˢ univ ∩ - ((equiv.sigma_equiv_prod H E).symm ∘ λ (p : E × E), ((I.symm) p.fst, p.snd)) ⁻¹' - (π (tangent_space I) ⁻¹' s) + ((total_space.to_prod H E).symm ∘ λ (p : E × E), ((I.symm) p.fst, p.snd)) ⁻¹' + (π E (tangent_space I) ⁻¹' s) = (range I ∩ I.symm ⁻¹' s) ×ˢ univ, - by { ext ⟨x, v⟩, simp only with mfld_simps }, + by { ext ⟨x, v⟩, simp only with mfld_simps, }, suffices h : cont_diff_on 𝕜 m (((λ (p : H' × E'), (I' p.fst, p.snd)) ∘ - (equiv.sigma_equiv_prod H' E')) ∘ tangent_map_within I I' f s ∘ - ((equiv.sigma_equiv_prod H E).symm) ∘ λ (p : E × E), (I.symm p.fst, p.snd)) + (total_space.to_prod H' E')) ∘ tangent_map_within I I' f s ∘ + ((total_space.to_prod H E).symm) ∘ λ (p : E × E), (I.symm p.fst, p.snd)) ((range ⇑I ∩ ⇑(I.symm) ⁻¹' s) ×ˢ univ), by simpa [A] using h, change cont_diff_on 𝕜 m (λ (p : E × E), @@ -369,7 +370,7 @@ is `C^m` when `m+1 ≤ n`. -/ theorem cont_mdiff_on.cont_mdiff_on_tangent_map_within (hf : cont_mdiff_on I I' n f s) (hmn : m + 1 ≤ n) (hs : unique_mdiff_on I s) : cont_mdiff_on I.tangent I'.tangent m (tangent_map_within I I' f s) - (π (tangent_space I) ⁻¹' s) := + (π E (tangent_space I) ⁻¹' s) := begin /- The strategy of the proof is to avoid unfolding the definitions, and reduce by functoriality to the case of functions on the model spaces, where we have already proved the result. @@ -407,20 +408,20 @@ begin let il := chart_at (model_prod H E) (tangent_map I I l p), let ir := chart_at (model_prod H' E') (tangent_map I I' (r ∘ f) p), let s' := f ⁻¹' r.source ∩ s ∩ l.source, - let s'_lift := π (tangent_space I) ⁻¹' s', + let s'_lift := π E (tangent_space I) ⁻¹' s', let s'l := l.target ∩ l.symm ⁻¹' s', - let s'l_lift := π (tangent_space I) ⁻¹' s'l, + let s'l_lift := π E (tangent_space I) ⁻¹' s'l, rcases continuous_on_iff'.1 hf'.1 r.source r.open_source with ⟨o, o_open, ho⟩, suffices h : cont_mdiff_on I.tangent I'.tangent m (tangent_map_within I I' f s) s'_lift, - { refine ⟨π (tangent_space I) ⁻¹' (o ∩ l.source), _, _, _⟩, - show is_open (π (tangent_space I) ⁻¹' (o ∩ l.source)), from + { refine ⟨π E (tangent_space I) ⁻¹' (o ∩ l.source), _, _, _⟩, + show is_open (π E (tangent_space I) ⁻¹' (o ∩ l.source)), from (is_open.inter o_open l.open_source).preimage (continuous_proj E _) , - show p ∈ π (tangent_space I) ⁻¹' (o ∩ l.source), + show p ∈ π E (tangent_space I) ⁻¹' (o ∩ l.source), { simp, have : p.proj ∈ f ⁻¹' r.source ∩ s, by simp [hp], rw ho at this, exact this.1 }, - { have : π (tangent_space I) ⁻¹' s ∩ π (tangent_space I) ⁻¹' (o ∩ l.source) = s'_lift, + { have : π E (tangent_space I) ⁻¹' s ∩ π E (tangent_space I) ⁻¹' (o ∩ l.source) = s'_lift, { dsimp only [s'_lift, s'], rw [ho], mfld_set_tac }, rw this, exact h } }, @@ -547,10 +548,10 @@ end derivative is continuous there. -/ theorem cont_mdiff_on.continuous_on_tangent_map_within (hf : cont_mdiff_on I I' n f s) (hmn : 1 ≤ n) (hs : unique_mdiff_on I s) : - continuous_on (tangent_map_within I I' f s) (π (tangent_space I) ⁻¹' s) := + continuous_on (tangent_map_within I I' f s) (π E (tangent_space I) ⁻¹' s) := begin have : cont_mdiff_on I.tangent I'.tangent 0 (tangent_map_within I I' f s) - (π (tangent_space I) ⁻¹' s) := + (π E (tangent_space I) ⁻¹' s) := hf.cont_mdiff_on_tangent_map_within hmn hs, exact this.continuous_on end @@ -599,15 +600,15 @@ may seem. TODO define splittings of vector bundles; state this result invariantly. -/ lemma tangent_map_tangent_bundle_pure (p : tangent_bundle I M) : - tangent_map I I.tangent (zero_section (tangent_space I)) p = ⟨⟨p.proj, 0⟩, ⟨p.2, 0⟩⟩ := + tangent_map I I.tangent (zero_section E (tangent_space I)) p = ⟨⟨p.proj, 0⟩, ⟨p.2, 0⟩⟩ := begin rcases p with ⟨x, v⟩, have N : I.symm ⁻¹' (chart_at H x).target ∈ 𝓝 (I ((chart_at H x) x)), { apply is_open.mem_nhds, apply (local_homeomorph.open_target _).preimage I.continuous_inv_fun, simp only with mfld_simps }, - have A : mdifferentiable_at I I.tangent (λ x, @total_space_mk M (tangent_space I) x 0) x, - { have : smooth I (I.prod 𝓘(𝕜, E)) (zero_section (tangent_space I : M → Type*)) := + have A : mdifferentiable_at I I.tangent (λ x, @total_space.mk M E (tangent_space I) x 0) x, + { have : smooth I (I.prod 𝓘(𝕜, E)) (zero_section E (tangent_space I : M → Type*)) := bundle.smooth_zero_section 𝕜 (tangent_space I : M → Type*), exact this.mdifferentiable_at }, have B : fderiv_within 𝕜 (λ (x' : E), (x', (0 : E))) (set.range ⇑I) (I ((chart_at H x) x)) v @@ -618,12 +619,12 @@ begin { exact differentiable_at_const _ }, { exact model_with_corners.unique_diff_at_image I }, { exact differentiable_at_id'.prod (differentiable_at_const _) } }, - simp only [bundle.zero_section, tangent_map, mfderiv, total_space.proj_mk, A, - if_pos, chart_at, fiber_bundle.charted_space_chart_at, tangent_bundle.trivialization_at_apply, + simp only [bundle.zero_section, tangent_map, mfderiv, A, if_pos, chart_at, + fiber_bundle.charted_space_chart_at, tangent_bundle.trivialization_at_apply, tangent_bundle_core, function.comp, continuous_linear_map.map_zero] with mfld_simps, rw [← fderiv_within_inter N] at B, rw [← fderiv_within_inter N, ← B], - congr' 2, + congr' 1, refine fderiv_within_congr (λ y hy, _) _, { simp only with mfld_simps at hy, simp only [hy, prod.mk.inj_iff] with mfld_simps }, diff --git a/src/geometry/manifold/mfderiv.lean b/src/geometry/manifold/mfderiv.lean index 30c9eb15fbf17..c621b85eee4fc 100644 --- a/src/geometry/manifold/mfderiv.lean +++ b/src/geometry/manifold/mfderiv.lean @@ -688,15 +688,9 @@ end @[simp, mfld_simps] lemma tangent_map_within_proj {p : tangent_bundle I M} : (tangent_map_within I I' f s p).proj = f p.proj := rfl -@[simp, mfld_simps] lemma tangent_map_within_fst {p : tangent_bundle I M} : - (tangent_map_within I I' f s p).1 = f p.1 := rfl - @[simp, mfld_simps] lemma tangent_map_proj {p : tangent_bundle I M} : (tangent_map I I' f p).proj = f p.proj := rfl -@[simp, mfld_simps] lemma tangent_map_fst {p : tangent_bundle I M} : - (tangent_map I I' f p).1 = f p.1 := rfl - omit Is I's lemma mdifferentiable_within_at.prod_mk {f : M → M'} {g : M → M''} @@ -841,8 +835,7 @@ lemma tangent_map_within_congr (h : ∀ x ∈ s, f x = f₁ x) (p : tangent_bundle I M) (hp : p.1 ∈ s) (hs : unique_mdiff_within_at I s p.1) : tangent_map_within I I' f s p = tangent_map_within I I' f₁ s p := begin - simp only [tangent_map_within, h p.fst hp, true_and, eq_self_iff_true, heq_iff_eq, - sigma.mk.inj_iff], + simp only [tangent_map_within, h p.1 hp, true_and, eq_self_iff_true, heq_iff_eq], congr' 1, exact mfderiv_within_congr hs h (h _ hp) end @@ -1344,17 +1337,17 @@ lemma mfderiv_within_fst {s : set (M × M')} {x : M × M'} by { rw mdifferentiable.mfderiv_within (mdifferentiable_at_fst I I') hxs, exact mfderiv_fst I I' } @[simp, mfld_simps] lemma tangent_map_prod_fst {p : tangent_bundle (I.prod I') (M × M')} : - tangent_map (I.prod I') I prod.fst p = total_space_mk p.proj.1 p.2.1 := + tangent_map (I.prod I') I prod.fst p = ⟨p.proj.1, p.2.1⟩ := by simp [tangent_map] lemma tangent_map_within_prod_fst {s : set (M × M')} {p : tangent_bundle (I.prod I') (M × M')} (hs : unique_mdiff_within_at (I.prod I') s p.proj) : - tangent_map_within (I.prod I') I prod.fst s p = total_space_mk p.proj.1 p.2.1 := + tangent_map_within (I.prod I') I prod.fst s p = ⟨p.proj.1, p.2.1⟩ := begin simp only [tangent_map_within], - rw mfderiv_within_fst, - { rcases p, refl }, - { exact hs } + rw mfderiv_within_fst _ _ hs, + rcases p, + exact ⟨rfl, heq.rfl⟩ end lemma has_mfderiv_at_snd (x : M × M') : @@ -1400,16 +1393,16 @@ lemma mfderiv_within_snd {s : set (M × M')} {x : M × M'} by { rw mdifferentiable.mfderiv_within (mdifferentiable_at_snd I I') hxs, exact mfderiv_snd I I' } @[simp, mfld_simps] lemma tangent_map_prod_snd {p : tangent_bundle (I.prod I') (M × M')} : - tangent_map (I.prod I') I' prod.snd p = total_space_mk p.proj.2 p.2.2 := + tangent_map (I.prod I') I' prod.snd p = ⟨p.proj.2, p.2.2⟩ := by simp [tangent_map] lemma tangent_map_within_prod_snd {s : set (M × M')} {p : tangent_bundle (I.prod I') (M × M')} (hs : unique_mdiff_within_at (I.prod I') s p.proj) : - tangent_map_within (I.prod I') I' prod.snd s p = total_space_mk p.proj.2 p.2.2 := + tangent_map_within (I.prod I') I' prod.snd s p = ⟨p.proj.2, p.2.2⟩ := begin simp only [tangent_map_within], rw mfderiv_within_snd, - { rcases p, refl }, + { rcases p, split; refl }, { exact hs } end @@ -1716,7 +1709,7 @@ mdifferentiable_of_mem_atlas _ (chart_mem_atlas _ _) the identification between the tangent bundle of the model space and the product space. -/ lemma tangent_map_chart {p q : tangent_bundle I M} (h : q.1 ∈ (chart_at H p.1).source) : tangent_map I I (chart_at H p.1) q = - (equiv.sigma_equiv_prod _ _).symm + (total_space.to_prod _ _).symm ((chart_at (model_prod H E) p : tangent_bundle I M → model_prod H E) q) := begin dsimp [tangent_map], @@ -1732,14 +1725,14 @@ lemma tangent_map_chart_symm {p : tangent_bundle I M} {q : tangent_bundle I H} (h : q.1 ∈ (chart_at H p.1).target) : tangent_map I I (chart_at H p.1).symm q = ((chart_at (model_prod H E) p).symm : model_prod H E → tangent_bundle I M) - ((equiv.sigma_equiv_prod H E) q) := + ((total_space.to_prod H E) q) := begin dsimp only [tangent_map], rw mdifferentiable_at.mfderiv (mdifferentiable_at_atlas_symm _ (chart_mem_atlas _ _) h), -- a trivial instance is needed after the rewrite, handle it right now. rotate, { apply_instance }, simp only [continuous_linear_map.coe_coe, tangent_bundle.chart_at, h, - tangent_bundle_core, chart_at, sigma.mk.inj_iff] with mfld_simps, + tangent_bundle_core, chart_at, total_space.to_prod_apply] with mfld_simps, end end charts @@ -2004,23 +1997,23 @@ end open bundle variables {F : Type*} [normed_add_comm_group F] [normed_space 𝕜 F] - (Z : M → Type*) [topological_space (total_space Z)] [∀ b, topological_space (Z b)] + (Z : M → Type*) [topological_space (total_space F Z)] [∀ b, topological_space (Z b)] [∀ b, add_comm_monoid (Z b)] [∀ b, module 𝕜 (Z b)] [fiber_bundle F Z] [vector_bundle 𝕜 F Z] [smooth_vector_bundle F Z I] /-- In a smooth fiber bundle, the preimage under the projection of a set with unique differential in the basis also has unique differential. -/ lemma unique_mdiff_on.smooth_bundle_preimage (hs : unique_mdiff_on I s) : - unique_mdiff_on (I.prod (𝓘(𝕜, F))) (π Z ⁻¹' s) := + unique_mdiff_on (I.prod (𝓘(𝕜, F))) (π F Z ⁻¹' s) := begin /- Using a chart (and the fact that unique differentiability is invariant under charts), we reduce the situation to the model space, where we can use the fact that products respect unique differentiability. -/ assume p hp, - replace hp : p.fst ∈ s, by simpa only with mfld_simps using hp, + replace hp : p.1 ∈ s, by simpa only with mfld_simps using hp, let e₀ := chart_at H p.1, let e := chart_at (model_prod H F) p, - have h2s : ∀ x, x ∈ e.target ∩ e.symm ⁻¹' (π Z ⁻¹' s) ↔ + have h2s : ∀ x, x ∈ e.target ∩ e.symm ⁻¹' (π F Z ⁻¹' s) ↔ (x.1 ∈ e₀.target ∧ (e₀.symm) x.1 ∈ (trivialization_at F Z p.1).base_set) ∧ (e₀.symm) x.1 ∈ s, { intro x, have A : x ∈ e.target ↔ x.1 ∈ e₀.target ∧ @@ -2032,13 +2025,13 @@ begin simp only [fiber_bundle.charted_space_chart_at_symm_fst p x hx] with mfld_simps }, -- It suffices to prove unique differentiability in a chart suffices h : unique_mdiff_on (I.prod (𝓘(𝕜, F))) - (e.target ∩ e.symm ⁻¹' (π Z ⁻¹' s)), + (e.target ∩ e.symm ⁻¹' (π F Z ⁻¹' s)), { have A : unique_mdiff_on (I.prod (𝓘(𝕜, F))) (e.symm.target ∩ - e.symm.symm ⁻¹' (e.target ∩ e.symm⁻¹' (π Z ⁻¹' s))), + e.symm.symm ⁻¹' (e.target ∩ e.symm⁻¹' (π F Z ⁻¹' s))), { apply h.unique_mdiff_on_preimage, exact (mdifferentiable_of_mem_atlas _ (chart_mem_atlas _ _)).symm, apply_instance }, - have : p ∈ e.symm.target ∩ e.symm.symm ⁻¹' (e.target ∩ e.symm⁻¹' (π Z ⁻¹' s)), + have : p ∈ e.symm.target ∩ e.symm.symm ⁻¹' (e.target ∩ e.symm⁻¹' (π F Z ⁻¹' s)), { simp only [e, hp] with mfld_simps }, apply (A _ this).mono, assume q hq, @@ -2047,7 +2040,7 @@ begin assume q hq, simp only [unique_mdiff_within_at, model_with_corners.prod, -preimage_inter] with mfld_simps, have : 𝓝[(I.symm ⁻¹' (e₀.target ∩ e₀.symm⁻¹' s) ∩ range I) ×ˢ univ] (I q.1, q.2) ≤ - 𝓝[(λ (p : E × F), (I.symm p.1, p.snd)) ⁻¹' (e.target ∩ e.symm ⁻¹' (π Z ⁻¹' s)) ∩ + 𝓝[(λ (p : E × F), (I.symm p.1, p.snd)) ⁻¹' (e.target ∩ e.symm ⁻¹' (π F Z ⁻¹' s)) ∩ (range I ×ˢ univ)] (I q.1, q.2), { rw [nhds_within_le_iff, mem_nhds_within], refine ⟨(λ (p : E × F), (I.symm p.1, p.snd)) ⁻¹' e.target, _, _, _⟩, @@ -2076,7 +2069,7 @@ begin end lemma unique_mdiff_on.tangent_bundle_proj_preimage (hs : unique_mdiff_on I s): - unique_mdiff_on I.tangent (π (tangent_space I) ⁻¹' s) := + unique_mdiff_on I.tangent (π E (tangent_space I) ⁻¹' s) := hs.smooth_bundle_preimage _ end unique_mdiff diff --git a/src/geometry/manifold/vector_bundle/basic.lean b/src/geometry/manifold/vector_bundle/basic.lean index 49445ccc558dc..2ee022460ef97 100644 --- a/src/geometry/manifold/vector_bundle/basic.lean +++ b/src/geometry/manifold/vector_bundle/basic.lean @@ -69,14 +69,14 @@ variables {𝕜 B B' F M : Type*} {E : B → Type*} /-! ### Charted space structure on a fiber bundle -/ section -variables [topological_space F] [topological_space (total_space E)] [∀ x, topological_space (E x)] +variables [topological_space F] [topological_space (total_space F E)] [∀ x, topological_space (E x)] {HB : Type*} [topological_space HB] [topological_space B] [charted_space HB B] [fiber_bundle F E] /-- A fiber bundle `E` over a base `B` with model fiber `F` is naturally a charted space modelled on `B × F`. -/ -instance fiber_bundle.charted_space : charted_space (B × F) (total_space E) := -{ atlas := (λ e : trivialization F (π E), e.to_local_homeomorph) '' trivialization_atlas F E, +instance fiber_bundle.charted_space : charted_space (B × F) (total_space F E) := +{ atlas := (λ e : trivialization F (π F E), e.to_local_homeomorph) '' trivialization_atlas F E, chart_at := λ x, (trivialization_at F E x.proj).to_local_homeomorph, mem_chart_source := λ x, (trivialization_at F E x.proj).mem_source.mpr (mem_base_set_trivialization_at F E x.proj), @@ -87,11 +87,11 @@ local attribute [reducible] model_prod /-- Let `B` be a charted space modelled on `HB`. Then a fiber bundle `E` over a base `B` with model fiber `F` is naturally a charted space modelled on `HB.prod F`. -/ -instance fiber_bundle.charted_space' : charted_space (model_prod HB F) (total_space E) := +instance fiber_bundle.charted_space' : charted_space (model_prod HB F) (total_space F E) := charted_space.comp _ (model_prod B F) _ end -lemma fiber_bundle.charted_space_chart_at (x : total_space E) : +lemma fiber_bundle.charted_space_chart_at (x : total_space F E) : chart_at (model_prod HB F) x = (trivialization_at F E x.proj).to_local_homeomorph ≫ₕ (chart_at HB x.proj).prod (local_homeomorph.refl F) := @@ -102,7 +102,7 @@ begin trivialization.coe_fst' _ (mem_base_set_trivialization_at F E x.proj)] end -lemma fiber_bundle.charted_space_chart_at_symm_fst (x : total_space E) (y : model_prod HB F) +lemma fiber_bundle.charted_space_chart_at_symm_fst (x : total_space F E) (y : model_prod HB F) (hy : y ∈ (chart_at (model_prod HB F) x).target) : ((chart_at (model_prod HB F) x).symm y).proj = (chart_at HB x.proj).symm y.1 := begin @@ -115,7 +115,7 @@ end section variables [nontrivially_normed_field 𝕜] [normed_add_comm_group F] [normed_space 𝕜 F] - [topological_space (total_space E)] [∀ x, topological_space (E x)] + [topological_space (total_space F E)] [∀ x, topological_space (E x)] {EB : Type*} [normed_add_comm_group EB] [normed_space 𝕜 EB] {HB : Type*} [topological_space HB] (IB : model_with_corners 𝕜 EB HB) @@ -127,7 +127,7 @@ variables [nontrivially_normed_field 𝕜] variables [topological_space B] [charted_space HB B] [fiber_bundle F E] -protected lemma fiber_bundle.ext_chart_at (x : total_space E) : +protected lemma fiber_bundle.ext_chart_at (x : total_space F E) : ext_chart_at (IB.prod 𝓘(𝕜, F)) x = (trivialization_at F E x.proj).to_local_equiv ≫ (ext_chart_at IB x.proj).prod (local_equiv.refl F) := @@ -146,7 +146,7 @@ namespace bundle variables {F E IB} /-- Characterization of C^n functions into a smooth vector bundle. -/ -lemma cont_mdiff_within_at_total_space (f : M → total_space E) {s : set M} {x₀ : M} : +lemma cont_mdiff_within_at_total_space (f : M → total_space F E) {s : set M} {x₀ : M} : cont_mdiff_within_at IM (IB.prod (𝓘(𝕜, F))) n f s x₀ ↔ cont_mdiff_within_at IM IB n (λ x, (f x).proj) s x₀ ∧ cont_mdiff_within_at IM 𝓘(𝕜, F) n (λ x, (trivialization_at F E (f x₀).proj (f x)).2) s x₀ := @@ -171,7 +171,7 @@ begin end /-- Characterization of C^n functions into a smooth vector bundle. -/ -lemma cont_mdiff_at_total_space (f : M → total_space E) (x₀ : M) : +lemma cont_mdiff_at_total_space (f : M → total_space F E) (x₀ : M) : cont_mdiff_at IM (IB.prod (𝓘(𝕜, F))) n f x₀ ↔ cont_mdiff_at IM IB n (λ x, (f x).proj) x₀ ∧ cont_mdiff_at IM 𝓘(𝕜, F) n (λ x, (trivialization_at F E (f x₀).proj (f x)).2) x₀ := @@ -179,12 +179,12 @@ by { simp_rw [← cont_mdiff_within_at_univ], exact cont_mdiff_within_at_total_s /-- Characterization of C^n sections of a smooth vector bundle. -/ lemma cont_mdiff_at_section (s : Π x, E x) (x₀ : B) : - cont_mdiff_at IB (IB.prod (𝓘(𝕜, F))) n (λ x, total_space_mk x (s x)) x₀ ↔ - cont_mdiff_at IB 𝓘(𝕜, F) n (λ x, (trivialization_at F E x₀ (total_space_mk x (s x))).2) x₀ := + cont_mdiff_at IB (IB.prod (𝓘(𝕜, F))) n (λ x, total_space.mk' F x (s x)) x₀ ↔ + cont_mdiff_at IB 𝓘(𝕜, F) n (λ x, (trivialization_at F E x₀ (total_space.mk' F x (s x))).2) x₀ := by { simp_rw [cont_mdiff_at_total_space, and_iff_right_iff_imp], intro x, exact cont_mdiff_at_id } variables (E) -lemma cont_mdiff_proj : cont_mdiff (IB.prod 𝓘(𝕜, F)) IB n (π E) := +lemma cont_mdiff_proj : cont_mdiff (IB.prod 𝓘(𝕜, F)) IB n (π F E) := begin intro x, rw [cont_mdiff_at, cont_mdiff_within_at_iff'], @@ -193,47 +193,47 @@ begin apply cont_diff_within_at_fst.congr, { rintros ⟨a, b⟩ hab, simp only with mfld_simps at hab, - have : ((chart_at HB x.1).symm (IB.symm a), b) ∈ (trivialization_at F E x.fst).target, + have : ((chart_at HB x.1).symm (IB.symm a), b) ∈ (trivialization_at F E x.proj).target, { simp only [hab] with mfld_simps }, simp only [trivialization.proj_symm_apply _ this, hab] with mfld_simps }, { simp only with mfld_simps } end -lemma smooth_proj : smooth (IB.prod 𝓘(𝕜, F)) IB (π E) := +lemma smooth_proj : smooth (IB.prod 𝓘(𝕜, F)) IB (π F E) := cont_mdiff_proj E -lemma cont_mdiff_on_proj {s : set (total_space E)} : - cont_mdiff_on (IB.prod 𝓘(𝕜, F)) IB n (π E) s := +lemma cont_mdiff_on_proj {s : set (total_space F E)} : + cont_mdiff_on (IB.prod 𝓘(𝕜, F)) IB n (π F E) s := (bundle.cont_mdiff_proj E).cont_mdiff_on -lemma smooth_on_proj {s : set (total_space E)} : - smooth_on (IB.prod 𝓘(𝕜, F)) IB (π E) s := +lemma smooth_on_proj {s : set (total_space F E)} : + smooth_on (IB.prod 𝓘(𝕜, F)) IB (π F E) s := cont_mdiff_on_proj E -lemma cont_mdiff_at_proj {p : total_space E} : +lemma cont_mdiff_at_proj {p : total_space F E} : cont_mdiff_at (IB.prod 𝓘(𝕜, F)) IB n - (π E) p := + (π F E) p := (bundle.cont_mdiff_proj E).cont_mdiff_at -lemma smooth_at_proj {p : total_space E} : - smooth_at (IB.prod 𝓘(𝕜, F)) IB (π E) p := +lemma smooth_at_proj {p : total_space F E} : + smooth_at (IB.prod 𝓘(𝕜, F)) IB (π F E) p := bundle.cont_mdiff_at_proj E lemma cont_mdiff_within_at_proj - {s : set (total_space E)} - {p : total_space E} : - cont_mdiff_within_at (IB.prod 𝓘(𝕜, F)) IB n (π E) s p := + {s : set (total_space F E)} + {p : total_space F E} : + cont_mdiff_within_at (IB.prod 𝓘(𝕜, F)) IB n (π F E) s p := (bundle.cont_mdiff_at_proj E).cont_mdiff_within_at lemma smooth_within_at_proj - {s : set (total_space E)} - {p : total_space E} : - smooth_within_at (IB.prod 𝓘(𝕜, F)) IB (π E) s p := + {s : set (total_space F E)} + {p : total_space F E} : + smooth_within_at (IB.prod 𝓘(𝕜, F)) IB (π F E) s p := bundle.cont_mdiff_within_at_proj E variables (𝕜 E) [∀ x, add_comm_monoid (E x)] [∀ x, module 𝕜 (E x)] [vector_bundle 𝕜 F E] -lemma smooth_zero_section : smooth IB (IB.prod 𝓘(𝕜, F)) (zero_section E) := +lemma smooth_zero_section : smooth IB (IB.prod 𝓘(𝕜, F)) (zero_section F E) := begin intro x, rw [bundle.cont_mdiff_at_total_space], @@ -263,7 +263,7 @@ variables [nontrivially_normed_field 𝕜] section with_topology -variables [topological_space (total_space E)] [∀ x, topological_space (E x)] +variables [topological_space (total_space F E)] [∀ x, topological_space (E x)] variables (F E) [fiber_bundle F E] [vector_bundle 𝕜 F E] @@ -272,7 +272,7 @@ topological vector bundle over `B` with fibers isomorphic to `F`, then `smooth_v registers that the bundle is smooth, in the sense of having smooth transition functions. This is a mixin, not carrying any new data`. -/ class smooth_vector_bundle : Prop := -(smooth_on_coord_change : ∀ (e e' : trivialization F (π E)) +(smooth_on_coord_change : ∀ (e e' : trivialization F (π F E)) [mem_trivialization_atlas e] [mem_trivialization_atlas e'], smooth_on IB 𝓘(𝕜, F →L[𝕜] F) (λ b : B, (e.coord_changeL 𝕜 e' b : F →L[𝕜] F)) (e.base_set ∩ e'.base_set)) @@ -284,7 +284,7 @@ variables [smooth_vector_bundle F E IB] /-- For a smooth vector bundle `E` over `B` with fiber modelled on `F`, the change-of-co-ordinates between two trivializations `e`, `e'` for `E`, considered as charts to `B × F`, is smooth and fiberwise linear. -/ -instance : has_groupoid (total_space E) (smooth_fiberwise_linear B F IB) := +instance : has_groupoid (total_space F E) (smooth_fiberwise_linear B F IB) := { compatible := begin rintros _ _ ⟨e, he, rfl⟩ ⟨e', he', rfl⟩, haveI : mem_trivialization_atlas e := ⟨he⟩, @@ -308,7 +308,7 @@ instance : has_groupoid (total_space E) (smooth_fiberwise_linear B F IB) := end } /-- A smooth vector bundle `E` is naturally a smooth manifold. -/ -instance : smooth_manifold_with_corners (IB.prod 𝓘(𝕜, F)) (total_space E) := +instance : smooth_manifold_with_corners (IB.prod 𝓘(𝕜, F)) (total_space F E) := begin refine { .. structure_groupoid.has_groupoid.comp (smooth_fiberwise_linear B F IB) _ }, intros e he, @@ -366,11 +366,11 @@ instance bundle.trivial.smooth_vector_bundle : smooth_vector_bundle F (bundle.tr section prod variables (F₁ : Type*) [normed_add_comm_group F₁] [normed_space 𝕜 F₁] - (E₁ : B → Type*) [topological_space (total_space E₁)] + (E₁ : B → Type*) [topological_space (total_space F₁ E₁)] [Π x, add_comm_monoid (E₁ x)] [Π x, module 𝕜 (E₁ x)] variables (F₂ : Type*) [normed_add_comm_group F₂] [normed_space 𝕜 F₂] - (E₂ : B → Type*) [topological_space (total_space E₂)] + (E₂ : B → Type*) [topological_space (total_space F₂ E₂)] [Π x, add_comm_monoid (E₂ x)] [Π x, module 𝕜 (E₂ x)] variables [Π x : B, topological_space (E₁ x)] [Π x : B, topological_space (E₂ x)] @@ -409,10 +409,9 @@ variables [∀ x, topological_space (E x)] {F E} class is_smooth (a : vector_prebundle 𝕜 F E) : Prop := (exists_smooth_coord_change : ∀ (e e' ∈ a.pretrivialization_atlas), ∃ f : B → F →L[𝕜] F, smooth_on IB 𝓘(𝕜, F →L[𝕜] F) f (e.base_set ∩ e'.base_set) ∧ - ∀ (b : B) (hb : b ∈ e.base_set ∩ e'.base_set) (v : F), - f b v = (e' (total_space_mk b (e.symm b v))).2) + ∀ (b : B) (hb : b ∈ e.base_set ∩ e'.base_set) (v : F), f b v = (e' ⟨b ,e.symm b v⟩).2) -variables (a : vector_prebundle 𝕜 F E) [ha : a.is_smooth IB] {e e' : pretrivialization F (π E)} +variables (a : vector_prebundle 𝕜 F E) [ha : a.is_smooth IB] {e e' : pretrivialization F (π F E)} include ha /-- A randomly chosen coordinate change on a `smooth_vector_prebundle`, given by @@ -430,12 +429,12 @@ lemma smooth_on_smooth_coord_change (he : e ∈ a.pretrivialization_atlas) lemma smooth_coord_change_apply (he : e ∈ a.pretrivialization_atlas) (he' : e' ∈ a.pretrivialization_atlas) {b : B} (hb : b ∈ e.base_set ∩ e'.base_set) (v : F) : - a.smooth_coord_change IB he he' b v = (e' (total_space_mk b (e.symm b v))).2 := + a.smooth_coord_change IB he he' b v = (e' ⟨b, e.symm b v⟩).2 := (classical.some_spec (ha.exists_smooth_coord_change e he e' he')).2 b hb v lemma mk_smooth_coord_change (he : e ∈ a.pretrivialization_atlas) (he' : e' ∈ a.pretrivialization_atlas) {b : B} (hb : b ∈ e.base_set ∩ e'.base_set) (v : F) : - (b, (a.smooth_coord_change IB he he' b v)) = e' (total_space_mk b (e.symm b v)) := + (b, (a.smooth_coord_change IB he he' b v)) = e' ⟨b, e.symm b v⟩ := begin ext, { rw [e.mk_symm hb.1 v, e'.coe_fst', e.proj_symm_apply' hb.1], diff --git a/src/geometry/manifold/vector_bundle/hom.lean b/src/geometry/manifold/vector_bundle/hom.lean index e0930af6bb21f..e83f070f9d9a6 100644 --- a/src/geometry/manifold/vector_bundle/hom.lean +++ b/src/geometry/manifold/vector_bundle/hom.lean @@ -25,13 +25,13 @@ variables {𝕜 B F F₁ F₂ M M₁ M₂ : Type*} [nontrivially_normed_field 𝕜] [∀ x, add_comm_group (E x)] [∀ x, module 𝕜 (E x)] [normed_add_comm_group F] [normed_space 𝕜 F] - [topological_space (total_space E)] [∀ x, topological_space (E x)] + [topological_space (total_space F E)] [∀ x, topological_space (E x)] [∀ x, add_comm_group (E₁ x)] [∀ x, module 𝕜 (E₁ x)] [normed_add_comm_group F₁] [normed_space 𝕜 F₁] - [topological_space (total_space E₁)] [∀ x, topological_space (E₁ x)] + [topological_space (total_space F₁ E₁)] [∀ x, topological_space (E₁ x)] [∀ x, add_comm_group (E₂ x)] [∀ x, module 𝕜 (E₂ x)] [normed_add_comm_group F₂] [normed_space 𝕜 F₂] - [topological_space (total_space E₂)] [∀ x, topological_space (E₂ x)] + [topological_space (total_space F₂ E₂)] [∀ x, topological_space (E₂ x)] [_i₁ : ∀ x, topological_add_group (E₂ x)] [_i₂ : ∀ x, has_continuous_smul 𝕜 (E₂ x)] {EB : Type*} [normed_add_comm_group EB] [normed_space 𝕜 EB] @@ -43,9 +43,10 @@ variables {𝕜 B F F₁ F₂ M M₁ M₂ : Type*} {n : ℕ∞} [fiber_bundle F₁ E₁] [vector_bundle 𝕜 F₁ E₁] [fiber_bundle F₂ E₂] [vector_bundle 𝕜 F₂ E₂] - {e₁ e₁' : trivialization F₁ (π E₁)} {e₂ e₂' : trivialization F₂ (π E₂)} + {e₁ e₁' : trivialization F₁ (π F₁ E₁)} {e₂ e₂' : trivialization F₂ (π F₂ E₂)} -local notation `LE₁E₂` := total_space (bundle.continuous_linear_map (ring_hom.id 𝕜) F₁ E₁ F₂ E₂) +local notation `LE₁E₂` := total_space (F₁ →L[𝕜] F₂) + (bundle.continuous_linear_map (ring_hom.id 𝕜) E₁ E₂) /- This proof is slow, especially the `simp only` and the elaboration of `h₂`. -/ lemma smooth_on_continuous_linear_map_coord_change @@ -110,6 +111,5 @@ instance bundle.continuous_linear_map.vector_prebundle.is_smooth : end } instance smooth_vector_bundle.continuous_linear_map : - smooth_vector_bundle (F₁ →L[𝕜] F₂) (bundle.continuous_linear_map (ring_hom.id 𝕜) F₁ E₁ F₂ E₂) - IB := + smooth_vector_bundle (F₁ →L[𝕜] F₂) (bundle.continuous_linear_map (ring_hom.id 𝕜) E₁ E₂) IB := (bundle.continuous_linear_map.vector_prebundle (ring_hom.id 𝕜) F₁ E₁ F₂ E₂).smooth_vector_bundle IB diff --git a/src/geometry/manifold/vector_bundle/pullback.lean b/src/geometry/manifold/vector_bundle/pullback.lean index 44c3c92709f7c..a68fa9c52a52a 100644 --- a/src/geometry/manifold/vector_bundle/pullback.lean +++ b/src/geometry/manifold/vector_bundle/pullback.lean @@ -27,7 +27,7 @@ variables {𝕜 B B' M : Type*} (F : Type*) (E : B → Type*) variables [nontrivially_normed_field 𝕜] [∀ x, add_comm_monoid (E x)] [∀ x, module 𝕜 (E x)] [normed_add_comm_group F] [normed_space 𝕜 F] - [topological_space (total_space E)] [∀ x, topological_space (E x)] + [topological_space (total_space F E)] [∀ x, topological_space (E x)] {EB : Type*} [normed_add_comm_group EB] [normed_space 𝕜 EB] {HB : Type*} [topological_space HB] (IB : model_with_corners 𝕜 EB HB) diff --git a/src/geometry/manifold/vector_bundle/smooth_section.lean b/src/geometry/manifold/vector_bundle/smooth_section.lean index 9190ea480622e..c7782d18cc4b6 100644 --- a/src/geometry/manifold/vector_bundle/smooth_section.lean +++ b/src/geometry/manifold/vector_bundle/smooth_section.lean @@ -15,7 +15,7 @@ In this file we define the type `cont_mdiff_section` of `n` times continuously d sections of a smooth vector bundle over a manifold `M` and prove that it's a module. -/ open bundle filter function -open_locale manifold +open_locale bundle manifold variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] {E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] @@ -34,7 +34,7 @@ variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] variables (F : Type*) [normed_add_comm_group F] [normed_space 𝕜 F] -- `F` model fiber (n : ℕ∞) - (V : M → Type*) [topological_space (total_space V)] -- `V` vector bundle + (V : M → Type*) [topological_space (total_space F V)] -- `V` vector bundle [Π x, add_comm_group (V x)] [Π x, module 𝕜 (V x)] variables [Π x : M, topological_space (V x)] [fiber_bundle F V] @@ -45,7 +45,7 @@ variables [Π x : M, topological_space (V x)] @[protect_proj] structure cont_mdiff_section := (to_fun : Π x, V x) -(cont_mdiff_to_fun : cont_mdiff I (I.prod 𝓘(𝕜, F)) n (λ x, total_space_mk x (to_fun x))) +(cont_mdiff_to_fun : cont_mdiff I (I.prod 𝓘(𝕜, F)) n (λ x, total_space.mk' F x (to_fun x))) /-- Bundled smooth sections of a vector bundle. -/ @[reducible] def smooth_section := cont_mdiff_section I F ⊤ V @@ -62,26 +62,26 @@ instance : has_coe_to_fun Cₛ^n⟮I; F, V⟯ (λ s, Π x, V x) := ⟨cont_mdiff variables {s t : Cₛ^n⟮I; F, V⟯} @[simp] lemma coe_fn_mk (s : Π x, V x) - (hs : cont_mdiff I (I.prod 𝓘(𝕜, F)) n (λ x, total_space_mk x (s x))) : + (hs : cont_mdiff I (I.prod 𝓘(𝕜, F)) n (λ x, total_space.mk x (s x))) : (mk s hs : Π x, V x) = s := rfl protected lemma cont_mdiff (s : Cₛ^n⟮I; F, V⟯) : - cont_mdiff I (I.prod 𝓘(𝕜, F)) n (λ x, total_space_mk x (s x : V x)) := s.cont_mdiff_to_fun + cont_mdiff I (I.prod 𝓘(𝕜, F)) n (λ x, total_space.mk' F x (s x : V x)) := s.cont_mdiff_to_fun protected lemma smooth (s : Cₛ^∞⟮I; F, V⟯) : - smooth I (I.prod 𝓘(𝕜, F)) (λ x, total_space_mk x (s x : V x)) := s.cont_mdiff_to_fun + smooth I (I.prod 𝓘(𝕜, F)) (λ x, total_space.mk' F x (s x : V x)) := s.cont_mdiff_to_fun protected lemma mdifferentiable' (s : Cₛ^n⟮I; F, V⟯) (hn : 1 ≤ n) : - mdifferentiable I (I.prod 𝓘(𝕜, F)) (λ x, total_space_mk x (s x : V x)) := + mdifferentiable I (I.prod 𝓘(𝕜, F)) (λ x, total_space.mk' F x (s x : V x)) := s.cont_mdiff.mdifferentiable hn protected lemma mdifferentiable (s : Cₛ^∞⟮I; F, V⟯) : - mdifferentiable I (I.prod 𝓘(𝕜, F)) (λ x, total_space_mk x (s x : V x)) := + mdifferentiable I (I.prod 𝓘(𝕜, F)) (λ x, total_space.mk' F x (s x : V x)) := s.cont_mdiff.mdifferentiable le_top protected lemma mdifferentiable_at (s : Cₛ^∞⟮I; F, V⟯) {x} : - mdifferentiable_at I (I.prod 𝓘(𝕜, F)) (λ x, total_space_mk x (s x : V x)) x := + mdifferentiable_at I (I.prod 𝓘(𝕜, F)) (λ x, total_space.mk' F x (s x : V x)) x := s.mdifferentiable x lemma coe_inj ⦃s t : Cₛ^n⟮I; F, V⟯⦄ (h : (s : Π x, V x) = t) : s = t := diff --git a/src/geometry/manifold/vector_bundle/tangent.lean b/src/geometry/manifold/vector_bundle/tangent.lean index 55d597435cb96..491f7107ff786 100644 --- a/src/geometry/manifold/vector_bundle/tangent.lean +++ b/src/geometry/manifold/vector_bundle/tangent.lean @@ -136,11 +136,10 @@ variable (M) /-- The tangent bundle to a smooth manifold, as a Sigma type. Defined in terms of `bundle.total_space` to be able to put a suitable topology on it. -/ @[nolint has_nonempty_instance, reducible] -- is empty if the base manifold is empty -def tangent_bundle := bundle.total_space (tangent_space I : M → Type*) +def tangent_bundle := bundle.total_space E (tangent_space I : M → Type*) local notation `TM` := tangent_bundle I M - section tangent_bundle_instances /- In general, the definition of tangent_space is not reducible, so that type class inference @@ -187,7 +186,8 @@ rfl @[simp, mfld_simps] lemma trivialization_at_source (x : M) : - (trivialization_at E (tangent_space I) x).source = π _ ⁻¹' (chart_at H x).source := + (trivialization_at E (tangent_space I) x).source = + π E (tangent_space I) ⁻¹' (chart_at H x).source := rfl @[simp, mfld_simps] @@ -287,7 +287,7 @@ end tangent_bundle_instances /-- In the tangent bundle to the model space, the charts are just the canonical identification between a product type and a sigma type, a.k.a. `equiv.sigma_equiv_prod`. -/ @[simp, mfld_simps] lemma tangent_bundle_model_space_chart_at (p : tangent_bundle I H) : - (chart_at (model_prod H E) p).to_local_equiv = (equiv.sigma_equiv_prod H E).to_local_equiv := + (chart_at (model_prod H E) p).to_local_equiv = (total_space.to_prod H E).to_local_equiv := begin ext x : 1, { ext, { refl }, @@ -303,12 +303,12 @@ begin end @[simp, mfld_simps] lemma tangent_bundle_model_space_coe_chart_at (p : tangent_bundle I H) : - ⇑(chart_at (model_prod H E) p) = equiv.sigma_equiv_prod H E := + ⇑(chart_at (model_prod H E) p) = total_space.to_prod H E := by { unfold_coes, simp_rw [tangent_bundle_model_space_chart_at], refl } @[simp, mfld_simps] lemma tangent_bundle_model_space_coe_chart_at_symm (p : tangent_bundle I H) : ((chart_at (model_prod H E) p).symm : model_prod H E → tangent_bundle I H) = - (equiv.sigma_equiv_prod H E).symm := + (total_space.to_prod H E).symm := by { unfold_coes, simp_rw [local_homeomorph.symm_to_local_equiv, tangent_bundle_model_space_chart_at], refl } @@ -339,16 +339,16 @@ def tangent_bundle_model_space_homeomorph : tangent_bundle I H ≃ₜ model_prod simp only with mfld_simps }, simpa only with mfld_simps using this, end, - .. equiv.sigma_equiv_prod H E } + .. total_space.to_prod H E } @[simp, mfld_simps] lemma tangent_bundle_model_space_homeomorph_coe : (tangent_bundle_model_space_homeomorph H I : tangent_bundle I H → model_prod H E) - = equiv.sigma_equiv_prod H E := + = total_space.to_prod H E := rfl @[simp, mfld_simps] lemma tangent_bundle_model_space_homeomorph_coe_symm : ((tangent_bundle_model_space_homeomorph H I).symm : model_prod H E → tangent_bundle I H) - = (equiv.sigma_equiv_prod H E).symm := + = (total_space.to_prod H E).symm := rfl section in_tangent_coordinates diff --git a/src/topology/covering.lean b/src/topology/covering.lean index 58f758a2ee4ed..7abc9ef9f9938 100644 --- a/src/topology/covering.lean +++ b/src/topology/covering.lean @@ -161,7 +161,7 @@ protected lemma is_fiber_bundle.is_covering_map {F : Type*} [topological_space F is_covering_map.mk f (λ x, F) (λ x, classical.some (hf x)) (λ x, classical.some_spec (hf x)) protected lemma fiber_bundle.is_covering_map {F : Type*} {E : X → Type*} [topological_space F] - [discrete_topology F] [topological_space (bundle.total_space E)] [Π x, topological_space (E x)] - [hf : fiber_bundle F E] : is_covering_map (π E) := + [discrete_topology F] [topological_space (bundle.total_space F E)] [Π x, topological_space (E x)] + [hf : fiber_bundle F E] : is_covering_map (π F E) := is_fiber_bundle.is_covering_map (λ x, ⟨trivialization_at F E x, mem_base_set_trivialization_at F E x ⟩) diff --git a/src/topology/fiber_bundle/basic.lean b/src/topology/fiber_bundle/basic.lean index f50b7b4f4e649..e855d01461f19 100644 --- a/src/topology/fiber_bundle/basic.lean +++ b/src/topology/fiber_bundle/basic.lean @@ -16,16 +16,14 @@ Mathematically, a (topological) fiber bundle with fiber `F` over a base `B` is a point is a direct product. In our formalism, a fiber bundle is by definition the type -`bundle.total_space E` where `E : B → Type*` is a function associating to -`x : B` the fiber over `x`. This type `bundle.total_space E` is just a type synonym for -`Σ (x : B), E x`, with the interest that one can put another topology than on `Σ (x : B), E x` -which has the disjoint union topology. +`bundle.total_space F E` where `E : B → Type*` is a function associating to `x : B` the fiber over +`x`. This type `bundle.total_space F E` is a type of pairs `(proj : B, snd : E proj)`. -To have a fiber bundle structure on `bundle.total_space E`, one should +To have a fiber bundle structure on `bundle.total_space F E`, one should additionally have the following data: * `F` should be a topological space; -* There should be a topology on `bundle.total_space E`, for which the projection to `B` is +* There should be a topology on `bundle.total_space F E`, for which the projection to `B` is a fiber bundle with fiber `F` (in particular, each fiber `E x` is homeomorphic to `F`); * For each `x`, the fiber `E x` should be a topological space, and the injection from `E x` to `bundle.total_space F E` should be an embedding; @@ -52,17 +50,16 @@ fiber bundle from trivializations given as local equivalences with minimum addit ### Construction of a bundle from trivializations -* `bundle.total_space E` is a type synonym for `Σ (x : B), E x`, that we can endow with a suitable - topology. +* `bundle.total_space F E` is the type of pairs `(proj : B, snd : E proj)`. We can use the extra + argument `F` to construct topology on the total space. * `fiber_bundle_core ι B F` : structure registering how changes of coordinates act on the fiber `F` above open subsets of `B`, where local trivializations are indexed by `ι`. Let `Z : fiber_bundle_core ι B F`. Then we define * `Z.fiber x` : the fiber above `x`, homeomorphic to `F` (and defeq to `F` as a type). -* `Z.total_space` : the total space of `Z`, defined as a `Type` as `Σ (b : B), F`, but with a - twisted topology coming from the fiber bundle structure. It is (reducibly) the same as - `bundle.total_space Z.fiber`. +* `Z.total_space` : the total space of `Z`, defined as a `Type*` as `bundle.total_space F Z.fiber` + with a custom topology. * `Z.proj` : projection from `Z.total_space` to `B`. It is continuous. * `Z.local_triv i`: for `i : ι`, bundle trivialization above the set `Z.base_set i`, which is an open set in `B`. @@ -153,8 +150,8 @@ choose for each `x` one specific trivialization around it. We include this choic of the `fiber_bundle_core`, as it makes some constructions more functorial and it is a nice way to say that the trivializations cover the whole space `B`. -With this definition, the type of the fiber bundle space constructed from the core data is just -`Σ (b : B), F `, but the topology is not the product one, in general. +With this definition, the type of the fiber bundle space constructed from the core data is +`bundle.total_space F (λ b : B, F)`, but the topology is not the product one, in general. We also take the indexing type (indexing all the trivializations) as a parameter to the fiber bundle core: it could always be taken as a subtype of all the maps from open subsets of `B` to continuous @@ -172,7 +169,7 @@ variables {ι B F X : Type*} [topological_space X] open topological_space filter set bundle open_locale topology classical bundle -attribute [mfld_simps] total_space_mk coe_fst coe_snd coe_snd_map_apply +attribute [mfld_simps] total_space.coe_proj total_space.coe_snd coe_snd_map_apply coe_snd_map_smul total_space.mk_cast /-! ### General definition of fiber bundles -/ @@ -180,15 +177,15 @@ attribute [mfld_simps] total_space_mk coe_fst coe_snd coe_snd_map_apply section fiber_bundle variables (F) [topological_space B] [topological_space F] (E : B → Type*) - [topological_space (total_space E)] [∀ b, topological_space (E b)] + [topological_space (total_space F E)] [∀ b, topological_space (E b)] /-- A (topological) fiber bundle with fiber `F` over a base `B` is a space projecting on `B` for which the fibers are all homeomorphic to `F`, such that the local situation around each point is a direct product. -/ class fiber_bundle := -(total_space_mk_inducing [] : ∀ (b : B), inducing (@total_space_mk B E b)) -(trivialization_atlas [] : set (trivialization F (π E))) -(trivialization_at [] : B → trivialization F (π E)) +(total_space_mk_inducing [] : ∀ (b : B), inducing (@total_space.mk B F E b)) +(trivialization_atlas [] : set (trivialization F (π F E))) +(trivialization_at [] : B → trivialization F (π F E)) (mem_base_set_trivialization_at [] : ∀ b : B, b ∈ (trivialization_at b).base_set) (trivialization_mem_atlas [] : ∀ b : B, trivialization_at b ∈ trivialization_atlas) @@ -202,7 +199,7 @@ bundle. This is needed because lemmas about the linearity of trivializations or functions to `F →L[R] F`, where `F` is the model fiber) of the transition functions are only expected to hold for trivializations in the designated atlas. -/ @[mk_iff] -class mem_trivialization_atlas [fiber_bundle F E] (e : trivialization F (π E)) : Prop := +class mem_trivialization_atlas [fiber_bundle F E] (e : trivialization F (π F E)) : Prop := (out : e ∈ trivialization_atlas F E) instance [fiber_bundle F E] (b : B) : mem_trivialization_atlas (trivialization_at F E b) := @@ -211,44 +208,44 @@ instance [fiber_bundle F E] (b : B) : mem_trivialization_atlas (trivialization_a namespace fiber_bundle variables (F) {E} [fiber_bundle F E] -lemma map_proj_nhds (x : total_space E) : - map (π E) (𝓝 x) = 𝓝 x.proj := +lemma map_proj_nhds (x : total_space F E) : + map (π F E) (𝓝 x) = 𝓝 x.proj := (trivialization_at F E x.proj).map_proj_nhds $ (trivialization_at F E x.proj).mem_source.2 $ mem_base_set_trivialization_at F E x.proj variables (E) /-- The projection from a fiber bundle to its base is continuous. -/ -@[continuity] lemma continuous_proj : continuous (π E) := +@[continuity] lemma continuous_proj : continuous (π F E) := continuous_iff_continuous_at.2 $ λ x, (map_proj_nhds F x).le /-- The projection from a fiber bundle to its base is an open map. -/ -lemma is_open_map_proj : is_open_map (π E) := +lemma is_open_map_proj : is_open_map (π F E) := is_open_map.of_nhds_le $ λ x, (map_proj_nhds F x).ge /-- The projection from a fiber bundle with a nonempty fiber to its base is a surjective map. -/ -lemma surjective_proj [nonempty F] : function.surjective (π E) := +lemma surjective_proj [nonempty F] : function.surjective (π F E) := λ b, let ⟨p, _, hpb⟩ := (trivialization_at F E b).proj_surj_on_base_set (mem_base_set_trivialization_at F E b) in ⟨p, hpb⟩ /-- The projection from a fiber bundle with a nonempty fiber to its base is a quotient map. -/ -lemma quotient_map_proj [nonempty F] : quotient_map (π E) := +lemma quotient_map_proj [nonempty F] : quotient_map (π F E) := (is_open_map_proj F E).to_quotient_map (continuous_proj F E) (surjective_proj F E) -lemma continuous_total_space_mk (x : B) : continuous (@total_space_mk B E x) := +lemma continuous_total_space_mk (x : B) : continuous (@total_space.mk B F E x) := (total_space_mk_inducing F E x).continuous variables {E F} @[simp, mfld_simps] -lemma mem_trivialization_at_proj_source {x : total_space E} : +lemma mem_trivialization_at_proj_source {x : total_space F E} : x ∈ (trivialization_at F E x.proj).source := (trivialization.mem_source _).mpr $ mem_base_set_trivialization_at F E x.proj @[simp, mfld_simps] -lemma trivialization_at_proj_fst {x : total_space E} : +lemma trivialization_at_proj_fst {x : total_space F E} : ((trivialization_at F E x.proj) x).1 = x.proj := trivialization.coe_fst' _ $ mem_base_set_trivialization_at F E x.proj @@ -256,7 +253,7 @@ variable (F) open trivialization /-- Characterization of continuous functions (at a point, within a set) into a fiber bundle. -/ -lemma continuous_within_at_total_space (f : X → total_space E) {s : set X} {x₀ : X} : +lemma continuous_within_at_total_space (f : X → total_space F E) {s : set X} {x₀ : X} : continuous_within_at f s x₀ ↔ continuous_within_at (λ x, (f x).proj) s x₀ ∧ continuous_within_at (λ x, ((trivialization_at F E (f x₀).proj) (f x)).2) s x₀ := @@ -276,7 +273,7 @@ begin end /-- Characterization of continuous functions (at a point) into a fiber bundle. -/ -lemma continuous_at_total_space (f : X → total_space E) {x₀ : X} : +lemma continuous_at_total_space (f : X → total_space F E) {x₀ : X} : continuous_at f x₀ ↔ continuous_at (λ x, (f x).proj) x₀ ∧ continuous_at (λ x, ((trivialization_at F E (f x₀).proj) (f x)).2) x₀ := by { simp_rw [← continuous_within_at_univ], exact continuous_within_at_total_space F f } @@ -289,16 +286,16 @@ variables (F E) then it is trivial over any closed interval. -/ lemma fiber_bundle.exists_trivialization_Icc_subset [conditionally_complete_linear_order B] [order_topology B] [fiber_bundle F E] (a b : B) : - ∃ e : trivialization F (π E), Icc a b ⊆ e.base_set := + ∃ e : trivialization F (π F E), Icc a b ⊆ e.base_set := begin classical, - obtain ⟨ea, hea⟩ : ∃ ea : trivialization F (π E), a ∈ ea.base_set := + obtain ⟨ea, hea⟩ : ∃ ea : trivialization F (π F E), a ∈ ea.base_set := ⟨trivialization_at F E a, mem_base_set_trivialization_at F E a⟩, -- If `a < b`, then `[a, b] = ∅`, and the statement is trivial cases le_or_lt a b with hab hab; [skip, exact ⟨ea, by simp *⟩], /- Let `s` be the set of points `x ∈ [a, b]` such that `E` is trivializable over `[a, x]`. We need to show that `b ∈ s`. Let `c = Sup s`. We will show that `c ∈ s` and `c = b`. -/ - set s : set B := {x ∈ Icc a b | ∃ e : trivialization F (π E), Icc a x ⊆ e.base_set}, + set s : set B := {x ∈ Icc a b | ∃ e : trivialization F (π F E), Icc a x ⊆ e.base_set}, have ha : a ∈ s, from ⟨left_mem_Icc.2 hab, ea, by simp [hea]⟩, have sne : s.nonempty := ⟨a, ha⟩, have hsb : b ∈ upper_bounds s, from λ x hx, hx.1.2, @@ -306,12 +303,12 @@ begin set c := Sup s, have hsc : is_lub s c, from is_lub_cSup sne sbd, have hc : c ∈ Icc a b, from ⟨hsc.1 ha, hsc.2 hsb⟩, - obtain ⟨-, ec : trivialization F (π E), hec : Icc a c ⊆ ec.base_set⟩ : c ∈ s, + obtain ⟨-, ec : trivialization F (π F E), hec : Icc a c ⊆ ec.base_set⟩ : c ∈ s, { cases hc.1.eq_or_lt with heq hlt, { rwa ← heq }, refine ⟨hc, _⟩, /- In order to show that `c ∈ s`, consider a trivialization `ec` of `proj` over a neighborhood of `c`. Its base set includes `(c', c]` for some `c' ∈ [a, c)`. -/ - obtain ⟨ec, hc⟩ : ∃ ec : trivialization F (π E), c ∈ ec.base_set := + obtain ⟨ec, hc⟩ : ∃ ec : trivialization F (π F E), c ∈ ec.base_set := ⟨trivialization_at F E c, mem_base_set_trivialization_at F E c⟩, obtain ⟨c', hc', hc'e⟩ : ∃ c' ∈ Ico a c, Ioc c' c ⊆ ec.base_set := (mem_nhds_within_Iic_iff_exists_mem_Ico_Ioc_subset hlt).1 @@ -325,7 +322,7 @@ begin done. Otherwise we show that `proj` can be trivialized over a larger interval `[a, d]`, `d ∈ (c, b]`, hence `c` is not an upper bound of `s`. -/ cases hc.2.eq_or_lt with heq hlt, { exact ⟨ec, heq ▸ hec⟩ }, - rsuffices ⟨d, hdcb, hd⟩ : ∃ (d ∈ Ioc c b) (e : trivialization F (π E)), Icc a d ⊆ e.base_set, + rsuffices ⟨d, hdcb, hd⟩ : ∃ (d ∈ Ioc c b) (e : trivialization F (π F E)), Icc a d ⊆ e.base_set, { exact ((hsc.1 ⟨⟨hc.1.trans hdcb.1.le, hdcb.2⟩, hd⟩).not_lt hdcb.1).elim }, /- Since the base set of `ec` is open, it includes `[c, d)` (hence, `[a, d)`) for some `d ∈ (c, b]`. -/ @@ -338,7 +335,7 @@ begin { /- If `(c, d) = ∅`, then let `ed` be a trivialization of `proj` over a neighborhood of `d`. Then the disjoint union of `ec` restricted to `(-∞, d)` and `ed` restricted to `(c, ∞)` is a trivialization over `[a, d]`. -/ - obtain ⟨ed, hed⟩ : ∃ ed : trivialization F (π E), d ∈ ed.base_set := + obtain ⟨ed, hed⟩ : ∃ ed : trivialization F (π F E), d ∈ ed.base_set := ⟨trivialization_at F E d, mem_base_set_trivialization_at F E d⟩, refine ⟨d, hdcb, (ec.restr_open (Iio d) is_open_Iio).disjoint_union (ed.restr_open (Ioi c) is_open_Ioi) (he.mono (inter_subset_right _ _) @@ -396,18 +393,13 @@ typeclass inference -/ @[nolint unused_arguments has_nonempty_instance] def fiber (x : B) := F -section fiber_instances -local attribute [reducible] fiber - -instance topological_space_fiber (x : B) : topological_space (Z.fiber x) := by apply_instance - -end fiber_instances +instance topological_space_fiber (x : B) : topological_space (Z.fiber x) := +‹topological_space F› /-- The total space of the fiber bundle, as a convenience function for dot notation. -It is by definition equal to `bundle.total_space Z.fiber`, a.k.a. `Σ x, Z.fiber x` but with a -different name for typeclass inference. -/ +It is by definition equal to `bundle.total_space Z.fiber` -/ @[nolint unused_arguments, reducible] -def total_space := bundle.total_space Z.fiber +def total_space := bundle.total_space F Z.fiber /-- The projection from the total space of a fiber bundle core, on its base. -/ @[reducible, simp, mfld_simps] def proj : Z.total_space → B := bundle.total_space.proj @@ -510,7 +502,7 @@ end /-- Topological structure on the total space of a fiber bundle created from core, designed so that all the local trivialization are continuous. -/ -instance to_topological_space : topological_space (bundle.total_space Z.fiber) := +instance to_topological_space : topological_space Z.total_space := topological_space.generate_from $ ⋃ (i : ι) (s : set (B × F)) (s_open : is_open s), {(Z.local_triv_as_local_equiv i).source ∩ (Z.local_triv_as_local_equiv i) ⁻¹' s} @@ -567,7 +559,7 @@ def local_triv (i : ι) : trivialization F Z.proj := /-- Preferred local trivialization of a fiber bundle constructed from core, at a given point, as a bundle trivialization -/ -def local_triv_at (b : B) : trivialization F (π Z.fiber) := +def local_triv_at (b : B) : trivialization F (π F Z.fiber) := Z.local_triv (Z.index_at b) @[simp, mfld_simps] lemma local_triv_at_def (b : B) : @@ -647,11 +639,11 @@ by { rw [local_triv_at, ←base_set_at], exact Z.mem_base_set_at b, } /-- The inclusion of a fiber into the total space is a continuous map. -/ @[continuity] lemma continuous_total_space_mk (b : B) : - continuous (total_space_mk b : Z.fiber b → bundle.total_space Z.fiber) := + continuous (total_space.mk b : Z.fiber b → Z.total_space) := begin rw [continuous_iff_le_induced, fiber_bundle_core.to_topological_space], apply le_induced_generate_from, - simp only [total_space_mk, mem_Union, mem_singleton_iff, local_triv_as_local_equiv_source, + simp only [mem_Union, mem_singleton_iff, local_triv_as_local_equiv_source, local_triv_as_local_equiv_coe], rintros s ⟨i, t, ht, rfl⟩, rw [←((Z.local_triv i).source_inter_preimage_target_inter t), preimage_inter, ←preimage_comp, @@ -683,7 +675,6 @@ instance fiber_bundle : fiber_bundle F Z.fiber := (Z.local_triv_at b).open_source).mp (Z.local_triv_at b).continuous_to_fun _ ((Z.local_triv_at b).open_base_set.prod h), _⟩, rw [preimage_inter, ←preimage_comp, function.comp], - simp only [total_space_mk], refine ext_iff.mpr (λ a, ⟨λ ha, _, λ ha, ⟨Z.mem_base_set_at b, _⟩⟩), { simp only [mem_prod, mem_preimage, mem_inter_iff, local_triv_at_apply_mk] at ha, exact ha.2.2, }, @@ -713,21 +704,21 @@ topology in such a way that there is a fiber bundle structure for which the loca are also local homeomorphism and hence local trivializations. -/ @[nolint has_nonempty_instance] structure fiber_prebundle := -(pretrivialization_atlas : set (pretrivialization F (π E))) -(pretrivialization_at : B → pretrivialization F (π E)) +(pretrivialization_atlas : set (pretrivialization F (π F E))) +(pretrivialization_at : B → pretrivialization F (π F E)) (mem_base_pretrivialization_at : ∀ x : B, x ∈ (pretrivialization_at x).base_set) (pretrivialization_mem_atlas : ∀ x : B, pretrivialization_at x ∈ pretrivialization_atlas) (continuous_triv_change : ∀ e e' ∈ pretrivialization_atlas, continuous_on (e ∘ e'.to_local_equiv.symm) (e'.target ∩ (e'.to_local_equiv.symm ⁻¹' e.source))) -(total_space_mk_inducing : ∀ (b : B), inducing ((pretrivialization_at b) ∘ (total_space_mk b))) +(total_space_mk_inducing : ∀ (b : B), inducing ((pretrivialization_at b) ∘ (total_space.mk b))) namespace fiber_prebundle -variables {F E} (a : fiber_prebundle F E) {e : pretrivialization F (π E)} +variables {F E} (a : fiber_prebundle F E) {e : pretrivialization F (π F E)} /-- Topology on the total space that will make the prebundle into a bundle. -/ -def total_space_topology (a : fiber_prebundle F E) : topological_space (total_space E) := -⨆ (e : pretrivialization F (π E)) (he : e ∈ a.pretrivialization_atlas), +def total_space_topology (a : fiber_prebundle F E) : topological_space (total_space F E) := +⨆ (e : pretrivialization F (π F E)) (he : e ∈ a.pretrivialization_atlas), coinduced e.set_symm (subtype.topological_space) lemma continuous_symm_of_mem_pretrivialization_atlas (he : e ∈ a.pretrivialization_atlas) : @@ -739,7 +730,7 @@ begin exact le_supr₂ e he, end -lemma is_open_source (e : pretrivialization F (π E)) : is_open[a.total_space_topology] e.source := +lemma is_open_source (e : pretrivialization F (π F E)) : is_open[a.total_space_topology] e.source := begin letI := a.total_space_topology, refine is_open_supr_iff.mpr (λ e', _), @@ -751,7 +742,7 @@ begin pretrivialization.preimage_symm_proj_inter], end -lemma is_open_target_of_mem_pretrivialization_atlas_inter (e e' : pretrivialization F (π E)) +lemma is_open_target_of_mem_pretrivialization_atlas_inter (e e' : pretrivialization F (π F E)) (he' : e' ∈ a.pretrivialization_atlas) : is_open (e'.to_local_equiv.target ∩ e'.to_local_equiv.symm ⁻¹' e.source) := begin @@ -764,7 +755,7 @@ end /-- Promotion from a `pretrivialization` to a `trivialization`. -/ def trivialization_of_mem_pretrivialization_atlas (he : e ∈ a.pretrivialization_atlas) : - @trivialization B F _ _ _ a.total_space_topology (π E) := + @trivialization B F _ _ _ a.total_space_topology (π F E) := { open_source := a.is_open_source e, continuous_to_fun := begin letI := a.total_space_topology, @@ -785,14 +776,14 @@ def trivialization_of_mem_pretrivialization_atlas (he : e ∈ a.pretrivializatio .. e } lemma mem_trivialization_at_source (b : B) (x : E b) : - total_space_mk b x ∈ (a.pretrivialization_at b).source := + total_space.mk b x ∈ (a.pretrivialization_at b).source := begin simp only [(a.pretrivialization_at b).source_eq, mem_preimage, total_space.proj], exact a.mem_base_pretrivialization_at b, end @[simp] lemma total_space_mk_preimage_source (b : B) : - (total_space_mk b) ⁻¹' (a.pretrivialization_at b).source = univ := + total_space.mk b ⁻¹' (a.pretrivialization_at b).source = univ := begin apply eq_univ_of_univ_subset, rw [(a.pretrivialization_at b).source_eq, ←preimage_comp, function.comp], @@ -802,7 +793,7 @@ begin end @[continuity] lemma continuous_total_space_mk (b : B) : - @continuous _ _ _ a.total_space_topology (total_space_mk b) := + @continuous _ _ _ a.total_space_topology (total_space.mk b) := begin letI := a.total_space_topology, let e := a.trivialization_of_mem_pretrivialization_atlas (a.pretrivialization_mem_atlas b), @@ -812,8 +803,8 @@ begin end lemma inducing_total_space_mk_of_inducing_comp (b : B) - (h : inducing ((a.pretrivialization_at b) ∘ (total_space_mk b))) : - @inducing _ _ _ a.total_space_topology (total_space_mk b) := + (h : inducing ((a.pretrivialization_at b) ∘ (total_space.mk b))) : + @inducing _ _ _ a.total_space_topology (total_space.mk b) := begin letI := a.total_space_topology, rw ←restrict_comp_cod_restrict (a.mem_trivialization_at_source b) at h, @@ -841,7 +832,7 @@ def to_fiber_bundle : mem_base_set_trivialization_at := a.mem_base_pretrivialization_at, trivialization_mem_atlas := λ x, ⟨_, a.pretrivialization_mem_atlas x, rfl⟩ } -lemma continuous_proj : @continuous _ _ a.total_space_topology _ (π E) := +lemma continuous_proj : @continuous _ _ a.total_space_topology _ (π F E) := begin letI := a.total_space_topology, letI := a.to_fiber_bundle, @@ -849,17 +840,17 @@ begin end /-- For a fiber bundle `E` over `B` constructed using the `fiber_prebundle` mechanism, -continuity of a function `total_space E → X` on an open set `s` can be checked by precomposing at +continuity of a function `total_space F E → X` on an open set `s` can be checked by precomposing at each point with the pretrivialization used for the construction at that point. -/ -lemma continuous_on_of_comp_right {X : Type*} [topological_space X] {f : total_space E → X} +lemma continuous_on_of_comp_right {X : Type*} [topological_space X] {f : total_space F E → X} {s : set B} (hs : is_open s) (hf : ∀ b ∈ s, continuous_on (f ∘ (a.pretrivialization_at b).to_local_equiv.symm) ((s ∩ (a.pretrivialization_at b).base_set) ×ˢ (set.univ : set F))) : - @continuous_on _ _ a.total_space_topology _ f ((π E) ⁻¹' s) := + @continuous_on _ _ a.total_space_topology _ f ((π F E) ⁻¹' s) := begin letI := a.total_space_topology, intros z hz, - let e : trivialization F (π E) := + let e : trivialization F (π F E) := a.trivialization_of_mem_pretrivialization_atlas (a.pretrivialization_mem_atlas z.proj), refine (e.continuous_at_of_comp_right _ ((hf z.proj hz).continuous_at (is_open.mem_nhds _ _))).continuous_within_at, diff --git a/src/topology/fiber_bundle/constructions.lean b/src/topology/fiber_bundle/constructions.lean index 4607bf248af1c..a4d859c6166d3 100644 --- a/src/topology/fiber_bundle/constructions.lean +++ b/src/topology/fiber_bundle/constructions.lean @@ -37,23 +37,21 @@ namespace trivial variables (B : Type*) (F : Type*) -instance [I : topological_space F] : ∀ x : B, topological_space (trivial B F x) := λ x, I - instance [t₁ : topological_space B] [t₂ : topological_space F] : - topological_space (total_space (trivial B F)) := -induced total_space.proj t₁ ⊓ induced (trivial.proj_snd B F) t₂ + topological_space (total_space F (trivial B F)) := +induced total_space.proj t₁ ⊓ induced (total_space.trivial_snd B F) t₂ variables [topological_space B] [topological_space F] /-- Local trivialization for trivial bundle. -/ -def trivialization : trivialization F (π (bundle.trivial B F)) := -{ to_fun := λ x, (x.fst, x.snd), +def trivialization : trivialization F (π F (λ _ : B, F)) := +{ to_fun := λ x, (x.proj, x.snd), inv_fun := λ y, ⟨y.fst, y.snd⟩, source := univ, target := univ, - map_source' := λ x h, mem_univ (x.fst, x.snd), - map_target' := λ y h, mem_univ ⟨y.fst, y.snd⟩, - left_inv' := λ x h, sigma.eq rfl rfl, + map_source' := λ x h, mem_univ _, + map_target' := λ y h, mem_univ _, + left_inv' := λ x h, total_space.ext _ _ rfl heq.rfl, right_inv' := λ x h, prod.ext rfl rfl, open_source := is_open_univ, open_target := is_open_univ, @@ -83,11 +81,10 @@ instance fiber_bundle : fiber_bundle F (bundle.trivial B F) := total_space_mk_inducing := λ b, ⟨begin have : (λ (x : trivial B F b), x) = @id F, by { ext x, refl }, simp only [total_space.topological_space, induced_inf, induced_compose, function.comp, - total_space.proj, induced_const, top_inf_eq, trivial.proj_snd, id.def, - trivial.topological_space, this, induced_id], + induced_const, top_inf_eq, total_space.trivial_snd, id.def, this, induced_id], end⟩ } -lemma eq_trivialization (e : _root_.trivialization F (π (bundle.trivial B F))) +lemma eq_trivialization (e : _root_.trivialization F (π F (bundle.trivial B F))) [i : mem_trivialization_atlas e] : e = trivialization B F := i.out @@ -102,21 +99,22 @@ section prod variables {B : Type*} section defs -variables (E₁ : B → Type*) (E₂ : B → Type*) -variables [topological_space (total_space E₁)] [topological_space (total_space E₂)] +variables (F₁ : Type*) (E₁ : B → Type*) (F₂ : Type*) (E₂ : B → Type*) +variables [topological_space (total_space F₁ E₁)] [topological_space (total_space F₂ E₂)] /-- Equip the total space of the fiberwise product of two fiber bundles `E₁`, `E₂` with the induced topology from the diagonal embedding into `total_space E₁ × total_space E₂`. -/ -instance fiber_bundle.prod.topological_space : topological_space (total_space (E₁ ×ᵇ E₂)) := +instance fiber_bundle.prod.topological_space : + topological_space (total_space (F₁ × F₂) (E₁ ×ᵇ E₂)) := topological_space.induced - (λ p, ((⟨p.1, p.2.1⟩ : total_space E₁), (⟨p.1, p.2.2⟩ : total_space E₂))) - (by apply_instance : topological_space (total_space E₁ × total_space E₂)) + (λ p, ((⟨p.1, p.2.1⟩ : total_space F₁ E₁), (⟨p.1, p.2.2⟩ : total_space F₂ E₂))) + (by apply_instance : topological_space (total_space F₁ E₁ × total_space F₂ E₂)) /-- The diagonal map from the total space of the fiberwise product of two fiber bundles `E₁`, `E₂` into `total_space E₁ × total_space E₂` is `inducing`. -/ lemma fiber_bundle.prod.inducing_diag : inducing (λ p, (⟨p.1, p.2.1⟩, ⟨p.1, p.2.2⟩) : - total_space (E₁ ×ᵇ E₂) → total_space E₁ × total_space E₂) := + total_space (F₁ × F₂) (E₁ ×ᵇ E₂) → total_space F₁ E₁ × total_space F₂ E₂) := ⟨rfl⟩ end defs @@ -124,28 +122,28 @@ end defs open fiber_bundle variables [topological_space B] - (F₁ : Type*) [topological_space F₁] (E₁ : B → Type*) [topological_space (total_space E₁)] - (F₂ : Type*) [topological_space F₂] (E₂ : B → Type*) [topological_space (total_space E₂)] + (F₁ : Type*) [topological_space F₁] (E₁ : B → Type*) [topological_space (total_space F₁ E₁)] + (F₂ : Type*) [topological_space F₂] (E₂ : B → Type*) [topological_space (total_space F₂ E₂)] namespace trivialization -variables {F₁ E₁ F₂ E₂} (e₁ : trivialization F₁ (π E₁)) (e₂ : trivialization F₂ (π E₂)) +variables {F₁ E₁ F₂ E₂} (e₁ : trivialization F₁ (π F₁ E₁)) (e₂ : trivialization F₂ (π F₂ E₂)) /-- Given trivializations `e₁`, `e₂` for fiber bundles `E₁`, `E₂` over a base `B`, the forward function for the construction `trivialization.prod`, the induced trivialization for the fiberwise product of `E₁` and `E₂`. -/ -def prod.to_fun' : total_space (E₁ ×ᵇ E₂) → B × (F₁ × F₂) := +def prod.to_fun' : total_space (F₁ × F₂) (E₁ ×ᵇ E₂) → B × (F₁ × F₂) := λ p, ⟨p.1, (e₁ ⟨p.1, p.2.1⟩).2, (e₂ ⟨p.1, p.2.2⟩).2⟩ variables {e₁ e₂} -lemma prod.continuous_to_fun : continuous_on (prod.to_fun' e₁ e₂) - (@total_space.proj B (E₁ ×ᵇ E₂) ⁻¹' (e₁.base_set ∩ e₂.base_set)) := +lemma prod.continuous_to_fun : + continuous_on (prod.to_fun' e₁ e₂) (π (F₁ × F₂) (E₁ ×ᵇ E₂) ⁻¹' (e₁.base_set ∩ e₂.base_set)) := begin - let f₁ : total_space (E₁ ×ᵇ E₂) → total_space E₁ × total_space E₂ := - λ p, ((⟨p.1, p.2.1⟩ : total_space E₁), (⟨p.1, p.2.2⟩ : total_space E₂)), - let f₂ : total_space E₁ × total_space E₂ → (B × F₁) × (B × F₂) := λ p, ⟨e₁ p.1, e₂ p.2⟩, + let f₁ : total_space (F₁ × F₂) (E₁ ×ᵇ E₂) → total_space F₁ E₁ × total_space F₂ E₂ := + λ p, ((⟨p.1, p.2.1⟩ : total_space F₁ E₁), (⟨p.1, p.2.2⟩ : total_space F₂ E₂)), + let f₂ : total_space F₁ E₁ × total_space F₂ E₂ → (B × F₁) × (B × F₂) := λ p, ⟨e₁ p.1, e₂ p.2⟩, let f₃ : (B × F₁) × (B × F₂) → B × F₁ × F₂ := λ p, ⟨p.1.1, p.1.2, p.2.2⟩, - have hf₁ : continuous f₁ := (prod.inducing_diag E₁ E₂).continuous, + have hf₁ : continuous f₁ := (prod.inducing_diag F₁ E₁ F₂ E₂).continuous, have hf₂ : continuous_on f₂ (e₁.source ×ˢ e₂.source) := e₁.to_local_homeomorph.continuous_on.prod_map e₂.to_local_homeomorph.continuous_on, have hf₃ : continuous f₃ := @@ -165,18 +163,19 @@ variables (e₁ e₂) [Π x, has_zero (E₁ x)] [∀ x, has_zero (E₂ x)] /-- Given trivializations `e₁`, `e₂` for fiber bundles `E₁`, `E₂` over a base `B`, the inverse function for the construction `trivialization.prod`, the induced trivialization for the fiberwise product of `E₁` and `E₂`. -/ -noncomputable def prod.inv_fun' (p : B × (F₁ × F₂)) : total_space (E₁ ×ᵇ E₂) := +noncomputable def prod.inv_fun' (p : B × (F₁ × F₂)) : total_space (F₁ × F₂) (E₁ ×ᵇ E₂) := ⟨p.1, e₁.symm p.1 p.2.1, e₂.symm p.1 p.2.2⟩ variables {e₁ e₂} -lemma prod.left_inv {x : total_space (E₁ ×ᵇ E₂)} - (h : x ∈ @total_space.proj B (E₁ ×ᵇ E₂) ⁻¹' (e₁.base_set ∩ e₂.base_set)) : +lemma prod.left_inv {x : total_space (F₁ × F₂) (E₁ ×ᵇ E₂)} + (h : x ∈ π (F₁ × F₂) (E₁ ×ᵇ E₂) ⁻¹' (e₁.base_set ∩ e₂.base_set)) : prod.inv_fun' e₁ e₂ (prod.to_fun' e₁ e₂ x) = x := begin obtain ⟨x, v₁, v₂⟩ := x, obtain ⟨h₁ : x ∈ e₁.base_set, h₂ : x ∈ e₂.base_set⟩ := h, - simp only [prod.to_fun', prod.inv_fun', symm_apply_apply_mk, h₁, h₂] + simp only [prod.to_fun', prod.inv_fun', symm_apply_apply_mk, h₁, h₂, + eq_self_iff_true, heq_iff_eq, and_self] end lemma prod.right_inv {x : B × F₁ × F₂} @@ -191,22 +190,22 @@ end lemma prod.continuous_inv_fun : continuous_on (prod.inv_fun' e₁ e₂) ((e₁.base_set ∩ e₂.base_set) ×ˢ univ) := begin - rw (prod.inducing_diag E₁ E₂).continuous_on_iff, + rw (prod.inducing_diag F₁ E₁ F₂ E₂).continuous_on_iff, have H₁ : continuous (λ p : B × F₁ × F₂, ((p.1, p.2.1), (p.1, p.2.2))) := (continuous_id.prod_map continuous_fst).prod_mk (continuous_id.prod_map continuous_snd), refine (e₁.continuous_on_symm.prod_map e₂.continuous_on_symm).comp H₁.continuous_on _, exact λ x h, ⟨⟨h.1.1, mem_univ _⟩, ⟨h.1.2, mem_univ _⟩⟩ end -variables (e₁ e₂ e₁ e₂) +variables (e₁ e₂) /-- Given trivializations `e₁`, `e₂` for bundle types `E₁`, `E₂` over a base `B`, the induced trivialization for the fiberwise product of `E₁` and `E₂`, whose base set is `e₁.base_set ∩ e₂.base_set`. -/ -noncomputable def prod : trivialization (F₁ × F₂) (π (E₁ ×ᵇ E₂)) := +noncomputable def prod : trivialization (F₁ × F₂) (π (F₁ × F₂) (E₁ ×ᵇ E₂)) := { to_fun := prod.to_fun' e₁ e₂, inv_fun := prod.inv_fun' e₁ e₂, - source := (@total_space.proj B (E₁ ×ᵇ E₂)) ⁻¹' (e₁.base_set ∩ e₂.base_set), + source := (π (F₁ × F₂) (E₁ ×ᵇ E₂)) ⁻¹' (e₁.base_set ∩ e₂.base_set), target := (e₁.base_set ∩ e₂.base_set) ×ˢ set.univ, map_source' := λ x h, ⟨h, set.mem_univ _⟩, map_target' := λ x h, h.1, @@ -214,7 +213,7 @@ noncomputable def prod : trivialization (F₁ × F₂) (π (E₁ ×ᵇ E₂)) := right_inv' := λ x, prod.right_inv, open_source := begin convert (e₁.open_source.prod e₂.open_source).preimage - (fiber_bundle.prod.inducing_diag E₁ E₂).continuous, + (fiber_bundle.prod.inducing_diag F₁ E₁ F₂ E₂).continuous, ext x, simp only [trivialization.source_eq] with mfld_simps, end, @@ -230,8 +229,8 @@ noncomputable def prod : trivialization (F₁ × F₂) (π (E₁ ×ᵇ E₂)) := @[simp] lemma base_set_prod : (prod e₁ e₂).base_set = e₁.base_set ∩ e₂.base_set := rfl -lemma prod_symm_apply (x : B) (w₁ : F₁) (w₂ : F₂) : (prod e₁ e₂).to_local_equiv.symm (x, w₁, w₂) - = ⟨x, e₁.symm x w₁, e₂.symm x w₂⟩ := +lemma prod_symm_apply (x : B) (w₁ : F₁) (w₂ : F₂) : + (prod e₁ e₂).to_local_equiv.symm (x, w₁, w₂) = ⟨x, e₁.symm x w₁, e₂.symm x w₂⟩ := rfl end trivialization @@ -246,11 +245,11 @@ variables [Π x, has_zero (E₁ x)] [∀ x, has_zero (E₂ x)] noncomputable instance fiber_bundle.prod : fiber_bundle (F₁ × F₂) (E₁ ×ᵇ E₂) := { total_space_mk_inducing := λ b, begin - rw (prod.inducing_diag E₁ E₂).inducing_iff, + rw (prod.inducing_diag F₁ E₁ F₂ E₂).inducing_iff, exact (total_space_mk_inducing F₁ E₁ b).prod_mk (total_space_mk_inducing F₂ E₂ b), end, trivialization_atlas := - {e | ∃ (e₁ : trivialization F₁ (π E₁)) (e₂ : trivialization F₂ (π E₂)) + {e | ∃ (e₁ : trivialization F₁ (π F₁ E₁)) (e₂ : trivialization F₂ (π F₂ E₂)) [mem_trivialization_atlas e₁] [mem_trivialization_atlas e₂], by exactI e = trivialization.prod e₁ e₂}, trivialization_at := λ b, (trivialization_at F₁ E₁ b).prod (trivialization_at F₂ E₂ b), @@ -259,9 +258,9 @@ noncomputable instance fiber_bundle.prod : fiber_bundle (F₁ × F₂) (E₁ × trivialization_mem_atlas := λ b, ⟨trivialization_at F₁ E₁ b, trivialization_at F₂ E₂ b, by apply_instance, by apply_instance, rfl⟩ } -instance {e₁ : trivialization F₁ (π E₁)} {e₂ : trivialization F₂ (π E₂)} +instance {e₁ : trivialization F₁ (π F₁ E₁)} {e₂ : trivialization F₂ (π F₂ E₂)} [mem_trivialization_atlas e₁] [mem_trivialization_atlas e₂] : - mem_trivialization_atlas (e₁.prod e₂ : trivialization (F₁ × F₂) (π (E₁ ×ᵇ E₂))) := + mem_trivialization_atlas (e₁.prod e₂ : trivialization (F₁ × F₂) (π (F₁ × F₂) (E₁ ×ᵇ E₂))) := { out := ⟨e₁, e₂, by apply_instance, by apply_instance, rfl⟩ } end prod @@ -274,32 +273,32 @@ variables {B : Type*} (F : Type*) (E : B → Type*) {B' : Type*} (f : B' → B) instance [∀ (x : B), topological_space (E x)] : ∀ (x : B'), topological_space ((f *ᵖ E) x) := by delta_instance bundle.pullback -variables [topological_space B'] [topological_space (total_space E)] +variables [topological_space B'] [topological_space (total_space F E)] /-- Definition of `pullback.total_space.topological_space`, which we make irreducible. -/ -@[irreducible] def pullback_topology : topological_space (total_space (f *ᵖ E)) := +@[irreducible] def pullback_topology : topological_space (total_space F (f *ᵖ E)) := induced total_space.proj ‹topological_space B'› ⊓ -induced (pullback.lift f) ‹topological_space (total_space E)› +induced (pullback.lift f) ‹topological_space (total_space F E)› /-- The topology on the total space of a pullback bundle is the coarsest topology for which both the projections to the base and the map to the original bundle are continuous. -/ -instance pullback.total_space.topological_space : topological_space (total_space (f *ᵖ E)) := -pullback_topology E f +instance pullback.total_space.topological_space : topological_space (total_space F (f *ᵖ E)) := +pullback_topology F E f -lemma pullback.continuous_proj (f : B' → B) : continuous (@total_space.proj _ (f *ᵖ E)) := +lemma pullback.continuous_proj (f : B' → B) : continuous (π F (f *ᵖ E)) := begin rw [continuous_iff_le_induced, pullback.total_space.topological_space, pullback_topology], exact inf_le_left, end -lemma pullback.continuous_lift (f : B' → B) : continuous (@pullback.lift B E B' f) := +lemma pullback.continuous_lift (f : B' → B) : continuous (@pullback.lift B F E B' f) := begin rw [continuous_iff_le_induced, pullback.total_space.topological_space, pullback_topology], exact inf_le_right, end lemma inducing_pullback_total_space_embedding (f : B' → B) : - inducing (@pullback_total_space_embedding B E B' f) := + inducing (@pullback_total_space_embedding B F E B' f) := begin constructor, simp_rw [prod.topological_space, induced_inf, induced_compose, @@ -312,44 +311,42 @@ variables (F) [topological_space F] [topological_space B] lemma pullback.continuous_total_space_mk [∀ x, topological_space (E x)] [fiber_bundle F E] {f : B' → B} {x : B'} : - continuous (@total_space_mk _ (f *ᵖ E) x) := + continuous (@total_space.mk _ F (f *ᵖ E) x) := begin simp only [continuous_iff_le_induced, pullback.total_space.topological_space, induced_compose, - induced_inf, function.comp, total_space_mk, total_space.proj, induced_const, top_inf_eq, - pullback_topology], + induced_inf, function.comp, induced_const, top_inf_eq, pullback_topology], exact le_of_eq (fiber_bundle.total_space_mk_inducing F E (f x)).induced, end variables {E F} [∀ b, has_zero (E b)] {K : Type*} [continuous_map_class K B' B] /-- A fiber bundle trivialization can be pulled back to a trivialization on the pullback bundle. -/ -noncomputable def trivialization.pullback (e : trivialization F (π E)) (f : K) : - trivialization F (π ((f : B' → B) *ᵖ E)) := +noncomputable def trivialization.pullback (e : trivialization F (π F E)) (f : K) : + trivialization F (π F ((f : B' → B) *ᵖ E)) := { to_fun := λ z, (z.proj, (e (pullback.lift f z)).2), - inv_fun := λ y, @total_space_mk _ (f *ᵖ E) y.1 (e.symm (f y.1) y.2), + inv_fun := λ y, @total_space.mk _ _ (f *ᵖ E) y.1 (e.symm (f y.1) y.2), source := pullback.lift f ⁻¹' e.source, base_set := f ⁻¹' e.base_set, target := (f ⁻¹' e.base_set) ×ˢ univ, - map_source' := λ x h, by { simp_rw [e.source_eq, mem_preimage, pullback.proj_lift] at h, + map_source' := λ x h, by { simp_rw [e.source_eq, mem_preimage, pullback.lift_proj] at h, simp_rw [prod_mk_mem_set_prod_eq, mem_univ, and_true, mem_preimage, h] }, map_target' := λ y h, by { rw [mem_prod, mem_preimage] at h, - simp_rw [e.source_eq, mem_preimage, pullback.proj_lift, h.1] }, - left_inv' := λ x h, by { simp_rw [mem_preimage, e.mem_source, pullback.proj_lift] at h, + simp_rw [e.source_eq, mem_preimage, pullback.lift_proj, h.1] }, + left_inv' := λ x h, by { simp_rw [mem_preimage, e.mem_source, pullback.lift_proj] at h, simp_rw [pullback.lift, e.symm_apply_apply_mk h, total_space.eta] }, right_inv' := λ x h, by { simp_rw [mem_prod, mem_preimage, mem_univ, and_true] at h, - simp_rw [total_space.proj_mk, pullback.lift_mk, e.apply_mk_symm h, prod.mk.eta] }, + simp_rw [pullback.lift_mk, e.apply_mk_symm h, prod.mk.eta] }, open_source := by { simp_rw [e.source_eq, ← preimage_comp], exact ((map_continuous f).comp $ - pullback.continuous_proj E f).is_open_preimage _ e.open_base_set }, + pullback.continuous_proj F E f).is_open_preimage _ e.open_base_set }, open_target := ((map_continuous f).is_open_preimage _ e.open_base_set).prod is_open_univ, open_base_set := (map_continuous f).is_open_preimage _ e.open_base_set, - continuous_to_fun := (pullback.continuous_proj E f).continuous_on.prod + continuous_to_fun := (pullback.continuous_proj F E f).continuous_on.prod (continuous_snd.comp_continuous_on $ - e.continuous_on.comp (pullback.continuous_lift E f).continuous_on subset.rfl), + e.continuous_on.comp (pullback.continuous_lift F E f).continuous_on subset.rfl), continuous_inv_fun := begin dsimp only, - simp_rw [(inducing_pullback_total_space_embedding E f).continuous_on_iff, function.comp, - pullback_total_space_embedding, total_space.proj_mk], - dsimp only [total_space.proj_mk], + simp_rw [(inducing_pullback_total_space_embedding F E f).continuous_on_iff, function.comp, + pullback_total_space_embedding], refine continuous_on_fst.prod (e.continuous_on_symm.comp ((map_continuous f).prod_map continuous_id).continuous_on subset.rfl) end, @@ -360,10 +357,10 @@ noncomputable def trivialization.pullback (e : trivialization F (π E)) (f : K) noncomputable instance fiber_bundle.pullback [∀ x, topological_space (E x)] [fiber_bundle F E] (f : K) : fiber_bundle F ((f : B' → B) *ᵖ E) := { total_space_mk_inducing := λ x, inducing_of_inducing_compose - (pullback.continuous_total_space_mk F E) (pullback.continuous_lift E f) + (pullback.continuous_total_space_mk F E) (pullback.continuous_lift F E f) (total_space_mk_inducing F E (f x)), trivialization_atlas := - {ef | ∃ (e : trivialization F (π E)) [mem_trivialization_atlas e], ef = e.pullback f}, + {ef | ∃ (e : trivialization F (π F E)) [mem_trivialization_atlas e], ef = e.pullback f}, trivialization_at := λ x, (trivialization_at F E (f x)).pullback f, mem_base_set_trivialization_at := λ x, mem_base_set_trivialization_at F E (f x), trivialization_mem_atlas := λ x, ⟨trivialization_at F E (f x), by apply_instance, rfl⟩ } diff --git a/src/topology/fiber_bundle/trivialization.lean b/src/topology/fiber_bundle/trivialization.lean index 84a15ebb82f14..fcde8b14a33e4 100644 --- a/src/topology/fiber_bundle/trivialization.lean +++ b/src/topology/fiber_bundle/trivialization.lean @@ -159,9 +159,9 @@ lemma symm_trans_target_eq (e e' : pretrivialization F proj) : (e.to_local_equiv.symm.trans e'.to_local_equiv).target = (e.base_set ∩ e'.base_set) ×ˢ univ := by rw [← local_equiv.symm_source, symm_trans_symm, symm_trans_source_eq, inter_comm] -variables {B F} (e' : pretrivialization F (π E)) {x' : total_space E} {b : B} {y : E b} +variables {B F} (e' : pretrivialization F (π F E)) {x' : total_space F E} {b : B} {y : E b} -lemma coe_mem_source : ↑y ∈ e'.source ↔ b ∈ e'.base_set := e'.mem_source +@[simp] theorem coe_mem_source : ↑y ∈ e'.source ↔ b ∈ e'.base_set := e'.mem_source @[simp, mfld_simps] lemma coe_coe_fst (hb : b ∈ e'.base_set) : (e' y).1 = b := e'.coe_fst (e'.mem_source.2 hb) @@ -169,7 +169,7 @@ e'.coe_fst (e'.mem_source.2 hb) lemma mk_mem_target {x : B} {y : F} : (x, y) ∈ e'.target ↔ x ∈ e'.base_set := e'.mem_target -lemma symm_coe_proj {x : B} {y : F} (e' : pretrivialization F (π E)) (h : x ∈ e'.base_set) : +lemma symm_coe_proj {x : B} {y : F} (e' : pretrivialization F (π F E)) (h : x ∈ e'.base_set) : (e'.to_local_equiv.symm (x, y)).1 = x := e'.proj_symm_apply' h @@ -177,46 +177,46 @@ section has_zero variables [∀ x, has_zero (E x)] /-- A fiberwise inverse to `e`. This is the function `F → E b` that induces a local inverse -`B × F → total_space E` of `e` on `e.base_set`. It is defined to be `0` outside `e.base_set`. -/ -protected noncomputable def symm (e : pretrivialization F (π E)) (b : B) (y : F) : E b := +`B × F → total_space F E` of `e` on `e.base_set`. It is defined to be `0` outside `e.base_set`. -/ +protected noncomputable def symm (e : pretrivialization F (π F E)) (b : B) (y : F) : E b := if hb : b ∈ e.base_set then cast (congr_arg E (e.proj_symm_apply' hb)) (e.to_local_equiv.symm (b, y)).2 else 0 -lemma symm_apply (e : pretrivialization F (π E)) {b : B} (hb : b ∈ e.base_set) (y : F) : +lemma symm_apply (e : pretrivialization F (π F E)) {b : B} (hb : b ∈ e.base_set) (y : F) : e.symm b y = cast (congr_arg E (e.symm_coe_proj hb)) (e.to_local_equiv.symm (b, y)).2 := dif_pos hb -lemma symm_apply_of_not_mem (e : pretrivialization F (π E)) {b : B} (hb : b ∉ e.base_set) (y : F) : - e.symm b y = 0 := +lemma symm_apply_of_not_mem (e : pretrivialization F (π F E)) {b : B} (hb : b ∉ e.base_set) + (y : F) : e.symm b y = 0 := dif_neg hb -lemma coe_symm_of_not_mem (e : pretrivialization F (π E)) {b : B} (hb : b ∉ e.base_set) : +lemma coe_symm_of_not_mem (e : pretrivialization F (π F E)) {b : B} (hb : b ∉ e.base_set) : (e.symm b : F → E b) = 0 := funext $ λ y, dif_neg hb -lemma mk_symm (e : pretrivialization F (π E)) {b : B} (hb : b ∈ e.base_set) (y : F) : - total_space_mk b (e.symm b y) = e.to_local_equiv.symm (b, y) := +lemma mk_symm (e : pretrivialization F (π F E)) {b : B} (hb : b ∈ e.base_set) (y : F) : + total_space.mk b (e.symm b y) = e.to_local_equiv.symm (b, y) := by rw [e.symm_apply hb, total_space.mk_cast, total_space.eta] -lemma symm_proj_apply (e : pretrivialization F (π E)) (z : total_space E) +lemma symm_proj_apply (e : pretrivialization F (π F E)) (z : total_space F E) (hz : z.proj ∈ e.base_set) : e.symm z.proj (e z).2 = z.2 := by rw [e.symm_apply hz, cast_eq_iff_heq, e.mk_proj_snd' hz, e.symm_apply_apply (e.mem_source.mpr hz)] -lemma symm_apply_apply_mk (e : pretrivialization F (π E)) {b : B} (hb : b ∈ e.base_set) (y : E b) : - e.symm b (e (total_space_mk b y)).2 = y := -e.symm_proj_apply (total_space_mk b y) hb +lemma symm_apply_apply_mk (e : pretrivialization F (π F E)) {b : B} (hb : b ∈ e.base_set) + (y : E b) : e.symm b (e ⟨b, y⟩).2 = y := +e.symm_proj_apply ⟨b, y⟩ hb -lemma apply_mk_symm (e : pretrivialization F (π E)) {b : B} (hb : b ∈ e.base_set) (y : F) : - e (total_space_mk b (e.symm b y)) = (b, y) := +lemma apply_mk_symm (e : pretrivialization F (π F E)) {b : B} (hb : b ∈ e.base_set) (y : F) : + e ⟨b, e.symm b y⟩ = (b, y) := by rw [e.mk_symm hb, e.apply_symm_apply (e.mk_mem_target.mpr hb)] end has_zero end pretrivialization -variables [topological_space Z] [topological_space (total_space E)] +variables [topological_space Z] [topological_space (total_space F E)] /-- A structure extending local homeomorphisms, defining a local trivialization of a projection @@ -403,7 +403,7 @@ begin exact hf_proj.preimage_mem_nhds (e.open_base_set.mem_nhds he), end -variables {E} (e' : trivialization F (π E)) {x' : total_space E} {b : B} {y : E b} +variables {E} (e' : trivialization F (π F E)) {x' : total_space F E} {b : B} {y : E b} protected lemma continuous_on : continuous_on e' e'.source := e'.continuous_to_fun @@ -418,51 +418,52 @@ e'.coe_fst (e'.mem_source.2 hb) lemma mk_mem_target {y : F} : (b, y) ∈ e'.target ↔ b ∈ e'.base_set := e'.to_pretrivialization.mem_target -lemma symm_apply_apply {x : total_space E} (hx : x ∈ e'.source) : +lemma symm_apply_apply {x : total_space F E} (hx : x ∈ e'.source) : e'.to_local_homeomorph.symm (e' x) = x := e'.to_local_equiv.left_inv hx @[simp, mfld_simps] lemma symm_coe_proj {x : B} {y : F} - (e : trivialization F (π E)) (h : x ∈ e.base_set) : + (e : trivialization F (π F E)) (h : x ∈ e.base_set) : (e.to_local_homeomorph.symm (x, y)).1 = x := e.proj_symm_apply' h section has_zero variables [∀ x, has_zero (E x)] /-- A fiberwise inverse to `e'`. The function `F → E x` that induces a local inverse -`B × F → total_space E` of `e'` on `e'.base_set`. It is defined to be `0` outside `e'.base_set`. -/ -protected noncomputable def symm (e : trivialization F (π E)) (b : B) (y : F) : E b := +`B × F → total_space F E` of `e'` on `e'.base_set`. It is defined to be `0` outside +`e'.base_set`. -/ +protected noncomputable def symm (e : trivialization F (π F E)) (b : B) (y : F) : E b := e.to_pretrivialization.symm b y -lemma symm_apply (e : trivialization F (π E)) {b : B} (hb : b ∈ e.base_set) (y : F) : +lemma symm_apply (e : trivialization F (π F E)) {b : B} (hb : b ∈ e.base_set) (y : F) : e.symm b y = cast (congr_arg E (e.symm_coe_proj hb)) (e.to_local_homeomorph.symm (b, y)).2 := dif_pos hb -lemma symm_apply_of_not_mem (e : trivialization F (π E)) {b : B} (hb : b ∉ e.base_set) (y : F) : +lemma symm_apply_of_not_mem (e : trivialization F (π F E)) {b : B} (hb : b ∉ e.base_set) (y : F) : e.symm b y = 0 := dif_neg hb -lemma mk_symm (e : trivialization F (π E)) {b : B} (hb : b ∈ e.base_set) (y : F) : - total_space_mk b (e.symm b y) = e.to_local_homeomorph.symm (b, y) := +lemma mk_symm (e : trivialization F (π F E)) {b : B} (hb : b ∈ e.base_set) (y : F) : + total_space.mk b (e.symm b y) = e.to_local_homeomorph.symm (b, y) := e.to_pretrivialization.mk_symm hb y -lemma symm_proj_apply (e : trivialization F (π E)) (z : total_space E) +lemma symm_proj_apply (e : trivialization F (π F E)) (z : total_space F E) (hz : z.proj ∈ e.base_set) : e.symm z.proj (e z).2 = z.2 := e.to_pretrivialization.symm_proj_apply z hz -lemma symm_apply_apply_mk (e : trivialization F (π E)) {b : B} (hb : b ∈ e.base_set) (y : E b) : - e.symm b (e (total_space_mk b y)).2 = y := -e.symm_proj_apply (total_space_mk b y) hb +lemma symm_apply_apply_mk (e : trivialization F (π F E)) {b : B} (hb : b ∈ e.base_set) (y : E b) : + e.symm b (e ⟨b, y⟩).2 = y := +e.symm_proj_apply ⟨b, y⟩ hb -lemma apply_mk_symm (e : trivialization F (π E)) {b : B} (hb : b ∈ e.base_set) (y : F) : - e (total_space_mk b (e.symm b y)) = (b, y) := +lemma apply_mk_symm (e : trivialization F (π F E)) {b : B} (hb : b ∈ e.base_set) (y : F) : + e ⟨b, e.symm b y⟩ = (b, y) := e.to_pretrivialization.apply_mk_symm hb y -lemma continuous_on_symm (e : trivialization F (π E)) : - continuous_on (λ z : B × F, total_space_mk z.1 (e.symm z.1 z.2)) (e.base_set ×ˢ univ) := +lemma continuous_on_symm (e : trivialization F (π F E)) : + continuous_on (λ z : B × F, total_space.mk' F z.1 (e.symm z.1 z.2)) (e.base_set ×ˢ univ) := begin have : ∀ (z : B × F) (hz : z ∈ e.base_set ×ˢ (univ : set F)), - total_space_mk z.1 (e.symm z.1 z.2) = e.to_local_homeomorph.symm z, + total_space.mk z.1 (e.symm z.1 z.2) = e.to_local_homeomorph.symm z, { rintro x ⟨hx : x.1 ∈ e.base_set, _⟩, simp_rw [e.mk_symm hx, prod.mk.eta] }, refine continuous_on.congr _ this, rw [← e.target_eq], diff --git a/src/topology/vector_bundle/basic.lean b/src/topology/vector_bundle/basic.lean index 6d1d563cfcd6c..5a703818a8660 100644 --- a/src/topology/vector_bundle/basic.lean +++ b/src/topology/vector_bundle/basic.lean @@ -19,7 +19,7 @@ Let `B` be the base space, let `F` be a normed space over a normed field `R`, an `E : B → Type*` be a `fiber_bundle` with fiber `F`, in which, for each `x`, the fiber `E x` is a topological vector space over `R`. -To have a vector bundle structure on `bundle.total_space E`, one should additionally have the +To have a vector bundle structure on `bundle.total_space F E`, one should additionally have the following properties: * The bundle trivializations in the trivialization atlas should be continuous linear equivs in the @@ -70,23 +70,23 @@ variables {B F E} [semiring R] /-- A mixin class for `pretrivialization`, stating that a pretrivialization is fiberwise linear with respect to given module structures on its fibers and the model fiber. -/ protected class pretrivialization.is_linear [add_comm_monoid F] [module R F] - [∀ x, add_comm_monoid (E x)] [∀ x, module R (E x)] (e : pretrivialization F (π E)) : + [∀ x, add_comm_monoid (E x)] [∀ x, module R (E x)] (e : pretrivialization F (π F E)) : Prop := -(linear : ∀ b ∈ e.base_set, is_linear_map R (λ x : E b, (e (total_space_mk b x)).2)) +(linear : ∀ b ∈ e.base_set, is_linear_map R (λ x : E b, (e ⟨b, x⟩).2)) namespace pretrivialization -variables {F E} (e : pretrivialization F (π E)) {x : total_space E} {b : B} {y : E b} +variables {F E} (e : pretrivialization F (π F E)) {x : total_space F E} {b : B} {y : E b} lemma linear [add_comm_monoid F] [module R F] [∀ x, add_comm_monoid (E x)] [∀ x, module R (E x)] [e.is_linear R] {b : B} (hb : b ∈ e.base_set) : - is_linear_map R (λ x : E b, (e (total_space_mk b x)).2) := + is_linear_map R (λ x : E b, (e ⟨b, x⟩).2) := pretrivialization.is_linear.linear b hb variables [add_comm_monoid F] [module R F] [∀ x, add_comm_monoid (E x)] [∀ x, module R (E x)] /-- A fiberwise linear inverse to `e`. -/ -@[simps] protected def symmₗ (e : pretrivialization F (π E)) [e.is_linear R] (b : B) : +@[simps] protected def symmₗ (e : pretrivialization F (π F E)) [e.is_linear R] (b : B) : F →ₗ[R] E b := begin refine is_linear_map.mk' (e.symm b) _, @@ -98,10 +98,10 @@ end /-- A pretrivialization for a vector bundle defines linear equivalences between the fibers and the model space. -/ -@[simps {fully_applied := ff}] def linear_equiv_at (e : pretrivialization F (π E)) [e.is_linear R] +@[simps {fully_applied := ff}] def linear_equiv_at (e : pretrivialization F (π F E)) [e.is_linear R] (b : B) (hb : b ∈ e.base_set) : E b ≃ₗ[R] F := -{ to_fun := λ y, (e (total_space_mk b y)).2, +{ to_fun := λ y, (e ⟨b, y⟩).2, inv_fun := e.symm b, left_inv := e.symm_apply_apply_mk hb, right_inv := λ v, by simp_rw [e.apply_mk_symm hb v], @@ -109,66 +109,67 @@ fibers and the model space. -/ map_smul' := λ c v, (e.linear R hb).map_smul c v } /-- A fiberwise linear map equal to `e` on `e.base_set`. -/ -protected def linear_map_at (e : pretrivialization F (π E)) [e.is_linear R] (b : B) : E b →ₗ[R] F := +protected def linear_map_at (e : pretrivialization F (π F E)) [e.is_linear R] (b : B) : + E b →ₗ[R] F := if hb : b ∈ e.base_set then e.linear_equiv_at R b hb else 0 variables {R} -lemma coe_linear_map_at (e : pretrivialization F (π E)) [e.is_linear R] (b : B) : - ⇑(e.linear_map_at R b) = λ y, if b ∈ e.base_set then (e (total_space_mk b y)).2 else 0 := +lemma coe_linear_map_at (e : pretrivialization F (π F E)) [e.is_linear R] (b : B) : + ⇑(e.linear_map_at R b) = λ y, if b ∈ e.base_set then (e ⟨b, y⟩).2 else 0 := by { rw [pretrivialization.linear_map_at], split_ifs; refl } -lemma coe_linear_map_at_of_mem (e : pretrivialization F (π E)) [e.is_linear R] {b : B} +lemma coe_linear_map_at_of_mem (e : pretrivialization F (π F E)) [e.is_linear R] {b : B} (hb : b ∈ e.base_set) : - ⇑(e.linear_map_at R b) = λ y, (e (total_space_mk b y)).2 := + ⇑(e.linear_map_at R b) = λ y, (e ⟨b, y⟩).2 := by simp_rw [coe_linear_map_at, if_pos hb] -lemma linear_map_at_apply (e : pretrivialization F (π E)) [e.is_linear R] {b : B} (y : E b) : - e.linear_map_at R b y = if b ∈ e.base_set then (e (total_space_mk b y)).2 else 0 := +lemma linear_map_at_apply (e : pretrivialization F (π F E)) [e.is_linear R] {b : B} (y : E b) : + e.linear_map_at R b y = if b ∈ e.base_set then (e ⟨b, y⟩).2 else 0 := by rw [coe_linear_map_at] -lemma linear_map_at_def_of_mem (e : pretrivialization F (π E)) [e.is_linear R] {b : B} +lemma linear_map_at_def_of_mem (e : pretrivialization F (π F E)) [e.is_linear R] {b : B} (hb : b ∈ e.base_set) : e.linear_map_at R b = e.linear_equiv_at R b hb := dif_pos hb -lemma linear_map_at_def_of_not_mem (e : pretrivialization F (π E)) [e.is_linear R] {b : B} +lemma linear_map_at_def_of_not_mem (e : pretrivialization F (π F E)) [e.is_linear R] {b : B} (hb : b ∉ e.base_set) : e.linear_map_at R b = 0 := dif_neg hb -lemma linear_map_at_eq_zero (e : pretrivialization F (π E)) [e.is_linear R] {b : B} +lemma linear_map_at_eq_zero (e : pretrivialization F (π F E)) [e.is_linear R] {b : B} (hb : b ∉ e.base_set) : e.linear_map_at R b = 0 := dif_neg hb -lemma symmₗ_linear_map_at (e : pretrivialization F (π E)) [e.is_linear R] {b : B} +lemma symmₗ_linear_map_at (e : pretrivialization F (π F E)) [e.is_linear R] {b : B} (hb : b ∈ e.base_set) (y : E b) : e.symmₗ R b (e.linear_map_at R b y) = y := by { rw [e.linear_map_at_def_of_mem hb], exact (e.linear_equiv_at R b hb).left_inv y } -lemma linear_map_at_symmₗ (e : pretrivialization F (π E)) [e.is_linear R] {b : B} +lemma linear_map_at_symmₗ (e : pretrivialization F (π F E)) [e.is_linear R] {b : B} (hb : b ∈ e.base_set) (y : F) : e.linear_map_at R b (e.symmₗ R b y) = y := by { rw [e.linear_map_at_def_of_mem hb], exact (e.linear_equiv_at R b hb).right_inv y } end pretrivialization -variables (R) [topological_space (total_space E)] +variables (R) [topological_space (total_space F E)] /-- A mixin class for `trivialization`, stating that a trivialization is fiberwise linear with respect to given module structures on its fibers and the model fiber. -/ protected class trivialization.is_linear [add_comm_monoid F] [module R F] - [∀ x, add_comm_monoid (E x)] [∀ x, module R (E x)] (e : trivialization F (π E)) : Prop := -(linear : ∀ b ∈ e.base_set, is_linear_map R (λ x : E b, (e (total_space_mk b x)).2)) + [∀ x, add_comm_monoid (E x)] [∀ x, module R (E x)] (e : trivialization F (π F E)) : Prop := +(linear : ∀ b ∈ e.base_set, is_linear_map R (λ x : E b, (e ⟨b, x⟩).2)) namespace trivialization -variables (e : trivialization F (π E)) {x : total_space E} {b : B} {y : E b} +variables (e : trivialization F (π F E)) {x : total_space F E} {b : B} {y : E b} protected lemma linear [add_comm_monoid F] [module R F] [∀ x, add_comm_monoid (E x)] [∀ x, module R (E x)] [e.is_linear R] {b : B} (hb : b ∈ e.base_set) : - is_linear_map R (λ y : E b, (e (total_space_mk b y)).2) := + is_linear_map R (λ y : E b, (e ⟨b, y⟩).2) := trivialization.is_linear.linear b hb instance to_pretrivialization.is_linear [add_comm_monoid F] [module R F] @@ -180,71 +181,72 @@ variables [add_comm_monoid F] [module R F] [∀ x, add_comm_monoid (E x)] [∀ x /-- A trivialization for a vector bundle defines linear equivalences between the fibers and the model space. -/ -def linear_equiv_at (e : trivialization F (π E)) [e.is_linear R] (b : B) (hb : b ∈ e.base_set) : +def linear_equiv_at (e : trivialization F (π F E)) [e.is_linear R] (b : B) (hb : b ∈ e.base_set) : E b ≃ₗ[R] F := e.to_pretrivialization.linear_equiv_at R b hb variables {R} @[simp] -lemma linear_equiv_at_apply (e : trivialization F (π E)) [e.is_linear R] (b : B) +lemma linear_equiv_at_apply (e : trivialization F (π F E)) [e.is_linear R] (b : B) (hb : b ∈ e.base_set) (v : E b) : - e.linear_equiv_at R b hb v = (e (total_space_mk b v)).2 := rfl + e.linear_equiv_at R b hb v = (e ⟨b, v⟩).2 := rfl @[simp] -lemma linear_equiv_at_symm_apply (e : trivialization F (π E)) [e.is_linear R] (b : B) +lemma linear_equiv_at_symm_apply (e : trivialization F (π F E)) [e.is_linear R] (b : B) (hb : b ∈ e.base_set) (v : F) : (e.linear_equiv_at R b hb).symm v = e.symm b v := rfl variables (R) /-- A fiberwise linear inverse to `e`. -/ -protected def symmₗ (e : trivialization F (π E)) [e.is_linear R] (b : B) : F →ₗ[R] E b := +protected def symmₗ (e : trivialization F (π F E)) [e.is_linear R] (b : B) : F →ₗ[R] E b := e.to_pretrivialization.symmₗ R b variables {R} -lemma coe_symmₗ (e : trivialization F (π E)) [e.is_linear R] (b : B) : ⇑(e.symmₗ R b) = e.symm b := +lemma coe_symmₗ (e : trivialization F (π F E)) [e.is_linear R] (b : B) : + ⇑(e.symmₗ R b) = e.symm b := rfl variables (R) /-- A fiberwise linear map equal to `e` on `e.base_set`. -/ -protected def linear_map_at (e : trivialization F (π E)) [e.is_linear R] (b : B) : E b →ₗ[R] F := +protected def linear_map_at (e : trivialization F (π F E)) [e.is_linear R] (b : B) : E b →ₗ[R] F := e.to_pretrivialization.linear_map_at R b variables {R} -lemma coe_linear_map_at (e : trivialization F (π E)) [e.is_linear R] (b : B) : - ⇑(e.linear_map_at R b) = λ y, if b ∈ e.base_set then (e (total_space_mk b y)).2 else 0 := +lemma coe_linear_map_at (e : trivialization F (π F E)) [e.is_linear R] (b : B) : + ⇑(e.linear_map_at R b) = λ y, if b ∈ e.base_set then (e ⟨b, y⟩).2 else 0 := e.to_pretrivialization.coe_linear_map_at b -lemma coe_linear_map_at_of_mem (e : trivialization F (π E)) [e.is_linear R] {b : B} +lemma coe_linear_map_at_of_mem (e : trivialization F (π F E)) [e.is_linear R] {b : B} (hb : b ∈ e.base_set) : - ⇑(e.linear_map_at R b) = λ y, (e (total_space_mk b y)).2 := + ⇑(e.linear_map_at R b) = λ y, (e ⟨b, y⟩).2 := by simp_rw [coe_linear_map_at, if_pos hb] -lemma linear_map_at_apply (e : trivialization F (π E)) [e.is_linear R] {b : B} (y : E b) : - e.linear_map_at R b y = if b ∈ e.base_set then (e (total_space_mk b y)).2 else 0 := +lemma linear_map_at_apply (e : trivialization F (π F E)) [e.is_linear R] {b : B} (y : E b) : + e.linear_map_at R b y = if b ∈ e.base_set then (e ⟨b, y⟩).2 else 0 := by rw [coe_linear_map_at] -lemma linear_map_at_def_of_mem (e : trivialization F (π E)) [e.is_linear R] {b : B} +lemma linear_map_at_def_of_mem (e : trivialization F (π F E)) [e.is_linear R] {b : B} (hb : b ∈ e.base_set) : e.linear_map_at R b = e.linear_equiv_at R b hb := dif_pos hb -lemma linear_map_at_def_of_not_mem (e : trivialization F (π E)) [e.is_linear R] {b : B} +lemma linear_map_at_def_of_not_mem (e : trivialization F (π F E)) [e.is_linear R] {b : B} (hb : b ∉ e.base_set) : e.linear_map_at R b = 0 := dif_neg hb -lemma symmₗ_linear_map_at (e : trivialization F (π E)) [e.is_linear R] {b : B} (hb : b ∈ e.base_set) - (y : E b) : +lemma symmₗ_linear_map_at (e : trivialization F (π F E)) [e.is_linear R] {b : B} + (hb : b ∈ e.base_set) (y : E b) : e.symmₗ R b (e.linear_map_at R b y) = y := e.to_pretrivialization.symmₗ_linear_map_at hb y -lemma linear_map_at_symmₗ (e : trivialization F (π E)) [e.is_linear R] {b : B} (hb : b ∈ e.base_set) - (y : F) : +lemma linear_map_at_symmₗ (e : trivialization F (π F E)) [e.is_linear R] {b : B} + (hb : b ∈ e.base_set) (y : F) : e.linear_map_at R b (e.symmₗ R b y) = y := e.to_pretrivialization.linear_map_at_symmₗ hb y @@ -252,7 +254,7 @@ variables (R) /-- A coordinate change function between two trivializations, as a continuous linear equivalence. Defined to be the identity when `b` does not lie in the base set of both trivializations. -/ -def coord_changeL (e e' : trivialization F (π E)) [e.is_linear R] [e'.is_linear R] (b : B) : +def coord_changeL (e e' : trivialization F (π F E)) [e.is_linear R] [e'.is_linear R] (b : B) : F ≃L[R] F := { continuous_to_fun := begin by_cases hb : b ∈ e.base_set ∩ e'.base_set, @@ -278,19 +280,19 @@ def coord_changeL (e e' : trivialization F (π E)) [e.is_linear R] [e'.is_linear variables {R} -lemma coe_coord_changeL (e e' : trivialization F (π E)) [e.is_linear R] [e'.is_linear R] {b : B} +lemma coe_coord_changeL (e e' : trivialization F (π F E)) [e.is_linear R] [e'.is_linear R] {b : B} (hb : b ∈ e.base_set ∩ e'.base_set) : ⇑(coord_changeL R e e' b) = (e.linear_equiv_at R b hb.1).symm.trans (e'.linear_equiv_at R b hb.2) := congr_arg linear_equiv.to_fun (dif_pos hb) -lemma coe_coord_changeL' (e e' : trivialization F (π E)) [e.is_linear R] [e'.is_linear R] {b : B} +lemma coe_coord_changeL' (e e' : trivialization F (π F E)) [e.is_linear R] [e'.is_linear R] {b : B} (hb : b ∈ e.base_set ∩ e'.base_set) : (coord_changeL R e e' b).to_linear_equiv = (e.linear_equiv_at R b hb.1).symm.trans (e'.linear_equiv_at R b hb.2) := linear_equiv.coe_injective (coe_coord_changeL _ _ _) -lemma symm_coord_changeL (e e' : trivialization F (π E)) [e.is_linear R] [e'.is_linear R] {b : B} +lemma symm_coord_changeL (e e' : trivialization F (π F E)) [e.is_linear R] [e'.is_linear R] {b : B} (hb : b ∈ e'.base_set ∩ e.base_set) : (e.coord_changeL R e' b).symm = e'.coord_changeL R e b := begin @@ -299,14 +301,14 @@ begin coe_coord_changeL' e e' hb.symm, linear_equiv.trans_symm, linear_equiv.symm_symm], end -lemma coord_changeL_apply (e e' : trivialization F (π E)) [e.is_linear R] [e'.is_linear R] {b : B} +lemma coord_changeL_apply (e e' : trivialization F (π F E)) [e.is_linear R] [e'.is_linear R] {b : B} (hb : b ∈ e.base_set ∩ e'.base_set) (y : F) : - coord_changeL R e e' b y = (e' (total_space_mk b (e.symm b y))).2 := + coord_changeL R e e' b y = (e' ⟨b, e.symm b y⟩).2 := congr_arg (λ f, linear_equiv.to_fun f y) (dif_pos hb) -lemma mk_coord_changeL (e e' : trivialization F (π E)) [e.is_linear R] [e'.is_linear R] {b : B} +lemma mk_coord_changeL (e e' : trivialization F (π F E)) [e.is_linear R] [e'.is_linear R] {b : B} (hb : b ∈ e.base_set ∩ e'.base_set) (y : F) : - (b, coord_changeL R e e' b y) = e' (total_space_mk b (e.symm b y)) := + (b, coord_changeL R e e' b y) = e' ⟨b, e.symm b y⟩ := begin ext, { rw [e.mk_symm hb.1 y, e'.coe_fst', e.proj_symm_apply' hb.1], @@ -314,19 +316,19 @@ begin { exact e.coord_changeL_apply e' hb y } end -lemma apply_symm_apply_eq_coord_changeL (e e' : trivialization F (π E)) [e.is_linear R] +lemma apply_symm_apply_eq_coord_changeL (e e' : trivialization F (π F E)) [e.is_linear R] [e'.is_linear R] {b : B} (hb : b ∈ e.base_set ∩ e'.base_set) (v : F) : e' (e.to_local_homeomorph.symm (b, v)) = (b, e.coord_changeL R e' b v) := by rw [e.mk_coord_changeL e' hb, e.mk_symm hb.1] /-- A version of `coord_change_apply` that fully unfolds `coord_change`. The right-hand side is ugly, but has good definitional properties for specifically defined trivializations. -/ -lemma coord_changeL_apply' (e e' : trivialization F (π E)) [e.is_linear R] [e'.is_linear R] {b : B} - (hb : b ∈ e.base_set ∩ e'.base_set) (y : F) : +lemma coord_changeL_apply' (e e' : trivialization F (π F E)) [e.is_linear R] [e'.is_linear R] + {b : B} (hb : b ∈ e.base_set ∩ e'.base_set) (y : F) : coord_changeL R e e' b y = (e' (e.to_local_homeomorph.symm (b, y))).2 := by rw [e.coord_changeL_apply e' hb, e.mk_symm hb.1] -lemma coord_changeL_symm_apply (e e' : trivialization F (π E)) [e.is_linear R] [e'.is_linear R] +lemma coord_changeL_symm_apply (e e' : trivialization F (π F E)) [e.is_linear R] [e'.is_linear R] {b : B} (hb : b ∈ e.base_set ∩ e'.base_set) : ⇑(coord_changeL R e e' b).symm = (e'.linear_equiv_at R b hb.2).symm.trans (e.linear_equiv_at R b hb.1) := @@ -341,29 +343,29 @@ section namespace bundle /-- The zero section of a vector bundle -/ -def zero_section [∀ x, has_zero (E x)] : B → total_space E := -λ x, total_space_mk x 0 +def zero_section [∀ x, has_zero (E x)] : B → total_space F E := +λ x, ⟨x, 0⟩ @[simp, mfld_simps] -lemma zero_section_proj [∀ x, has_zero (E x)] (x : B) : (zero_section E x).proj = x := rfl +lemma zero_section_proj [∀ x, has_zero (E x)] (x : B) : (zero_section F E x).proj = x := rfl @[simp, mfld_simps] -lemma zero_section_snd [∀ x, has_zero (E x)] (x : B) : (zero_section E x).2 = 0 := rfl +lemma zero_section_snd [∀ x, has_zero (E x)] (x : B) : (zero_section F E x).2 = 0 := rfl end bundle open bundle variables [nontrivially_normed_field R] [∀ x, add_comm_monoid (E x)] [∀ x, module R (E x)] [normed_add_comm_group F] [normed_space R F] [topological_space B] - [topological_space (total_space E)] [∀ x, topological_space (E x)] [fiber_bundle F E] + [topological_space (total_space F E)] [∀ x, topological_space (E x)] [fiber_bundle F E] -/-- The space `total_space E` (for `E : B → Type*` such that each `E x` is a topological vector +/-- The space `total_space F E` (for `E : B → Type*` such that each `E x` is a topological vector space) has a topological vector space structure with fiber `F` (denoted with `vector_bundle R F E`) if around every point there is a fiber bundle trivialization which is linear in the fibers. -/ class vector_bundle : Prop := -(trivialization_linear' : ∀ (e : trivialization F (π E)) [mem_trivialization_atlas e], +(trivialization_linear' : ∀ (e : trivialization F (π F E)) [mem_trivialization_atlas e], e.is_linear R) -(continuous_on_coord_change' [] : ∀ (e e' : trivialization F (π E)) [mem_trivialization_atlas e] +(continuous_on_coord_change' [] : ∀ (e e' : trivialization F (π F E)) [mem_trivialization_atlas e] [mem_trivialization_atlas e'], continuous_on (λ b, by exactI trivialization.coord_changeL R e e' b : B → F →L[R] F) (e.base_set ∩ e'.base_set)) @@ -371,12 +373,12 @@ class vector_bundle : Prop := variables {F E} @[priority 100] -instance trivialization_linear [vector_bundle R F E] (e : trivialization F (π E)) +instance trivialization_linear [vector_bundle R F E] (e : trivialization F (π F E)) [mem_trivialization_atlas e] : e.is_linear R := vector_bundle.trivialization_linear' e -lemma continuous_on_coord_change [vector_bundle R F E] (e e' : trivialization F (π E)) +lemma continuous_on_coord_change [vector_bundle R F E] (e e' : trivialization F (π F E)) [he : mem_trivialization_atlas e] [he' : mem_trivialization_atlas e'] : continuous_on @@ -388,7 +390,7 @@ namespace trivialization /-- Forward map of `continuous_linear_equiv_at` (only propositionally equal), defined everywhere (`0` outside domain). -/ @[simps apply {fully_applied := ff}] -def continuous_linear_map_at (e : trivialization F (π E)) [e.is_linear R] (b : B) : +def continuous_linear_map_at (e : trivialization F (π F E)) [e.is_linear R] (b : B) : E b →L[R] F := { to_fun := e.linear_map_at R b, -- given explicitly to help `simps` cont := begin @@ -403,7 +405,7 @@ def continuous_linear_map_at (e : trivialization F (π E)) [e.is_linear R] (b : /-- Backwards map of `continuous_linear_equiv_at`, defined everywhere. -/ @[simps apply {fully_applied := ff}] -def symmL (e : trivialization F (π E)) [e.is_linear R] (b : B) : F →L[R] E b := +def symmL (e : trivialization F (π F E)) [e.is_linear R] (b : B) : F →L[R] E b := { to_fun := e.symm b, -- given explicitly to help `simps` cont := begin by_cases hb : b ∈ e.base_set, @@ -416,12 +418,12 @@ def symmL (e : trivialization F (π E)) [e.is_linear R] (b : B) : F →L[R] E b variables {R} -lemma symmL_continuous_linear_map_at (e : trivialization F (π E)) [e.is_linear R] {b : B} +lemma symmL_continuous_linear_map_at (e : trivialization F (π F E)) [e.is_linear R] {b : B} (hb : b ∈ e.base_set) (y : E b) : e.symmL R b (e.continuous_linear_map_at R b y) = y := e.symmₗ_linear_map_at hb y -lemma continuous_linear_map_at_symmL (e : trivialization F (π E)) [e.is_linear R] {b : B} +lemma continuous_linear_map_at_symmL (e : trivialization F (π F E)) [e.is_linear R] {b : B} (hb : b ∈ e.base_set) (y : F) : e.continuous_linear_map_at R b (e.symmL R b y) = y := e.linear_map_at_symmₗ hb y @@ -431,9 +433,9 @@ variables (R) /-- In a vector bundle, a trivialization in the fiber (which is a priori only linear) is in fact a continuous linear equiv between the fibers and the model fiber. -/ @[simps apply symm_apply {fully_applied := ff}] -def continuous_linear_equiv_at (e : trivialization F (π E)) [e.is_linear R] (b : B) +def continuous_linear_equiv_at (e : trivialization F (π F E)) [e.is_linear R] (b : B) (hb : b ∈ e.base_set) : E b ≃L[R] F := -{ to_fun := λ y, (e (total_space_mk b y)).2, -- given explicitly to help `simps` +{ to_fun := λ y, (e ⟨b, y⟩).2, -- given explicitly to help `simps` inv_fun := e.symm b, -- given explicitly to help `simps` continuous_to_fun := continuous_snd.comp (e.continuous_on.comp_continuous (fiber_bundle.total_space_mk_inducing F E b).continuous @@ -443,24 +445,24 @@ def continuous_linear_equiv_at (e : trivialization F (π E)) [e.is_linear R] (b variables {R} -lemma coe_continuous_linear_equiv_at_eq (e : trivialization F (π E)) [e.is_linear R] {b : B} +lemma coe_continuous_linear_equiv_at_eq (e : trivialization F (π F E)) [e.is_linear R] {b : B} (hb : b ∈ e.base_set) : (e.continuous_linear_equiv_at R b hb : E b → F) = e.continuous_linear_map_at R b := (e.coe_linear_map_at_of_mem hb).symm -lemma symm_continuous_linear_equiv_at_eq (e : trivialization F (π E)) [e.is_linear R] {b : B} +lemma symm_continuous_linear_equiv_at_eq (e : trivialization F (π F E)) [e.is_linear R] {b : B} (hb : b ∈ e.base_set) : ((e.continuous_linear_equiv_at R b hb).symm : F → E b) = e.symmL R b := rfl -@[simp] lemma continuous_linear_equiv_at_apply' (e : trivialization F (π E)) [e.is_linear R] - (x : total_space E) (hx : x ∈ e.source) : +@[simp] lemma continuous_linear_equiv_at_apply' (e : trivialization F (π F E)) [e.is_linear R] + (x : total_space F E) (hx : x ∈ e.source) : e.continuous_linear_equiv_at R x.proj (e.mem_source.1 hx) x.2 = (e x).2 := by { cases x, refl } variables (R) -lemma apply_eq_prod_continuous_linear_equiv_at (e : trivialization F (π E)) [e.is_linear R] (b : B) - (hb : b ∈ e.base_set) (z : E b) : +lemma apply_eq_prod_continuous_linear_equiv_at (e : trivialization F (π F E)) [e.is_linear R] + (b : B) (hb : b ∈ e.base_set) (z : E b) : e ⟨b, z⟩ = (b, e.continuous_linear_equiv_at R b hb z) := begin ext, @@ -470,17 +472,17 @@ begin { simp only [coe_coe, continuous_linear_equiv_at_apply] } end -protected lemma zero_section (e : trivialization F (π E)) [e.is_linear R] - {x : B} (hx : x ∈ e.base_set) : e (zero_section E x) = (x, 0) := -by simp_rw [zero_section, total_space_mk, e.apply_eq_prod_continuous_linear_equiv_at R x hx 0, +protected lemma zero_section (e : trivialization F (π F E)) [e.is_linear R] + {x : B} (hx : x ∈ e.base_set) : e (zero_section F E x) = (x, 0) := +by simp_rw [zero_section, e.apply_eq_prod_continuous_linear_equiv_at R x hx 0, map_zero] variables {R} -lemma symm_apply_eq_mk_continuous_linear_equiv_at_symm (e : trivialization F (π E)) [e.is_linear R] - (b : B) (hb : b ∈ e.base_set) (z : F) : +lemma symm_apply_eq_mk_continuous_linear_equiv_at_symm (e : trivialization F (π F E)) + [e.is_linear R] (b : B) (hb : b ∈ e.base_set) (z : F) : e.to_local_homeomorph.symm ⟨b, z⟩ - = total_space_mk b ((e.continuous_linear_equiv_at R b hb).symm z) := + = ⟨b, (e.continuous_linear_equiv_at R b hb).symm z⟩ := begin have h : (b, z) ∈ e.target, { rw e.target_eq, @@ -491,7 +493,7 @@ begin continuous_linear_equiv.apply_symm_apply], end -lemma comp_continuous_linear_equiv_at_eq_coord_change (e e' : trivialization F (π E)) +lemma comp_continuous_linear_equiv_at_eq_coord_change (e e' : trivialization F (π F E)) [e.is_linear R] [e'.is_linear R] {b : B} (hb : b ∈ e.base_set ∩ e'.base_set) : (e.continuous_linear_equiv_at R b hb.1).symm.trans (e'.continuous_linear_equiv_at R b hb.2) = coord_changeL R e e' b := @@ -578,13 +580,12 @@ instance add_comm_group_fiber [add_comm_group F] : ∀ (x : B), add_comm_group ( by dsimp [vector_bundle_core.fiber]; delta_instance fiber_bundle_core.fiber /-- The projection from the total space of a fiber bundle core, on its base. -/ -@[reducible, simp, mfld_simps] protected def proj : total_space Z.fiber → B := total_space.proj +@[reducible, simp, mfld_simps] protected def proj : total_space F Z.fiber → B := total_space.proj /-- The total space of the vector bundle, as a convenience function for dot notation. -It is by definition equal to `bundle.total_space Z.fiber`, a.k.a. `Σ x, Z.fiber x` but with a -different name for typeclass inference. -/ +It is by definition equal to `bundle.total_space Z.fiber`. -/ @[nolint unused_arguments, reducible] -protected def total_space := bundle.total_space Z.fiber +protected def total_space := bundle.total_space F Z.fiber /-- Local homeomorphism version of the trivialization change. -/ def triv_change (i j : ι) : local_homeomorph (B × F) (B × F) := @@ -606,7 +607,7 @@ variables (b : B) (a : F) /-- One of the standard local trivializations of a vector bundle constructed from core, taken by considering this in particular as a fiber bundle constructed from core. -/ -def local_triv (i : ι) : trivialization F (π Z.fiber) := +def local_triv (i : ι) : trivialization F (π F Z.fiber) := by dsimp [vector_bundle_core.total_space, vector_bundle_core.fiber]; exact Z.to_fiber_bundle_core.local_triv i @@ -650,7 +651,7 @@ end /-- Preferred local trivialization of a vector bundle constructed from core, at a given point, as a bundle trivialization -/ -def local_triv_at (b : B) : trivialization F (π Z.fiber) := +def local_triv_at (b : B) : trivialization F (π F Z.fiber) := Z.local_triv (Z.index_at b) @[simp, mfld_simps] lemma local_triv_at_def : @@ -750,17 +751,17 @@ This makes it inconvenient to explicitly define a `coord_change` function when c `vector_prebundle`. -/ @[nolint has_nonempty_instance] structure vector_prebundle := -(pretrivialization_atlas : set (pretrivialization F (π E))) -(pretrivialization_linear' : ∀ (e : pretrivialization F (π E)) (he : e ∈ pretrivialization_atlas), +(pretrivialization_atlas : set (pretrivialization F (π F E))) +(pretrivialization_linear' : ∀ (e : pretrivialization F (π F E)) (he : e ∈ pretrivialization_atlas), e.is_linear R) -(pretrivialization_at : B → pretrivialization F (π E)) +(pretrivialization_at : B → pretrivialization F (π F E)) (mem_base_pretrivialization_at : ∀ x : B, x ∈ (pretrivialization_at x).base_set) (pretrivialization_mem_atlas : ∀ x : B, pretrivialization_at x ∈ pretrivialization_atlas) (exists_coord_change : ∀ (e e' ∈ pretrivialization_atlas), ∃ f : B → F →L[R] F, continuous_on f (e.base_set ∩ e'.base_set) ∧ ∀ (b : B) (hb : b ∈ e.base_set ∩ e'.base_set) (v : F), - f b v = (e' (total_space_mk b (e.symm b v))).2) -(total_space_mk_inducing : ∀ (b : B), inducing ((pretrivialization_at b) ∘ (total_space_mk b))) + f b v = (e' ⟨b, e.symm b v⟩).2) +(total_space_mk_inducing : ∀ (b : B), inducing ((pretrivialization_at b) ∘ (total_space.mk b))) namespace vector_prebundle @@ -769,26 +770,26 @@ variables {R E F} /-- A randomly chosen coordinate change on a `vector_prebundle`, given by the field `exists_coord_change`. -/ def coord_change (a : vector_prebundle R F E) - {e e' : pretrivialization F (π E)} (he : e ∈ a.pretrivialization_atlas) + {e e' : pretrivialization F (π F E)} (he : e ∈ a.pretrivialization_atlas) (he' : e' ∈ a.pretrivialization_atlas) (b : B) : F →L[R] F := classical.some (a.exists_coord_change e he e' he') b lemma continuous_on_coord_change (a : vector_prebundle R F E) - {e e' : pretrivialization F (π E)} (he : e ∈ a.pretrivialization_atlas) + {e e' : pretrivialization F (π F E)} (he : e ∈ a.pretrivialization_atlas) (he' : e' ∈ a.pretrivialization_atlas) : continuous_on (a.coord_change he he') (e.base_set ∩ e'.base_set) := (classical.some_spec (a.exists_coord_change e he e' he')).1 lemma coord_change_apply (a : vector_prebundle R F E) - {e e' : pretrivialization F (π E)} (he : e ∈ a.pretrivialization_atlas) + {e e' : pretrivialization F (π F E)} (he : e ∈ a.pretrivialization_atlas) (he' : e' ∈ a.pretrivialization_atlas) {b : B} (hb : b ∈ e.base_set ∩ e'.base_set) (v : F) : - a.coord_change he he' b v = (e' (total_space_mk b (e.symm b v))).2 := + a.coord_change he he' b v = (e' ⟨b, e.symm b v⟩).2 := (classical.some_spec (a.exists_coord_change e he e' he')).2 b hb v lemma mk_coord_change (a : vector_prebundle R F E) - {e e' : pretrivialization F (π E)} (he : e ∈ a.pretrivialization_atlas) + {e e' : pretrivialization F (π F E)} (he : e ∈ a.pretrivialization_atlas) (he' : e' ∈ a.pretrivialization_atlas) {b : B} (hb : b ∈ e.base_set ∩ e'.base_set) (v : F) : - (b, a.coord_change he he' b v) = e' (total_space_mk b (e.symm b v)) := + (b, a.coord_change he he' b v) = e' ⟨b, e.symm b v⟩ := begin ext, { rw [e.mk_symm hb.1 v, e'.coe_fst', e.proj_symm_apply' hb.1], @@ -820,18 +821,18 @@ def to_fiber_prebundle (a : vector_prebundle R F E) : /-- Topology on the total space that will make the prebundle into a bundle. -/ def total_space_topology (a : vector_prebundle R F E) : - topological_space (total_space E) := + topological_space (total_space F E) := a.to_fiber_prebundle.total_space_topology /-- Promotion from a `trivialization` in the `pretrivialization_atlas` of a `vector_prebundle` to a `trivialization`. -/ def trivialization_of_mem_pretrivialization_atlas (a : vector_prebundle R F E) - {e : pretrivialization F (π E)} (he : e ∈ a.pretrivialization_atlas) : - @trivialization B F _ _ _ a.total_space_topology (π E) := + {e : pretrivialization F (π F E)} (he : e ∈ a.pretrivialization_atlas) : + @trivialization B F _ _ _ a.total_space_topology (π F E) := a.to_fiber_prebundle.trivialization_of_mem_pretrivialization_atlas he lemma linear_of_mem_pretrivialization_atlas (a : vector_prebundle R F E) - {e : pretrivialization F (π E)} (he : e ∈ a.pretrivialization_atlas) : + {e : pretrivialization F (π F E)} (he : e ∈ a.pretrivialization_atlas) : @trivialization.is_linear R B F _ _ _ _ a.total_space_topology _ _ _ _ (trivialization_of_mem_pretrivialization_atlas a he) := { linear := (a.pretrivialization_linear' e he).linear } @@ -839,15 +840,15 @@ lemma linear_of_mem_pretrivialization_atlas (a : vector_prebundle R F E) variable (a : vector_prebundle R F E) lemma mem_trivialization_at_source (b : B) (x : E b) : - total_space_mk b x ∈ (a.pretrivialization_at b).source := + total_space.mk b x ∈ (a.pretrivialization_at b).source := a.to_fiber_prebundle.mem_trivialization_at_source b x @[simp] lemma total_space_mk_preimage_source (b : B) : - (total_space_mk b) ⁻¹' (a.pretrivialization_at b).source = univ := + (total_space.mk b) ⁻¹' (a.pretrivialization_at b).source = univ := a.to_fiber_prebundle.total_space_mk_preimage_source b @[continuity] lemma continuous_total_space_mk (b : B) : - @continuous _ _ _ a.total_space_topology (total_space_mk b) := + @continuous _ _ _ a.total_space_topology (total_space.mk b) := a.to_fiber_prebundle.continuous_total_space_mk b /-- Make a `fiber_bundle` from a `vector_prebundle`; auxiliary construction for @@ -884,10 +885,10 @@ variables {𝕜₁ 𝕜₂ : Type*} [nontrivially_normed_field 𝕜₁] [nontriv variables {σ : 𝕜₁ →+* 𝕜₂} variables {B' : Type*} [topological_space B'] -variables [normed_space 𝕜₁ F] [Π x, module 𝕜₁ (E x)] [topological_space (total_space E)] +variables [normed_space 𝕜₁ F] [Π x, module 𝕜₁ (E x)] [topological_space (total_space F E)] variables {F' : Type*} [normed_add_comm_group F'] [normed_space 𝕜₂ F'] {E' : B' → Type*} [Π x, add_comm_monoid (E' x)] [Π x, module 𝕜₂ (E' x)] - [topological_space (total_space E')] + [topological_space (total_space F' E')] variables [fiber_bundle F E] [vector_bundle 𝕜₁ F E] variables [Π x, topological_space (E' x)] [fiber_bundle F' E'] [vector_bundle 𝕜₂ F' E'] variables (F E F' E') diff --git a/src/topology/vector_bundle/constructions.lean b/src/topology/vector_bundle/constructions.lean index 08067b91091e8..210ae668d9ec6 100644 --- a/src/topology/vector_bundle/constructions.lean +++ b/src/topology/vector_bundle/constructions.lean @@ -75,15 +75,15 @@ end bundle.trivial section variables (𝕜 : Type*) {B : Type*} [nontrivially_normed_field 𝕜] [topological_space B] (F₁ : Type*) [normed_add_comm_group F₁] [normed_space 𝕜 F₁] - (E₁ : B → Type*) [topological_space (total_space E₁)] + (E₁ : B → Type*) [topological_space (total_space F₁ E₁)] (F₂ : Type*) [normed_add_comm_group F₂] [normed_space 𝕜 F₂] - (E₂ : B → Type*) [topological_space (total_space E₂)] + (E₂ : B → Type*) [topological_space (total_space F₂ E₂)] namespace trivialization variables {F₁ E₁ F₂ E₂} [Π x, add_comm_monoid (E₁ x)] [Π x, module 𝕜 (E₁ x)] [Π x, add_comm_monoid (E₂ x)] [Π x, module 𝕜 (E₂ x)] - (e₁ e₁' : trivialization F₁ (π E₁)) (e₂ e₂' : trivialization F₂ (π E₂)) + (e₁ e₁' : trivialization F₁ (π F₁ E₁)) (e₂ e₂' : trivialization F₂ (π F₂ E₂)) instance prod.is_linear [e₁.is_linear 𝕜] [e₂.is_linear 𝕜] : (e₁.prod e₂).is_linear 𝕜 := { linear := λ x ⟨h₁, h₂⟩, (((e₁.linear 𝕜 h₁).mk' _).prod_map ((e₂.linear 𝕜 h₂).mk' _)).is_linear } @@ -146,9 +146,9 @@ instance vector_bundle.prod [vector_bundle 𝕜 F₁ E₁] [vector_bundle 𝕜 variables {𝕜 F₁ E₁ F₂ E₂} -@[simp] lemma trivialization.continuous_linear_equiv_at_prod {e₁ : trivialization F₁ (π E₁)} - {e₂ : trivialization F₂ (π E₂)} [e₁.is_linear 𝕜] [e₂.is_linear 𝕜] {x : B} (hx₁ : x ∈ e₁.base_set) - (hx₂ : x ∈ e₂.base_set) : +@[simp] lemma trivialization.continuous_linear_equiv_at_prod {e₁ : trivialization F₁ (π F₁ E₁)} + {e₂ : trivialization F₂ (π F₂ E₂)} [e₁.is_linear 𝕜] [e₂.is_linear 𝕜] {x : B} + (hx₁ : x ∈ e₁.base_set) (hx₂ : x ∈ e₂.base_set) : (e₁.prod e₂).continuous_linear_equiv_at 𝕜 x ⟨hx₁, hx₂⟩ = (e₁.continuous_linear_equiv_at 𝕜 x hx₁).prod (e₂.continuous_linear_equiv_at 𝕜 x hx₂) := begin @@ -172,12 +172,12 @@ instance [semiring R] [∀ (x : B), add_comm_monoid (E x)] [∀ x, module R (E x ∀ (x : B'), module R ((f *ᵖ E) x) := by delta_instance bundle.pullback -variables {E F} [topological_space B'] [topological_space (total_space E)] +variables {E F} [topological_space B'] [topological_space (total_space F E)] [nontrivially_normed_field 𝕜] [normed_add_comm_group F] [normed_space 𝕜 F] [topological_space B] [∀ x, add_comm_monoid (E x)] [∀ x, module 𝕜 (E x)] {K : Type*} [continuous_map_class K B' B] -instance trivialization.pullback_linear (e : trivialization F (π E)) [e.is_linear 𝕜] (f : K) : +instance trivialization.pullback_linear (e : trivialization F (π F E)) [e.is_linear 𝕜] (f : K) : (@trivialization.pullback _ _ _ B' _ _ _ _ _ _ _ e f).is_linear 𝕜 := { linear := λ x h, e.linear 𝕜 h } diff --git a/src/topology/vector_bundle/hom.lean b/src/topology/vector_bundle/hom.lean index 68258be0aa117..7e09babbf0ef9 100644 --- a/src/topology/vector_bundle/hom.lean +++ b/src/topology/vector_bundle/hom.lean @@ -40,62 +40,39 @@ noncomputable theory open_locale bundle open bundle set continuous_linear_map -section defs -variables {𝕜₁ 𝕜₂ : Type*} [normed_field 𝕜₁] [normed_field 𝕜₂] -variables (σ : 𝕜₁ →+* 𝕜₂) -variables {B : Type*} -variables (F₁ : Type*) (E₁ : B → Type*) [Π x, add_comm_group (E₁ x)] [Π x, module 𝕜₁ (E₁ x)] -variables [Π x, topological_space (E₁ x)] -variables (F₂ : Type*) (E₂ : B → Type*) [Π x, add_comm_group (E₂ x)] [Π x, module 𝕜₂ (E₂ x)] -variables [Π x, topological_space (E₂ x)] - -include F₁ F₂ - --- In this definition we require the scalar rings `𝕜₁` and `𝕜₂` to be normed fields, although --- something much weaker (maybe `comm_semiring`) would suffice mathematically -- this is because of --- a typeclass inference bug with pi-types: --- https://leanprover.zulipchat.com/#narrow/stream/116395-maths/topic/vector.20bundles.20--.20typeclass.20inference.20issue -/-- The bundle of continuous `σ`-semilinear maps between the topological vector bundles `E₁` and -`E₂`. This is a type synonym for `λ x, E₁ x →SL[σ] E₂ x`. - -We intentionally add `F₁` and `F₂` as arguments to this type, so that instances on this type -(that depend on `F₁` and `F₂`) actually refer to `F₁` and `F₂`. -/ -@[derive inhabited, nolint unused_arguments] -protected def bundle.continuous_linear_map (x : B) : Type* := -E₁ x →SL[σ] E₂ x - -instance bundle.continuous_linear_map.add_monoid_hom_class (x : B) : - add_monoid_hom_class (bundle.continuous_linear_map σ F₁ E₁ F₂ E₂ x) (E₁ x) (E₂ x) := -by delta_instance bundle.continuous_linear_map - -variables [Π x, topological_add_group (E₂ x)] - -instance (x : B) : topological_space (bundle.continuous_linear_map σ F₁ E₁ F₂ E₂ x) := -by delta_instance bundle.continuous_linear_map - -instance (x : B) : add_comm_monoid (bundle.continuous_linear_map σ F₁ E₁ F₂ E₂ x) := -by delta_instance bundle.continuous_linear_map - -variables [∀ x, has_continuous_smul 𝕜₂ (E₂ x)] - -instance (x : B) : module 𝕜₂ (bundle.continuous_linear_map σ F₁ E₁ F₂ E₂ x) := -by delta_instance bundle.continuous_linear_map - -end defs - variables {𝕜₁ : Type*} [nontrivially_normed_field 𝕜₁] {𝕜₂ : Type*} [nontrivially_normed_field 𝕜₂] (σ : 𝕜₁ →+* 𝕜₂) [iσ : ring_hom_isometric σ] -variables {B : Type*} [topological_space B] +variables {B : Type*} -variables (F₁ : Type*) [normed_add_comm_group F₁] [normed_space 𝕜₁ F₁] +variables {F₁ : Type*} [normed_add_comm_group F₁] [normed_space 𝕜₁ F₁] (E₁ : B → Type*) [Π x, add_comm_group (E₁ x)] [Π x, module 𝕜₁ (E₁ x)] - [topological_space (total_space E₁)] -variables (F₂ : Type*) [normed_add_comm_group F₂][normed_space 𝕜₂ F₂] + [topological_space (total_space F₁ E₁)] +variables {F₂ : Type*} [normed_add_comm_group F₂] [normed_space 𝕜₂ F₂] (E₂ : B → Type*) [Π x, add_comm_group (E₂ x)] [Π x, module 𝕜₂ (E₂ x)] - [topological_space (total_space E₂)] + [topological_space (total_space F₂ E₂)] + +/-- A reducible type synonym for the bundle of continuous (semi)linear maps. For some reason, it +helps with instance search. + +Porting note: after the port is done, we may want to remove this definition. +-/ +@[reducible] +protected def bundle.continuous_linear_map [∀ x, topological_space (E₁ x)] + [∀ x, topological_space (E₂ x)] : Π x : B, Type* := +λ x, E₁ x →SL[σ] E₂ x + +-- Porting note: possibly remove after the port +instance bundle.continuous_linear_map.module [∀ x, topological_space (E₁ x)] + [∀ x, topological_space (E₂ x)] [∀ x, topological_add_group (E₂ x)] + [∀ x, has_continuous_const_smul 𝕜₂ (E₂ x)] : + ∀ x, module 𝕜₂ (bundle.continuous_linear_map σ E₁ E₂ x) := +λ _, infer_instance + +variables {E₁ E₂} -variables {F₁ E₁ F₂ E₂} (e₁ e₁' : trivialization F₁ (π E₁)) (e₂ e₂' : trivialization F₂ (π E₂)) +variables [topological_space B] (e₁ e₁' : trivialization F₁ (π F₁ E₁)) + (e₂ e₂' : trivialization F₂ (π F₂ E₂)) namespace pretrivialization @@ -149,7 +126,7 @@ continuous `σ`-semilinear maps from `E₁` to `E₂`. That is, the map which wi trivialization, after the bundle of continuous semilinear maps is equipped with the right topological vector bundle structure. -/ def continuous_linear_map : - pretrivialization (F₁ →SL[σ] F₂) (π (bundle.continuous_linear_map σ F₁ E₁ F₂ E₂)) := + pretrivialization (F₁ →SL[σ] F₂) (π (F₁ →SL[σ] F₂) (bundle.continuous_linear_map σ E₁ E₂)) := { to_fun := λ p, ⟨p.1, continuous_linear_map.comp (e₂.continuous_linear_map_at 𝕜₂ p.1) (p.2.comp (e₁.symmL 𝕜₁ p.1 : F₁ →L[𝕜₁] E₁ p.1) : F₁ →SL[σ] E₂ p.1)⟩, inv_fun := λ p, ⟨p.1, continuous_linear_map.comp (e₂.symmL 𝕜₂ p.1) @@ -179,6 +156,7 @@ def continuous_linear_map : include ita +-- porting note: todo: see if Lean 4 can generate this instance without a hint instance continuous_linear_map.is_linear [Π x, has_continuous_add (E₂ x)] [Π x, has_continuous_smul 𝕜₂ (E₂ x)] : (pretrivialization.continuous_linear_map σ e₁ e₂).is_linear 𝕜₂ := @@ -199,7 +177,7 @@ instance continuous_linear_map.is_linear omit ita lemma continuous_linear_map_apply - (p : total_space (bundle.continuous_linear_map σ F₁ E₁ F₂ E₂)) : + (p : total_space (F₁ →SL[σ] F₂) (λ x, E₁ x →SL[σ] E₂ x)) : (continuous_linear_map σ e₁ e₂) p = ⟨p.1, continuous_linear_map.comp (e₂.continuous_linear_map_at 𝕜₂ p.1) (p.2.comp (e₁.symmL 𝕜₁ p.1 : F₁ →L[𝕜₁] E₁ p.1) : F₁ →SL[σ] E₂ p.1)⟩ := @@ -224,8 +202,7 @@ end lemma continuous_linear_map_coord_change_apply (b : B) (hb : b ∈ (e₁.base_set ∩ e₂.base_set) ∩ (e₁'.base_set ∩ e₂'.base_set)) (L : F₁ →SL[σ] F₂) : continuous_linear_map_coord_change σ e₁ e₁' e₂ e₂' b L = - (continuous_linear_map σ e₁' e₂' - (total_space_mk b ((continuous_linear_map σ e₁ e₂).symm b L))).2 := + (continuous_linear_map σ e₁' e₂' ⟨b, ((continuous_linear_map σ e₁ e₂).symm b L)⟩).2 := begin ext v, simp_rw [continuous_linear_map_coord_change, continuous_linear_equiv.coe_coe, @@ -234,7 +211,6 @@ begin continuous_linear_map_apply, continuous_linear_map_symm_apply' σ e₁ e₂ hb.1, comp_apply, continuous_linear_equiv.coe_coe, continuous_linear_equiv.symm_symm, trivialization.continuous_linear_map_at_apply, trivialization.symmL_apply], - dsimp only [total_space_mk], rw [e₂.coord_changeL_apply e₂', e₁'.coord_changeL_apply e₁, e₁.coe_linear_map_at_of_mem hb.1.1, e₂'.coe_linear_map_at_of_mem hb.2.2], exacts [⟨hb.2.1, hb.1.1⟩, ⟨hb.1.2, hb.2.2⟩] @@ -255,10 +231,9 @@ include iσ `vector_bundle` instance, in which the pretrivializations are collated but no topology on the total space is yet provided). -/ def _root_.bundle.continuous_linear_map.vector_prebundle : - vector_prebundle 𝕜₂ (F₁ →SL[σ] F₂) - (bundle.continuous_linear_map σ F₁ E₁ F₂ E₂) := + vector_prebundle 𝕜₂ (F₁ →SL[σ] F₂) (bundle.continuous_linear_map σ E₁ E₂) := { pretrivialization_atlas := - {e | ∃ (e₁ : trivialization F₁ (π E₁)) (e₂ : trivialization F₂ (π E₂)) + {e | ∃ (e₁ : trivialization F₁ (π F₁ E₁)) (e₂ : trivialization F₂ (π F₂ E₂)) [mem_trivialization_atlas e₁] [mem_trivialization_atlas e₂], by exactI e = pretrivialization.continuous_linear_map σ e₁ e₂}, pretrivialization_linear' := begin @@ -279,7 +254,6 @@ def _root_.bundle.continuous_linear_map.vector_prebundle : total_space_mk_inducing := begin intros b, - dsimp [bundle.continuous_linear_map.topological_space, bundle.continuous_linear_map], let L₁ : E₁ b ≃L[𝕜₁] F₁ := (trivialization_at F₁ E₁ b).continuous_linear_equiv_at 𝕜₁ b (mem_base_set_trivialization_at _ _ _), let L₂ : E₂ b ≃L[𝕜₂] F₂ := (trivialization_at F₂ E₂ b).continuous_linear_equiv_at 𝕜₂ b @@ -298,19 +272,19 @@ def _root_.bundle.continuous_linear_map.vector_prebundle : /-- Topology on the total space of the continuous `σ`-semilinear_maps between two "normable" vector bundles over the same base. -/ instance bundle.continuous_linear_map.topological_space_total_space : - topological_space (total_space (bundle.continuous_linear_map σ F₁ E₁ F₂ E₂)) := + topological_space (total_space (F₁ →SL[σ] F₂) (bundle.continuous_linear_map σ E₁ E₂)) := (bundle.continuous_linear_map.vector_prebundle σ F₁ E₁ F₂ E₂).total_space_topology /-- The continuous `σ`-semilinear_maps between two vector bundles form a fiber bundle. -/ instance _root_.bundle.continuous_linear_map.fiber_bundle : - fiber_bundle (F₁ →SL[σ] F₂) (bundle.continuous_linear_map σ F₁ E₁ F₂ E₂) := + fiber_bundle (F₁ →SL[σ] F₂) (λ x, E₁ x →SL[σ] E₂ x) := (bundle.continuous_linear_map.vector_prebundle σ F₁ E₁ F₂ E₂).to_fiber_bundle /-- The continuous `σ`-semilinear_maps between two vector bundles form a vector bundle. -/ instance _root_.bundle.continuous_linear_map.vector_bundle : - vector_bundle 𝕜₂ (F₁ →SL[σ] F₂) (bundle.continuous_linear_map σ F₁ E₁ F₂ E₂) := + vector_bundle 𝕜₂ (F₁ →SL[σ] F₂) (bundle.continuous_linear_map σ E₁ E₂) := (bundle.continuous_linear_map.vector_prebundle σ F₁ E₁ F₂ E₂).to_vector_bundle @@ -323,12 +297,12 @@ include he₁ he₂ the induced trivialization for the continuous `σ`-semilinear maps from `E₁` to `E₂`, whose base set is `e₁.base_set ∩ e₂.base_set`. -/ def trivialization.continuous_linear_map : - trivialization (F₁ →SL[σ] F₂) (π (bundle.continuous_linear_map σ F₁ E₁ F₂ E₂)) := + trivialization (F₁ →SL[σ] F₂) (π (F₁ →SL[σ] F₂) (bundle.continuous_linear_map σ E₁ E₂)) := vector_prebundle.trivialization_of_mem_pretrivialization_atlas _ ⟨e₁, e₂, he₁, he₂, rfl⟩ instance _root_.bundle.continuous_linear_map.mem_trivialization_atlas : mem_trivialization_atlas (e₁.continuous_linear_map σ e₂ : - trivialization (F₁ →SL[σ] F₂) (π (bundle.continuous_linear_map σ F₁ E₁ F₂ E₂))) := + trivialization (F₁ →SL[σ] F₂) (π (F₁ →SL[σ] F₂) (bundle.continuous_linear_map σ E₁ E₂))) := { out := ⟨_, ⟨e₁, e₂, by apply_instance, by apply_instance, rfl⟩, rfl⟩ } variables {e₁ e₂} @@ -338,7 +312,7 @@ variables {e₁ e₂} rfl lemma trivialization.continuous_linear_map_apply - (p : total_space (bundle.continuous_linear_map σ F₁ E₁ F₂ E₂)) : + (p : total_space (F₁ →SL[σ] F₂) (bundle.continuous_linear_map σ E₁ E₂)) : e₁.continuous_linear_map σ e₂ p = ⟨p.1, (e₂.continuous_linear_map_at 𝕜₂ p.1 : _ →L[𝕜₂] _).comp (p.2.comp (e₁.symmL 𝕜₁ p.1 : F₁ →L[𝕜₁] E₁ p.1) : F₁ →SL[σ] E₂ p.1)⟩ := @@ -347,20 +321,20 @@ rfl omit he₁ he₂ lemma hom_trivialization_at_apply (x₀ : B) - (x : total_space (bundle.continuous_linear_map σ F₁ E₁ F₂ E₂)) : - trivialization_at (F₁ →SL[σ] F₂) (bundle.continuous_linear_map σ F₁ E₁ F₂ E₂) x₀ x = + (x : total_space (F₁ →SL[σ] F₂) (bundle.continuous_linear_map σ E₁ E₂)) : + trivialization_at (F₁ →SL[σ] F₂) (λ x, E₁ x →SL[σ] E₂ x) x₀ x = ⟨x.1, in_coordinates F₁ E₁ F₂ E₂ x₀ x.1 x₀ x.1 x.2⟩ := rfl @[simp, mfld_simps] lemma hom_trivialization_at_source (x₀ : B) : - (trivialization_at (F₁ →SL[σ] F₂) (bundle.continuous_linear_map σ F₁ E₁ F₂ E₂) x₀).source = - π (bundle.continuous_linear_map σ F₁ E₁ F₂ E₂) ⁻¹' + (trivialization_at (F₁ →SL[σ] F₂) (bundle.continuous_linear_map σ E₁ E₂) x₀).source = + π (F₁ →SL[σ] F₂) (bundle.continuous_linear_map σ E₁ E₂) ⁻¹' ((trivialization_at F₁ E₁ x₀).base_set ∩ (trivialization_at F₂ E₂ x₀).base_set) := rfl @[simp, mfld_simps] lemma hom_trivialization_at_target (x₀ : B) : - (trivialization_at (F₁ →SL[σ] F₂) (bundle.continuous_linear_map σ F₁ E₁ F₂ E₂) x₀).target = + (trivialization_at (F₁ →SL[σ] F₂) (λ x, E₁ x →SL[σ] E₂ x) x₀).target = ((trivialization_at F₁ E₁ x₀).base_set ∩ (trivialization_at F₂ E₂ x₀).base_set) ×ˢ set.univ := rfl From 728ef9dbb281241906f25cbeb30f90d83e0bb451 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Mon, 3 Jul 2023 06:18:05 +0000 Subject: [PATCH 04/61] chore(*): add mathlib4 synchronization comments (#19226) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.category.Mon.adjunctions` * `algebra.category.Semigroup.basic` * `algebraic_geometry.function_field` * `algebraic_geometry.morphisms.basic` * `algebraic_geometry.morphisms.open_immersion` * `algebraic_geometry.morphisms.universally_closed` * `analysis.complex.upper_half_plane.manifold` * `analysis.convex.cone.proper` * `category_theory.limits.constructions.over.basic` * `geometry.manifold.algebra.smooth_functions` * `geometry.manifold.derivation_bundle` * `number_theory.modular_forms.jacobi_theta.manifold` * `representation_theory.fdRep` * `representation_theory.invariants` * `ring_theory.witt_vector.isocrystal` * `set_theory.game.domineering` * `set_theory.game.short` * `set_theory.game.state` * `topology.homotopy.product` --- src/algebra/category/Mon/adjunctions.lean | 3 +++ src/algebra/category/Semigroup/basic.lean | 3 +++ src/algebraic_geometry/function_field.lean | 3 +++ src/algebraic_geometry/morphisms/basic.lean | 3 +++ src/algebraic_geometry/morphisms/open_immersion.lean | 3 +++ src/algebraic_geometry/morphisms/universally_closed.lean | 3 +++ src/analysis/complex/upper_half_plane/manifold.lean | 3 +++ src/analysis/convex/cone/proper.lean | 3 +++ src/category_theory/limits/constructions/over/basic.lean | 3 +++ src/geometry/manifold/algebra/smooth_functions.lean | 3 +++ src/geometry/manifold/derivation_bundle.lean | 3 +++ src/number_theory/modular_forms/jacobi_theta/manifold.lean | 3 +++ src/representation_theory/fdRep.lean | 3 +++ src/representation_theory/invariants.lean | 3 +++ src/ring_theory/witt_vector/isocrystal.lean | 3 +++ src/set_theory/game/domineering.lean | 3 +++ src/set_theory/game/short.lean | 3 +++ src/set_theory/game/state.lean | 3 +++ src/topology/homotopy/product.lean | 3 +++ 19 files changed, 57 insertions(+) diff --git a/src/algebra/category/Mon/adjunctions.lean b/src/algebra/category/Mon/adjunctions.lean index 54e756448f059..d3e500532d249 100644 --- a/src/algebra/category/Mon/adjunctions.lean +++ b/src/algebra/category/Mon/adjunctions.lean @@ -11,6 +11,9 @@ import algebra.free_monoid.basic /-! # Adjunctions regarding the category of monoids +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves the adjunction between adjoining a unit to a semigroup and the forgetful functor from monoids to semigroups. diff --git a/src/algebra/category/Semigroup/basic.lean b/src/algebra/category/Semigroup/basic.lean index 65bb2df05eb96..7e8666e59eda3 100644 --- a/src/algebra/category/Semigroup/basic.lean +++ b/src/algebra/category/Semigroup/basic.lean @@ -12,6 +12,9 @@ import category_theory.elementwise /-! # Category instances for has_mul, has_add, semigroup and add_semigroup +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We introduce the bundled categories: * `Magma` * `AddMagma` diff --git a/src/algebraic_geometry/function_field.lean b/src/algebraic_geometry/function_field.lean index 6e8f90ac361b0..628ccd684a434 100644 --- a/src/algebraic_geometry/function_field.lean +++ b/src/algebraic_geometry/function_field.lean @@ -8,6 +8,9 @@ import algebraic_geometry.properties /-! # Function field of integral schemes +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define the function field of an irreducible scheme as the stalk of the generic point. This is a field when the scheme is integral. diff --git a/src/algebraic_geometry/morphisms/basic.lean b/src/algebraic_geometry/morphisms/basic.lean index 322daeb6e9816..5ff036cda8b73 100644 --- a/src/algebraic_geometry/morphisms/basic.lean +++ b/src/algebraic_geometry/morphisms/basic.lean @@ -10,6 +10,9 @@ import category_theory.morphism_property /-! # Properties of morphisms between Schemes +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We provide the basic framework for talking about properties of morphisms between Schemes. A `morphism_property Scheme` is a predicate on morphisms between schemes, and an diff --git a/src/algebraic_geometry/morphisms/open_immersion.lean b/src/algebraic_geometry/morphisms/open_immersion.lean index 62465fea56bb8..4ae87dfbdce6e 100644 --- a/src/algebraic_geometry/morphisms/open_immersion.lean +++ b/src/algebraic_geometry/morphisms/open_immersion.lean @@ -10,6 +10,9 @@ import algebraic_geometry.morphisms.basic # Open immersions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A morphism is an open immersions if the underlying map of spaces is an open embedding `f : X ⟶ U ⊆ Y`, and the sheaf map `Y(V) ⟶ f _* X(V)` is an iso for each `V ⊆ U`. diff --git a/src/algebraic_geometry/morphisms/universally_closed.lean b/src/algebraic_geometry/morphisms/universally_closed.lean index 036ff1844d1d9..1d53bcffeb4ca 100644 --- a/src/algebraic_geometry/morphisms/universally_closed.lean +++ b/src/algebraic_geometry/morphisms/universally_closed.lean @@ -9,6 +9,9 @@ import topology.local_at_target /-! # Universally closed morphism +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A morphism of schemes `f : X ⟶ Y` is universally closed if `X ×[Y] Y' ⟶ Y'` is a closed map for all base change `Y' ⟶ Y`. diff --git a/src/analysis/complex/upper_half_plane/manifold.lean b/src/analysis/complex/upper_half_plane/manifold.lean index b5a8b7ec6de4b..0c070195cc5dd 100644 --- a/src/analysis/complex/upper_half_plane/manifold.lean +++ b/src/analysis/complex/upper_half_plane/manifold.lean @@ -8,6 +8,9 @@ import geometry.manifold.cont_mdiff_mfderiv /-! # Manifold structure on the upper half plane. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define the complex manifold structure on the upper half-plane. -/ diff --git a/src/analysis/convex/cone/proper.lean b/src/analysis/convex/cone/proper.lean index 9babbda4c6ef3..8a76a074f3ae9 100644 --- a/src/analysis/convex/cone/proper.lean +++ b/src/analysis/convex/cone/proper.lean @@ -9,6 +9,9 @@ import analysis.inner_product_space.adjoint /-! # Proper cones +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define a proper cone as a nonempty, closed, convex cone. Proper cones are used in defining conic programs which generalize linear programs. A linear program is a conic program for the positive cone. We then prove Farkas' lemma for conic programs following the proof in the reference below. diff --git a/src/category_theory/limits/constructions/over/basic.lean b/src/category_theory/limits/constructions/over/basic.lean index e4085355f2824..9e71f7f3ed156 100644 --- a/src/category_theory/limits/constructions/over/basic.lean +++ b/src/category_theory/limits/constructions/over/basic.lean @@ -12,6 +12,9 @@ import category_theory.limits.constructions.equalizers /-! # Limits in the over category +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Declare instances for limits in the over category: If `C` has finite wide pullbacks, `over B` has finite limits, and if `C` has arbitrary wide pullbacks then `over B` has limits. -/ diff --git a/src/geometry/manifold/algebra/smooth_functions.lean b/src/geometry/manifold/algebra/smooth_functions.lean index 932b3e776ebf1..47e08e8c88e99 100644 --- a/src/geometry/manifold/algebra/smooth_functions.lean +++ b/src/geometry/manifold/algebra/smooth_functions.lean @@ -9,6 +9,9 @@ import geometry.manifold.algebra.structures /-! # Algebraic structures over smooth functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, we define instances of algebraic structures over smooth functions. -/ diff --git a/src/geometry/manifold/derivation_bundle.lean b/src/geometry/manifold/derivation_bundle.lean index 71f103747b268..d4ad6affb65b6 100644 --- a/src/geometry/manifold/derivation_bundle.lean +++ b/src/geometry/manifold/derivation_bundle.lean @@ -11,6 +11,9 @@ import ring_theory.derivation.basic # Derivation bundle +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define the derivations at a point of a manifold on the algebra of smooth fuctions. Moreover, we define the differential of a function in terms of derivations. diff --git a/src/number_theory/modular_forms/jacobi_theta/manifold.lean b/src/number_theory/modular_forms/jacobi_theta/manifold.lean index 67094c9757bca..81a9584cd7556 100644 --- a/src/number_theory/modular_forms/jacobi_theta/manifold.lean +++ b/src/number_theory/modular_forms/jacobi_theta/manifold.lean @@ -9,6 +9,9 @@ import analysis.complex.upper_half_plane.manifold /-! # Manifold differentiability of the Jacobi's theta function +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we reformulate differentiability of the Jacobi's theta function in terms of manifold differentiability. diff --git a/src/representation_theory/fdRep.lean b/src/representation_theory/fdRep.lean index 213f589b738dd..70cef6ac6b148 100644 --- a/src/representation_theory/fdRep.lean +++ b/src/representation_theory/fdRep.lean @@ -11,6 +11,9 @@ import representation_theory.basic /-! # `fdRep k G` is the category of finite dimensional `k`-linear representations of `G`. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + If `V : fdRep k G`, there is a coercion that allows you to treat `V` as a type, and this type comes equipped with `module k V` and `finite_dimensional k V` instances. Also `V.ρ` gives the homomorphism `G →* (V →ₗ[k] V)`. diff --git a/src/representation_theory/invariants.lean b/src/representation_theory/invariants.lean index 465fc5df768ae..5a9380d2cc6c3 100644 --- a/src/representation_theory/invariants.lean +++ b/src/representation_theory/invariants.lean @@ -9,6 +9,9 @@ import representation_theory.fdRep /-! # Subspace of invariants a group representation +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file introduces the subspace of invariants of a group representation and proves basic results about it. The main tool used is the average of all elements of the group, seen as an element of diff --git a/src/ring_theory/witt_vector/isocrystal.lean b/src/ring_theory/witt_vector/isocrystal.lean index bc06cef87421a..02374a808e1f4 100644 --- a/src/ring_theory/witt_vector/isocrystal.lean +++ b/src/ring_theory/witt_vector/isocrystal.lean @@ -10,6 +10,9 @@ import ring_theory.witt_vector.frobenius_fraction_field ## F-isocrystals over a perfect field +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + When `k` is an integral domain, so is `𝕎 k`, and we can consider its field of fractions `K(p, k)`. The endomorphism `witt_vector.frobenius` lifts to `φ : K(p, k) → K(p, k)`; if `k` is perfect, `φ` is an automorphism. diff --git a/src/set_theory/game/domineering.lean b/src/set_theory/game/domineering.lean index 3190d0fa99854..43958be820b70 100644 --- a/src/set_theory/game/domineering.lean +++ b/src/set_theory/game/domineering.lean @@ -8,6 +8,9 @@ import set_theory.game.state /-! # Domineering as a combinatorial game. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define the game of Domineering, played on a chessboard of arbitrary shape (possibly even disconnected). Left moves by placing a domino vertically, while Right moves by placing a domino horizontally. diff --git a/src/set_theory/game/short.lean b/src/set_theory/game/short.lean index 8d2bb319530e5..6b5867a25770c 100644 --- a/src/set_theory/game/short.lean +++ b/src/set_theory/game/short.lean @@ -10,6 +10,9 @@ import set_theory.game.birthday /-! # Short games +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A combinatorial game is `short` [Conway, ch.9][conway2001] if it has only finitely many positions. In particular, this means there is a finite set of moves at every point. diff --git a/src/set_theory/game/state.lean b/src/set_theory/game/state.lean index 14f118bc5685d..27f9ca055dce5 100644 --- a/src/set_theory/game/state.lean +++ b/src/set_theory/game/state.lean @@ -8,6 +8,9 @@ import set_theory.game.short /-! # Games described via "the state of the board". +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We provide a simple mechanism for constructing combinatorial (pre-)games, by describing "the state of the board", and providing an upper bound on the number of turns remaining. diff --git a/src/topology/homotopy/product.lean b/src/topology/homotopy/product.lean index 06b9d0449ed71..bc5ccd0df73f3 100644 --- a/src/topology/homotopy/product.lean +++ b/src/topology/homotopy/product.lean @@ -9,6 +9,9 @@ import topology.homotopy.path /-! # Product of homotopies +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, we introduce definitions for the product of homotopies. We show that the products of relative homotopies are still relative homotopies. Finally, we specialize to the case From 5dc6092d09e5e489106865241986f7f2ad28d4c8 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Tue, 4 Jul 2023 07:03:10 +0000 Subject: [PATCH 05/61] chore(topology/sheaves): revert universe generalizations from #19153 (#19230) This reverts commit 13361559d66b84f80b6d5a1c4a26aa5054766725. These are just too difficult to forward port as is because of the `max u v =?= max u ?v` issue https://github.com/leanprover/lean4/issues/2297. We have another candidate approach to this, using a new `UnivLE` typeclass, and I would prefer if we investigated that without the pressure of the port at the same time. This will delay @hrmacbeth's plans to define meromorphic functions, perhaps. Co-authored-by: Scott Morrison --- .../morphisms/quasi_compact.lean | 4 +- .../morphisms/quasi_separated.lean | 6 +-- src/algebraic_geometry/ringed_space.lean | 4 +- src/algebraic_geometry/structure_sheaf.lean | 4 +- src/category_theory/limits/shapes/types.lean | 19 +-------- src/topology/sheaves/forget.lean | 28 ++++++------- src/topology/sheaves/local_predicate.lean | 40 ++++++++----------- src/topology/sheaves/presheaf.lean | 20 +++++----- .../sheaves/presheaf_of_functions.lean | 8 ++-- .../sheaf_condition/unique_gluing.lean | 29 +++++++------- src/topology/sheaves/sheaf_of_functions.lean | 8 ++-- src/topology/sheaves/stalks.lean | 6 +-- 12 files changed, 74 insertions(+), 102 deletions(-) diff --git a/src/algebraic_geometry/morphisms/quasi_compact.lean b/src/algebraic_geometry/morphisms/quasi_compact.lean index 5da4cb3432cfd..b3281898a3059 100644 --- a/src/algebraic_geometry/morphisms/quasi_compact.lean +++ b/src/algebraic_geometry/morphisms/quasi_compact.lean @@ -298,7 +298,7 @@ end /-- If `x : Γ(X, U)` is zero on `D(f)` for some `f : Γ(X, U)`, and `U` is quasi-compact, then `f ^ n * x = 0` for some `n`. -/ -lemma exists_pow_mul_eq_zero_of_res_basic_open_eq_zero_of_is_compact (X : Scheme.{u}) +lemma exists_pow_mul_eq_zero_of_res_basic_open_eq_zero_of_is_compact (X : Scheme) {U : opens X.carrier} (hU : is_compact U.1) (x f : X.presheaf.obj (op U)) (H : x |_ X.basic_open f = 0) : ∃ n : ℕ, f ^ n * x = 0 := @@ -322,7 +322,7 @@ begin use finset.univ.sup n, suffices : ∀ (i : s), X.presheaf.map (hom_of_le (h₁ i)).op (f ^ (finset.univ.sup n) * x) = 0, { subst e, - apply Top.sheaf.eq_of_locally_eq.{(u+1) u} X.sheaf (λ (i : s), (i : opens X.carrier)), + apply X.sheaf.eq_of_locally_eq (λ (i : s), (i : opens X.carrier)), intro i, rw map_zero, apply this }, diff --git a/src/algebraic_geometry/morphisms/quasi_separated.lean b/src/algebraic_geometry/morphisms/quasi_separated.lean index f00888c66f0ba..8b75d385dcf82 100644 --- a/src/algebraic_geometry/morphisms/quasi_separated.lean +++ b/src/algebraic_geometry/morphisms/quasi_separated.lean @@ -345,7 +345,7 @@ begin exact e end -lemma exists_eq_pow_mul_of_is_compact_of_is_quasi_separated (X : Scheme.{u}) +lemma exists_eq_pow_mul_of_is_compact_of_is_quasi_separated (X : Scheme) (U : opens X.carrier) (hU : is_compact U.1) (hU' : is_quasi_separated U.1) (f : X.presheaf.obj (op U)) (x : X.presheaf.obj (op $ X.basic_open f)) : ∃ (n : ℕ) (y : X.presheaf.obj (op U)), y |_ X.basic_open f = (f |_ X.basic_open f) ^ n * x := @@ -401,7 +401,7 @@ begin ((X.presheaf.map (hom_of_le le_sup_left).op f) ^ (finset.univ.sup n + n₂) * y₁) = X.presheaf.map (hom_of_le $ inf_le_right).op ((X.presheaf.map (hom_of_le le_sup_right).op f) ^ (finset.univ.sup n + n₁) * y₂), - { fapply Top.sheaf.eq_of_locally_eq'.{(u+1) u} X.sheaf (λ i : s, i.1.1), + { fapply X.sheaf.eq_of_locally_eq' (λ i : s, i.1.1), { refine λ i, hom_of_le _, erw hs, exact le_supr _ _ }, { exact le_of_eq hs }, { intro i, @@ -419,7 +419,7 @@ begin -- By the sheaf condition, since `f ^ (n + n₂) * y₁ = f ^ (n + n₁) * y₂`, it can be glued into -- the desired section on `S ∪ U`. use (X.sheaf.obj_sup_iso_prod_eq_locus S U.1).inv ⟨⟨_ * _, _ * _⟩, this⟩, - refine Top.sheaf.eq_of_locally_eq₂.{(u+1) u} X.sheaf + refine X.sheaf.eq_of_locally_eq₂ (hom_of_le (_ : X.basic_open (X.presheaf.map (hom_of_le le_sup_left).op f) ≤ _)) (hom_of_le (_ : X.basic_open (X.presheaf.map (hom_of_le le_sup_right).op f) ≤ _)) _ _ _ _ _, { rw X.basic_open_res, exact inf_le_right }, diff --git a/src/algebraic_geometry/ringed_space.lean b/src/algebraic_geometry/ringed_space.lean index 268e8c7f83277..ff9354559ca56 100644 --- a/src/algebraic_geometry/ringed_space.lean +++ b/src/algebraic_geometry/ringed_space.lean @@ -76,7 +76,7 @@ begin -- Let `g x` denote the inverse of `f` in `U x`. choose g hg using λ x : U, is_unit.exists_right_inv (h_unit x), -- We claim that these local inverses glue together to a global inverse of `f`. - obtain ⟨gl, gl_spec, -⟩ := Top.sheaf.exists_unique_gluing'.{(v+1) v} X.sheaf V U iVU hcover g _, + obtain ⟨gl, gl_spec, -⟩ := X.sheaf.exists_unique_gluing' V U iVU hcover g _, swap, { intros x y, apply section_ext X.sheaf (V x ⊓ V y), @@ -89,7 +89,7 @@ begin congr_arg (X.presheaf.germ (⟨z, hzVy⟩ : V y)) (hg y), ring_hom.map_one, ring_hom.map_one] }, apply is_unit_of_mul_eq_one f gl, - apply Top.sheaf.eq_of_locally_eq'.{(v+1) v} X.sheaf V U iVU hcover, + apply X.sheaf.eq_of_locally_eq' V U iVU hcover, intro i, rw [ring_hom.map_one, ring_hom.map_mul, gl_spec], exact hg i, diff --git a/src/algebraic_geometry/structure_sheaf.lean b/src/algebraic_geometry/structure_sheaf.lean index 799927233cb81..1f2848001ed88 100644 --- a/src/algebraic_geometry/structure_sheaf.lean +++ b/src/algebraic_geometry/structure_sheaf.lean @@ -772,10 +772,10 @@ begin rw to_basic_open_mk', -- Since the structure sheaf is a sheaf, we can show the desired equality locally. - -- Annoyingly, `sheaf.eq_of_locally_eq'` requires an open cover indexed by a *type*, so we need to + -- Annoyingly, `sheaf.eq_of_locally_eq` requires an open cover indexed by a *type*, so we need to -- coerce our finset `t` to a type first. let tt := ((t : set (basic_open f)) : Type u), - apply Top.sheaf.eq_of_locally_eq'.{(u+1) u} (structure_sheaf R) + apply (structure_sheaf R).eq_of_locally_eq' (λ i : tt, basic_open (h i)) (basic_open f) (λ i : tt, iDh i), { -- This feels a little redundant, since already have `ht_cover` as a hypothesis -- Unfortunately, `ht_cover` uses a bounded union over the set `t`, while here we have the diff --git a/src/category_theory/limits/shapes/types.lean b/src/category_theory/limits/shapes/types.lean index 6439eb4210fb4..4666f03ce1a24 100644 --- a/src/category_theory/limits/shapes/types.lean +++ b/src/category_theory/limits/shapes/types.lean @@ -51,28 +51,13 @@ local attribute [tidy] tactic.discrete_cases /-- A restatement of `types.lift_π_apply` that uses `pi.π` and `pi.lift`. -/ @[simp] lemma pi_lift_π_apply - {β : Type v} (f : β → Type max v u) {P : Type max v u} (s : Π b, P ⟶ f b) (b : β) (x : P) : - (pi.π f b : (∏ f) → f b) (@pi.lift β _ _ f _ P s x) = s b x := -congr_fun (limit.lift_π (fan.mk P s) ⟨b⟩) x - -/-- A restatement of `types.lift_π_apply` that uses `pi.π` and `pi.lift`, -with specialized universes. -/ -@[simp] -lemma pi_lift_π_apply' - {β : Type v} (f : β → Type v) {P : Type v} (s : Π b, P ⟶ f b) (b : β) (x : P) : + {β : Type u} (f : β → Type u) {P : Type u} (s : Π b, P ⟶ f b) (b : β) (x : P) : (pi.π f b : (∏ f) → f b) (@pi.lift β _ _ f _ P s x) = s b x := congr_fun (limit.lift_π (fan.mk P s) ⟨b⟩) x /-- A restatement of `types.map_π_apply` that uses `pi.π` and `pi.map`. -/ @[simp] -lemma pi_map_π_apply {β : Type v} {f g : β → Type max v u} (α : Π j, f j ⟶ g j) (b : β) (x) : - (pi.π g b : (∏ g) → g b) (pi.map α x) = α b ((pi.π f b : (∏ f) → f b) x) := -limit.map_π_apply _ _ _ - -/-- A restatement of `types.map_π_apply` that uses `pi.π` and `pi.map`, -with specialized universes. -/ -@[simp] -lemma pi_map_π_apply' {β : Type v} {f g : β → Type v} (α : Π j, f j ⟶ g j) (b : β) (x) : +lemma pi_map_π_apply {β : Type u} {f g : β → Type u} (α : Π j, f j ⟶ g j) (b : β) (x) : (pi.π g b : (∏ g) → g b) (pi.map α x) = α b ((pi.π f b : (∏ f) → f b) x) := limit.map_π_apply _ _ _ diff --git a/src/topology/sheaves/forget.lean b/src/topology/sheaves/forget.lean index f6fbde37a2ce5..3fb52ea35b299 100644 --- a/src/topology/sheaves/forget.lean +++ b/src/topology/sheaves/forget.lean @@ -42,13 +42,13 @@ namespace sheaf_condition open sheaf_condition_equalizer_products -universes v u₁ u₂ w +universes v u₁ u₂ -variables {C : Type u₁} [category.{v} C] [has_limits_of_size.{w w} C] -variables {D : Type u₂} [category.{v} D] [has_limits_of_size.{w w} D] -variables (G : C ⥤ D) [preserves_limits_of_size.{w w} G] -variables {X : Top.{w}} (F : presheaf C X) -variables {ι : Type w} (U : ι → opens X) +variables {C : Type u₁} [category.{v} C] [has_limits C] +variables {D : Type u₂} [category.{v} D] [has_limits D] +variables (G : C ⥤ D) [preserves_limits G] +variables {X : Top.{v}} (F : presheaf C X) +variables {ι : Type v} (U : ι → opens X) local attribute [reducible] diagram left_res right_res @@ -57,7 +57,7 @@ When `G` preserves limits, the sheaf condition diagram for `F` composed with `G` naturally isomorphic to the sheaf condition diagram for `F ⋙ G`. -/ def diagram_comp_preserves_limits : - diagram F U ⋙ G ≅ diagram.{w} (F ⋙ G) U := + diagram F U ⋙ G ≅ diagram.{v} (F ⋙ G) U := begin fapply nat_iso.of_components, rintro ⟨j⟩, @@ -85,7 +85,7 @@ When `G` preserves limits, the image under `G` of the sheaf condition fork for ` is the sheaf condition fork for `F ⋙ G`, postcomposed with the inverse of the natural isomorphism `diagram_comp_preserves_limits`. -/ -def map_cone_fork : G.map_cone (fork.{w} F U) ≅ +def map_cone_fork : G.map_cone (fork.{v} F U) ≅ (cones.postcompose (diagram_comp_preserves_limits G F U).inv).obj (fork (F ⋙ G) U) := cones.ext (iso.refl _) (λ j, begin @@ -102,17 +102,16 @@ end) end sheaf_condition -universes v u₁ u₂ w +universes v u₁ u₂ open sheaf_condition sheaf_condition_equalizer_products variables {C : Type u₁} [category.{v} C] {D : Type u₂} [category.{v} D] variables (G : C ⥤ D) variables [reflects_isomorphisms G] -variables [has_limits_of_size.{w w} C] [has_limits_of_size.{w w} D] - [preserves_limits_of_size.{w w} G] +variables [has_limits C] [has_limits D] [preserves_limits G] -variables {X : Top.{w}} (F : presheaf C X) +variables {X : Top.{v}} (F : presheaf C X) /-- If `G : C ⥤ D` is a functor which reflects isomorphisms and preserves limits @@ -172,7 +171,7 @@ begin -- image under `G` of the equalizer cone for the sheaf condition diagram. let c := fork (F ⋙ G) U, obtain ⟨hc⟩ := S U, - let d := G.map_cone (equalizer.fork (left_res.{w} F U) (right_res F U)), + let d := G.map_cone (equalizer.fork (left_res.{v} F U) (right_res F U)), letI := preserves_smallest_limits_of_preserves_limits G, have hd : is_limit d := preserves_limit.preserves (limit.is_limit _), -- Since both of these are limit cones @@ -183,8 +182,7 @@ begin -- introduced above. let d' := (cones.postcompose (diagram_comp_preserves_limits G F U).hom).obj d, have hd' : is_limit d' := - (is_limit.postcompose_hom_equiv - (diagram_comp_preserves_limits.{v u₁ u₂ w} G F U : _) d).symm hd, + (is_limit.postcompose_hom_equiv (diagram_comp_preserves_limits G F U : _) d).symm hd, -- Now everything works: we verify that `f` really is a morphism between these cones: let f' : c ⟶ d' := fork.mk_hom (G.map f) diff --git a/src/topology/sheaves/local_predicate.lean b/src/topology/sheaves/local_predicate.lean index 3b682723589bf..04659ff17fd1b 100644 --- a/src/topology/sheaves/local_predicate.lean +++ b/src/topology/sheaves/local_predicate.lean @@ -38,12 +38,12 @@ to the types in the ambient type family. We give conditions sufficient to show that this map is injective and/or surjective. -/ -universes v w +universe v noncomputable theory variables {X : Top.{v}} -variables (T : X → Type w) +variables (T : X → Type v) open topological_space open opposite @@ -69,12 +69,12 @@ variables (X) Continuity is a "prelocal" predicate on functions to a fixed topological space `T`. -/ @[simps] -def continuous_prelocal (T : Top.{w}) : prelocal_predicate (λ x : X, T) := +def continuous_prelocal (T : Top.{v}) : prelocal_predicate (λ x : X, T) := { pred := λ U f, continuous f, res := λ U V i f h, continuous.comp h (opens.open_embedding_of_le i.le).continuous, } /-- Satisfying the inhabited linter. -/ -instance inhabited_prelocal_predicate (T : Top.{w}) : inhabited (prelocal_predicate (λ x : X, T)) := +instance inhabited_prelocal_predicate (T : Top.{v}) : inhabited (prelocal_predicate (λ x : X, T)) := ⟨continuous_prelocal X T⟩ variables {X} @@ -150,7 +150,7 @@ lemma prelocal_predicate.sheafify_of {T : X → Type v} {P : prelocal_predicate The subpresheaf of dependent functions on `X` satisfying the "pre-local" predicate `P`. -/ @[simps] -def subpresheaf_to_Types (P : prelocal_predicate T) : presheaf (Type (max v w)) X := +def subpresheaf_to_Types (P : prelocal_predicate T) : presheaf (Type v) X := { obj := λ U, { f : Π x : unop U, T x // P.pred f }, map := λ U V i f, ⟨λ x, f.1 (i.unop x), P.res i.unop f.1 f.2⟩ }. @@ -211,27 +211,22 @@ end subpresheaf_to_Types The subsheaf of the sheaf of all dependently typed functions satisfying the local predicate `P`. -/ @[simps] -def subsheaf_to_Types (P : local_predicate T) : sheaf (Type max v w) X := +def subsheaf_to_Types (P : local_predicate T) : sheaf (Type v) X := ⟨subpresheaf_to_Types P.to_prelocal_predicate, subpresheaf_to_Types.is_sheaf P⟩ --- TODO There is further universe generalization to do here, --- but it depends on more difficult universe generalization in `topology.sheaves.stalks`. --- The aim is to replace `T'` with the more general `T` below. -variables {T' : X → Type v} - /-- There is a canonical map from the stalk to the original fiber, given by evaluating sections. -/ -def stalk_to_fiber (P : local_predicate T') (x : X) : - (subsheaf_to_Types P).presheaf.stalk x ⟶ T' x := +def stalk_to_fiber (P : local_predicate T) (x : X) : + (subsheaf_to_Types P).presheaf.stalk x ⟶ T x := begin refine colimit.desc _ - { X := T' x, ι := { app := λ U f, _, naturality' := _ } }, + { X := T x, ι := { app := λ U f, _, naturality' := _ } }, { exact f.1 ⟨x, (unop U).2⟩, }, { tidy, } end -@[simp] lemma stalk_to_fiber_germ (P : local_predicate T') (U : opens X) (x : U) (f) : +@[simp] lemma stalk_to_fiber_germ (P : local_predicate T) (U : opens X) (x : U) (f) : stalk_to_fiber P x ((subsheaf_to_Types P).presheaf.germ x f) = f.1 x := begin dsimp [presheaf.germ, stalk_to_fiber], @@ -244,8 +239,8 @@ end The `stalk_to_fiber` map is surjective at `x` if every point in the fiber `T x` has an allowed section passing through it. -/ -lemma stalk_to_fiber_surjective (P : local_predicate T') (x : X) - (w : ∀ (t : T' x), ∃ (U : open_nhds x) (f : Π y : U.1, T' y) (h : P.pred f), f ⟨x, U.2⟩ = t) : +lemma stalk_to_fiber_surjective (P : local_predicate T) (x : X) + (w : ∀ (t : T x), ∃ (U : open_nhds x) (f : Π y : U.1, T y) (h : P.pred f), f ⟨x, U.2⟩ = t) : function.surjective (stalk_to_fiber P x) := λ t, begin @@ -259,16 +254,16 @@ end The `stalk_to_fiber` map is injective at `x` if any two allowed sections which agree at `x` agree on some neighborhood of `x`. -/ -lemma stalk_to_fiber_injective (P : local_predicate T') (x : X) - (w : ∀ (U V : open_nhds x) (fU : Π y : U.1, T' y) (hU : P.pred fU) - (fV : Π y : V.1, T' y) (hV : P.pred fV) (e : fU ⟨x, U.2⟩ = fV ⟨x, V.2⟩), +lemma stalk_to_fiber_injective (P : local_predicate T) (x : X) + (w : ∀ (U V : open_nhds x) (fU : Π y : U.1, T y) (hU : P.pred fU) + (fV : Π y : V.1, T y) (hV : P.pred fV) (e : fU ⟨x, U.2⟩ = fV ⟨x, V.2⟩), ∃ (W : open_nhds x) (iU : W ⟶ U) (iV : W ⟶ V), ∀ (w : W.1), fU (iU w : U.1) = fV (iV w : V.1)) : function.injective (stalk_to_fiber P x) := λ tU tV h, begin -- We promise to provide all the ingredients of the proof later: let Q : - ∃ (W : (open_nhds x)ᵒᵖ) (s : Π w : (unop W).1, T' w) (hW : P.pred s), + ∃ (W : (open_nhds x)ᵒᵖ) (s : Π w : (unop W).1, T w) (hW : P.pred s), tU = (subsheaf_to_Types P).presheaf.germ ⟨x, (unop W).2⟩ ⟨s, hW⟩ ∧ tV = (subsheaf_to_Types P).presheaf.germ ⟨x, (unop W).2⟩ ⟨s, hW⟩ := _, { choose W s hW e using Q, @@ -288,9 +283,6 @@ begin colimit_sound iV.op (subtype.eq (funext w).symm)⟩, }, end --- TODO to continue generalizing universes here, we will need to deal with the issue noted --- at `presheaf_to_Top` (universe generalization for the Yoneda embedding). - /-- Some repackaging: the presheaf of functions satisfying `continuous_prelocal` is just the same thing as diff --git a/src/topology/sheaves/presheaf.lean b/src/topology/sheaves/presheaf.lean index ee4af1ab86ca9..e4c312be03265 100644 --- a/src/topology/sheaves/presheaf.lean +++ b/src/topology/sheaves/presheaf.lean @@ -221,7 +221,7 @@ def pushforward_map {X Y : Top.{w}} (f : X ⟶ Y) {ℱ 𝒢 : X.presheaf C} (α open category_theory.limits section pullback -variable [has_colimits_of_size.{w w} C] +variable [has_colimits C] noncomputable theory /-- @@ -231,17 +231,17 @@ This is defined in terms of left Kan extensions, which is just a fancy way of sa "take the colimits over the open sets whose preimage contains U". -/ @[simps] -def pullback_obj {X Y : Top.{w}} (f : X ⟶ Y) (ℱ : Y.presheaf C) : X.presheaf C := +def pullback_obj {X Y : Top.{v}} (f : X ⟶ Y) (ℱ : Y.presheaf C) : X.presheaf C := (Lan (opens.map f).op).obj ℱ /-- Pulling back along continuous maps is functorial. -/ -def pullback_map {X Y : Top.{w}} (f : X ⟶ Y) {ℱ 𝒢 : Y.presheaf C} (α : ℱ ⟶ 𝒢) : +def pullback_map {X Y : Top.{v}} (f : X ⟶ Y) {ℱ 𝒢 : Y.presheaf C} (α : ℱ ⟶ 𝒢) : pullback_obj f ℱ ⟶ pullback_obj f 𝒢 := (Lan (opens.map f).op).map α /-- If `f '' U` is open, then `f⁻¹ℱ U ≅ ℱ (f '' U)`. -/ @[simps] -def pullback_obj_obj_of_image_open {X Y : Top.{w}} (f : X ⟶ Y) (ℱ : Y.presheaf C) (U : opens X) +def pullback_obj_obj_of_image_open {X Y : Top.{v}} (f : X ⟶ Y) (ℱ : Y.presheaf C) (U : opens X) (H : is_open (f '' U)) : (pullback_obj f ℱ).obj (op U) ≅ ℱ.obj (op ⟨_, H⟩) := begin let x : costructured_arrow (opens.map f).op (op U) := begin @@ -263,7 +263,7 @@ begin end namespace pullback -variables {X Y : Top.{w}} (ℱ : Y.presheaf C) +variables {X Y : Top.{v}} (ℱ : Y.presheaf C) /-- The pullback along the identity is isomorphic to the original presheaf. -/ def id : pullback_obj (𝟙 _) ℱ ≅ ℱ := @@ -366,30 +366,30 @@ by simpa [pushforward_to_of_iso, equivalence.to_adjunction] end iso -variables (C) [has_colimits_of_size.{w w} C] +variables (C) [has_colimits C] /-- Pullback a presheaf on `Y` along a continuous map `f : X ⟶ Y`, obtaining a presheaf on `X`. -/ @[simps map_app] -def pullback {X Y : Top.{w}} (f : X ⟶ Y) : Y.presheaf C ⥤ X.presheaf C := Lan (opens.map f).op +def pullback {X Y : Top.{v}} (f : X ⟶ Y) : Y.presheaf C ⥤ X.presheaf C := Lan (opens.map f).op @[simp] lemma pullback_obj_eq_pullback_obj {C} [category C] [has_colimits C] {X Y : Top.{w}} (f : X ⟶ Y) (ℱ : Y.presheaf C) : (pullback C f).obj ℱ = pullback_obj f ℱ := rfl /-- The pullback and pushforward along a continuous map are adjoint to each other. -/ @[simps unit_app_app counit_app_app] -def pushforward_pullback_adjunction {X Y : Top.{w}} (f : X ⟶ Y) : +def pushforward_pullback_adjunction {X Y : Top.{v}} (f : X ⟶ Y) : pullback C f ⊣ pushforward C f := Lan.adjunction _ _ /-- Pulling back along a homeomorphism is the same as pushing forward along its inverse. -/ -def pullback_hom_iso_pushforward_inv {X Y : Top.{w}} (H : X ≅ Y) : +def pullback_hom_iso_pushforward_inv {X Y : Top.{v}} (H : X ≅ Y) : pullback C H.hom ≅ pushforward C H.inv := adjunction.left_adjoint_uniq (pushforward_pullback_adjunction C H.hom) (presheaf_equiv_of_iso C H.symm).to_adjunction /-- Pulling back along the inverse of a homeomorphism is the same as pushing forward along it. -/ -def pullback_inv_iso_pushforward_hom {X Y : Top.{w}} (H : X ≅ Y) : +def pullback_inv_iso_pushforward_hom {X Y : Top.{v}} (H : X ≅ Y) : pullback C H.inv ≅ pushforward C H.hom := adjunction.left_adjoint_uniq (pushforward_pullback_adjunction C H.inv) diff --git a/src/topology/sheaves/presheaf_of_functions.lean b/src/topology/sheaves/presheaf_of_functions.lean index a6c432a1dea51..07212f6643439 100644 --- a/src/topology/sheaves/presheaf_of_functions.lean +++ b/src/topology/sheaves/presheaf_of_functions.lean @@ -28,7 +28,7 @@ We construct some simple examples of presheaves of functions on a topological sp is the presheaf of rings of continuous complex-valued functions on `X`. -/ -universes v u w +universes v u open category_theory open topological_space @@ -42,7 +42,7 @@ variables (X : Top.{v}) The presheaf of dependently typed functions on `X`, with fibres given by a type family `T`. There is no requirement that the functions are continuous, here. -/ -def presheaf_to_Types (T : X → Type w) : X.presheaf (Type (max v w)) := +def presheaf_to_Types (T : X → Type v) : X.presheaf (Type v) := { obj := λ U, Π x : (unop U), T x, map := λ U V i g, λ (x : unop V), g (i.unop x), map_id' := λ U, by { ext g ⟨x, hx⟩, refl }, @@ -68,7 +68,7 @@ There is no requirement that the functions are continuous, here. -- We don't use `@[simps]` to generate the projection lemmas here, -- as it turns out to be useful to have `presheaf_to_Type_map` -- written as an equality of functions (rather than being applied to some argument). -def presheaf_to_Type (T : Type w) : X.presheaf (Type max v w) := +def presheaf_to_Type (T : Type v) : X.presheaf (Type v) := { obj := λ U, (unop U) → T, map := λ U V i g, g ∘ i.unop, map_id' := λ U, by { ext g ⟨x, hx⟩, refl }, @@ -86,8 +86,6 @@ rfl /-- The presheaf of continuous functions on `X` with values in fixed target topological space `T`. -/ --- TODO it may prove useful to generalize the universes here, --- but the definition would need to change. def presheaf_to_Top (T : Top.{v}) : X.presheaf (Type v) := (opens.to_Top X).op ⋙ (yoneda.obj T) diff --git a/src/topology/sheaves/sheaf_condition/unique_gluing.lean b/src/topology/sheaves/sheaf_condition/unique_gluing.lean index 543920fcc9b69..2cbe750fc2433 100644 --- a/src/topology/sheaves/sheaf_condition/unique_gluing.lean +++ b/src/topology/sheaves/sheaf_condition/unique_gluing.lean @@ -48,9 +48,9 @@ open topological_space open topological_space.opens open opposite -universes u v w +universes u v -variables {C : Type u} [category.{max w v} C] [concrete_category.{max w v} C] +variables {C : Type u} [category.{v} C] [concrete_category.{v} C] namespace Top @@ -60,7 +60,7 @@ section local attribute [instance] concrete_category.has_coe_to_sort concrete_category.has_coe_to_fun -variables {X : Top.{w}} (F : presheaf C X) {ι : Type w} (U : ι → opens X) +variables {X : Top.{v}} (F : presheaf C X) {ι : Type v} (U : ι → opens X) /-- A family of sections `sf` is compatible, if the restrictions of `sf i` and `sf j` to `U i ⊓ U j` @@ -85,14 +85,14 @@ We prove this to be equivalent to the usual one below in `is_sheaf_iff_is_sheaf_unique_gluing` -/ def is_sheaf_unique_gluing : Prop := -∀ ⦃ι : Type w⦄ (U : ι → opens X) (sf : Π i : ι, F.obj (op (U i))), +∀ ⦃ι : Type v⦄ (U : ι → opens X) (sf : Π i : ι, F.obj (op (U i))), is_compatible F U sf → ∃! s : F.obj (op (supr U)), is_gluing F U sf s end section type_valued -variables {X : Top.{w}} (F : presheaf (Type max w v) X) {ι : Type w} (U : ι → opens X) +variables {X : Top.{v}} (F : presheaf (Type v) X) {ι : Type v} (U : ι → opens X) /-- For presheaves of types, terms of `pi_opens F U` are just families of sections. @@ -100,7 +100,7 @@ For presheaves of types, terms of `pi_opens F U` are just families of sections. def pi_opens_iso_sections_family : pi_opens F U ≅ Π i : ι, F.obj (op (U i)) := limits.is_limit.cone_point_unique_up_to_iso (limit.is_limit (discrete.functor (λ i : ι, F.obj (op (U i))))) - ((types.product_limit_cone.{w (max w v)} (λ i : ι, F.obj (op (U i)))).is_limit) + ((types.product_limit_cone.{v v} (λ i : ι, F.obj (op (U i)))).is_limit) /-- Under the isomorphism `pi_opens_iso_sections_family`, compatibility of sections is the same @@ -112,8 +112,8 @@ lemma compatible_iff_left_res_eq_right_res (sf : pi_opens F U) : begin split ; intros h, { ext ⟨i, j⟩, - rw [left_res, types.limit.lift_π_apply, fan.mk_π_app, - right_res, types.limit.lift_π_apply, fan.mk_π_app], + rw [left_res, types.limit.lift_π_apply', fan.mk_π_app, + right_res, types.limit.lift_π_apply', fan.mk_π_app], exact h i j, }, { intros i j, convert congr_arg (limits.pi.π (λ p : ι × ι, F.obj (op (U p.1 ⊓ U p.2))) (i,j)) h, @@ -132,7 +132,7 @@ lemma is_gluing_iff_eq_res (sf : pi_opens F U) (s : F.obj (op (supr U))): begin split ; intros h, { ext ⟨i⟩, - rw [res, types.limit.lift_π_apply, fan.mk_π_app], + rw [res, types.limit.lift_π_apply', fan.mk_π_app], exact h i, }, { intro i, convert congr_arg (limits.pi.π (λ i : ι, F.obj (op (U i))) i) h, @@ -209,9 +209,8 @@ section local attribute [instance] concrete_category.has_coe_to_sort concrete_category.has_coe_to_fun -variables [has_limits_of_size.{w w} C] [reflects_isomorphisms (forget C)] - [preserves_limits_of_size.{w w} (forget C)] -variables {X : Top.{w}} (F : presheaf C X) {ι : Type w} (U : ι → opens X) +variables [has_limits C] [reflects_isomorphisms (forget C)] [preserves_limits (forget C)] +variables {X : Top.{v}} (F : presheaf C X) {ι : Type v} (U : ι → opens X) /-- For presheaves valued in a concrete category, whose forgetful functor reflects isomorphisms and @@ -236,10 +235,10 @@ section local attribute [instance] concrete_category.has_coe_to_sort concrete_category.has_coe_to_fun -variables [has_limits_of_size.{w w} C] [reflects_isomorphisms (concrete_category.forget C)] -variables [preserves_limits_of_size.{w w} (concrete_category.forget C)] +variables [has_limits C] [reflects_isomorphisms (concrete_category.forget C)] +variables [preserves_limits (concrete_category.forget C)] -variables {X : Top.{w}} (F : sheaf C X) {ι : Type w} (U : ι → opens X) +variables {X : Top.{v}} (F : sheaf C X) {ι : Type v} (U : ι → opens X) /-- A more convenient way of obtaining a unique gluing of sections for a sheaf. diff --git a/src/topology/sheaves/sheaf_of_functions.lean b/src/topology/sheaves/sheaf_of_functions.lean index 1ffac5ae6b1d9..8b3e8bcc0e660 100644 --- a/src/topology/sheaves/sheaf_of_functions.lean +++ b/src/topology/sheaves/sheaf_of_functions.lean @@ -33,11 +33,11 @@ open category_theory.limits open topological_space open topological_space.opens -universes v u +universe u noncomputable theory -variables (X : Top.{v}) +variables (X : Top.{u}) open Top @@ -101,13 +101,13 @@ namespace Top The sheaf of not-necessarily-continuous functions on `X` with values in type family `T : X → Type u`. -/ -def sheaf_to_Types (T : X → Type u) : sheaf (Type max v u) X := +def sheaf_to_Types (T : X → Type u) : sheaf (Type u) X := ⟨presheaf_to_Types X T, presheaf.to_Types_is_sheaf _ _⟩ /-- The sheaf of not-necessarily-continuous functions on `X` with values in a type `T`. -/ -def sheaf_to_Type (T : Type u) : sheaf (Type max v u) X := +def sheaf_to_Type (T : Type u) : sheaf (Type u) X := ⟨presheaf_to_Type X T, presheaf.to_Type_is_sheaf _ _⟩ end Top diff --git a/src/topology/sheaves/stalks.lean b/src/topology/sheaves/stalks.lean index c8564219b13bb..ea3a7426b187a 100644 --- a/src/topology/sheaves/stalks.lean +++ b/src/topology/sheaves/stalks.lean @@ -419,7 +419,7 @@ begin choose V m i₁ i₂ heq using λ x : U, F.presheaf.germ_eq x.1 x.2 x.2 s t (h x), -- Since `F` is a sheaf, we can prove the equality locally, if we can show that these -- neighborhoods form a cover of `U`. - apply Top.sheaf.eq_of_locally_eq'.{u v v} F V U i₁, + apply F.eq_of_locally_eq' V U i₁, { intros x hxU, rw [opens.mem_supr], exact ⟨⟨x, hxU⟩, m ⟨x, hxU⟩⟩ }, @@ -488,9 +488,9 @@ begin rw [opens.mem_supr], exact ⟨⟨x, hxU⟩, mV ⟨x, hxU⟩⟩ }, -- Since `F` is a sheaf, we can glue all the local preimages together to get a global preimage. - obtain ⟨s, s_spec, -⟩ := Top.sheaf.exists_unique_gluing'.{u v v} F V U iVU V_cover sf _, + obtain ⟨s, s_spec, -⟩ := F.exists_unique_gluing' V U iVU V_cover sf _, { use s, - apply Top.sheaf.eq_of_locally_eq'.{u v v} G V U iVU V_cover, + apply G.eq_of_locally_eq' V U iVU V_cover, intro x, rw [← comp_apply, ← f.1.naturality, comp_apply, s_spec, heq] }, { intros x y, From 8905e5ed90859939681a725b00f6063e65096d95 Mon Sep 17 00:00:00 2001 From: Anatole Dedecker Date: Thu, 6 Jul 2023 13:05:39 +0000 Subject: [PATCH 06/61] feat(topology/algebra/module/strong_topology): golf `arrow_congrSL` introduced in #19107 (#19128) I added more general definitions `precomp` and `postcomp` for expressing that (pre/post)-composing by a *fixed* continuous linear maps is continuous. These were planned about a year ago when I defined the strong topology and follow from [uniform_on_fun.precomp_uniform_continuous](https://leanprover-community.github.io/mathlib_docs/topology/uniform_space/uniform_convergence_topology.html#uniform_on_fun.precomp_uniform_continuous) and [uniform_on_fun.postcomp_uniform_continuous](https://leanprover-community.github.io/mathlib_docs/topology/uniform_space/uniform_convergence_topology.html#uniform_on_fun.postcomp_uniform_continuous). The proof of continuity of `arrow_congrSL` is a direct consequence of these, so we don't have to do it by hand. This is not really a "golf" since I added more lines than I removed, but these more general constructions will be needed at some point anyway (my use case was distribution theory) so I'm doing some proactive golfing :smile:. --- src/analysis/convolution.lean | 4 +- src/geometry/manifold/vector_bundle/hom.lean | 2 +- .../algebra/module/strong_topology.lean | 88 ++++++++++++------- src/topology/vector_bundle/hom.lean | 8 +- 4 files changed, 62 insertions(+), 40 deletions(-) diff --git a/src/analysis/convolution.lean b/src/analysis/convolution.lean index 74b30a8db71c0..199a713352ea1 100644 --- a/src/analysis/convolution.lean +++ b/src/analysis/convolution.lean @@ -1500,9 +1500,7 @@ begin simp only [ef, eg, comp_app, continuous_linear_equiv.apply_symm_apply, coe_comp', continuous_linear_equiv.prod_apply, continuous_linear_equiv.map_sub, continuous_linear_equiv.arrow_congr, continuous_linear_equiv.arrow_congrSL_symm_apply, - continuous_linear_equiv.coe_coe, comp_app, continuous_linear_equiv.apply_symm_apply, - linear_equiv.inv_fun_eq_symm, continuous_linear_equiv.arrow_congrₛₗ_symm_apply, - eq_self_iff_true] }, + continuous_linear_equiv.coe_coe, comp_app, continuous_linear_equiv.apply_symm_apply ] }, simp_rw [this] at A, exact A, end diff --git a/src/geometry/manifold/vector_bundle/hom.lean b/src/geometry/manifold/vector_bundle/hom.lean index e83f070f9d9a6..9dff42177cf64 100644 --- a/src/geometry/manifold/vector_bundle/hom.lean +++ b/src/geometry/manifold/vector_bundle/hom.lean @@ -70,7 +70,7 @@ begin simp only [continuous_linear_map_coord_change, continuous_linear_equiv.coe_coe, continuous_linear_equiv.arrow_congrSL_apply, comp_apply, function.comp, compL_apply, flip_apply, continuous_linear_equiv.symm_symm, linear_equiv.to_fun_eq_coe, - continuous_linear_equiv.arrow_congrₛₗ_apply, continuous_linear_map.coe_comp'] }, + continuous_linear_map.coe_comp'] }, end include _i₁ _i₂ diff --git a/src/topology/algebra/module/strong_topology.lean b/src/topology/algebra/module/strong_topology.lean index 706f6818c70fe..78e2334d92150 100644 --- a/src/topology/algebra/module/strong_topology.lean +++ b/src/topology/algebra/module/strong_topology.lean @@ -175,9 +175,12 @@ end general section bounded_sets -variables {𝕜₁ 𝕜₂ : Type*} [normed_field 𝕜₁] [normed_field 𝕜₂] {σ : 𝕜₁ →+* 𝕜₂} {E E' F F' : Type*} +variables {𝕜₁ 𝕜₂ 𝕜₃ : Type*} [normed_field 𝕜₁] [normed_field 𝕜₂] [normed_field 𝕜₃] + {σ : 𝕜₁ →+* 𝕜₂} {τ : 𝕜₂ →+* 𝕜₃} {ρ : 𝕜₁ →+* 𝕜₃} [ring_hom_comp_triple σ τ ρ] + {E E' F F' G : Type*} [add_comm_group E] [module 𝕜₁ E] [add_comm_group E'] [module ℝ E'] [add_comm_group F] [module 𝕜₂ F] [add_comm_group F'] [module ℝ F'] + [add_comm_group G] [module 𝕜₃ G] [topological_space E] /-- The topology of bounded convergence on `E →L[𝕜] F`. This coincides with the topology induced by @@ -224,6 +227,49 @@ protected lemma has_basis_nhds_zero [topological_space F] (λ SV, {f : E →SL[σ] F | ∀ x ∈ SV.1, f x ∈ SV.2}) := continuous_linear_map.has_basis_nhds_zero_of_basis (𝓝 0).basis_sets +variables (G) [topological_space F] [topological_space G] + +/-- Pre-composition by a *fixed* continuous linear map as a continuous linear map. +Note that in non-normed space it is not always true that composition is continuous +in both variables, so we have to fix one of them. -/ +@[simps] def precomp [topological_add_group G] [has_continuous_const_smul 𝕜₃ G] + [ring_hom_surjective σ] [ring_hom_isometric σ] (L : E →SL[σ] F) : + (F →SL[τ] G) →L[𝕜₃] (E →SL[ρ] G) := +{ to_fun := λ f, f.comp L, + map_add' := λ f g, add_comp f g L, + map_smul' := λ a f, smul_comp a f L, + cont := + begin + letI : uniform_space G := topological_add_group.to_uniform_space G, + haveI : uniform_add_group G := topological_add_comm_group_is_uniform, + rw (strong_topology.embedding_coe_fn _ _ _).continuous_iff, + refine (uniform_on_fun.precomp_uniform_continuous _).continuous.comp + (strong_topology.embedding_coe_fn _ _ _).continuous, + exact λ S hS, hS.image L, + end } + +variables (E) {G} + +/-- Post-composition by a *fixed* continuous linear map as a continuous linear map. +Note that in non-normed space it is not always true that composition is continuous +in both variables, so we have to fix one of them. -/ +@[simps] def postcomp [topological_add_group F] [topological_add_group G] + [has_continuous_const_smul 𝕜₃ G] [has_continuous_const_smul 𝕜₂ F] (L : F →SL[τ] G) : + (E →SL[σ] F) →SL[τ] (E →SL[ρ] G) := +{ to_fun := λ f, L.comp f, + map_add' := comp_add L, + map_smul' := comp_smulₛₗ L, + cont := + begin + letI : uniform_space G := topological_add_group.to_uniform_space G, + haveI : uniform_add_group G := topological_add_comm_group_is_uniform, + letI : uniform_space F := topological_add_group.to_uniform_space F, + haveI : uniform_add_group F := topological_add_comm_group_is_uniform, + rw (strong_topology.embedding_coe_fn _ _ _).continuous_iff, + exact (uniform_on_fun.postcomp_uniform_continuous L.uniform_continuous).continuous.comp + (strong_topology.embedding_coe_fn _ _ _).continuous + end } + end bounded_sets end continuous_linear_map @@ -249,49 +295,29 @@ variables {𝕜 : Type*} {𝕜₂ : Type*} {𝕜₃ : Type*} {𝕜₄ : Type*} [ring_hom_inv_pair σ₄₃ σ₃₄] [ring_hom_comp_triple σ₂₁ σ₁₄ σ₂₄] [ring_hom_comp_triple σ₂₄ σ₄₃ σ₂₃] [ring_hom_comp_triple σ₁₂ σ₂₃ σ₁₃] [ring_hom_comp_triple σ₁₃ σ₃₄ σ₁₄] + [ring_hom_comp_triple σ₂₃ σ₃₄ σ₂₄] [ring_hom_comp_triple σ₁₂ σ₂₄ σ₁₄] + [ring_hom_isometric σ₁₂] [ring_hom_isometric σ₂₁] include σ₁₄ σ₂₄ σ₁₃ σ₃₄ σ₂₁ σ₂₃ /-- A pair of continuous (semi)linear equivalences generates a (semi)linear equivalence between the spaces of continuous (semi)linear maps. -/ -@[simps] def arrow_congrₛₗ (e₁₂ : E ≃SL[σ₁₂] F) (e₄₃ : H ≃SL[σ₄₃] G) : - (E →SL[σ₁₄] H) ≃ₛₗ[σ₄₃] (F →SL[σ₂₃] G) := +@[simps] def arrow_congrSL (e₁₂ : E ≃SL[σ₁₂] F) (e₄₃ : H ≃SL[σ₄₃] G) : + (E →SL[σ₁₄] H) ≃SL[σ₄₃] (F →SL[σ₂₃] G) := { -- given explicitly to help `simps` to_fun := λ L, (e₄₃ : H →SL[σ₄₃] G).comp (L.comp (e₁₂.symm : F →SL[σ₂₁] E)), -- given explicitly to help `simps` inv_fun := λ L, (e₄₃.symm : G →SL[σ₃₄] H).comp (L.comp (e₁₂ : E →SL[σ₁₂] F)), map_add' := λ f g, by rw [add_comp, comp_add], map_smul' := λ t f, by rw [smul_comp, comp_smulₛₗ], + continuous_to_fun := + ((postcomp F e₄₃.to_continuous_linear_map).comp + (precomp H e₁₂.symm.to_continuous_linear_map)).continuous, + continuous_inv_fun := + ((precomp H e₁₂.to_continuous_linear_map).comp + (postcomp F e₄₃.symm.to_continuous_linear_map)).continuous, .. e₁₂.arrow_congr_equiv e₄₃, } -variables [ring_hom_isometric σ₂₁] - -lemma arrow_congrₛₗ_continuous (e₁₂ : E ≃SL[σ₁₂] F) (e₄₃ : H ≃SL[σ₄₃] G) : - continuous (id (e₁₂.arrow_congrₛₗ e₄₃ : (E →SL[σ₁₄] H) ≃ₛₗ[σ₄₃] (F →SL[σ₂₃] G))) := -begin - apply continuous_of_continuous_at_zero, - show filter.tendsto _ _ _, - simp_rw [(e₁₂.arrow_congrₛₗ e₄₃).map_zero], - rw continuous_linear_map.has_basis_nhds_zero.tendsto_iff - continuous_linear_map.has_basis_nhds_zero, - rintros ⟨sF, sG⟩ ⟨h1 : bornology.is_vonN_bounded 𝕜₂ sF, h2 : sG ∈ nhds (0:G)⟩, - dsimp, - refine ⟨(e₁₂.symm '' sF, e₄₃ ⁻¹' sG), ⟨h1.image (e₁₂.symm : F →SL[σ₂₁] E), _⟩, - λ _ h _ hx, h _ (set.mem_image_of_mem _ hx)⟩, - apply e₄₃.continuous.continuous_at, - simpa using h2, -end - -variables [ring_hom_isometric σ₁₂] - -/-- A pair of continuous (semi)linear equivalences generates an continuous (semi)linear equivalence -between the spaces of continuous (semi)linear maps. -/ -@[simps] def arrow_congrSL (e₁₂ : E ≃SL[σ₁₂] F) (e₄₃ : H ≃SL[σ₄₃] G) : - (E →SL[σ₁₄] H) ≃SL[σ₄₃] (F →SL[σ₂₃] G) := -{ continuous_to_fun := e₁₂.arrow_congrₛₗ_continuous e₄₃, - continuous_inv_fun := e₁₂.symm.arrow_congrₛₗ_continuous e₄₃.symm, - .. e₁₂.arrow_congrₛₗ e₄₃, } - end semilinear section linear diff --git a/src/topology/vector_bundle/hom.lean b/src/topology/vector_bundle/hom.lean index 7e09babbf0ef9..3ebe5d8e1e42c 100644 --- a/src/topology/vector_bundle/hom.lean +++ b/src/topology/vector_bundle/hom.lean @@ -110,9 +110,8 @@ begin { mfld_set_tac }, { intros b hb, ext L v, simp only [continuous_linear_map_coord_change, continuous_linear_equiv.coe_coe, - continuous_linear_equiv.arrow_congrₛₗ_apply, linear_equiv.to_fun_eq_coe, coe_comp', - continuous_linear_equiv.arrow_congrSL_apply, comp_apply, function.comp, compSL_apply, - flip_apply, continuous_linear_equiv.symm_symm] }, + continuous_linear_equiv.arrow_congrSL_apply, + comp_apply, function.comp, compSL_apply, flip_apply, continuous_linear_equiv.symm_symm] }, end omit iσ @@ -206,8 +205,7 @@ lemma continuous_linear_map_coord_change_apply (b : B) begin ext v, simp_rw [continuous_linear_map_coord_change, continuous_linear_equiv.coe_coe, - continuous_linear_equiv.arrow_congrSL_apply, linear_equiv.to_fun_eq_coe, - continuous_linear_equiv.arrow_congrₛₗ_apply, + continuous_linear_equiv.arrow_congrSL_apply, continuous_linear_map_apply, continuous_linear_map_symm_apply' σ e₁ e₂ hb.1, comp_apply, continuous_linear_equiv.coe_coe, continuous_linear_equiv.symm_symm, trivialization.continuous_linear_map_at_apply, trivialization.symmL_apply], From 1a51edf13debfcbe223fa06b1cb353b9ed9751cc Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Sat, 8 Jul 2023 12:34:58 +0000 Subject: [PATCH 07/61] chore(*): add mathlib4 synchronization comments (#19229) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.category.Algebra.basic` * `algebra.category.Algebra.limits` * `algebra.category.Module.change_of_rings` * `algebra.module.pid` * `algebraic_geometry.elliptic_curve.point` * `algebraic_geometry.morphisms.quasi_compact` * `algebraic_geometry.morphisms.quasi_separated` * `algebraic_geometry.projective_spectrum.structure_sheaf` * `algebraic_topology.fundamental_groupoid.induced_maps` * `algebraic_topology.fundamental_groupoid.product` * `algebraic_topology.fundamental_groupoid.simply_connected` * `analysis.normed.group.SemiNormedGroup.kernels` * `category_theory.adjunction.lifting` * `category_theory.monoidal.internal.Module` * `data.qpf.multivariate.constructions.cofix` * `geometry.manifold.diffeomorph` * `geometry.manifold.vector_bundle.smooth_section` * `geometry.manifold.whitney_embedding` * `group_theory.finite_abelian` * `linear_algebra.clifford_algebra.contraction` * `linear_algebra.exterior_algebra.grading` * `linear_algebra.matrix.schur_complement` * `linear_algebra.tensor_algebra.to_tensor_power` * `model_theory.direct_limit` * `model_theory.fraisse` * `number_theory.modular_forms.basic` * `representation_theory.character` * `representation_theory.group_cohomology.basic` * `ring_theory.jacobson` * `ring_theory.nullstellensatz` * `topology.metric_space.gromov_hausdorff` * `topology.vector_bundle.hom` --- src/algebra/category/Algebra/basic.lean | 3 +++ src/algebra/category/Algebra/limits.lean | 3 +++ src/algebra/category/Module/change_of_rings.lean | 3 +++ src/algebra/module/pid.lean | 3 +++ src/algebraic_geometry/elliptic_curve/point.lean | 3 +++ src/algebraic_geometry/morphisms/quasi_compact.lean | 3 +++ src/algebraic_geometry/morphisms/quasi_separated.lean | 3 +++ .../projective_spectrum/structure_sheaf.lean | 3 +++ src/algebraic_topology/fundamental_groupoid/induced_maps.lean | 3 +++ src/algebraic_topology/fundamental_groupoid/product.lean | 3 +++ .../fundamental_groupoid/simply_connected.lean | 3 +++ src/analysis/normed/group/SemiNormedGroup/kernels.lean | 3 +++ src/category_theory/adjunction/lifting.lean | 3 +++ src/category_theory/monoidal/internal/Module.lean | 3 +++ src/data/qpf/multivariate/constructions/cofix.lean | 3 +++ src/geometry/manifold/diffeomorph.lean | 3 +++ src/geometry/manifold/vector_bundle/smooth_section.lean | 3 +++ src/geometry/manifold/whitney_embedding.lean | 3 +++ src/group_theory/finite_abelian.lean | 3 +++ src/linear_algebra/clifford_algebra/contraction.lean | 3 +++ src/linear_algebra/exterior_algebra/grading.lean | 3 +++ src/linear_algebra/matrix/schur_complement.lean | 3 +++ src/linear_algebra/tensor_algebra/to_tensor_power.lean | 3 +++ src/model_theory/direct_limit.lean | 3 +++ src/model_theory/fraisse.lean | 3 +++ src/number_theory/modular_forms/basic.lean | 3 +++ src/representation_theory/character.lean | 3 +++ src/representation_theory/group_cohomology/basic.lean | 3 +++ src/ring_theory/jacobson.lean | 3 +++ src/ring_theory/nullstellensatz.lean | 3 +++ src/topology/metric_space/gromov_hausdorff.lean | 3 +++ src/topology/vector_bundle/hom.lean | 3 +++ 32 files changed, 96 insertions(+) diff --git a/src/algebra/category/Algebra/basic.lean b/src/algebra/category/Algebra/basic.lean index b0630ea5f9992..f8162553f0040 100644 --- a/src/algebra/category/Algebra/basic.lean +++ b/src/algebra/category/Algebra/basic.lean @@ -11,6 +11,9 @@ import algebra.category.Module.basic /-! # Category instance for algebras over a commutative ring +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We introduce the bundled category `Algebra` of algebras over a fixed commutative ring `R ` along with the forgetful functors to `Ring` and `Module`. We furthermore show that the functor associating to a type the free `R`-algebra on that type is left adjoint to the forgetful functor. diff --git a/src/algebra/category/Algebra/limits.lean b/src/algebra/category/Algebra/limits.lean index 19c423b60363e..abcf7968145a5 100644 --- a/src/algebra/category/Algebra/limits.lean +++ b/src/algebra/category/Algebra/limits.lean @@ -10,6 +10,9 @@ import algebra.category.Ring.limits /-! # The category of R-algebras has all limits +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Further, these limits are preserved by the forgetful functor --- that is, the underlying types are just the limits in the category of types. -/ diff --git a/src/algebra/category/Module/change_of_rings.lean b/src/algebra/category/Module/change_of_rings.lean index b52900dc79917..327a5e779e98a 100644 --- a/src/algebra/category/Module/change_of_rings.lean +++ b/src/algebra/category/Module/change_of_rings.lean @@ -9,6 +9,9 @@ import ring_theory.tensor_product /-! # Change Of Rings +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main definitions * `category_theory.Module.restrict_scalars`: given rings `R, S` and a ring homomorphism `R ⟶ S`, diff --git a/src/algebra/module/pid.lean b/src/algebra/module/pid.lean index fe80596510b4a..adc0435a263cd 100644 --- a/src/algebra/module/pid.lean +++ b/src/algebra/module/pid.lean @@ -11,6 +11,9 @@ import algebra.category.Module.biproducts /-! # Structure of finitely generated modules over a PID +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main statements * `module.equiv_direct_sum_of_is_torsion` : A finitely generated torsion module over a PID is diff --git a/src/algebraic_geometry/elliptic_curve/point.lean b/src/algebraic_geometry/elliptic_curve/point.lean index e347de9ad650d..f8e6785033fa0 100644 --- a/src/algebraic_geometry/elliptic_curve/point.lean +++ b/src/algebraic_geometry/elliptic_curve/point.lean @@ -11,6 +11,9 @@ import ring_theory.class_group /-! # Nonsingular rational points on Weierstrass curves +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the type of nonsingular rational points on a Weierstrass curve over a field and proves that it forms an abelian group under a geometric secant-and-tangent process. diff --git a/src/algebraic_geometry/morphisms/quasi_compact.lean b/src/algebraic_geometry/morphisms/quasi_compact.lean index b3281898a3059..b0914f8516fd2 100644 --- a/src/algebraic_geometry/morphisms/quasi_compact.lean +++ b/src/algebraic_geometry/morphisms/quasi_compact.lean @@ -10,6 +10,9 @@ import algebraic_geometry.limits /-! # Quasi-compact morphisms +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A morphism of schemes is quasi-compact if the preimages of quasi-compact open sets are quasi-compact. diff --git a/src/algebraic_geometry/morphisms/quasi_separated.lean b/src/algebraic_geometry/morphisms/quasi_separated.lean index 8b75d385dcf82..33569b42b09e2 100644 --- a/src/algebraic_geometry/morphisms/quasi_separated.lean +++ b/src/algebraic_geometry/morphisms/quasi_separated.lean @@ -9,6 +9,9 @@ import topology.quasi_separated /-! # Quasi-separated morphisms +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A morphism of schemes `f : X ⟶ Y` is quasi-separated if the diagonal morphism `X ⟶ X ×[Y] X` is quasi-compact. diff --git a/src/algebraic_geometry/projective_spectrum/structure_sheaf.lean b/src/algebraic_geometry/projective_spectrum/structure_sheaf.lean index d6d743bec77d1..9003dcd0fa7d1 100644 --- a/src/algebraic_geometry/projective_spectrum/structure_sheaf.lean +++ b/src/algebraic_geometry/projective_spectrum/structure_sheaf.lean @@ -11,6 +11,9 @@ import algebraic_geometry.locally_ringed_space /-! # The structure sheaf on `projective_spectrum 𝒜`. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In `src/algebraic_geometry/topology.lean`, we have given a topology on `projective_spectrum 𝒜`; in this file we will construct a sheaf on `projective_spectrum 𝒜`. diff --git a/src/algebraic_topology/fundamental_groupoid/induced_maps.lean b/src/algebraic_topology/fundamental_groupoid/induced_maps.lean index a1c1a2b4c194c..1536082abcf65 100644 --- a/src/algebraic_topology/fundamental_groupoid/induced_maps.lean +++ b/src/algebraic_topology/fundamental_groupoid/induced_maps.lean @@ -10,6 +10,9 @@ import algebraic_topology.fundamental_groupoid.product /-! # Homotopic maps induce naturally isomorphic functors +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main definitions - `fundamental_groupoid_functor.homotopic_maps_nat_iso H` The natural isomorphism diff --git a/src/algebraic_topology/fundamental_groupoid/product.lean b/src/algebraic_topology/fundamental_groupoid/product.lean index 8ecf0078ff45d..7d8a7c033710a 100644 --- a/src/algebraic_topology/fundamental_groupoid/product.lean +++ b/src/algebraic_topology/fundamental_groupoid/product.lean @@ -11,6 +11,9 @@ import topology.homotopy.product /-! # Fundamental groupoid preserves products + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. In this file, we give the following definitions/theorems: - `fundamental_groupoid_functor.pi_iso` An isomorphism between Π i, (π Xᵢ) and π (Πi, Xᵢ), whose diff --git a/src/algebraic_topology/fundamental_groupoid/simply_connected.lean b/src/algebraic_topology/fundamental_groupoid/simply_connected.lean index e04bbe10c63b6..05e4e9f523fa4 100644 --- a/src/algebraic_topology/fundamental_groupoid/simply_connected.lean +++ b/src/algebraic_topology/fundamental_groupoid/simply_connected.lean @@ -10,6 +10,9 @@ import algebraic_topology.fundamental_groupoid.punit /-! # Simply connected spaces + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. This file defines simply connected spaces. A topological space is simply connected if its fundamental groupoid is equivalent to `unit`. diff --git a/src/analysis/normed/group/SemiNormedGroup/kernels.lean b/src/analysis/normed/group/SemiNormedGroup/kernels.lean index 0dc24df7e08d0..b1822fbc5011c 100644 --- a/src/analysis/normed/group/SemiNormedGroup/kernels.lean +++ b/src/analysis/normed/group/SemiNormedGroup/kernels.lean @@ -10,6 +10,9 @@ import category_theory.limits.shapes.kernels /-! # Kernels and cokernels in SemiNormedGroup₁ and SemiNormedGroup +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We show that `SemiNormedGroup₁` has cokernels (for which of course the `cokernel.π f` maps are norm non-increasing), as well as the easier result that `SemiNormedGroup` has cokernels. We also show that diff --git a/src/category_theory/adjunction/lifting.lean b/src/category_theory/adjunction/lifting.lean index f8fa2872d8234..d1b70e8462e1a 100644 --- a/src/category_theory/adjunction/lifting.lean +++ b/src/category_theory/adjunction/lifting.lean @@ -11,6 +11,9 @@ import category_theory.monad.coequalizer /-! # Adjoint lifting +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file gives two constructions for building left adjoints: the adjoint triangle theorem and the adjoint lifting theorem. The adjoint triangle theorem says that given a functor `U : B ⥤ C` with a left adjoint `F` such diff --git a/src/category_theory/monoidal/internal/Module.lean b/src/category_theory/monoidal/internal/Module.lean index 27071854b7d83..a8dcbc32b37ac 100644 --- a/src/category_theory/monoidal/internal/Module.lean +++ b/src/category_theory/monoidal/internal/Module.lean @@ -11,6 +11,9 @@ import category_theory.monoidal.Mon_ /-! # `Mon_ (Module R) ≌ Algebra R` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The category of internal monoid objects in `Module R` is equivalent to the category of "native" bundled `R`-algebras. diff --git a/src/data/qpf/multivariate/constructions/cofix.lean b/src/data/qpf/multivariate/constructions/cofix.lean index ba67e60e0df8b..93bc6f9098c1e 100644 --- a/src/data/qpf/multivariate/constructions/cofix.lean +++ b/src/data/qpf/multivariate/constructions/cofix.lean @@ -12,6 +12,9 @@ import data.qpf.multivariate.basic /-! # The final co-algebra of a multivariate qpf is again a qpf. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + For a `(n+1)`-ary QPF `F (α₀,..,αₙ)`, we take the least fixed point of `F` with regards to its last argument `αₙ`. The result is a `n`-ary functor: `fix F (α₀,..,αₙ₋₁)`. Making `fix F` into a functor allows us to take the fixed point, compose with other functors diff --git a/src/geometry/manifold/diffeomorph.lean b/src/geometry/manifold/diffeomorph.lean index 4a1a820e6dfda..f7eef23e11ca7 100644 --- a/src/geometry/manifold/diffeomorph.lean +++ b/src/geometry/manifold/diffeomorph.lean @@ -9,6 +9,9 @@ import geometry.manifold.cont_mdiff_mfderiv /-! # Diffeomorphisms + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. This file implements diffeomorphisms. ## Definitions diff --git a/src/geometry/manifold/vector_bundle/smooth_section.lean b/src/geometry/manifold/vector_bundle/smooth_section.lean index c7782d18cc4b6..d6c5115c08f70 100644 --- a/src/geometry/manifold/vector_bundle/smooth_section.lean +++ b/src/geometry/manifold/vector_bundle/smooth_section.lean @@ -11,6 +11,9 @@ import geometry.manifold.algebra.lie_group /-! # Smooth sections +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define the type `cont_mdiff_section` of `n` times continuously differentiable sections of a smooth vector bundle over a manifold `M` and prove that it's a module. -/ diff --git a/src/geometry/manifold/whitney_embedding.lean b/src/geometry/manifold/whitney_embedding.lean index 736ba534e91f6..c47bcd133c47e 100644 --- a/src/geometry/manifold/whitney_embedding.lean +++ b/src/geometry/manifold/whitney_embedding.lean @@ -10,6 +10,9 @@ import geometry.manifold.partition_of_unity /-! # Whitney embedding theorem +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove a version of the Whitney embedding theorem: for any compact real manifold `M`, for sufficiently large `n` there exists a smooth embedding `M → ℝ^n`. diff --git a/src/group_theory/finite_abelian.lean b/src/group_theory/finite_abelian.lean index 0a844cd0e78bc..bf23854627a70 100644 --- a/src/group_theory/finite_abelian.lean +++ b/src/group_theory/finite_abelian.lean @@ -9,6 +9,9 @@ import data.zmod.quotient /-! # Structure of finite(ly generated) abelian groups +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + * `add_comm_group.equiv_free_prod_direct_sum_zmod` : Any finitely generated abelian group is the product of a power of `ℤ` and a direct sum of some `zmod (p i ^ e i)` for some prime powers `p i ^ e i`. diff --git a/src/linear_algebra/clifford_algebra/contraction.lean b/src/linear_algebra/clifford_algebra/contraction.lean index 152ebc5bd63b3..6384379fe7097 100644 --- a/src/linear_algebra/clifford_algebra/contraction.lean +++ b/src/linear_algebra/clifford_algebra/contraction.lean @@ -11,6 +11,9 @@ import linear_algebra.clifford_algebra.conjugation /-! # Contraction in Clifford Algebras +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains some of the results from [grinberg_clifford_2016][]. The key result is `clifford_algebra.equiv_exterior`. diff --git a/src/linear_algebra/exterior_algebra/grading.lean b/src/linear_algebra/exterior_algebra/grading.lean index eee9bc304edb2..0f5089ccbee07 100644 --- a/src/linear_algebra/exterior_algebra/grading.lean +++ b/src/linear_algebra/exterior_algebra/grading.lean @@ -9,6 +9,9 @@ import ring_theory.graded_algebra.basic /-! # Results about the grading structure of the exterior algebra +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Many of these results are copied with minimal modification from the tensor algebra. The main result is `exterior_algebra.graded_algebra`, which says that the exterior algebra is a diff --git a/src/linear_algebra/matrix/schur_complement.lean b/src/linear_algebra/matrix/schur_complement.lean index 1c354e0f09131..4e3cc43ea1180 100644 --- a/src/linear_algebra/matrix/schur_complement.lean +++ b/src/linear_algebra/matrix/schur_complement.lean @@ -9,6 +9,9 @@ import linear_algebra.matrix.pos_def /-! # 2×2 block matrices and the Schur complement +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves properties of 2×2 block matrices `[A B; C D]` that relate to the Schur complement `D - C⬝A⁻¹⬝B`. diff --git a/src/linear_algebra/tensor_algebra/to_tensor_power.lean b/src/linear_algebra/tensor_algebra/to_tensor_power.lean index 5d619d69bba0b..d79a8e426472b 100644 --- a/src/linear_algebra/tensor_algebra/to_tensor_power.lean +++ b/src/linear_algebra/tensor_algebra/to_tensor_power.lean @@ -8,6 +8,9 @@ import linear_algebra.tensor_power /-! # Tensor algebras as direct sums of tensor powers +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we show that `tensor_algebra R M` is isomorphic to a direct sum of tensor powers, as `tensor_algebra.equiv_direct_sum`. -/ diff --git a/src/model_theory/direct_limit.lean b/src/model_theory/direct_limit.lean index b228a91cfa896..7833595665f85 100644 --- a/src/model_theory/direct_limit.lean +++ b/src/model_theory/direct_limit.lean @@ -10,6 +10,9 @@ import model_theory.finitely_generated /-! # Direct Limits of First-Order Structures + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. This file constructs the direct limit of a directed system of first-order embeddings. ## Main Definitions diff --git a/src/model_theory/fraisse.lean b/src/model_theory/fraisse.lean index d819e48cea930..2f372ad156a27 100644 --- a/src/model_theory/fraisse.lean +++ b/src/model_theory/fraisse.lean @@ -11,6 +11,9 @@ import model_theory.bundled /-! # Fraïssé Classes and Fraïssé Limits +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file pertains to the ages of countable first-order structures. The age of a structure is the class of all finitely-generated structures that embed into it. diff --git a/src/number_theory/modular_forms/basic.lean b/src/number_theory/modular_forms/basic.lean index 1d1346c5dfee0..01c52760e61b2 100644 --- a/src/number_theory/modular_forms/basic.lean +++ b/src/number_theory/modular_forms/basic.lean @@ -9,6 +9,9 @@ import number_theory.modular_forms.slash_invariant_forms /-! # Modular forms +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines modular forms and proves some basic properties about them. We begin by defining modular forms and cusp forms as extension of `slash_invariant_forms` then we diff --git a/src/representation_theory/character.lean b/src/representation_theory/character.lean index e346efe45599a..07ddd2fbf264d 100644 --- a/src/representation_theory/character.lean +++ b/src/representation_theory/character.lean @@ -10,6 +10,9 @@ import representation_theory.invariants /-! # Characters of representations +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file introduces characters of representation and proves basic lemmas about how characters behave under various operations on representations. diff --git a/src/representation_theory/group_cohomology/basic.lean b/src/representation_theory/group_cohomology/basic.lean index a972d9f61f849..b2860c0192415 100644 --- a/src/representation_theory/group_cohomology/basic.lean +++ b/src/representation_theory/group_cohomology/basic.lean @@ -10,6 +10,9 @@ import representation_theory.group_cohomology.resolution /-! # The group cohomology of a `k`-linear `G`-representation +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Let `k` be a commutative ring and `G` a group. This file defines the group cohomology of `A : Rep k G` to be the cohomology of the complex $$0 \to \mathrm{Fun}(G^0, A) \to \mathrm{Fun}(G^1, A) \to \mathrm{Fun}(G^2, A) \to \dots$$ diff --git a/src/ring_theory/jacobson.lean b/src/ring_theory/jacobson.lean index a44c12838c5c5..8c0a791739917 100644 --- a/src/ring_theory/jacobson.lean +++ b/src/ring_theory/jacobson.lean @@ -9,6 +9,9 @@ import ring_theory.jacobson_ideal /-! # Jacobson Rings + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. The following conditions are equivalent for a ring `R`: 1. Every radical ideal `I` is equal to its Jacobson radical 2. Every radical ideal `I` can be written as an intersection of maximal ideals diff --git a/src/ring_theory/nullstellensatz.lean b/src/ring_theory/nullstellensatz.lean index ab70ecadfb579..def2196661005 100644 --- a/src/ring_theory/nullstellensatz.lean +++ b/src/ring_theory/nullstellensatz.lean @@ -10,6 +10,9 @@ import algebraic_geometry.prime_spectrum.basic /-! # Nullstellensatz + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. This file establishes a version of Hilbert's classical Nullstellensatz for `mv_polynomial`s. The main statement of the theorem is `vanishing_ideal_zero_locus_eq_radical`. diff --git a/src/topology/metric_space/gromov_hausdorff.lean b/src/topology/metric_space/gromov_hausdorff.lean index d6b068acf48f9..f3d546bb9f5dc 100644 --- a/src/topology/metric_space/gromov_hausdorff.lean +++ b/src/topology/metric_space/gromov_hausdorff.lean @@ -12,6 +12,9 @@ import topology.metric_space.kuratowski /-! # Gromov-Hausdorff distance +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the Gromov-Hausdorff distance on the space of nonempty compact metric spaces up to isometry. diff --git a/src/topology/vector_bundle/hom.lean b/src/topology/vector_bundle/hom.lean index 3ebe5d8e1e42c..96ea682a0be2b 100644 --- a/src/topology/vector_bundle/hom.lean +++ b/src/topology/vector_bundle/hom.lean @@ -10,6 +10,9 @@ import analysis.normed_space.operator_norm /-! # The vector bundle of continuous (semi)linear maps +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define the (topological) vector bundle of continuous (semi)linear maps between two vector bundles over the same base. From bd365b1a4901dbd878e86cb146c2bd86533df468 Mon Sep 17 00:00:00 2001 From: b-reinke Date: Tue, 11 Jul 2023 12:40:13 +0000 Subject: [PATCH 08/61] feat(group_theory/sylow): add inverse to card_eq_multiplicity (#18300) The lemma `card_eq_multiplicity` states that a Sylow group of a finite group has cardinality p^n, where n is the multiplicity of p in the group order. This PR adds an inverse definition `card_eq_multiplicity_to_sylow`, promoting a subgroup of the right cardinality to a Sylow group, and a simplification lemma `coe_card_eq_multiplicity_to_sylow` for the coercion of the resulting Sylow group back to a subgroup. --- src/group_theory/sylow.lean | 27 +++++++++++++++++++++------ 1 file changed, 21 insertions(+), 6 deletions(-) diff --git a/src/group_theory/sylow.lean b/src/group_theory/sylow.lean index 91d962c8014b9..8b32b29b66805 100644 --- a/src/group_theory/sylow.lean +++ b/src/group_theory/sylow.lean @@ -36,7 +36,7 @@ The Sylow theorems are the following results for every finite group `G` and ever there exists a subgroup of `G` of order `pⁿ`. * `is_p_group.exists_le_sylow`: A generalization of Sylow's first theorem: Every `p`-subgroup is contained in a Sylow `p`-subgroup. -* `sylow.card_eq_multiplicity`: The cardinality of a Sylow group is `p ^ n` +* `sylow.card_eq_multiplicity`: The cardinality of a Sylow subgroup is `p ^ n` where `n` is the multiplicity of `p` in the group order. * `sylow_conjugate`: A generalization of Sylow's second theorem: If the number of Sylow `p`-subgroups is finite, then all Sylow `p`-subgroups are conjugate. @@ -592,7 +592,7 @@ begin rwa [h, card_bot] at key, end -/-- The cardinality of a Sylow group is `p ^ n` +/-- The cardinality of a Sylow subgroup is `p ^ n` where `n` is the multiplicity of `p` in the group order. -/ lemma card_eq_multiplicity [fintype G] {p : ℕ} [hp : fact p.prime] (P : sylow p G) : card P = p ^ nat.factorization (card G) p := @@ -603,6 +603,21 @@ begin exact P.1.card_subgroup_dvd_card, end +/-- A subgroup with cardinality `p ^ n` is a Sylow subgroup + where `n` is the multiplicity of `p` in the group order. -/ +def of_card [fintype G] {p : ℕ} [hp : fact p.prime] (H : subgroup G) [fintype H] + (card_eq : card H = p ^ (card G).factorization p) : sylow p G := +{ to_subgroup := H, + is_p_group' := is_p_group.of_card card_eq, + is_maximal' := begin + obtain ⟨P, hHP⟩ := (is_p_group.of_card card_eq).exists_le_sylow, + exact set_like.ext' (set.eq_of_subset_of_card_le hHP + (P.card_eq_multiplicity.trans card_eq.symm).le).symm ▸ λ _, P.3, + end } + +@[simp, norm_cast] lemma coe_of_card [fintype G] {p : ℕ} [hp : fact p.prime] (H : subgroup G) + [fintype H] (card_eq : card H = p ^ (card G).factorization p) : ↑(of_card H card_eq) = H := rfl + lemma subsingleton_of_normal {p : ℕ} [fact p.prime] [finite (sylow p G)] (P : sylow p G) (h : (P : subgroup G).normal) : subsingleton (sylow p G) := begin @@ -667,8 +682,8 @@ normalizer_eq_top.mp $ normalizer_condition_iff_only_full_group_self_normalizing open_locale big_operators -/-- If all its sylow groups are normal, then a finite group is isomorphic to the direct product -of these sylow groups. +/-- If all its Sylow subgroups are normal, then a finite group is isomorphic to the direct product +of these Sylow subgroups. -/ noncomputable def direct_product_of_normal [fintype G] @@ -677,7 +692,7 @@ def direct_product_of_normal [fintype G] begin set ps := (fintype.card G).factorization.support, - -- “The” sylow group for p + -- “The” Sylow subgroup for p let P : Π p, sylow p G := default, have hcomm : pairwise (λ (p₁ p₂ : ps), ∀ (x y : G), x ∈ P p₁ → y ∈ P p₂ → commute x y), @@ -689,7 +704,7 @@ begin apply is_p_group.disjoint_of_ne p₁ p₂ hne' _ _ (P p₁).is_p_group' (P p₂).is_p_group', }, refine mul_equiv.trans _ _, - -- There is only one sylow group for each p, so the inner product is trivial + -- There is only one Sylow subgroup for each p, so the inner product is trivial show (Π p : ps, Π P : sylow p G, P) ≃* (Π p : ps, P p), { -- here we need to help the elaborator with an explicit instantiation apply @mul_equiv.Pi_congr_right ps (λ p, (Π P : sylow p G, P)) (λ p, P p) _ _ , From 3ed3f98a1e836241990d3d308f1577e434977130 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 11 Jul 2023 12:40:14 +0000 Subject: [PATCH 09/61] =?UTF-8?q?feat(data/nat/order/basic):=20`a=20+=20b?= =?UTF-8?q?=20-=201=20=E2=89=A4=20a=20*=20b`=20(#18737)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit and golf `max_eq_zero_iff`/`min_eq_zero_iff` To be used for Kneser's addition theorem. Co-authored-by: Yury G. Kudryashov --- src/data/nat/order/basic.lean | 32 ++++++++++---------------------- 1 file changed, 10 insertions(+), 22 deletions(-) diff --git a/src/data/nat/order/basic.lean b/src/data/nat/order/basic.lean index fcd6864736ece..c0913019067e2 100644 --- a/src/data/nat/order/basic.lean +++ b/src/data/nat/order/basic.lean @@ -94,29 +94,9 @@ lemma eq_zero_of_mul_le (hb : 2 ≤ n) (h : n * m ≤ m) : m = 0 := eq_zero_of_double_le $ le_trans (nat.mul_le_mul_right _ hb) h lemma zero_max : max 0 n = n := max_eq_right (zero_le _) -@[simp] lemma min_eq_zero_iff : min m n = 0 ↔ m = 0 ∨ n = 0 := -begin - split, - { intro h, - cases le_total m n with H H, - { simpa [H] using or.inl h }, - { simpa [H] using or.inr h } }, - { rintro (rfl|rfl); - simp } -end -@[simp] lemma max_eq_zero_iff : max m n = 0 ↔ m = 0 ∧ n = 0 := -begin - split, - { intro h, - cases le_total m n with H H, - { simp only [H, max_eq_right] at h, - exact ⟨le_antisymm (H.trans h.le) (zero_le _), h⟩ }, - { simp only [H, max_eq_left] at h, - exact ⟨h, le_antisymm (H.trans h.le) (zero_le _)⟩ } }, - { rintro ⟨rfl, rfl⟩, - simp } -end +@[simp] lemma min_eq_zero_iff : min m n = 0 ↔ m = 0 ∨ n = 0 := min_eq_bot +@[simp] lemma max_eq_zero_iff : max m n = 0 ↔ m = 0 ∧ n = 0 := max_eq_bot lemma add_eq_max_iff : m + n = max m n ↔ m = 0 ∨ n = 0 := begin @@ -286,6 +266,14 @@ end | 0 := iff_of_false (lt_irrefl _) zero_le_one.not_lt | (n + 1) := lt_mul_iff_one_lt_left n.succ_pos +lemma add_sub_one_le_mul (hm : m ≠ 0) (hn : n ≠ 0) : m + n - 1 ≤ m * n := +begin + cases m, + { cases hm rfl }, + { rw [succ_add, succ_sub_one, succ_mul], + exact add_le_add_right (le_mul_of_one_le_right' $ pos_iff_ne_zero.2 hn) _ } +end + /-! ### Recursion and induction principles From 4e24c4bfcff371c71f7ba22050308aa17815626c Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 11 Jul 2023 15:46:03 +0000 Subject: [PATCH 10/61] feat(data/set/intervals/proj_Icc): Extending from `Ici a` (#18795) Define the `Ici` and `Iic` versions of `set.proj_Icc`/`set.Icc_extend`. Slightly reorder the lemmas to allow better overall structure. This is useful for extending convex functions by a junk value. --- src/data/set/intervals/proj_Icc.lean | 152 ++++++++++++++++++++++----- 1 file changed, 127 insertions(+), 25 deletions(-) diff --git a/src/data/set/intervals/proj_Icc.lean b/src/data/set/intervals/proj_Icc.lean index e6cb3cff298e2..41a033fc0ed3a 100644 --- a/src/data/set/intervals/proj_Icc.lean +++ b/src/data/set/intervals/proj_Icc.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury G. Kudryashov, Patrick Massot -/ import data.set.function -import data.set.intervals.basic +import data.set.intervals.ord_connected /-! # Projection of a line onto a closed interval @@ -14,8 +14,16 @@ import data.set.intervals.basic Given a linearly ordered type `α`, in this file we define +* `set.proj_Ici (a : α)` to be the map `α → [a, ∞[` sending `]-∞, a]` to `a`, + and each point `x ∈ [a, ∞[` to itself; +* `set.proj_Iic (b : α)` to be the map `α → ]-∞, b[` sending `[b, ∞[` to `b`, + and each point `x ∈ ]-∞, b]` to itself; * `set.proj_Icc (a b : α) (h : a ≤ b)` to be the map `α → [a, b]` sending `(-∞, a]` to `a`, `[b, ∞)` to `b`, and each point `x ∈ [a, b]` to itself; +* `set.Ici_extend {a : α} (f : Ici a → β)` to be the extension of `f` to `α` defined + as `f ∘ proj_Ici a`. +* `set.Iic_extend {b : α} (f : Iic b → β)` to be the extension of `f` to `α` defined + as `f ∘ proj_Iic b`. * `set.Icc_extend {a b : α} (h : a ≤ b) (f : Icc a b → β)` to be the extension of `f` to `α` defined as `f ∘ proj_Icc a b h`. @@ -28,88 +36,156 @@ open function namespace set +/-- Projection of `α` to the closed interval `[a, ∞[`. -/ +def proj_Ici (a x : α) : Ici a := ⟨max a x, le_max_left _ _⟩ + +/-- Projection of `α` to the closed interval `]-∞, b]`. -/ +def proj_Iic (b x : α) : Iic b := ⟨min b x, min_le_left _ _⟩ + /-- Projection of `α` to the closed interval `[a, b]`. -/ def proj_Icc (a b : α) (h : a ≤ b) (x : α) : Icc a b := ⟨max a (min b x), le_max_left _ _, max_le h (min_le_left _ _)⟩ variables {a b : α} (h : a ≤ b) {x : α} +@[norm_cast] lemma coe_proj_Ici (a x : α) : (proj_Ici a x : α) = max a x := rfl +@[norm_cast] lemma coe_proj_Iic (b x : α) : (proj_Iic b x : α) = min b x := rfl +@[norm_cast] lemma coe_proj_Icc (a b : α) (h : a ≤ b) (x : α) : + (proj_Icc a b h x : α) = max a (min b x) := rfl + +lemma proj_Ici_of_le (hx : x ≤ a) : proj_Ici a x = ⟨a, le_rfl⟩ := subtype.ext $ max_eq_left hx +lemma proj_Iic_of_le (hx : b ≤ x) : proj_Iic b x = ⟨b, le_rfl⟩ := subtype.ext $ min_eq_left hx + lemma proj_Icc_of_le_left (hx : x ≤ a) : proj_Icc a b h x = ⟨a, left_mem_Icc.2 h⟩ := by simp [proj_Icc, hx, hx.trans h] -@[simp] lemma proj_Icc_left : proj_Icc a b h a = ⟨a, left_mem_Icc.2 h⟩ := -proj_Icc_of_le_left h le_rfl - lemma proj_Icc_of_right_le (hx : b ≤ x) : proj_Icc a b h x = ⟨b, right_mem_Icc.2 h⟩ := by simp [proj_Icc, hx, h] +@[simp] lemma proj_Ici_self (a : α) : proj_Ici a a = ⟨a, le_rfl⟩ := proj_Ici_of_le le_rfl +@[simp] lemma proj_Iic_self (b : α) : proj_Iic b b = ⟨b, le_rfl⟩ := proj_Iic_of_le le_rfl + +@[simp] lemma proj_Icc_left : proj_Icc a b h a = ⟨a, left_mem_Icc.2 h⟩ := +proj_Icc_of_le_left h le_rfl + @[simp] lemma proj_Icc_right : proj_Icc a b h b = ⟨b, right_mem_Icc.2 h⟩ := proj_Icc_of_right_le h le_rfl +lemma proj_Ici_eq_self : proj_Ici a x = ⟨a, le_rfl⟩ ↔ x ≤ a := by simp [proj_Ici, subtype.ext_iff] +lemma proj_Iic_eq_self : proj_Iic b x = ⟨b, le_rfl⟩ ↔ b ≤ x := by simp [proj_Iic, subtype.ext_iff] + lemma proj_Icc_eq_left (h : a < b) : proj_Icc a b h.le x = ⟨a, left_mem_Icc.mpr h.le⟩ ↔ x ≤ a := -begin - refine ⟨λ h', _, proj_Icc_of_le_left _⟩, - simp_rw [subtype.ext_iff_val, proj_Icc, max_eq_left_iff, min_le_iff, h.not_le, false_or] at h', - exact h' -end +by simp [proj_Icc, subtype.ext_iff, h.not_le] lemma proj_Icc_eq_right (h : a < b) : proj_Icc a b h.le x = ⟨b, right_mem_Icc.mpr h.le⟩ ↔ b ≤ x := -begin - refine ⟨λ h', _, proj_Icc_of_right_le _⟩, - simp_rw [subtype.ext_iff_val, proj_Icc] at h', - have := ((max_choice _ _).resolve_left (by simp [h.ne', h'])).symm.trans h', - exact min_eq_left_iff.mp this -end +by simp [proj_Icc, subtype.ext_iff, max_min_distrib_left, h.le, h.not_le] +lemma proj_Ici_of_mem (hx : x ∈ Ici a) : proj_Ici a x = ⟨x, hx⟩ := by simpa [proj_Ici] +lemma proj_Iic_of_mem (hx : x ∈ Iic b) : proj_Iic b x = ⟨x, hx⟩ := by simpa [proj_Iic] lemma proj_Icc_of_mem (hx : x ∈ Icc a b) : proj_Icc a b h x = ⟨x, hx⟩ := by simp [proj_Icc, hx.1, hx.2] +@[simp] lemma proj_Ici_coe (x : Ici a) : proj_Ici a x = x := by { cases x, apply proj_Ici_of_mem } +@[simp] lemma proj_Iic_coe (x : Iic b) : proj_Iic b x = x := by { cases x, apply proj_Iic_of_mem } @[simp] lemma proj_Icc_coe (x : Icc a b) : proj_Icc a b h x = x := by { cases x, apply proj_Icc_of_mem } +lemma proj_Ici_surj_on : surj_on (proj_Ici a) (Ici a) univ := λ x _, ⟨x, x.2, proj_Ici_coe x⟩ +lemma proj_Iic_surj_on : surj_on (proj_Iic b) (Iic b) univ := λ x _, ⟨x, x.2, proj_Iic_coe x⟩ lemma proj_Icc_surj_on : surj_on (proj_Icc a b h) (Icc a b) univ := λ x _, ⟨x, x.2, proj_Icc_coe h x⟩ -lemma proj_Icc_surjective : surjective (proj_Icc a b h) := -λ x, ⟨x, proj_Icc_coe h x⟩ +lemma proj_Ici_surjective : surjective (proj_Ici a) := λ x, ⟨x, proj_Ici_coe x⟩ +lemma proj_Iic_surjective : surjective (proj_Iic b) := λ x, ⟨x, proj_Iic_coe x⟩ +lemma proj_Icc_surjective : surjective (proj_Icc a b h) := λ x, ⟨x, proj_Icc_coe h x⟩ -@[simp] lemma range_proj_Icc : range (proj_Icc a b h) = univ := -(proj_Icc_surjective h).range_eq +@[simp] lemma range_proj_Ici : range (proj_Ici a) = univ := proj_Ici_surjective.range_eq +@[simp] lemma range_proj_Iic : range (proj_Iic a) = univ := proj_Iic_surjective.range_eq +@[simp] lemma range_proj_Icc : range (proj_Icc a b h) = univ := (proj_Icc_surjective h).range_eq +lemma monotone_proj_Ici : monotone (proj_Ici a) := λ x y, max_le_max le_rfl +lemma monotone_proj_Iic : monotone (proj_Iic a) := λ x y, min_le_min le_rfl lemma monotone_proj_Icc : monotone (proj_Icc a b h) := λ x y hxy, max_le_max le_rfl $ min_le_min le_rfl hxy +lemma strict_mono_on_proj_Ici : strict_mono_on (proj_Ici a) (Ici a) := +λ x hx y hy hxy, by simpa only [proj_Ici_of_mem, hx, hy] +lemma strict_mono_on_proj_Iic : strict_mono_on (proj_Iic b) (Iic b) := +λ x hx y hy hxy, by simpa only [proj_Iic_of_mem, hx, hy] lemma strict_mono_on_proj_Icc : strict_mono_on (proj_Icc a b h) (Icc a b) := λ x hx y hy hxy, by simpa only [proj_Icc_of_mem, hx, hy] +/-- Extend a function `[a, ∞[ → β` to a map `α → β`. -/ +def Ici_extend (f : Ici a → β) : α → β := f ∘ proj_Ici a + +/-- Extend a function `]-∞, b] → β` to a map `α → β`. -/ +def Iic_extend (f : Iic b → β) : α → β := f ∘ proj_Iic b + /-- Extend a function `[a, b] → β` to a map `α → β`. -/ def Icc_extend {a b : α} (h : a ≤ b) (f : Icc a b → β) : α → β := f ∘ proj_Icc a b h +@[simp] lemma Ici_extend_apply (f : Ici a → β) (x : α) : + Ici_extend f x = f ⟨max a x, le_max_left _ _⟩ := rfl +@[simp] lemma Iic_extend_apply (f : Iic b → β) (x : α) : + Iic_extend f x = f ⟨min b x, min_le_left _ _⟩ := rfl +lemma Icc_extend_apply (h : a ≤ b) (f : Icc a b → β) (x : α) : + Icc_extend h f x = f ⟨max a (min b x), le_max_left _ _, max_le h (min_le_left _ _)⟩ := rfl + +@[simp] lemma range_Ici_extend (f : Ici a → β) : range (Ici_extend f) = range f := +by simp only [Ici_extend, range_comp f, range_proj_Ici, range_id'] + +@[simp] lemma range_Iic_extend (f : Iic b → β) : range (Iic_extend f) = range f := +by simp only [Iic_extend, range_comp f, range_proj_Iic, range_id'] + @[simp] lemma Icc_extend_range (f : Icc a b → β) : range (Icc_extend h f) = range f := by simp only [Icc_extend, range_comp f, range_proj_Icc, range_id'] +lemma Ici_extend_of_le (f : Ici a → β) (hx : x ≤ a) : Ici_extend f x = f ⟨a, le_rfl⟩ := +congr_arg f $ proj_Ici_of_le hx + +lemma Iic_extend_of_le (f : Iic b → β) (hx : b ≤ x) : Iic_extend f x = f ⟨b, le_rfl⟩ := +congr_arg f $ proj_Iic_of_le hx + lemma Icc_extend_of_le_left (f : Icc a b → β) (hx : x ≤ a) : Icc_extend h f x = f ⟨a, left_mem_Icc.2 h⟩ := congr_arg f $ proj_Icc_of_le_left h hx -@[simp] lemma Icc_extend_left (f : Icc a b → β) : - Icc_extend h f a = f ⟨a, left_mem_Icc.2 h⟩ := -Icc_extend_of_le_left h f le_rfl - lemma Icc_extend_of_right_le (f : Icc a b → β) (hx : b ≤ x) : Icc_extend h f x = f ⟨b, right_mem_Icc.2 h⟩ := congr_arg f $ proj_Icc_of_right_le h hx +@[simp] lemma Ici_extend_self (f : Ici a → β) : Ici_extend f a = f ⟨a, le_rfl⟩ := +Ici_extend_of_le f le_rfl + +@[simp] lemma Iic_extend_self (f : Iic b → β) : Iic_extend f b = f ⟨b, le_rfl⟩ := +Iic_extend_of_le f le_rfl + +@[simp] lemma Icc_extend_left (f : Icc a b → β) : + Icc_extend h f a = f ⟨a, left_mem_Icc.2 h⟩ := +Icc_extend_of_le_left h f le_rfl + @[simp] lemma Icc_extend_right (f : Icc a b → β) : Icc_extend h f b = f ⟨b, right_mem_Icc.2 h⟩ := Icc_extend_of_right_le h f le_rfl +lemma Ici_extend_of_mem (f : Ici a → β) (hx : x ∈ Ici a) : Ici_extend f x = f ⟨x, hx⟩ := +congr_arg f $ proj_Ici_of_mem hx + +lemma Iic_extend_of_mem (f : Iic b → β) (hx : x ∈ Iic b) : Iic_extend f x = f ⟨x, hx⟩ := +congr_arg f $ proj_Iic_of_mem hx + lemma Icc_extend_of_mem (f : Icc a b → β) (hx : x ∈ Icc a b) : Icc_extend h f x = f ⟨x, hx⟩ := congr_arg f $ proj_Icc_of_mem h hx +@[simp] lemma Ici_extend_coe (f : Ici a → β) (x : Ici a) : Ici_extend f x = f x := +congr_arg f $ proj_Ici_coe x + +@[simp] lemma Iic_extend_coe (f : Iic b → β) (x : Iic b) : Iic_extend f x = f x := +congr_arg f $ proj_Iic_coe x + @[simp] lemma Icc_extend_coe (f : Icc a b → β) (x : Icc a b) : Icc_extend h f x = f x := congr_arg f $ proj_Icc_coe h x @@ -132,11 +208,37 @@ end set open set -variables [preorder β] {a b : α} (h : a ≤ b) {f : Icc a b → β} +variables [preorder β] {s t : set α} {a b : α} (h : a ≤ b) {f : Icc a b → β} + +protected lemma monotone.Ici_extend {f : Ici a → β} (hf : monotone f) : monotone (Ici_extend f) := +hf.comp monotone_proj_Ici + +protected lemma monotone.Iic_extend {f : Iic b → β} (hf : monotone f) : monotone (Iic_extend f) := +hf.comp monotone_proj_Iic -lemma monotone.Icc_extend (hf : monotone f) : monotone (Icc_extend h f) := +protected lemma monotone.Icc_extend (hf : monotone f) : monotone (Icc_extend h f) := hf.comp $ monotone_proj_Icc h +lemma strict_mono.strict_mono_on_Ici_extend {f : Ici a → β} (hf : strict_mono f) : + strict_mono_on (Ici_extend f) (Ici a) := +hf.comp_strict_mono_on strict_mono_on_proj_Ici + +lemma strict_mono.strict_mono_on_Iic_extend {f : Iic b → β} (hf : strict_mono f) : + strict_mono_on (Iic_extend f) (Iic b) := +hf.comp_strict_mono_on strict_mono_on_proj_Iic + lemma strict_mono.strict_mono_on_Icc_extend (hf : strict_mono f) : strict_mono_on (Icc_extend h f) (Icc a b) := hf.comp_strict_mono_on (strict_mono_on_proj_Icc h) + +protected lemma set.ord_connected.Ici_extend {s : set (Ici a)} (hs : s.ord_connected) : + {x | Ici_extend (∈ s) x}.ord_connected := +⟨λ x hx y hy z hz, hs.out hx hy ⟨max_le_max le_rfl hz.1, max_le_max le_rfl hz.2⟩⟩ + +protected lemma set.ord_connected.Iic_extend {s : set (Iic b)} (hs : s.ord_connected) : + {x | Iic_extend (∈ s) x}.ord_connected := +⟨λ x hx y hy z hz, hs.out hx hy ⟨min_le_min le_rfl hz.1, min_le_min le_rfl hz.2⟩⟩ + +protected lemma set.ord_connected.restrict (hs : s.ord_connected) : + {x | restrict t (∈ s) x}.ord_connected := +⟨λ x hx y hy z hz, hs.out hx hy hz⟩ From 3a2b5524a138b5d0b818b858b516d4ac8a484b03 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 12 Jul 2023 12:15:34 +0000 Subject: [PATCH 11/61] feat(data/fin/basic): extra instances that cover `fin 0` (#18970) These apply to `fin 0`, unlike the `comm_ring` instance which needs `ne_zero n`. Co-authored-by: Yury G. Kudryashov --- src/data/fin/basic.lean | 26 ++++++++++++++++++++++++-- src/data/zmod/defs.lean | 18 +++++++++++++++--- 2 files changed, 39 insertions(+), 5 deletions(-) diff --git a/src/data/fin/basic.lean b/src/data/fin/basic.lean index a5ba4089bbf05..1579e1a3837ba 100644 --- a/src/data/fin/basic.lean +++ b/src/data/fin/basic.lean @@ -423,6 +423,11 @@ by rcases n with _|_|n; simp [is_empty.subsingleton, unique.subsingleton, not_su section monoid +instance add_comm_semigroup (n : ℕ) : add_comm_semigroup (fin n) := +{ add := (+), + add_assoc := by simp [eq_iff_veq, add_def, add_assoc], + add_comm := by simp [eq_iff_veq, add_def, add_comm] } + @[simp] protected lemma add_zero [ne_zero n] (k : fin n) : k + 0 = k := by simp [eq_iff_veq, add_def, mod_eq_of_lt (is_lt k)] @@ -431,11 +436,10 @@ by simp [eq_iff_veq, add_def, mod_eq_of_lt (is_lt k)] instance add_comm_monoid (n : ℕ) [ne_zero n] : add_comm_monoid (fin n) := { add := (+), - add_assoc := by simp [eq_iff_veq, add_def, add_assoc], zero := 0, zero_add := fin.zero_add, add_zero := fin.add_zero, - add_comm := by simp [eq_iff_veq, add_def, add_comm] } + ..fin.add_comm_semigroup n } instance [ne_zero n] : add_monoid_with_one (fin n) := { one := 1, @@ -1363,6 +1367,24 @@ instance (n : ℕ) [ne_zero n] : add_comm_group (fin n) := ..fin.add_comm_monoid n, ..fin.has_neg n } +/-- Note this is more general than `fin.add_comm_group` as it applies (vacuously) to `fin 0` too. -/ +instance (n : ℕ) : has_involutive_neg (fin n) := +{ neg := has_neg.neg, + neg_neg := nat.cases_on n fin_zero_elim (λ i, neg_neg) } + +/-- Note this is more general than `fin.add_comm_group` as it applies (vacuously) to `fin 0` too. -/ +instance (n : ℕ) : is_cancel_add (fin n) := +{ add_left_cancel := nat.cases_on n fin_zero_elim (λ i _ _ _, add_left_cancel), + add_right_cancel := nat.cases_on n fin_zero_elim (λ i _ _ _, add_right_cancel) } + +/-- Note this is more general than `fin.add_comm_group` as it applies (vacuously) to `fin 0` too. -/ +instance (n : ℕ) : add_left_cancel_semigroup (fin n) := +{ ..fin.add_comm_semigroup n, .. fin.is_cancel_add n } + +/-- Note this is more general than `fin.add_comm_group` as it applies (vacuously) to `fin 0` too. -/ +instance (n : ℕ) : add_right_cancel_semigroup (fin n) := +{ ..fin.add_comm_semigroup n, .. fin.is_cancel_add n } + protected lemma coe_neg (a : fin n) : ((-a : fin n) : ℕ) = (n - a) % n := rfl protected lemma coe_sub (a b : fin n) : ((a - b : fin n) : ℕ) = (a + (n - b)) % n := diff --git a/src/data/zmod/defs.lean b/src/data/zmod/defs.lean index a8a6a3732836c..918a0caf9ab08 100644 --- a/src/data/zmod/defs.lean +++ b/src/data/zmod/defs.lean @@ -57,15 +57,27 @@ private lemma left_distrib_aux (n : ℕ) : ∀ a b c : fin n, a * (b + c) = a * ... ≡ (a * b) % n + (a * c) % n [MOD n] : (nat.mod_modeq _ _).symm.add (nat.mod_modeq _ _).symm) +instance (n : ℕ) : distrib (fin n) := +{ left_distrib := left_distrib_aux n, + right_distrib := λ a b c, by rw [mul_comm, left_distrib_aux, mul_comm _ b, mul_comm]; refl, + ..fin.add_comm_semigroup n, + ..fin.comm_semigroup n } + /-- Commutative ring structure on `fin n`. -/ instance (n : ℕ) [ne_zero n] : comm_ring (fin n) := { one_mul := fin.one_mul, mul_one := fin.mul_one, - left_distrib := left_distrib_aux n, - right_distrib := λ a b c, by rw [mul_comm, left_distrib_aux, mul_comm _ b, mul_comm]; refl, ..fin.add_monoid_with_one, ..fin.add_comm_group n, - ..fin.comm_semigroup n } + ..fin.comm_semigroup n, + ..fin.distrib n } + +/-- Note this is more general than `fin.comm_ring` as it applies (vacuously) to `fin 0` too. -/ +instance (n : ℕ) : has_distrib_neg (fin n) := +{ neg := has_neg.neg, + mul_neg := nat.cases_on n fin_zero_elim $ λ i, mul_neg, + neg_mul := nat.cases_on n fin_zero_elim $ λ i, neg_mul, + ..fin.has_involutive_neg n } end fin From 92bd7b1ffeb306a89f450bee126ddd8a284c259d Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 12 Jul 2023 15:28:30 +0000 Subject: [PATCH 12/61] feat(analysis/convex): convexity of n-ary sums (#18943) --- src/analysis/convex/basic.lean | 30 +++++++++++++++++++ src/analysis/convex/combination.lean | 20 +++++++++++++ src/analysis/convex/hull.lean | 3 ++ src/measure_theory/measure/haar/of_basis.lean | 15 ++-------- 4 files changed, 55 insertions(+), 13 deletions(-) diff --git a/src/analysis/convex/basic.lean b/src/analysis/convex/basic.lean index 6825f10078981..d1520b6a6ab74 100644 --- a/src/analysis/convex/basic.lean +++ b/src/analysis/convex/basic.lean @@ -152,6 +152,9 @@ convex_iff_pairwise_pos.mpr (h.pairwise _) lemma convex_singleton (c : E) : convex 𝕜 ({c} : set E) := subsingleton_singleton.convex +lemma convex_zero : convex 𝕜 (0 : set E) := +convex_singleton _ + lemma convex_segment (x y : E) : convex 𝕜 [x -[𝕜] y] := begin rintro p ⟨ap, bp, hap, hbp, habp, rfl⟩ q ⟨aq, bq, haq, hbq, habq, rfl⟩ a b ha hb hab, @@ -190,6 +193,33 @@ hs.linear_preimage $ hf.mk' f lemma convex.add {t : set E} (hs : convex 𝕜 s) (ht : convex 𝕜 t) : convex 𝕜 (s + t) := by { rw ← add_image_prod, exact (hs.prod ht).is_linear_image is_linear_map.is_linear_map_add } +variables (𝕜 E) + +/-- The convex sets form an additive submonoid under pointwise addition. -/ +def convex_add_submonoid : add_submonoid (set E) := +{ carrier := {s : set E | convex 𝕜 s}, + zero_mem' := convex_zero, + add_mem' := λ s t, convex.add } + +@[simp, norm_cast] +lemma coe_convex_add_submonoid : ↑(convex_add_submonoid 𝕜 E) = {s : set E | convex 𝕜 s} := rfl + +variables {𝕜 E} + +@[simp] lemma mem_convex_add_submonoid {s : set E} : + s ∈ convex_add_submonoid 𝕜 E ↔ convex 𝕜 s := +iff.rfl + +lemma convex_list_sum {l : list (set E)} (h : ∀ i ∈ l, convex 𝕜 i) : convex 𝕜 l.sum := +(convex_add_submonoid 𝕜 E).list_sum_mem h + +lemma convex_multiset_sum {s : multiset (set E)} (h : ∀ i ∈ s, convex 𝕜 i) : convex 𝕜 s.sum := +(convex_add_submonoid 𝕜 E).multiset_sum_mem _ h + +lemma convex_sum {ι} {s : finset ι} (t : ι → set E) (h : ∀ i ∈ s, convex 𝕜 (t i)) : + convex 𝕜 (∑ i in s, t i) := +(convex_add_submonoid 𝕜 E).sum_mem h + lemma convex.vadd (hs : convex 𝕜 s) (z : E) : convex 𝕜 (z +ᵥ s) := by { simp_rw [←image_vadd, vadd_eq_add, ←singleton_add], exact (convex_singleton _).add hs } diff --git a/src/analysis/convex/combination.lean b/src/analysis/convex/combination.lean index 8cc8ae375c102..15443b1b1fa4f 100644 --- a/src/analysis/convex/combination.lean +++ b/src/analysis/convex/combination.lean @@ -418,9 +418,29 @@ lemma convex_hull_add (s t : set E) : convex_hull R (s + t) = convex_hull R s + by simp_rw [←image2_add, ←image_prod, is_linear_map.is_linear_map_add.convex_hull_image, convex_hull_prod] +variables (R E) +/-- `convex_hull` is an additive monoid morphism under pointwise addition. -/ +@[simps] +def convex_hull_add_monoid_hom : set E →+ set E := +{ to_fun := convex_hull R, + map_add' := convex_hull_add, + map_zero' := convex_hull_zero } +variables {R E} + lemma convex_hull_sub (s t : set E) : convex_hull R (s - t) = convex_hull R s - convex_hull R t := by simp_rw [sub_eq_add_neg, convex_hull_add, convex_hull_neg] +lemma convex_hull_list_sum (l : list (set E)) : convex_hull R l.sum = (l.map $ convex_hull R).sum := +map_list_sum (convex_hull_add_monoid_hom R E) l + +lemma convex_hull_multiset_sum (s : multiset (set E)) : + convex_hull R s.sum = (s.map $ convex_hull R).sum := +map_multiset_sum (convex_hull_add_monoid_hom R E) s + +lemma convex_hull_sum {ι} (s : finset ι) (t : ι → set E) : + convex_hull R (∑ i in s, t i) = ∑ i in s, convex_hull R (t i):= +map_sum (convex_hull_add_monoid_hom R E) _ _ + /-! ### `std_simplex` -/ variables (ι) [fintype ι] {f : ι → R} diff --git a/src/analysis/convex/hull.lean b/src/analysis/convex/hull.lean index efbc727029664..535a17885c56f 100644 --- a/src/analysis/convex/hull.lean +++ b/src/analysis/convex/hull.lean @@ -101,6 +101,9 @@ lemma segment_subset_convex_hull (hx : x ∈ s) (hy : y ∈ s) : segment 𝕜 x @[simp] lemma convex_hull_singleton (x : E) : convex_hull 𝕜 ({x} : set E) = {x} := (convex_singleton x).convex_hull_eq +@[simp] lemma convex_hull_zero : convex_hull 𝕜 (0 : set E) = 0 := +convex_hull_singleton 0 + @[simp] lemma convex_hull_pair (x y : E) : convex_hull 𝕜 {x, y} = segment 𝕜 x y := begin refine (convex_hull_min _ $ convex_segment _ _).antisymm diff --git a/src/measure_theory/measure/haar/of_basis.lean b/src/measure_theory/measure/haar/of_basis.lean index d8aed76d5d53d..dd1d618b23b18 100644 --- a/src/measure_theory/measure/haar/of_basis.lean +++ b/src/measure_theory/measure/haar/of_basis.lean @@ -126,24 +126,13 @@ end lemma convex_parallelepiped (v : ι → E) : convex ℝ (parallelepiped v) := begin rw parallelepiped_eq_sum_segment, - -- TODO: add `convex.sum` to match `convex.add` - let : add_submonoid (set E) := - { carrier := { s | convex ℝ s}, zero_mem' := convex_singleton _, add_mem' := λ x y, convex.add }, - exact this.sum_mem (λ i hi, convex_segment _ _), + exact convex_sum _ (λ i hi, convex_segment _ _), end /-- A `parallelepiped` is the convex hull of its vertices -/ lemma parallelepiped_eq_convex_hull (v : ι → E) : parallelepiped v = convex_hull ℝ (∑ i, {(0 : E), v i}) := -begin - -- TODO: add `convex_hull_sum` to match `convex_hull_add` - let : set E →+ set E := - { to_fun := convex_hull ℝ, - map_zero' := convex_hull_singleton _, - map_add' := convex_hull_add }, - simp_rw [parallelepiped_eq_sum_segment, ←convex_hull_pair], - exact (this.map_sum _ _).symm, -end +by simp_rw [convex_hull_sum, convex_hull_pair, parallelepiped_eq_sum_segment] /-- The axis aligned parallelepiped over `ι → ℝ` is a cuboid. -/ lemma parallelepiped_single [decidable_eq ι] (a : ι → ℝ) : From d9e96a3e3e0894e93e10aff5244f4c96655bac1c Mon Sep 17 00:00:00 2001 From: Devon Tuma Date: Wed, 12 Jul 2023 15:28:31 +0000 Subject: [PATCH 13/61] feat(data/list/dedup): Lemmas about `list.dedup` (#19142) Basic lemmas about `dedup` applied with `cons`, `nil`, `head`, and `tail`. Co-authored-by: Devon Tuma --- src/data/finset/basic.lean | 6 ++++++ src/data/list/dedup.lean | 37 +++++++++++++++++++++++++++++++++++++ 2 files changed, 43 insertions(+) diff --git a/src/data/finset/basic.lean b/src/data/finset/basic.lean index 43bca0b7aca26..ae0db3da0b9b8 100644 --- a/src/data/finset/basic.lean +++ b/src/data/finset/basic.lean @@ -2162,6 +2162,9 @@ by ext; simp @[simp] theorem to_finset_eq_empty {m : multiset α} : m.to_finset = ∅ ↔ m = 0 := finset.val_inj.symm.trans multiset.dedup_eq_zero +@[simp] lemma to_finset_nonempty : s.to_finset.nonempty ↔ s ≠ 0 := +by simp only [to_finset_eq_empty, ne.def, finset.nonempty_iff_ne_empty] + @[simp] lemma to_finset_subset : s.to_finset ⊆ t.to_finset ↔ s ⊆ t := by simp only [finset.subset_iff, multiset.subset_iff, multiset.mem_to_finset] @@ -2253,6 +2256,9 @@ by { ext, simp } @[simp] lemma to_finset_eq_empty_iff (l : list α) : l.to_finset = ∅ ↔ l = nil := by cases l; simp +@[simp] lemma to_finset_nonempty_iff (l : list α) : l.to_finset.nonempty ↔ l ≠ [] := +by simp [finset.nonempty_iff_ne_empty] + end list namespace finset diff --git a/src/data/list/dedup.lean b/src/data/list/dedup.lean index 2bfdd74a3ad16..4b42facffb4c6 100644 --- a/src/data/list/dedup.lean +++ b/src/data/list/dedup.lean @@ -57,8 +57,45 @@ theorem subset_dedup (l : list α) : l ⊆ dedup l := theorem nodup_dedup : ∀ l : list α, nodup (dedup l) := pairwise_pw_filter +theorem head_dedup [inhabited α] (l : list α) : + l.dedup.head = if l.head ∈ l.tail then l.tail.dedup.head else l.head := +match l with +| [] := rfl +| (a :: l) := by { by_cases ha : a ∈ l; simp [ha, list.dedup_cons_of_mem] } +end + +theorem tail_dedup [inhabited α] (l : list α) : + l.dedup.tail = if l.head ∈ l.tail then l.tail.dedup.tail else l.tail.dedup := +match l with +| [] := rfl +| (a :: l) := by { by_cases ha : a ∈ l; simp [ha, list.dedup_cons_of_mem] } +end + theorem dedup_eq_self {l : list α} : dedup l = l ↔ nodup l := pw_filter_eq_self +theorem dedup_eq_cons (l : list α) (a : α) (l' : list α) : + l.dedup = a :: l' ↔ a ∈ l ∧ a ∉ l' ∧ l.dedup.tail = l' := +begin + refine ⟨λ h, _, λ h, _⟩, + { refine ⟨mem_dedup.1 (h.symm ▸ mem_cons_self _ _), λ ha, _, by rw [h, tail_cons]⟩, + have : count a l.dedup ≤ 1 := nodup_iff_count_le_one.1 (nodup_dedup l) a, + rw [h, count_cons_self, add_le_iff_nonpos_left] at this, + exact (not_le_of_lt (count_pos.2 ha) this) }, + { have := @cons_head_tail α ⟨a⟩ _ (ne_nil_of_mem (mem_dedup.2 h.1)), + have hal : a ∈ l.dedup := mem_dedup.2 h.1, + rw [← this, mem_cons_iff, or_iff_not_imp_right] at hal, + exact this ▸ h.2.2.symm ▸ (cons_eq_cons.2 ⟨(hal (h.2.2.symm ▸ h.2.1)).symm, rfl⟩) } +end + +@[simp] theorem dedup_eq_nil (l : list α) : l.dedup = [] ↔ l = [] := +begin + induction l with a l hl, + { exact iff.rfl }, + { by_cases h : a ∈ l, + { simp only [list.dedup_cons_of_mem h, hl, list.ne_nil_of_mem h] }, + { simp only [list.dedup_cons_of_not_mem h, list.cons_ne_nil] } } +end + protected lemma nodup.dedup {l : list α} (h : l.nodup) : l.dedup = l := list.dedup_eq_self.2 h From d3b54a9f9613af6ae818c7fee31872db11c4d928 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 12 Jul 2023 15:28:32 +0000 Subject: [PATCH 14/61] fix(tactic/interval_cases): instantiate metavars (#19232) The test included in this commit fails without this change. This does not need forward-porting, the test already passes in Lean 4. --- src/tactic/interval_cases.lean | 12 ++++++------ test/interval_cases.lean | 7 +++++++ 2 files changed, 13 insertions(+), 6 deletions(-) diff --git a/src/tactic/interval_cases.lean b/src/tactic/interval_cases.lean index 9cffa15e8d41a..5fd7bfac4410f 100644 --- a/src/tactic/interval_cases.lean +++ b/src/tactic/interval_cases.lean @@ -47,14 +47,14 @@ return that proof. -- We use `expr.to_rat` merely to decide if an `expr` is an explicit number. -- It would be more natural to use `expr.to_int`, but that hasn't been implemented. meta def gives_upper_bound (n e : expr) : tactic expr := -do t ← infer_type e, +do t ← infer_type e >>= instantiate_mvars, match t with | `(%%n' < %%b) := do guard (n = n'), b ← b.to_rat, return e | `(%%b > %%n') := do guard (n = n'), b ← b.to_rat, return e | `(%%n' ≤ %%b) := do guard (n = n'), b ← b.to_rat, - tn ← infer_type n, + tn ← infer_type n >>= instantiate_mvars, match tn with | `(ℕ) := to_expr ``(nat.lt_add_one_iff.mpr %%e) | `(ℕ+) := to_expr ``(pnat.lt_add_one_iff.mpr %%e) @@ -64,7 +64,7 @@ do t ← infer_type e, | `(%%b ≥ %%n') := do guard (n = n'), b ← b.to_rat, - tn ← infer_type n, + tn ← infer_type n >>= instantiate_mvars, match tn with | `(ℕ) := to_expr ``(nat.lt_add_one_iff.mpr %%e) | `(ℕ+) := to_expr ``(pnat.lt_add_one_iff.mpr %%e) @@ -80,14 +80,14 @@ for some explicit `b`, return that proof. -/ meta def gives_lower_bound (n e : expr) : tactic expr := -do t ← infer_type e, +do t ← infer_type e >>= instantiate_mvars, match t with | `(%%n' ≥ %%b) := do guard (n = n'), b ← b.to_rat, return e | `(%%b ≤ %%n') := do guard (n = n'), b ← b.to_rat, return e | `(%%n' > %%b) := do guard (n = n'), b ← b.to_rat, - tn ← infer_type n, + tn ← infer_type n >>= instantiate_mvars, match tn with | `(ℕ) := to_expr ``(nat.add_one_le_iff.mpr %%e) | `(ℕ+) := to_expr ``(pnat.add_one_le_iff.mpr %%e) @@ -97,7 +97,7 @@ do t ← infer_type e, | `(%%b < %%n') := do guard (n = n'), b ← b.to_rat, - tn ← infer_type n, + tn ← infer_type n >>= instantiate_mvars, match tn with | `(ℕ) := to_expr ``(nat.add_one_le_iff.mpr %%e) | `(ℕ+) := to_expr ``(pnat.add_one_le_iff.mpr %%e) diff --git a/test/interval_cases.lean b/test/interval_cases.lean index 7e9c8e528b0a1..5c9f6074378f2 100644 --- a/test/interval_cases.lean +++ b/test/interval_cases.lean @@ -136,6 +136,13 @@ begin exact h, end +example : ∀ y, y ≤ 3 → true := +begin + refine λ y hy, _, + interval_cases y, + all_goals { trivial }, +end + /- Sadly, this one doesn't work, reporting: `deep recursion was detected at 'expression equality test'` From 41bef4ae1254365bc190aee63b947674d2977f01 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 12 Jul 2023 18:27:42 +0000 Subject: [PATCH 15/61] feat(analysis/normed/group/basic): norm lemmas for lipschitz and antilipschitz (#19103) This also corrects some nonsense names produced by to_additive. --- src/analysis/calculus/fderiv/basic.lean | 2 +- src/analysis/normed/group/basic.lean | 30 ++++++++++++++++--- .../normed_space/continuous_linear_map.lean | 2 +- 3 files changed, 28 insertions(+), 6 deletions(-) diff --git a/src/analysis/calculus/fderiv/basic.lean b/src/analysis/calculus/fderiv/basic.lean index b7fcebc366f5d..ca62479ba25cc 100644 --- a/src/analysis/calculus/fderiv/basic.lean +++ b/src/analysis/calculus/fderiv/basic.lean @@ -693,7 +693,7 @@ lemma has_fderiv_at_filter.is_O_sub_rev (hf : has_fderiv_at_filter f f' x L) {C} (λ x', x' - x) =O[L] (λ x', f x' - f x) := have (λ x', x' - x) =O[L] (λ x', f' (x' - x)), from is_O_iff.2 ⟨C, eventually_of_forall $ λ x', - add_monoid_hom_class.bound_of_antilipschitz f' hf' _⟩, + zero_hom_class.bound_of_antilipschitz f' hf' _⟩, (this.trans (hf.trans_is_O this).right_is_O_add).congr (λ _, rfl) (λ _, sub_add_cancel _ _) end continuous diff --git a/src/analysis/normed/group/basic.lean b/src/analysis/normed/group/basic.lean index 7b25f1a4fc002..6fe45df80f9c3 100644 --- a/src/analysis/normed/group/basic.lean +++ b/src/analysis/normed/group/basic.lean @@ -387,7 +387,8 @@ by simpa [dist_eq_norm_div] using dist_triangle a 1 b ‖a₁ / a₂‖ ≤ r₁ + r₂ := (norm_div_le a₁ a₂).trans $ add_le_add H₁ H₂ -@[to_additive] lemma dist_le_norm_mul_norm (a b : E) : dist a b ≤ ‖a‖ + ‖b‖ := +@[to_additive dist_le_norm_add_norm] lemma dist_le_norm_add_norm' (a b : E) : + dist a b ≤ ‖a‖ + ‖b‖ := by { rw dist_eq_norm_div, apply norm_div_le } @[to_additive abs_norm_sub_norm_le] lemma abs_norm_sub_norm_le' (a b : E) : |‖a‖ - ‖b‖| ≤ ‖a / b‖ := @@ -656,9 +657,29 @@ by rw [emetric.mem_ball, edist_eq_coe_nnnorm'] antilipschitz_with K f := antilipschitz_with.of_le_mul_dist $ λ x y, by simpa only [dist_eq_norm_div, map_div] using h (x / y) -@[to_additive] lemma monoid_hom_class.bound_of_antilipschitz [monoid_hom_class 𝓕 E F] (f : 𝓕) +@[to_additive lipschitz_with.norm_le_mul] +lemma lipschitz_with.norm_le_mul' {f : E → F} + {K : ℝ≥0} (h : lipschitz_with K f) (hf : f 1 = 1) (x) : ‖f x‖ ≤ K * ‖x‖ := +by simpa only [dist_one_right, hf] using h.dist_le_mul x 1 + +@[to_additive lipschitz_with.nnorm_le_mul] +lemma lipschitz_with.nnorm_le_mul' {f : E → F} + {K : ℝ≥0} (h : lipschitz_with K f) (hf : f 1 = 1) (x) : ‖f x‖₊ ≤ K * ‖x‖₊ := +h.norm_le_mul' hf x + +@[to_additive antilipschitz_with.le_mul_norm] +lemma antilipschitz_with.le_mul_norm' {f : E → F} + {K : ℝ≥0} (h : antilipschitz_with K f) (hf : f 1 = 1) (x) : ‖x‖ ≤ K * ‖f x‖ := +by simpa only [dist_one_right, hf] using h.le_mul_dist x 1 + +@[to_additive antilipschitz_with.le_mul_nnnorm] +lemma antilipschitz_with.le_mul_nnnorm' {f : E → F} + {K : ℝ≥0} (h : antilipschitz_with K f) (hf : f 1 = 1) (x) : ‖x‖₊ ≤ K * ‖f x‖₊ := +h.le_mul_norm' hf x + +@[to_additive] lemma one_hom_class.bound_of_antilipschitz [one_hom_class 𝓕 E F] (f : 𝓕) {K : ℝ≥0} (h : antilipschitz_with K f) (x) : ‖x‖ ≤ K * ‖f x‖ := -by simpa only [dist_one_right, map_one] using h.le_mul_dist x 1 +h.le_mul_nnnorm' (map_one f) x end nnnorm @@ -1285,7 +1306,8 @@ end (hg : lipschitz_with Kg (g / f)) (hK : Kg < Kf⁻¹) : antilipschitz_with (Kf⁻¹ - Kg)⁻¹ g := by simpa only [pi.div_apply, mul_div_cancel'_right] using hf.mul_lipschitz_with hg hK -@[to_additive] lemma le_mul_norm_div {f : E → F} (hf : antilipschitz_with K f) (x y : E) : +@[to_additive le_mul_norm_sub] +lemma le_mul_norm_div {f : E → F} (hf : antilipschitz_with K f) (x y : E) : ‖x / y‖ ≤ K * ‖f x / f y‖ := by simp [← dist_eq_norm_div, hf.le_mul_dist x y] diff --git a/src/analysis/normed_space/continuous_linear_map.lean b/src/analysis/normed_space/continuous_linear_map.lean index 87ad5e2c41002..a46af72230230 100644 --- a/src/analysis/normed_space/continuous_linear_map.lean +++ b/src/analysis/normed_space/continuous_linear_map.lean @@ -105,7 +105,7 @@ add_monoid_hom_class.antilipschitz_of_bound _ h lemma bound_of_antilipschitz (f : E →SL[σ] F) {K : ℝ≥0} (h : antilipschitz_with K f) (x) : ‖x‖ ≤ K * ‖f x‖ := -add_monoid_hom_class.bound_of_antilipschitz _ h x +zero_hom_class.bound_of_antilipschitz _ h x end continuous_linear_map From b01d6eb9d0a308807af54319b264d0994b91774b Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Pol=27tta=20/=20Miyahara=20K=C5=8D?= Date: Wed, 12 Jul 2023 18:27:44 +0000 Subject: [PATCH 16/61] fix(control/traversable/derive): the `functor` deriving handler makes weird definitions for inductive types which have recursive arguments separated by a non-recursive argument (#19228) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit I found this bug during the porting. The current deriving handler doesn't work for this. ```lean @[derive [traversable,is_lawful_traversable]] inductive my_tree' (α : Type) | leaf : my_tree' | node : my_tree' → α → my_tree' → my_tree' ``` This is because the `functor` deriving handler makes weird definitions for inductive types which have recursive arguments separated by a non-recursive argument. Fortunatelly, the cause of this bug is just a mistake of the argument in `control.traversable.derive`. --- src/control/traversable/derive.lean | 2 +- test/traversable.lean | 5 +++++ 2 files changed, 6 insertions(+), 1 deletion(-) diff --git a/src/control/traversable/derive.lean b/src/control/traversable/derive.lean index 3bd145a69f05f..3545bc23feb9c 100644 --- a/src/control/traversable/derive.lean +++ b/src/control/traversable/derive.lean @@ -50,7 +50,7 @@ meta def map_constructor (c n : name) (f α β : expr) do g ← target, (_, args') ← mmap_accuml (λ (x : list expr) (y : bool × expr), if y.1 then pure (x.tail,x.head) - else prod.mk rec_call <$> map_field n g.app_fn f α β y.2) rec_call args₁, + else prod.mk x <$> map_field n g.app_fn f α β y.2) rec_call args₁, constr ← mk_const c, let r := constr.mk_app (args₀ ++ args'), return r diff --git a/test/traversable.lean b/test/traversable.lean index 6adc03f87b411..9751f19392291 100644 --- a/test/traversable.lean +++ b/test/traversable.lean @@ -46,6 +46,11 @@ inductive my_tree (α : Type) | leaf : my_tree | node : my_tree → my_tree → α → my_tree +@[derive [traversable,is_lawful_traversable]] +inductive my_tree' (α : Type) +| leaf : my_tree' +| node : my_tree' → α → my_tree' → my_tree' + section open my_tree (hiding traverse) From baa88307f3e699fa7054ef04ec79fa4f056169cb Mon Sep 17 00:00:00 2001 From: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> Date: Wed, 12 Jul 2023 21:38:36 +0000 Subject: [PATCH 17/61] feat(analysis/inner_product_space/of_norm): Create an inner product from a norm (#4798) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit A normed space respecting the polarization identity is an inner product space. Co-authored-by: Frédéric Dupuis Co-authored-by: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Co-authored-by: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Co-authored-by: Yaël Dillies --- docs/references.bib | 14 + src/analysis/inner_product_space/of_norm.lean | 379 ++++++++++++++++++ src/data/is_R_or_C/basic.lean | 6 + 3 files changed, 399 insertions(+) create mode 100644 src/analysis/inner_product_space/of_norm.lean diff --git a/docs/references.bib b/docs/references.bib index d97389ddd156d..03a341208d7bb 100644 --- a/docs/references.bib +++ b/docs/references.bib @@ -1406,6 +1406,20 @@ @Book{ james1999 url = {https://doi.org/10.1007/978-1-4471-3994-2} } +@Article{ Jordan1935, + title = "On inner products in linear, metric spaces", + author = "Jordan, P. and von Neumann, J.", + fjournal = {{Annals of Mathematics}}, + journal = "Ann. Math.", + volume = 36, + number = 3, + pages = "719-723", + month = jul, + year = 1935, + url = "http://www.mathematik.uni-muenchen.de/~michel/jordan-von_neumann_-_parallelogram_identity.pdf", + doi = {10.2307/1968653} +} + @Article{ joyal1977, author = {André Joyal}, title = {Remarques sur la théorie des jeux à deux personnes}, diff --git a/src/analysis/inner_product_space/of_norm.lean b/src/analysis/inner_product_space/of_norm.lean new file mode 100644 index 0000000000000..be0bf32167923 --- /dev/null +++ b/src/analysis/inner_product_space/of_norm.lean @@ -0,0 +1,379 @@ +/- +Copyright (c) 2020 Heather Macbeth. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Heather Macbeth +-/ + +import topology.algebra.algebra +import analysis.inner_product_space.basic + +/-! +# Inner product space derived from a norm + +This file defines an `inner_product_space` instance from a norm that respects the +parallellogram identity. The parallelogram identity is a way to express the inner product of `x` and +`y` in terms of the norms of `x`, `y`, `x + y`, `x - y`. + +## Main results + +- `inner_product_space.of_norm`: a normed space whose norm respects the parallellogram identity, + can be seen as an inner product space. + +## Implementation notes + +We define `inner_` + +$$\langle x, y \rangle := \frac{1}{4} (‖x + y‖^2 - ‖x - y‖^2 + i ‖ix + y‖ ^ 2 - i ‖ix - y‖^2)$$ + +and use the parallelogram identity + +$$‖x + y‖^2 + ‖x - y‖^2 = 2 (‖x‖^2 + ‖y‖^2)$$ + +to prove it is an inner product, i.e., that it is conjugate-symmetric (`inner_.conj_symm`) and +linear in the first argument. `add_left` is proved by judicious application of the parallelogram +identity followed by tedious arithmetic. `smul_left` is proved step by step, first noting that +$\langle λ x, y \rangle = λ \langle x, y \rangle$ for $λ ∈ ℕ$, $λ = -1$, hence $λ ∈ ℤ$ and $λ ∈ ℚ$ +by arithmetic. Then by continuity and the fact that ℚ is dense in ℝ, the same is true for ℝ. +The case of ℂ then follows by applying the result for ℝ and more arithmetic. + +## TODO + +Move upstream to `analysis.inner_product_space.basic`. + +## References + +- [Jordan, P. and von Neumann, J., *On inner products in linear, metric spaces*][Jordan1935] +- https://math.stackexchange.com/questions/21792/norms-induced-by-inner-products-and-the-parallelogram-law +- https://math.dartmouth.edu/archive/m113w10/public_html/jordan-vneumann-thm.pdf + +## Tags + +inner product space, Hilbert space, norm +-/ + +open is_R_or_C +open_locale complex_conjugate + +variables {𝕜 : Type*} [is_R_or_C 𝕜] (E : Type*) [normed_add_comm_group E] + +/-- Predicate for the parallelogram identity to hold in a normed group. This is a scalar-less +version of `inner_product_space`. If you have an `inner_product_spaceable` assumption, you can +locally upgrade that to `inner_product_space 𝕜 E` using `casesI nonempty_inner_product_space 𝕜 E`. +-/ +class inner_product_spaceable : Prop := +(parallelogram_identity : + ∀ x y : E, ‖x + y‖ * ‖x + y‖ + ‖x - y‖ * ‖x - y‖ = 2 * (‖x‖ * ‖x‖ + ‖y‖ * ‖y‖)) + +variables (𝕜) {E} + +lemma inner_product_space.to_inner_product_spaceable [inner_product_space 𝕜 E] : + inner_product_spaceable E := +⟨parallelogram_law_with_norm 𝕜⟩ + +@[priority 100] -- See note [lower instance priority] +instance inner_product_space.to_inner_product_spaceable_of_real [inner_product_space ℝ E] : + inner_product_spaceable E := +⟨parallelogram_law_with_norm ℝ⟩ + +variables [normed_space 𝕜 E] + +local notation `𝓚` := algebra_map ℝ 𝕜 + +/-- Auxiliary definition of the inner product derived from the norm. -/ +private noncomputable def inner_ (x y : E) : 𝕜 := +4⁻¹ * ((𝓚 ‖x + y‖) * (𝓚 ‖x + y‖) - (𝓚 ‖x - y‖) * (𝓚 ‖x - y‖) + + (I:𝕜) * (𝓚 ‖(I:𝕜) • x + y‖) * (𝓚 ‖(I:𝕜) • x + y‖) + - (I:𝕜) * (𝓚 ‖(I:𝕜) • x - y‖) * (𝓚 ‖(I:𝕜) • x - y‖)) + +namespace inner_product_spaceable + +variables {𝕜} (E) + +/-- Auxiliary definition for the `add_left` property -/ +private def inner_prop (r : 𝕜) : Prop := ∀ x y : E, inner_ 𝕜 (r • x) y = conj r * inner_ 𝕜 x y + +variables {E} + +lemma inner_prop_neg_one : inner_prop E ((-1 : ℤ) : 𝕜) := +begin + intros x y, + simp only [inner_, neg_mul_eq_neg_mul, one_mul, int.cast_one, one_smul, ring_hom.map_one, + map_neg, int.cast_neg, neg_smul, neg_one_mul], + rw neg_mul_comm, + congr' 1, + have h₁ : ‖-x - y‖ = ‖x + y‖, + { rw [←neg_add', norm_neg], }, + have h₂ : ‖-x + y‖ = ‖x - y‖, + { rw [←neg_sub, norm_neg, sub_eq_neg_add], }, + have h₃ : ‖(I : 𝕜) • (-x) + y‖ = ‖(I : 𝕜) • x - y‖, + { rw [←neg_sub, norm_neg, sub_eq_neg_add, ←smul_neg], }, + have h₄ : ‖(I : 𝕜) • (-x) - y‖ = ‖(I : 𝕜) • x + y‖, + { rw [smul_neg, ←neg_add', norm_neg] }, + rw [h₁, h₂, h₃, h₄], + ring, +end + +lemma continuous.inner_ {f g : ℝ → E} (hf : continuous f) (hg : continuous g) : + continuous (λ x, inner_ 𝕜 (f x) (g x)) := +by { unfold inner_, continuity } + +lemma inner_.norm_sq (x : E) : ‖x‖ ^ 2 = re (inner_ 𝕜 x x) := +begin + simp only [inner_], + have h₁ : norm_sq (4 : 𝕜) = 16, + { have : ((4 : ℝ) : 𝕜) = (4 : 𝕜), + { simp only [of_real_one, of_real_bit0] }, + rw [←this, norm_sq_eq_def', + is_R_or_C.norm_of_nonneg (by norm_num : (0 : ℝ) ≤ 4)], + norm_num }, + have h₂ : ‖x + x‖ = 2 * ‖x‖, + { rw [←two_smul 𝕜, norm_smul, is_R_or_C.norm_two] }, + simp only [inner, h₁, h₂, one_im, bit0_zero, add_zero, norm_zero, I_re, of_real_im, map_add, + bit0_im, zero_div, zero_mul, add_monoid_hom.map_neg, of_real_re, map_sub, sub_zero, inv_re, + one_re, inv_im, bit0_re, mul_re, mul_zero, sub_self, neg_zero, algebra_map_eq_of_real], + ring, +end + +lemma inner_.conj_symm (x y : E) : conj (inner_ 𝕜 y x) = inner_ 𝕜 x y := +begin + simp only [inner_], + have h4 : conj (4⁻¹ : 𝕜) = 4⁻¹, + { rw [is_R_or_C.conj_inv, ←of_real_one, ←of_real_bit0, ←of_real_bit0, conj_of_real] }, + rw [map_mul, h4], + congr' 1, + simp only [map_sub, map_add, algebra_map_eq_of_real, ←of_real_mul, conj_of_real, map_mul, conj_I], + rw [add_comm y x, norm_sub_rev], + by_cases hI : (I : 𝕜) = 0, + { simp only [hI, neg_zero, zero_mul] }, + have h₁ : ‖(I : 𝕜) • y - x‖ = ‖(I : 𝕜) • x + y‖, + { transitivity ‖(I : 𝕜) • ((I : 𝕜) • y - x)‖, + { rw [norm_smul, norm_I_of_ne_zero hI, one_mul] }, + { rw [smul_sub, smul_smul, I_mul_I_of_nonzero hI, neg_one_smul, ←neg_add', add_comm, + norm_neg] } }, + have h₂ : ‖(I : 𝕜) • y + x‖ = ‖(I : 𝕜) • x - y‖, + { transitivity ‖(I : 𝕜) • ((I : 𝕜) • y + x)‖, + { rw [norm_smul, norm_I_of_ne_zero hI, one_mul] }, + { rw [smul_add, smul_smul, I_mul_I_of_nonzero hI, neg_one_smul, ←neg_add_eq_sub] }}, + rw [h₁, h₂, ←sub_add_eq_add_sub], + simp only [neg_mul, sub_eq_add_neg, neg_neg], +end + +variables [inner_product_spaceable E] + +private lemma add_left_aux1 (x y z : E) : + ‖x + y + z‖ * ‖x + y + z‖ = + (‖2 • x + y‖ * ‖2 • x + y‖ + ‖2 • z + y‖ * ‖2 • z + y‖) / 2 - ‖x - z‖ * ‖x - z‖ := +begin + rw [eq_sub_iff_add_eq, eq_div_iff (two_ne_zero' ℝ), mul_comm _ (2 : ℝ), eq_comm], + convert parallelogram_identity (x + y + z) (x - z) using 4; { rw two_smul, abel } +end + +private lemma add_left_aux2 (x y z : E) : + ‖x + y - z‖ * ‖x + y - z‖ = + (‖2 • x + y‖ * ‖2 • x + y‖ + ‖y - 2 • z‖ * ‖y - 2 • z‖) / 2 - ‖x + z‖ * ‖x + z‖ := +begin + rw [eq_sub_iff_add_eq, eq_div_iff (two_ne_zero' ℝ), mul_comm _ (2 : ℝ), eq_comm], + have h₀ := parallelogram_identity (x + y - z) (x + z), + convert h₀ using 4; { rw two_smul, abel } +end + +private lemma add_left_aux2' (x y z : E) : + ‖x + y + z‖ * ‖x + y + z‖ - ‖x + y - z‖ * ‖x + y - z‖ = + ‖x + z‖ * ‖x + z‖ - ‖x - z‖ * ‖x - z‖ + + (‖2 • z + y‖ * ‖2 • z + y‖ - ‖y - 2 • z‖ * ‖y - 2 • z‖) / 2 := +by { rw [add_left_aux1 , add_left_aux2], ring } + +private lemma add_left_aux3 (y z : E) : + ‖2 • z + y‖ * ‖2 • z + y‖ = 2 * (‖y + z‖ * ‖y + z‖ + ‖z‖ * ‖z‖) - ‖y‖ * ‖y‖ := +begin + apply eq_sub_of_add_eq, + convert parallelogram_identity (y + z) z using 4; { try { rw two_smul }, abel } +end + +private lemma add_left_aux4 (y z : E) : + ‖y - 2 • z‖ * ‖y - 2 • z‖ = 2 * (‖y - z‖ * ‖y - z‖ + ‖z‖ * ‖z‖) - ‖y‖ * ‖y‖ := +begin + apply eq_sub_of_add_eq', + have h₀ := parallelogram_identity (y - z) z, + convert h₀ using 4; { try { rw two_smul }, abel } +end + +private lemma add_left_aux4' (y z : E) : + (‖2 • z + y‖ * ‖2 • z + y‖ - ‖y - 2 • z‖ * ‖y - 2 • z‖) / 2 = + (‖y + z‖ * ‖y + z‖) - (‖y - z‖ * ‖y - z‖) := +by { rw [add_left_aux3 , add_left_aux4], ring } + +private lemma add_left_aux5 (x y z : E) : + ‖(I : 𝕜) • (x + y) + z‖ * ‖(I : 𝕜) • (x + y) + z‖ = + (‖(I : 𝕜) • (2 • x + y)‖ * ‖(I : 𝕜) • (2 • x + y)‖ + + ‖(I : 𝕜) • y + 2 • z‖ * ‖(I : 𝕜) • y + 2 • z‖) / 2 - ‖(I : 𝕜) • x - z‖ * ‖(I : 𝕜) • x - z‖ := +begin + rw [eq_sub_iff_add_eq, eq_div_iff (two_ne_zero' ℝ), mul_comm _ (2 : ℝ), eq_comm], + have h₀ := parallelogram_identity ((I : 𝕜) • (x + y) + z) ((I : 𝕜) • x - z), + convert h₀ using 4; { try { simp only [two_smul, smul_add] }, abel } +end + +private lemma add_left_aux6 (x y z : E) : + ‖(I : 𝕜) • (x + y) - z‖ * ‖(I : 𝕜) • (x + y) - z‖ = + (‖(I : 𝕜) • (2 • x + y)‖ * ‖(I : 𝕜) • (2 • x + y)‖ + + ‖(I : 𝕜) • y - 2 • z‖ * ‖(I : 𝕜) • y - 2 • z‖) / 2 - + ‖(I : 𝕜) • x + z‖ * ‖(I : 𝕜) • x + z‖ := +begin + rw [eq_sub_iff_add_eq, eq_div_iff (two_ne_zero' ℝ), mul_comm _ (2 : ℝ), eq_comm], + have h₀ := parallelogram_identity ((I : 𝕜) • (x + y) - z) ((I : 𝕜) • x + z), + convert h₀ using 4; { try { simp only [two_smul, smul_add] }, abel } +end + +private lemma add_left_aux7 (y z : E) : + ‖(I : 𝕜) • y + 2 • z‖ * ‖(I : 𝕜) • y + 2 • z‖ = + 2 * (‖(I : 𝕜) • y + z‖ * ‖(I : 𝕜) • y + z‖ + ‖z‖ * ‖z‖) - ‖(I : 𝕜) • y‖ * ‖(I : 𝕜) • y‖ := +begin + apply eq_sub_of_add_eq, + have h₀ := parallelogram_identity ((I : 𝕜) • y + z) z, + convert h₀ using 4; { try { simp only [two_smul, smul_add] }, abel } +end + +private lemma add_left_aux8 (y z : E) : + ‖(I : 𝕜) • y - 2 • z‖ * ‖(I : 𝕜) • y - 2 • z‖ = + 2 * (‖(I : 𝕜) • y - z‖ * ‖(I : 𝕜) • y - z‖ + ‖z‖ * ‖z‖) - ‖(I : 𝕜) • y‖ * ‖(I : 𝕜) • y‖ := +begin + apply eq_sub_of_add_eq', + have h₀ := parallelogram_identity ((I : 𝕜) • y - z) z, + convert h₀ using 4; { try { simp only [two_smul, smul_add] }, abel } +end + +lemma add_left (x y z : E) : inner_ 𝕜 (x + y) z = inner_ 𝕜 x z + inner_ 𝕜 y z := +begin + simp only [inner_, ←mul_add], + congr, + simp only [mul_assoc, ←map_mul, add_sub_assoc, ←mul_sub, ←map_sub], + rw add_add_add_comm, + simp only [←map_add, ←mul_add], + congr, + { rw [←add_sub_assoc, add_left_aux2', add_left_aux4'] }, + { rw [add_left_aux5, add_left_aux6, add_left_aux7, add_left_aux8], + simp only [map_sub, map_mul, map_add, div_eq_mul_inv], + ring } +end + +lemma nat (n : ℕ) (x y : E) : inner_ 𝕜 ((n : 𝕜) • x) y = (n : 𝕜) * inner_ 𝕜 x y := +begin + induction n with n ih, + { simp only [inner_, nat.nat_zero_eq_zero, zero_sub, nat.cast_zero, zero_mul, eq_self_iff_true, + zero_smul, zero_add, mul_zero, sub_self, norm_neg, smul_zero], }, + { simp only [nat.cast_succ, add_smul, one_smul], + rw [add_left, ih, add_mul, one_mul] } +end + +private lemma nat_prop (r : ℕ) : inner_prop E (r : 𝕜) := +λ x y, by { simp only [map_nat_cast], exact nat r x y } + +private lemma int_prop (n : ℤ) : inner_prop E (n : 𝕜) := +begin + intros x y, + rw ←n.sign_mul_nat_abs, + simp only [int.cast_coe_nat, map_nat_cast, map_int_cast, int.cast_mul, map_mul, mul_smul], + obtain hn | rfl | hn := lt_trichotomy n 0, + { rw [int.sign_eq_neg_one_of_neg hn, inner_prop_neg_one ((n.nat_abs : 𝕜) • x), nat], + simp only [map_neg, neg_mul, one_mul, mul_eq_mul_left_iff, true_or, + int.nat_abs_eq_zero, eq_self_iff_true, int.cast_one, map_one, neg_inj, nat.cast_eq_zero, + int.cast_neg] }, + { simp only [inner_, int.cast_zero, zero_sub, nat.cast_zero, zero_mul, eq_self_iff_true, + int.sign_zero, zero_smul, zero_add, mul_zero, smul_zero, sub_self, norm_neg, + int.nat_abs_zero] }, + { rw int.sign_eq_one_of_pos hn, + simp only [one_mul, mul_eq_mul_left_iff, true_or, int.nat_abs_eq_zero, eq_self_iff_true, + int.cast_one, one_smul, nat.cast_eq_zero, nat] } +end + +private lemma rat_prop (r : ℚ) : inner_prop E (r : 𝕜) := +begin + intros x y, + have : (r.denom : 𝕜) ≠ 0, + { haveI : char_zero 𝕜 := is_R_or_C.char_zero_R_or_C, + exact_mod_cast r.pos.ne' }, + rw [←r.num_div_denom, ←mul_right_inj' this, ←nat r.denom _ y, smul_smul, rat.cast_div], + simp only [map_nat_cast, rat.cast_coe_nat, map_int_cast, rat.cast_coe_int, map_div₀], + rw [←mul_assoc, mul_div_cancel' _ this, int_prop _ x, map_int_cast], +end + +private lemma real_prop (r : ℝ) : inner_prop E (r : 𝕜) := +begin + intros x y, + revert r, + rw ←function.funext_iff, + refine rat.dense_embedding_coe_real.dense.equalizer _ _ (funext $ λ X, _), + { exact (continuous_of_real.smul continuous_const).inner_ continuous_const }, + { exact (continuous_conj.comp continuous_of_real).mul continuous_const }, + { simp only [function.comp_app, is_R_or_C.of_real_rat_cast, rat_prop _ _] } +end + +private lemma I_prop : inner_prop E (I : 𝕜) := +begin + by_cases hI : (I : 𝕜) = 0, + { rw [hI, ←nat.cast_zero], exact nat_prop _ }, + intros x y, + have hI' : (-I : 𝕜) * I = 1, + { rw [←inv_I, inv_mul_cancel hI], }, + rw [conj_I, inner_, inner_, mul_left_comm], + congr' 1, + rw [smul_smul, I_mul_I_of_nonzero hI, neg_one_smul], + rw [mul_sub, mul_add, mul_sub, + mul_assoc I (𝓚 ‖I • x - y‖), ←mul_assoc (-I) I, hI', one_mul, + mul_assoc I (𝓚 ‖I • x + y‖), ←mul_assoc (-I) I, hI', one_mul], + have h₁ : ‖-x - y‖ = ‖x + y‖, + { rw [←neg_add', norm_neg], }, + have h₂ : ‖-x + y‖ = ‖x - y‖, + { rw [←neg_sub, norm_neg, sub_eq_neg_add], }, + rw [h₁, h₂], + simp only [sub_eq_add_neg, mul_assoc], + rw [←neg_mul_eq_neg_mul, ←neg_mul_eq_neg_mul], + abel +end + +lemma inner_prop (r : 𝕜) : inner_prop E r := +begin + intros x y, + rw [←re_add_im r, add_smul, add_left, real_prop _ x, ←smul_smul, real_prop _ _ y, I_prop, + map_add, map_mul, conj_of_real, conj_of_real, conj_I], + ring, +end + +end inner_product_spaceable + +open inner_product_spaceable + +/-- **Fréchet–von Neumann–Jordan Theorem**. A normed space `E` whose norm satisfies the +parallelogram identity can be given a compatible inner product. -/ +noncomputable def inner_product_space.of_norm + (h : ∀ x y : E, ‖x + y‖ * ‖x + y‖ + ‖x - y‖ * ‖x - y‖ = 2 * (‖x‖ * ‖x‖ + ‖y‖ * ‖y‖)) : + inner_product_space 𝕜 E := +begin + haveI : inner_product_spaceable E := ⟨h⟩, + exact + { inner := inner_ 𝕜, + norm_sq_eq_inner := inner_.norm_sq, + conj_symm := inner_.conj_symm, + add_left := add_left, + smul_left := λ _ _ _, inner_prop _ _ _ } +end + +variables (𝕜 E) [inner_product_spaceable E] + +/-- **Fréchet–von Neumann–Jordan Theorem**. A normed space `E` whose norm satisfies the +parallelogram identity can be given a compatible inner product. Do +`casesI nonempty_inner_product_space 𝕜 E` to locally upgrade `inner_product_spaceable E` to +`inner_product_space 𝕜 E`. -/ +lemma nonempty_inner_product_space : nonempty (inner_product_space 𝕜 E) := +⟨{ inner := inner_ 𝕜, + norm_sq_eq_inner := inner_.norm_sq, + conj_symm := inner_.conj_symm, + add_left := add_left, + smul_left := λ _ _ _, inner_prop _ _ _ }⟩ + +variables {𝕜 E} [normed_space ℝ E] + +-- TODO: Replace `inner_product_space.to_uniform_convex_space` +@[priority 100] -- See note [lower instance priority] +instance inner_product_spaceable.to_uniform_convex_space : uniform_convex_space E := +by { casesI nonempty_inner_product_space ℝ E, apply_instance } diff --git a/src/data/is_R_or_C/basic.lean b/src/data/is_R_or_C/basic.lean index c747fef34bf60..71714a3f5cf1c 100644 --- a/src/data/is_R_or_C/basic.lean +++ b/src/data/is_R_or_C/basic.lean @@ -491,6 +491,12 @@ begin exact div_le_one_of_le (abs_im_le_norm _) (norm_nonneg _) end +lemma norm_I_of_ne_zero (hI : (I : K) ≠ 0) : ‖(I : K)‖ = 1 := +begin + rw [← mul_self_inj_of_nonneg (norm_nonneg I) zero_le_one, one_mul, ← norm_mul, + I_mul_I_of_nonzero hI, norm_neg, norm_one], +end + lemma re_eq_norm_of_mul_conj (x : K) : re (x * conj x) = ‖x * conj x‖ := by rw [mul_conj, of_real_re, norm_of_real, abs_of_nonneg (norm_sq_nonneg _)] From cd391184c85986113f8c00844cfe6dda1d34be3d Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 13 Jul 2023 10:41:08 +0000 Subject: [PATCH 18/61] feat(*/prod): `prod_prod_prod` equivs (#19235) These send `((a, b), (c, d))` to `((a, c), (b, d))`, and this commit provides this bundled as `equiv`, `add_equiv`, `mul_equiv`, `ring_equiv`, and `linear_equiv`. We already have something analogous for `tensor_product`. --- src/algebra/group/prod.lean | 20 ++++++++++++++++++++ src/algebra/ring/prod.lean | 29 ++++++++++++++++++++++++++++- src/linear_algebra/prod.lean | 24 +++++++++++++++++++++++- src/logic/equiv/basic.lean | 11 +++++++++++ 4 files changed, 82 insertions(+), 2 deletions(-) diff --git a/src/algebra/group/prod.lean b/src/algebra/group/prod.lean index 3fabcc474d8b4..2546fc56e9435 100644 --- a/src/algebra/group/prod.lean +++ b/src/algebra/group/prod.lean @@ -465,6 +465,26 @@ def prod_comm : M × N ≃* N × M := variables {M' N' : Type*} [mul_one_class M'] [mul_one_class N'] +section +variables (M N M' N') + +/-- Four-way commutativity of `prod`. The name matches `mul_mul_mul_comm`. -/ +@[to_additive prod_prod_prod_comm "Four-way commutativity of `prod`. +The name matches `mul_mul_mul_comm`", simps apply] +def prod_prod_prod_comm : (M × N) × (M' × N') ≃* (M × M') × (N × N') := +{ to_fun := λ mnmn, ((mnmn.1.1, mnmn.2.1), (mnmn.1.2, mnmn.2.2)), + inv_fun := λ mmnn, ((mmnn.1.1, mmnn.2.1), (mmnn.1.2, mmnn.2.2)), + map_mul' := λ mnmn mnmn', rfl, + ..equiv.prod_prod_prod_comm M N M' N', } + +@[simp, to_additive] lemma prod_prod_prod_comm_to_equiv : + (prod_prod_prod_comm M N M' N').to_equiv = equiv.prod_prod_prod_comm M N M' N' := rfl + +@[simp] lemma prod_prod_prod_comm_symm : + (prod_prod_prod_comm M N M' N').symm = prod_prod_prod_comm M M' N N' := rfl + +end + /--Product of multiplicative isomorphisms; the maps come from `equiv.prod_congr`.-/ @[to_additive prod_congr "Product of additive isomorphisms; the maps come from `equiv.prod_congr`."] def prod_congr (f : M ≃* M') (g : N ≃* N') : M × N ≃* M' × N' := diff --git a/src/algebra/ring/prod.lean b/src/algebra/ring/prod.lean index 1ac110c8c81a5..5f5c2da349c5e 100644 --- a/src/algebra/ring/prod.lean +++ b/src/algebra/ring/prod.lean @@ -212,7 +212,9 @@ end prod_map end ring_hom namespace ring_equiv -variables {R S} [non_assoc_semiring R] [non_assoc_semiring S] +variables {R S R' S'} +variables [non_assoc_semiring R] [non_assoc_semiring S] +variables [non_assoc_semiring R'] [non_assoc_semiring S'] /-- Swapping components as an equivalence of (semi)rings. -/ def prod_comm : R × S ≃+* S × R := @@ -229,6 +231,31 @@ ring_hom.ext $ λ _, rfl (ring_hom.snd S R).comp ↑(prod_comm : R × S ≃+* S × R) = ring_hom.fst R S := ring_hom.ext $ λ _, rfl +section +variables (R R' S S') + +/-- Four-way commutativity of `prod`. The name matches `mul_mul_mul_comm`. -/ +@[simps apply] +def prod_prod_prod_comm : (R × R') × (S × S') ≃+* (R × S) × (R' × S') := +{ to_fun := λ rrss, ((rrss.1.1, rrss.2.1), (rrss.1.2, rrss.2.2)), + inv_fun := λ rsrs, ((rsrs.1.1, rsrs.2.1), (rsrs.1.2, rsrs.2.2)), + .. add_equiv.prod_prod_prod_comm R R' S S', + .. mul_equiv.prod_prod_prod_comm R R' S S' } + +@[simp] lemma prod_prod_prod_comm_symm : + (prod_prod_prod_comm R R' S S').symm = prod_prod_prod_comm R S R' S' := rfl + +@[simp] lemma prod_prod_prod_comm_to_add_equiv : + (prod_prod_prod_comm R R' S S').to_add_equiv = add_equiv.prod_prod_prod_comm R R' S S' := rfl + +@[simp] lemma prod_prod_prod_comm_to_mul_equiv : + (prod_prod_prod_comm R R' S S').to_mul_equiv = mul_equiv.prod_prod_prod_comm R R' S S' := rfl + +@[simp] lemma prod_prod_prod_comm_to_equiv : + (prod_prod_prod_comm R R' S S').to_equiv = equiv.prod_prod_prod_comm R R' S S' := rfl + +end + variables (R S) [subsingleton S] /-- A ring `R` is isomorphic to `R × S` when `S` is the zero ring -/ diff --git a/src/linear_algebra/prod.lean b/src/linear_algebra/prod.lean index 2b190f6abffca..546d089982d5b 100644 --- a/src/linear_algebra/prod.lean +++ b/src/linear_algebra/prod.lean @@ -227,7 +227,7 @@ def prod_map (f : M →ₗ[R] M₃) (g : M₂ →ₗ[R] M₄) : (M × M₂) → lemma coe_prod_map (f : M →ₗ[R] M₃) (g : M₂ →ₗ[R] M₄) : ⇑(f.prod_map g) = prod.map f g := rfl - + @[simp] theorem prod_map_apply (f : M →ₗ[R] M₃) (g : M₂ →ₗ[R] M₄) (x) : f.prod_map g x = (f x.1, g x.2) := rfl @@ -559,6 +559,28 @@ def prod_comm (R M N : Type*) [semiring R] [add_comm_monoid M] [add_comm_monoid map_smul' := λ r ⟨m, n⟩, rfl, ..add_equiv.prod_comm } +section +variables (R M M₂ M₃ M₄) +variables [semiring R] +variables [add_comm_monoid M] [add_comm_monoid M₂] [add_comm_monoid M₃] [add_comm_monoid M₄] +variables [module R M] [module R M₂] [module R M₃] [module R M₄] + +/-- Four-way commutativity of `prod`. The name matches `mul_mul_mul_comm`. -/ +@[simps apply] +def prod_prod_prod_comm : ((M × M₂) × (M₃ × M₄)) ≃ₗ[R] (M × M₃) × (M₂ × M₄) := +{ to_fun := λ mnmn, ((mnmn.1.1, mnmn.2.1), (mnmn.1.2, mnmn.2.2)), + inv_fun := λ mmnn, ((mmnn.1.1, mmnn.2.1), (mmnn.1.2, mmnn.2.2)), + map_smul' := λ c mnmn, rfl, + ..add_equiv.prod_prod_prod_comm M M₂ M₃ M₄ } + +@[simp] lemma prod_prod_prod_comm_symm : + (prod_prod_prod_comm R M M₂ M₃ M₄).symm = prod_prod_prod_comm R M M₃ M₂ M₄ := rfl + +@[simp] lemma prod_prod_prod_comm_to_add_equiv : + (prod_prod_prod_comm R M M₂ M₃ M₄).to_add_equiv = add_equiv.prod_prod_prod_comm M M₂ M₃ M₄ := rfl + +end + section variables [semiring R] diff --git a/src/logic/equiv/basic.lean b/src/logic/equiv/basic.lean index 6b6601d3d857c..84745166d6e14 100644 --- a/src/logic/equiv/basic.lean +++ b/src/logic/equiv/basic.lean @@ -105,6 +105,17 @@ def prod_comm (α β : Type*) : α × β ≃ β × α := @[simps] def prod_assoc (α β γ : Sort*) : (α × β) × γ ≃ α × (β × γ) := ⟨λ p, (p.1.1, p.1.2, p.2), λ p, ((p.1, p.2.1), p.2.2), λ ⟨⟨a, b⟩, c⟩, rfl, λ ⟨a, ⟨b, c⟩⟩, rfl⟩ +/-- Four-way commutativity of `prod`. The name matches `mul_mul_mul_comm`. -/ +@[simps apply] +def prod_prod_prod_comm (α β γ δ : Type*) : (α × β) × (γ × δ) ≃ (α × γ) × (β × δ) := +{ to_fun := λ abcd, ((abcd.1.1, abcd.2.1), (abcd.1.2, abcd.2.2)), + inv_fun := λ acbd, ((acbd.1.1, acbd.2.1), (acbd.1.2, acbd.2.2)), + left_inv := λ ⟨⟨a, b⟩, ⟨c, d⟩⟩, rfl, + right_inv := λ ⟨⟨a, c⟩, ⟨b, d⟩⟩, rfl, } + +@[simp] lemma prod_prod_prod_comm_symm (α β γ δ : Type*) : + (prod_prod_prod_comm α β γ δ).symm = prod_prod_prod_comm α γ β δ := rfl + /-- Functions on `α × β` are equivalent to functions `α → β → γ`. -/ @[simps {fully_applied := ff}] def curry (α β γ : Type*) : (α × β → γ) ≃ (α → β → γ) := From d608fc5d4e69d4cc21885913fb573a88b0deb521 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 13 Jul 2023 13:57:44 +0000 Subject: [PATCH 19/61] feat(analysis/calculus/fderiv/mul): derivative of inverse in division rings (#19127) --- src/analysis/calculus/fderiv/mul.lean | 90 ++++++++++++++++++++++++- src/analysis/normed_space/spectrum.lean | 2 +- 2 files changed, 89 insertions(+), 3 deletions(-) diff --git a/src/analysis/calculus/fderiv/mul.lean b/src/analysis/calculus/fderiv/mul.lean index f52bc203dcd4d..12ccaad983130 100644 --- a/src/analysis/calculus/fderiv/mul.lean +++ b/src/analysis/calculus/fderiv/mul.lean @@ -456,13 +456,99 @@ begin units.inv_mul, add_sub_cancel'_right, mul_sub, sub_mul, one_mul, sub_neg_eq_add] end -lemma differentiable_at_inverse (x : Rˣ) : differentiable_at 𝕜 (@ring.inverse R _) x := -(has_fderiv_at_ring_inverse x).differentiable_at +lemma differentiable_at_inverse {x : R} (hx : is_unit x) : + differentiable_at 𝕜 (@ring.inverse R _) x := +let ⟨u, hu⟩ := hx in hu ▸ (has_fderiv_at_ring_inverse u).differentiable_at + +lemma differentiable_within_at_inverse {x : R} (hx : is_unit x) (s : set R): + differentiable_within_at 𝕜 (@ring.inverse R _) s x := +(differentiable_at_inverse hx).differentiable_within_at + +lemma differentiable_on_inverse : differentiable_on 𝕜 (@ring.inverse R _) {x | is_unit x} := +λ x hx, differentiable_within_at_inverse hx _ lemma fderiv_inverse (x : Rˣ) : fderiv 𝕜 (@ring.inverse R _) x = - mul_left_right 𝕜 R ↑x⁻¹ ↑x⁻¹ := (has_fderiv_at_ring_inverse x).fderiv +variables {h : E → R} {z : E} {S : set E} + +lemma differentiable_within_at.inverse (hf : differentiable_within_at 𝕜 h S z) + (hz : is_unit (h z)) : + differentiable_within_at 𝕜 (λ x, ring.inverse (h x)) S z := +(differentiable_at_inverse hz).comp_differentiable_within_at z hf + +@[simp] lemma differentiable_at.inverse (hf : differentiable_at 𝕜 h z) (hz : is_unit (h z)) : + differentiable_at 𝕜 (λ x, ring.inverse (h x)) z := +(differentiable_at_inverse hz).comp z hf + +lemma differentiable_on.inverse (hf : differentiable_on 𝕜 h S) (hz : ∀ x ∈ S, is_unit (h x)) : + differentiable_on 𝕜 (λ x, ring.inverse (h x)) S := +λ x h, (hf x h).inverse (hz x h) + +@[simp] lemma differentiable.inverse (hf : differentiable 𝕜 h) (hz : ∀ x, is_unit (h x)) : + differentiable 𝕜 (λ x, ring.inverse (h x)) := +λ x, (hf x).inverse (hz x) + end algebra_inverse +/-! ### Derivative of the inverse in a division ring + +Note these lemmas are primed as they need `complete_space R`, whereas the other lemmas in +`deriv/inv.lean` do not, but instead need `nontrivially_normed_field R`. +-/ + +section division_ring_inverse +variables {R : Type*} [normed_division_ring R] [normed_algebra 𝕜 R] [complete_space R] +open normed_ring continuous_linear_map ring + +/-- At an invertible element `x` of a normed division algebra `R`, the Fréchet derivative of the +inversion operation is the linear map `λ t, - x⁻¹ * t * x⁻¹`. -/ +lemma has_fderiv_at_inv' {x : R} (hx : x ≠ 0) : + has_fderiv_at has_inv.inv (-mul_left_right 𝕜 R x⁻¹ x⁻¹) x := +by simpa using has_fderiv_at_ring_inverse (units.mk0 _ hx) + +lemma differentiable_at_inv' {x : R} (hx : x ≠ 0) : differentiable_at 𝕜 has_inv.inv x := +(has_fderiv_at_inv' hx).differentiable_at + +lemma differentiable_within_at_inv' {x : R} (hx : x ≠ 0) (s : set R): + differentiable_within_at 𝕜 (λx, x⁻¹) s x := +(differentiable_at_inv' hx).differentiable_within_at + +lemma differentiable_on_inv' : differentiable_on 𝕜 (λ x : R, x⁻¹) {x | x ≠ 0} := +λ x hx, differentiable_within_at_inv' hx _ + +/-- Non-commutative version of `fderiv_inv` -/ +lemma fderiv_inv' {x : R} (hx : x ≠ 0) : + fderiv 𝕜 has_inv.inv x = - mul_left_right 𝕜 R x⁻¹ x⁻¹ := +(has_fderiv_at_inv' hx).fderiv + +/-- Non-commutative version of `fderiv_within_inv` -/ +lemma fderiv_within_inv' {s : set R} {x : R} (hx : x ≠ 0) (hxs : unique_diff_within_at 𝕜 s x) : + fderiv_within 𝕜 (λ x, x⁻¹) s x = - mul_left_right 𝕜 R x⁻¹ x⁻¹ := +begin + rw differentiable_at.fderiv_within (differentiable_at_inv' hx) hxs, + exact fderiv_inv' hx +end + +variables {h : E → R} {z : E} {S : set E} + +lemma differentiable_within_at.inv' (hf : differentiable_within_at 𝕜 h S z) (hz : h z ≠ 0) : + differentiable_within_at 𝕜 (λ x, (h x)⁻¹) S z := +(differentiable_at_inv' hz).comp_differentiable_within_at z hf + +@[simp] lemma differentiable_at.inv' (hf : differentiable_at 𝕜 h z) (hz : h z ≠ 0) : + differentiable_at 𝕜 (λ x, (h x)⁻¹) z := +(differentiable_at_inv' hz).comp z hf + +lemma differentiable_on.inv' (hf : differentiable_on 𝕜 h S) (hz : ∀ x ∈ S, h x ≠ 0) : + differentiable_on 𝕜 (λ x, (h x)⁻¹) S := +λ x h, (hf x h).inv' (hz x h) + +@[simp] lemma differentiable.inv' (hf : differentiable 𝕜 h) (hz : ∀ x, h x ≠ 0) : + differentiable 𝕜 (λ x, (h x)⁻¹) := +λ x, (hf x).inv' (hz x) + +end division_ring_inverse + end diff --git a/src/analysis/normed_space/spectrum.lean b/src/analysis/normed_space/spectrum.lean index 09dbf5e2d5006..d222373613d9b 100644 --- a/src/analysis/normed_space/spectrum.lean +++ b/src/analysis/normed_space/spectrum.lean @@ -314,7 +314,7 @@ begin simpa only [norm_to_nnreal, real.to_nnreal_coe] using real.to_nnreal_mono (mem_closed_ball_zero_iff.mp z_mem) }, have H₁ : differentiable 𝕜 (λ w : 𝕜, 1 - w • a) := (differentiable_id.smul_const a).const_sub 1, - exact differentiable_at.comp z (differentiable_at_inverse hu.unit) (H₁.differentiable_at), + exact differentiable_at.comp z (differentiable_at_inverse hu) (H₁.differentiable_at), end end one_sub_smul From d07245fd37786daa997af4f1a73a49fa3b748408 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 14 Jul 2023 16:30:09 +0000 Subject: [PATCH 20/61] =?UTF-8?q?feat(group=5Ftheory/order=5Fof=5Felement)?= =?UTF-8?q?:=20Order=20in=20`=CE=B1=20=C3=97=20=CE=B2`=20(#18719)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The order of `(a, b)` is the lcm of the orders of `a` and `b`. Match `pow` and `zpow` lemmas. Also some `variables` noise because I could not use `x` to mean what I wanted, and incidentally the type `A` was mostly unused. --- src/data/prod/basic.lean | 2 +- src/dynamics/periodic_pts.lean | 25 +++++++ src/group_theory/order_of_element.lean | 90 ++++++++++++++++++++------ 3 files changed, 95 insertions(+), 22 deletions(-) diff --git a/src/data/prod/basic.lean b/src/data/prod/basic.lean index a2a159aced62e..eb8f6034c7058 100644 --- a/src/data/prod/basic.lean +++ b/src/data/prod/basic.lean @@ -97,7 +97,7 @@ funext (λ p, ext (map_fst f g p) (map_snd f g p)) lemma id_prod : (λ (p : α × β), (p.1, p.2)) = id := funext $ λ ⟨a, b⟩, rfl -lemma map_id : (prod.map (@id α) (@id β)) = id := +@[simp] lemma map_id : (prod.map (@id α) (@id β)) = id := id_prod lemma fst_surjective [h : nonempty β] : function.surjective (@fst α β) := diff --git a/src/dynamics/periodic_pts.lean b/src/dynamics/periodic_pts.lean index 30e6211e65126..353804e7661ea 100644 --- a/src/dynamics/periodic_pts.lean +++ b/src/dynamics/periodic_pts.lean @@ -519,6 +519,31 @@ end end function +namespace function +variables {α β : Type*} {f : α → α} {g : β → β} {x : α × β} {a : α} {b : β} {m n : ℕ} + +@[simp] lemma iterate_prod_map (f : α → α) (g : β → β) (n : ℕ) : + (prod.map f g)^[n] = prod.map (f^[n]) (g^[n]) := by induction n; simp [*, prod.map_comp_map] + +@[simp] lemma is_fixed_pt_prod_map (x : α × β) : + is_fixed_pt (prod.map f g) x ↔ is_fixed_pt f x.1 ∧ is_fixed_pt g x.2 := prod.ext_iff + +@[simp] lemma is_periodic_pt_prod_map (x : α × β) : + is_periodic_pt (prod.map f g) n x ↔ is_periodic_pt f n x.1 ∧ is_periodic_pt g n x.2 := +by simp [is_periodic_pt] + +lemma minimal_period_prod_map (f : α → α) (g : β → β) (x : α × β) : + minimal_period (prod.map f g) x = (minimal_period f x.1).lcm (minimal_period g x.2) := +eq_of_forall_dvd $ by cases x; simp [←is_periodic_pt_iff_minimal_period_dvd, nat.lcm_dvd_iff] + +lemma minimal_period_fst_dvd : minimal_period f x.1 ∣ minimal_period (prod.map f g) x := +by { rw minimal_period_prod_map, exact nat.dvd_lcm_left _ _ } + +lemma minimal_period_snd_dvd : minimal_period g x.2 ∣ minimal_period (prod.map f g) x := +by { rw minimal_period_prod_map, exact nat.dvd_lcm_right _ _ } + +end function + namespace mul_action open function diff --git a/src/group_theory/order_of_element.lean b/src/group_theory/order_of_element.lean index 44eb664b3460a..2f1065af425db 100644 --- a/src/group_theory/order_of_element.lean +++ b/src/group_theory/order_of_element.lean @@ -3,8 +3,9 @@ Copyright (c) 2018 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Julian Kuelshammer -/ +import algebra.gcd_monoid.finset import algebra.hom.iterate -import data.nat.modeq +import data.int.modeq import data.set.pointwise.basic import data.set.intervals.infinite import dynamics.periodic_pts @@ -35,14 +36,11 @@ order of an element open function nat open_locale pointwise -universes u v - -variables {G : Type u} {A : Type v} -variables {x y : G} {a b : A} {n m : ℕ} +variables {G H A α β : Type*} section monoid_add_monoid -variables [monoid G] [add_monoid A] +variables [monoid G] [add_monoid A] {x y : G} {a b : A} {n m : ℕ} section is_of_fin_order @@ -95,8 +93,7 @@ by { rw [is_of_fin_order_iff_pow_eq_one, is_of_fin_order_iff_pow_eq_one], norm_c /-- The image of an element of finite order has finite order. -/ @[to_additive add_monoid_hom.is_of_fin_order "The image of an element of finite additive order has finite additive order."] -lemma monoid_hom.is_of_fin_order - {H : Type v} [monoid H] (f : G →* H) {x : G} (h : is_of_fin_order x) : +lemma monoid_hom.is_of_fin_order [monoid H] (f : G →* H) {x : G} (h : is_of_fin_order x) : is_of_fin_order $ f x := (is_of_fin_order_iff_pow_eq_one _).mpr $ begin rcases (is_of_fin_order_iff_pow_eq_one _).mp h with ⟨n, npos, hn⟩, @@ -171,6 +168,11 @@ end lemma order_of_pos_iff : 0 < order_of x ↔ is_of_fin_order x := by rwa [iff_not_comm.mp order_of_eq_zero_iff, pos_iff_ne_zero] +@[to_additive is_of_fin_add_order.mono] +lemma is_of_fin_order.mono [monoid β] {y : β} (hx : is_of_fin_order x) + (h : order_of y ∣ order_of x) : is_of_fin_order y := +by { rw ←order_of_pos_iff at ⊢ hx, exact nat.pos_of_dvd_of_pos h hx } + @[to_additive nsmul_ne_zero_of_lt_add_order_of'] lemma pow_ne_one_of_lt_order_of' (n0 : n ≠ 0) (h : n < order_of x) : x ^ n ≠ 1 := λ j, not_is_periodic_pt_of_pos_of_lt_minimal_period n0 h @@ -311,11 +313,11 @@ end lemma order_of_dvd_lcm_mul : order_of y ∣ nat.lcm (order_of x) (order_of (x * y)) := begin by_cases h0 : order_of x = 0, - { rw [h0, lcm_zero_left], apply dvd_zero }, + { rw [h0, nat.lcm_zero_left], apply dvd_zero }, conv_lhs { rw [← one_mul y, ← pow_order_of_eq_one x, ← succ_pred_eq_of_pos (nat.pos_of_ne_zero h0), pow_succ', mul_assoc] }, exact (((commute.refl x).mul_right h).pow_left _).order_of_mul_dvd_lcm.trans - (lcm_dvd_iff.2 ⟨trans (order_of_pow_dvd _) (dvd_lcm_left _ _), dvd_lcm_right _ _⟩), + (nat.lcm_dvd_iff.2 ⟨trans (order_of_pow_dvd _) (dvd_lcm_left _ _), dvd_lcm_right _ _⟩), end @[to_additive add_order_of_add_dvd_mul_add_order_of] @@ -394,7 +396,7 @@ end p_prime end monoid_add_monoid section cancel_monoid -variables [left_cancel_monoid G] (x y) +variables [left_cancel_monoid G] (x y : G) {m n : ℕ} @[to_additive nsmul_injective_of_lt_add_order_of] lemma pow_injective_of_lt_order_of @@ -451,7 +453,7 @@ end end cancel_monoid section group -variables [group G] [add_group A] {x a} {i : ℤ} +variables [group G] {x y : G} {i : ℤ} /-- Inverses of elements of finite order have finite order. -/ @[to_additive "Inverses of elements of finite additive order have finite additive order."] @@ -560,7 +562,7 @@ end group section comm_monoid -variables [comm_monoid G] +variables [comm_monoid G] {x y : G} /-- Elements of finite order are closed under multiplication. -/ @[to_additive "Elements of finite additive order are closed under addition."] @@ -571,7 +573,7 @@ lemma is_of_fin_order.mul (hx : is_of_fin_order x) (hy : is_of_fin_order y) : end comm_monoid section finite_monoid -variables [monoid G] +variables [monoid G] {n : ℕ} open_locale big_operators @[to_additive sum_card_add_order_of_eq_card_nsmul_eq_zero] @@ -592,7 +594,7 @@ end finite_monoid section finite_cancel_monoid -- TODO: Of course everything also works for right_cancel_monoids. -variables [left_cancel_monoid G] [add_left_cancel_monoid A] +variables [left_cancel_monoid G] {x y : G} {n : ℕ} -- TODO: Use this to show that a finite left cancellative monoid is a group. @[to_additive] @@ -682,7 +684,7 @@ lemma order_eq_card_powers [fintype G] : order_of x = fintype.card (submonoid.po end finite_cancel_monoid section finite_group -variables [group G] [add_group A] +variables [group G] {x y : G} {n : ℕ} @[to_additive] lemma exists_zpow_eq_one [finite G] (x : G) : ∃ (i : ℤ) (H : i ≠ 0), x ^ (i : ℤ) = 1 := @@ -712,6 +714,23 @@ lemma mem_zpowers_iff_mem_range_order_of [finite G] [decidable_eq G] : y ∈ subgroup.zpowers x ↔ y ∈ (finset.range (order_of x)).image ((^) x : ℕ → G) := by rw [← mem_powers_iff_mem_zpowers, mem_powers_iff_mem_range_order_of] +@[to_additive] lemma zpow_eq_one_iff_modeq {n : ℤ} : x ^ n = 1 ↔ n ≡ 0 [ZMOD (order_of x)] := +by rw [int.modeq_zero_iff_dvd, order_of_dvd_iff_zpow_eq_one] + +@[to_additive] lemma zpow_eq_zpow_iff_modeq {m n : ℤ} : x ^ m = x ^ n ↔ m ≡ n [ZMOD (order_of x)] := +by rw [←mul_inv_eq_one, ←zpow_sub, zpow_eq_one_iff_modeq, int.modeq_iff_dvd, int.modeq_iff_dvd, + zero_sub, neg_sub] + +@[simp, to_additive] lemma injective_zpow_iff_not_is_of_fin_order : + injective (λ n : ℤ, x ^ n) ↔ ¬ is_of_fin_order x := +begin + refine ⟨_, λ h n m hnm, _⟩, + { simp_rw is_of_fin_order_iff_pow_eq_one, + rintro h ⟨n, hn, hx⟩, + exact nat.cast_ne_zero.2 hn.ne' (h $ by simpa using hx) }, + rwa [zpow_eq_zpow_iff_modeq, order_of_eq_zero_iff.2 h, nat.cast_zero, int.modeq_zero_iff] at hnm, +end + @[to_additive decidable_zmultiples] noncomputable instance decidable_zpowers : decidable_pred (∈ subgroup.zpowers x) := classical.dec_pred _ @@ -755,8 +774,8 @@ end variables [fintype G] -/-- See also `order_eq_card_zpowers'`. -/ -@[to_additive add_order_eq_card_zmultiples "See also `add_order_eq_card_zmultiples'`."] +/-- See also `nat.card_zpowers'`. -/ +@[to_additive add_order_eq_card_zmultiples "See also `nat.card_zmultiples`."] lemma order_eq_card_zpowers : order_of x = fintype.card (zpowers x) := (fintype.card_fin (order_of x)).symm.trans (fintype.card_eq.2 ⟨fin_equiv_zpowers x⟩) @@ -846,8 +865,6 @@ begin (congr_arg (∣ fintype.card K) (order_of_subgroup ⟨x, hx.2⟩)).mpr order_of_dvd_card_univ⟩, end -variable (a) - /-- TODO: Generalise to `submonoid.powers`.-/ @[to_additive image_range_add_order_of, nolint to_additive_doc] lemma image_range_order_of [decidable_eq G] : @@ -909,7 +926,7 @@ end pow_is_subgroup section linear_ordered_ring -variable [linear_ordered_ring G] +variables [linear_ordered_ring G] {x : G} lemma order_of_abs_ne_one (h : |x| ≠ 1) : order_of x = 0 := begin @@ -931,3 +948,34 @@ begin end end linear_ordered_ring + +section prod +variables [monoid α] [monoid β] {x : α × β} {a : α} {b : β} + +@[to_additive prod.add_order_of] protected lemma prod.order_of (x : α × β) : + order_of x = (order_of x.1).lcm (order_of x.2) := +minimal_period_prod_map _ _ _ + +@[to_additive add_order_of_fst_dvd_add_order_of] lemma order_of_fst_dvd_order_of : + order_of x.1 ∣ order_of x := +minimal_period_fst_dvd + +@[to_additive add_order_of_snd_dvd_add_order_of] lemma order_of_snd_dvd_order_of : + order_of x.2 ∣ order_of x := +minimal_period_snd_dvd + +@[to_additive is_of_fin_add_order.fst] +lemma is_of_fin_order.fst {x : α × β} (hx : is_of_fin_order x) : is_of_fin_order x.1 := +hx.mono order_of_fst_dvd_order_of + +@[to_additive is_of_fin_add_order.snd] +lemma is_of_fin_order.snd {x : α × β} (hx : is_of_fin_order x) : is_of_fin_order x.2 := +hx.mono order_of_snd_dvd_order_of + +@[to_additive is_of_fin_add_order.prod_mk] +lemma is_of_fin_order.prod_mk : is_of_fin_order a → is_of_fin_order b → is_of_fin_order (a, b) := +by simpa only [←order_of_pos_iff, prod.order_of] using nat.lcm_pos + +end prod + +-- TODO: Corresponding `pi` lemmas. We cannot currently state them here because of import cycles From 2c84c2c5496117349007d97104e7bbb471381592 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 14 Jul 2023 19:58:48 +0000 Subject: [PATCH 21/61] feat(order/well_founded_set): `prod.lex` is well-founded (#18665) Co-authored-by: Junyan Xu Co-authored-by: Eric Wieser --- src/order/well_founded.lean | 27 +++++----- src/order/well_founded_set.lean | 96 ++++++++++++++++++++++++++++++++- 2 files changed, 108 insertions(+), 15 deletions(-) diff --git a/src/order/well_founded.lean b/src/order/well_founded.lean index fb82de79a2a55..40aaca54853a1 100644 --- a/src/order/well_founded.lean +++ b/src/order/well_founded.lean @@ -21,21 +21,24 @@ and provide a few new definitions: `well_founded.min`, `well_founded.sup`, and ` and an induction principle `well_founded.induction_bot`. -/ -variables {α : Type*} +variables {α β γ : Type*} namespace well_founded +variables {r r' : α → α → Prop} -protected theorem is_asymm {α : Sort*} {r : α → α → Prop} (h : well_founded r) : is_asymm α r := -⟨h.asymmetric⟩ +protected theorem is_asymm (h : well_founded r) : is_asymm α r := ⟨h.asymmetric⟩ -instance {α : Sort*} [has_well_founded α] : is_asymm α has_well_founded.r := -has_well_founded.wf.is_asymm - -protected theorem is_irrefl {α : Sort*} {r : α → α → Prop} (h : well_founded r) : is_irrefl α r := +protected theorem is_irrefl (h : well_founded r) : is_irrefl α r := (@is_asymm.is_irrefl α r h.is_asymm) -instance {α : Sort*} [has_well_founded α] : is_irrefl α has_well_founded.r := -is_asymm.is_irrefl +instance [has_well_founded α] : is_asymm α has_well_founded.r := has_well_founded.wf.is_asymm +instance [has_well_founded α] : is_irrefl α has_well_founded.r := is_asymm.is_irrefl + +lemma mono (hr : well_founded r) (h : ∀ a b, r' a b → r a b) : well_founded r' := +subrelation.wf h hr + +lemma on_fun {α β : Sort*} {r : β → β → Prop} {f : α → β} : + well_founded r → well_founded (r on f) := inv_image.wf _ /-- If `r` is a well-founded relation, then any nonempty set has a minimal element with respect to `r`. -/ @@ -110,8 +113,7 @@ end section linear_order -variables {β : Type*} [linear_order β] (h : well_founded ((<) : β → β → Prop)) - {γ : Type*} [partial_order γ] +variables [linear_order β] (h : well_founded ((<) : β → β → Prop)) [partial_order γ] theorem min_le {x : β} {s : set β} (hx : x ∈ s) (hne : s.nonempty := ⟨x, hx⟩) : h.min s hne ≤ x := @@ -149,8 +151,7 @@ end linear_order end well_founded namespace function - -variables {β : Type*} (f : α → β) +variables (f : α → β) section has_lt diff --git a/src/order/well_founded_set.lean b/src/order/well_founded_set.lean index 37a675eb498cb..d3c122a549c21 100644 --- a/src/order/well_founded_set.lean +++ b/src/order/well_founded_set.lean @@ -3,6 +3,7 @@ Copyright (c) 2021 Aaron Anderson. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Aaron Anderson -/ +import data.sigma.lex import order.antichain import order.order_iso_nat import order.well_founded @@ -45,7 +46,7 @@ Prove that `s` is partial well ordered iff it has no infinite descending chain o * [Nash-Williams, *On Well-Quasi-Ordering Finite Trees*][Nash-Williams63] -/ -variables {ι α β : Type*} +variables {ι α β γ : Type*} {π : ι → Type*} namespace set @@ -61,7 +62,7 @@ section well_founded_on variables {r r' : α → α → Prop} section any_rel -variables {s t : set α} {x y : α} +variables {f : β → α} {s t : set α} {x y : α} lemma well_founded_on_iff : s.well_founded_on r ↔ well_founded (λ (a b : α), r a b ∧ a ∈ s ∧ b ∈ s) := @@ -79,6 +80,26 @@ begin exact ⟨m, mt, λ x xt ⟨xm, xs, ms⟩, hst ⟨m, ⟨ms, mt⟩⟩⟩ } end +@[simp] lemma well_founded_on_univ : (univ : set α).well_founded_on r ↔ well_founded r := +by simp [well_founded_on_iff] + +lemma _root_.well_founded.well_founded_on : well_founded r → s.well_founded_on r := inv_image.wf _ + +@[simp] lemma well_founded_on_range : (range f).well_founded_on r ↔ well_founded (r on f) := +begin + let f' : β → range f := λ c, ⟨f c, c, rfl⟩, + refine ⟨λ h, (inv_image.wf f' h).mono $ λ c c', id, λ h, ⟨_⟩⟩, + rintro ⟨_, c, rfl⟩, + refine acc.of_downward_closed f' _ _ (_), + { rintro _ ⟨_, c', rfl⟩ -, + exact ⟨c', rfl⟩ }, + { exact h.apply _ } +end + +@[simp] lemma well_founded_on_image {s : set β} : + (f '' s).well_founded_on r ↔ s.well_founded_on (r on f) := +by { rw image_eq_range, exact well_founded_on_range } + namespace well_founded_on protected lemma induction (hs : s.well_founded_on r) (hx : x ∈ s) {P : α → Prop} @@ -98,6 +119,9 @@ begin exact ⟨hle _ _ xy.1, hst xy.2.1, hst xy.2.2⟩ end +lemma mono' (h : ∀ a b ∈ s, r' a b → r a b) : s.well_founded_on r → s.well_founded_on r' := +subrelation.wf $ λ a b, h _ a.2 _ b.2 + lemma subset (h : t.well_founded_on r) (hst : s ⊆ t) : s.well_founded_on r := h.mono le_rfl hst open relation @@ -707,3 +731,71 @@ begin refine ⟨g'.trans g, λ a b hab, (finset.forall_mem_cons _ _).2 _⟩, exact ⟨hg (order_hom_class.mono g' hab), hg' hab⟩ } end + +section prod_lex +variables {rα : α → α → Prop} {rβ : β → β → Prop} {f : γ → α} {g : γ → β} {s : set γ} + +/-- Stronger version of `prod.lex_wf`. Instead of requiring `rβ on g` to be well-founded, we only +require it to be well-founded on fibers of `f`.-/ +lemma well_founded.prod_lex_of_well_founded_on_fiber (hα : well_founded (rα on f)) + (hβ : ∀ a, (f ⁻¹' {a}).well_founded_on (rβ on g)) : + well_founded (prod.lex rα rβ on λ c, (f c, g c)) := +begin + refine ((psigma.lex_wf (well_founded_on_range.2 hα) $ λ a, hβ a).on_fun).mono (λ c c' h, _), + exact λ c, ⟨⟨_, c, rfl⟩, c, rfl⟩, + obtain h' | h' := prod.lex_iff.1 h, + { exact psigma.lex.left _ _ h' }, + { dsimp only [inv_image, (on)] at h' ⊢, + convert psigma.lex.right (⟨_, c', rfl⟩ : range f) _ using 1, swap, + exacts [⟨c, h'.1⟩, psigma.subtype_ext (subtype.ext h'.1) rfl, h'.2] } +end + +lemma set.well_founded_on.prod_lex_of_well_founded_on_fiber (hα : s.well_founded_on (rα on f)) + (hβ : ∀ a, (s ∩ f ⁻¹' {a}).well_founded_on (rβ on g)) : + s.well_founded_on (prod.lex rα rβ on λ c, (f c, g c)) := +begin + refine well_founded.prod_lex_of_well_founded_on_fiber hα + (λ a, subrelation.wf (λ b c h, _) (hβ a).on_fun), + exact λ x, ⟨x, x.1.2, x.2⟩, + assumption, +end + +end prod_lex + +section sigma_lex +variables {rι : ι → ι → Prop} {rπ : Π i, π i → π i → Prop} {f : γ → ι} {g : Π i, γ → π i} + {s : set γ} + +/-- Stronger version of `psigma.lex_wf`. Instead of requiring `rπ on g` to be well-founded, we only +require it to be well-founded on fibers of `f`.-/ +lemma well_founded.sigma_lex_of_well_founded_on_fiber (hι : well_founded (rι on f)) + (hπ : ∀ i, (f ⁻¹' {i}).well_founded_on (rπ i on g i)) : + well_founded (sigma.lex rι rπ on λ c, ⟨f c, g (f c) c⟩) := +begin + refine ((psigma.lex_wf (well_founded_on_range.2 hι) $ λ a, hπ a).on_fun).mono (λ c c' h, _), + exact λ c, ⟨⟨_, c, rfl⟩, c, rfl⟩, + obtain h' | ⟨h', h''⟩ := sigma.lex_iff.1 h, + { exact psigma.lex.left _ _ h' }, + { dsimp only [inv_image, (on)] at h' ⊢, + convert psigma.lex.right (⟨_, c', rfl⟩ : range f) _ using 1, swap, + { exact ⟨c, h'⟩ }, + { exact psigma.subtype_ext (subtype.ext h') rfl }, + { dsimp only [subtype.coe_mk] at *, + revert h', + generalize : f c = d, + rintro rfl _, + exact h'' } } +end + +lemma set.well_founded_on.sigma_lex_of_well_founded_on_fiber (hι : s.well_founded_on (rι on f)) + (hπ : ∀ i, (s ∩ f ⁻¹' {i}).well_founded_on (rπ i on g i)) : + s.well_founded_on (sigma.lex rι rπ on λ c, ⟨f c, g (f c) c⟩) := +begin + show well_founded (sigma.lex rι rπ on λ (c : s), ⟨f c, g (f c) c⟩), + refine @well_founded.sigma_lex_of_well_founded_on_fiber _ s _ _ rπ (λ c, f c) (λ i c, g _ c) hι + (λ i, subrelation.wf (λ b c h, _) (hπ i).on_fun), + exact λ x, ⟨x, x.1.2, x.2⟩, + assumption, +end + +end sigma_lex From 1d29de43a5ba4662dd33b5cfeecfc2a27a5a8a29 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 14 Jul 2023 19:58:49 +0000 Subject: [PATCH 22/61] feat(data/*/interval): `finset.uIcc` on concrete structures (#18838) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Calculate the size of `finset.uIcc` in `ℕ`, `ℤ`, `fin`, `prod`, `pi`, `multiset`, `finset`... Co-authored-by: Eric Wieser --- src/data/dfinsupp/interval.lean | 13 ++++++-- src/data/dfinsupp/multiset.lean | 52 ++++++++++++++++++++++++++--- src/data/dfinsupp/order.lean | 29 ++++++++++------ src/data/fin/interval.lean | 20 +++++++++++ src/data/finset/locally_finite.lean | 17 +++++++--- src/data/finsupp/interval.lean | 9 +++++ src/data/finsupp/order.lean | 14 ++++++-- src/data/int/interval.lean | 23 +++++++------ src/data/multiset/interval.lean | 34 ++++++++++++------- src/data/nat/interval.lean | 14 ++++++++ src/data/pi/interval.lean | 43 +++++++++++++----------- src/data/pnat/interval.lean | 20 ++++++++--- src/order/locally_finite.lean | 38 ++++++++++----------- 13 files changed, 234 insertions(+), 92 deletions(-) diff --git a/src/data/dfinsupp/interval.lean b/src/data/dfinsupp/interval.lean index e010360860147..990341f78ec32 100644 --- a/src/data/dfinsupp/interval.lean +++ b/src/data/dfinsupp/interval.lean @@ -140,7 +140,7 @@ end end pi -section locally_finite +section partial_order variables [decidable_eq ι] [Π i, decidable_eq (α i)] variables [Π i, partial_order (α i)] [Π i, has_zero (α i)] [Π i, locally_finite_order (α i)] @@ -169,7 +169,16 @@ by rw [card_Ioc_eq_card_Icc_sub_one, card_Icc] lemma card_Ioo : (Ioo f g).card = ∏ i in f.support ∪ g.support, (Icc (f i) (g i)).card - 2 := by rw [card_Ioo_eq_card_Icc_sub_two, card_Icc] -end locally_finite +end partial_order + +section lattice +variables [decidable_eq ι] [Π i, decidable_eq (α i)] [Π i, lattice (α i)] [Π i, has_zero (α i)] + [Π i, locally_finite_order (α i)] (f g : Π₀ i, α i) + +lemma card_uIcc : (uIcc f g).card = ∏ i in f.support ∪ g.support, (uIcc (f i) (g i)).card := +by { rw ←support_inf_union_support_sup, exact card_Icc _ _ } + +end lattice section canonically_ordered variables [decidable_eq ι] [Π i, decidable_eq (α i)] diff --git a/src/data/dfinsupp/multiset.lean b/src/data/dfinsupp/multiset.lean index de23e647e818c..de2ed3c8f1031 100644 --- a/src/data/dfinsupp/multiset.lean +++ b/src/data/dfinsupp/multiset.lean @@ -17,6 +17,8 @@ with `multiset.to_dfinsupp` the reverse equivalence. Note that this provides a computable alternative to `finsupp.to_multiset`. -/ +open function + variables {α : Type*} {β : α → Type*} namespace dfinsupp @@ -38,7 +40,7 @@ dfinsupp.sum_add_hom_single _ _ _ end dfinsupp namespace multiset -variables [decidable_eq α] +variables [decidable_eq α] {s t : multiset α} /-- A computable version of `multiset.to_finsupp` -/ def to_dfinsupp : multiset α →+ Π₀ a : α, ℕ := @@ -78,12 +80,52 @@ add_monoid_hom.to_add_equiv @[simp] lemma to_dfinsupp_to_multiset (s : multiset α) : s.to_dfinsupp.to_multiset = s := equiv_dfinsupp.symm_apply_apply s -@[simp] lemma to_dfinsupp_le_to_dfinsupp (s t : multiset α) : - to_dfinsupp s ≤ to_dfinsupp t ↔ s ≤ t := +lemma to_dfinsupp_injective : injective (to_dfinsupp : multiset α → Π₀ a, ℕ) := +equiv_dfinsupp.injective + +@[simp] lemma to_dfinsupp_inj : to_dfinsupp s = to_dfinsupp t ↔ s = t := +to_dfinsupp_injective.eq_iff + +@[simp] lemma to_dfinsupp_le_to_dfinsupp : to_dfinsupp s ≤ to_dfinsupp t ↔ s ≤ t := by simp [multiset.le_iff_count, dfinsupp.le_def] +@[simp] lemma to_dfinsupp_lt_to_dfinsupp : to_dfinsupp s < to_dfinsupp t ↔ s < t := +lt_iff_lt_of_le_iff_le' to_dfinsupp_le_to_dfinsupp to_dfinsupp_le_to_dfinsupp + +@[simp] lemma to_dfinsupp_inter (s t : multiset α) : + to_dfinsupp (s ∩ t) = s.to_dfinsupp ⊓ t.to_dfinsupp := +by { ext i, simp [inf_eq_min] } + +@[simp] lemma to_dfinsupp_union (s t : multiset α) : + to_dfinsupp (s ∪ t) = s.to_dfinsupp ⊔ t.to_dfinsupp := +by { ext i, simp [sup_eq_max] } + end multiset -@[simp] lemma dfinsupp.to_multiset_to_dfinsupp [decidable_eq α] (f : Π₀ a : α, ℕ) : - f.to_multiset.to_dfinsupp = f := +namespace dfinsupp +variables [decidable_eq α] {f g : Π₀ a : α, ℕ} + +@[simp] lemma to_multiset_to_dfinsupp : f.to_multiset.to_dfinsupp = f := multiset.equiv_dfinsupp.apply_symm_apply f + +lemma to_multiset_injective : injective (to_multiset : (Π₀ a, ℕ) → multiset α) := +multiset.equiv_dfinsupp.symm.injective + +@[simp] lemma to_multiset_inj : to_multiset f = to_multiset g ↔ f = g := +to_multiset_injective.eq_iff + +@[simp] lemma to_multiset_le_to_multiset : to_multiset f ≤ to_multiset g ↔ f ≤ g := +by simp_rw [←multiset.to_dfinsupp_le_to_dfinsupp, to_multiset_to_dfinsupp] + +@[simp] lemma to_multiset_lt_to_multiset : to_multiset f < to_multiset g ↔ f < g := +by simp_rw [←multiset.to_dfinsupp_lt_to_dfinsupp, to_multiset_to_dfinsupp] + +variables (f g) + +@[simp] lemma to_multiset_inf : to_multiset (f ⊓ g) = f.to_multiset ∩ g.to_multiset := +multiset.to_dfinsupp_injective $ by simp + +@[simp] lemma to_multiset_sup : to_multiset (f ⊔ g) = f.to_multiset ∪ g.to_multiset := +multiset.to_dfinsupp_injective $ by simp + +end dfinsupp diff --git a/src/data/dfinsupp/order.lean b/src/data/dfinsupp/order.lean index 0c6344b670dda..c5a90bc29224c 100644 --- a/src/data/dfinsupp/order.lean +++ b/src/data/dfinsupp/order.lean @@ -31,15 +31,13 @@ namespace dfinsupp /-! ### Order structures -/ section has_zero -variables (α) [Π i, has_zero (α i)] +variables [Π i, has_zero (α i)] section has_le variables [Π i, has_le (α i)] instance : has_le (Π₀ i, α i) := ⟨λ f g, ∀ i, f i ≤ g i⟩ -variables {α} - lemma le_def {f g : Π₀ i, α i} : f ≤ g ↔ ∀ i, f i ≤ g i := iff.rfl /-- The order on `dfinsupp`s over a partial order embeds into the order on functions -/ @@ -59,7 +57,7 @@ variables [Π i, preorder (α i)] instance : preorder (Π₀ i, α i) := { le_refl := λ f i, le_rfl, le_trans := λ f g h hfg hgh i, (hfg i).trans (hgh i), - .. dfinsupp.has_le α } + .. dfinsupp.has_le } lemma coe_fn_mono : monotone (coe_fn : (Π₀ i, α i) → Π i, α i) := λ f g, le_def.1 @@ -67,14 +65,14 @@ end preorder instance [Π i, partial_order (α i)] : partial_order (Π₀ i, α i) := { le_antisymm := λ f g hfg hgf, ext $ λ i, (hfg i).antisymm (hgf i), - .. dfinsupp.preorder α} + .. dfinsupp.preorder } instance [Π i, semilattice_inf (α i)] : semilattice_inf (Π₀ i, α i) := { inf := zip_with (λ _, (⊓)) (λ _, inf_idem), inf_le_left := λ f g i, by { rw zip_with_apply, exact inf_le_left }, inf_le_right := λ f g i, by { rw zip_with_apply, exact inf_le_right }, le_inf := λ f g h hf hg i, by { rw zip_with_apply, exact le_inf (hf i) (hg i) }, - ..dfinsupp.partial_order α } + ..dfinsupp.partial_order } @[simp] lemma inf_apply [Π i, semilattice_inf (α i)] (f g : Π₀ i, α i) (i : ι) : (f ⊓ g) i = f i ⊓ g i := @@ -85,15 +83,26 @@ instance [Π i, semilattice_sup (α i)] : semilattice_sup (Π₀ i, α i) := le_sup_left := λ f g i, by { rw zip_with_apply, exact le_sup_left }, le_sup_right := λ f g i, by { rw zip_with_apply, exact le_sup_right }, sup_le := λ f g h hf hg i, by { rw zip_with_apply, exact sup_le (hf i) (hg i) }, - ..dfinsupp.partial_order α } + ..dfinsupp.partial_order } @[simp] lemma sup_apply [Π i, semilattice_sup (α i)] (f g : Π₀ i, α i) (i : ι) : (f ⊔ g) i = f i ⊔ g i := zip_with_apply _ _ _ _ _ -instance lattice [Π i, lattice (α i)] : lattice (Π₀ i, α i) := -{ .. dfinsupp.semilattice_inf α, .. dfinsupp.semilattice_sup α } +section lattice +variables [Π i, lattice (α i)] (f g : Π₀ i, α i) + +instance lattice : lattice (Π₀ i, α i) := { ..dfinsupp.semilattice_inf, ..dfinsupp.semilattice_sup } + +variables [decidable_eq ι] [Π i (x : α i), decidable (x ≠ 0)] + +lemma support_inf_union_support_sup : (f ⊓ g).support ∪ (f ⊔ g).support = f.support ∪ g.support := +coe_injective $ compl_injective $ by { ext, simp [inf_eq_and_sup_eq_iff] } + +lemma support_sup_union_support_inf : (f ⊔ g).support ∪ (f ⊓ g).support = f.support ∪ g.support := +(union_comm _ _).trans $ support_inf_union_support_sup _ _ +end lattice end has_zero /-! ### Algebraic order structures -/ @@ -102,7 +111,7 @@ instance (α : ι → Type*) [Π i, ordered_add_comm_monoid (α i)] : ordered_add_comm_monoid (Π₀ i, α i) := { add_le_add_left := λ a b h c i, by { rw [add_apply, add_apply], exact add_le_add_left (h i) (c i) }, - .. dfinsupp.add_comm_monoid, .. dfinsupp.partial_order α } + .. dfinsupp.add_comm_monoid, .. dfinsupp.partial_order } instance (α : ι → Type*) [Π i, ordered_cancel_add_comm_monoid (α i)] : ordered_cancel_add_comm_monoid (Π₀ i, α i) := diff --git a/src/data/fin/interval.lean b/src/data/fin/interval.lean index 41a0f2409e519..46e3a23c11bdf 100644 --- a/src/data/fin/interval.lean +++ b/src/data/fin/interval.lean @@ -16,6 +16,16 @@ This file proves that `fin n` is a `locally_finite_order` and calculates the car intervals as finsets and fintypes. -/ +namespace fin +variables {n : ℕ} (a b : fin n) + +@[simp, norm_cast] lemma coe_sup : (↑(a ⊔ b) : ℕ) = a ⊔ b := rfl +@[simp, norm_cast] lemma coe_inf : (↑(a ⊓ b) : ℕ) = a ⊓ b := rfl +@[simp, norm_cast] lemma coe_max : (↑(max a b) : ℕ) = max a b := rfl +@[simp, norm_cast] lemma coe_min : (↑(min a b) : ℕ) = min a b := rfl + +end fin + open finset fin function open_locale big_operators @@ -39,6 +49,7 @@ lemma Icc_eq_finset_subtype : Icc a b = (Icc (a : ℕ) b).fin n := rfl lemma Ico_eq_finset_subtype : Ico a b = (Ico (a : ℕ) b).fin n := rfl lemma Ioc_eq_finset_subtype : Ioc a b = (Ioc (a : ℕ) b).fin n := rfl lemma Ioo_eq_finset_subtype : Ioo a b = (Ioo (a : ℕ) b).fin n := rfl +lemma uIcc_eq_finset_subtype : uIcc a b = (uIcc (a : ℕ) b).fin n := rfl @[simp] lemma map_subtype_embedding_Icc : (Icc a b).map fin.coe_embedding = Icc a b := by simp [Icc_eq_finset_subtype, finset.fin, finset.map_map, Icc_filter_lt_of_lt_right] @@ -52,6 +63,9 @@ by simp [Ioc_eq_finset_subtype, finset.fin, finset.map_map, Ioc_filter_lt_of_lt_ @[simp] lemma map_subtype_embedding_Ioo : (Ioo a b).map fin.coe_embedding = Ioo a b := by simp [Ioo_eq_finset_subtype, finset.fin, finset.map_map] +@[simp] lemma map_subtype_embedding_uIcc : (uIcc a b).map coe_embedding = uIcc a b := +map_subtype_embedding_Icc _ _ + @[simp] lemma card_Icc : (Icc a b).card = b + 1 - a := by rw [←nat.card_Icc, ←map_subtype_embedding_Icc, card_map] @@ -64,6 +78,9 @@ by rw [←nat.card_Ioc, ←map_subtype_embedding_Ioc, card_map] @[simp] lemma card_Ioo : (Ioo a b).card = b - a - 1 := by rw [←nat.card_Ioo, ←map_subtype_embedding_Ioo, card_map] +@[simp] lemma card_uIcc : (uIcc a b).card = (b - a : ℤ).nat_abs + 1 := +by rw [coe_coe, coe_coe, ←nat.card_uIcc, ←map_subtype_embedding_uIcc, card_map] + @[simp] lemma card_fintype_Icc : fintype.card (set.Icc a b) = b + 1 - a := by rw [←card_Icc, fintype.card_of_finset] @@ -76,6 +93,9 @@ by rw [←card_Ioc, fintype.card_of_finset] @[simp] lemma card_fintype_Ioo : fintype.card (set.Ioo a b) = b - a - 1 := by rw [←card_Ioo, fintype.card_of_finset] +@[simp] lemma card_fintype_uIcc : fintype.card (set.uIcc a b) = (b - a : ℤ).nat_abs + 1 := +by rw [←card_uIcc, fintype.card_of_finset] + lemma Ici_eq_finset_subtype : Ici a = (Icc (a : ℕ) n).fin n := by { ext, simp } lemma Ioi_eq_finset_subtype : Ioi a = (Ioc (a : ℕ) n).fin n := by { ext, simp } lemma Iic_eq_finset_subtype : Iic b = (Iic (b : ℕ)).fin n := rfl diff --git a/src/data/finset/locally_finite.lean b/src/data/finset/locally_finite.lean index 8dc178e34bc1d..f8b6e5b429cc9 100644 --- a/src/data/finset/locally_finite.lean +++ b/src/data/finset/locally_finite.lean @@ -564,8 +564,11 @@ lemma Icc_subset_uIcc' : Icc b a ⊆ [a, b] := Icc_subset_Icc inf_le_right le_su @[simp] lemma left_mem_uIcc : a ∈ [a, b] := mem_Icc.2 ⟨inf_le_left, le_sup_left⟩ @[simp] lemma right_mem_uIcc : b ∈ [a, b] := mem_Icc.2 ⟨inf_le_right, le_sup_right⟩ -lemma mem_uIcc_of_le (ha : a ≤ x) (hb : x ≤ b) : x ∈ [a, b] := Icc_subset_uIcc $ mem_Icc.2 ⟨ha, hb⟩ -lemma mem_uIcc_of_ge (hb : b ≤ x) (ha : x ≤ a) : x ∈ [a, b] := Icc_subset_uIcc' $ mem_Icc.2 ⟨hb, ha⟩ +lemma mem_uIcc_of_le (ha : a ≤ x) (hb : x ≤ b) : x ∈ [a, b] := +Icc_subset_uIcc $ mem_Icc.2 ⟨ha, hb⟩ + +lemma mem_uIcc_of_ge (hb : b ≤ x) (ha : x ≤ a) : x ∈ [a, b] := +Icc_subset_uIcc' $ mem_Icc.2 ⟨hb, ha⟩ lemma uIcc_subset_uIcc (h₁ : a₁ ∈ [a₂, b₂]) (h₂ : b₁ ∈ [a₂, b₂]) : [a₁, b₁] ⊆ [a₂, b₂] := by { rw mem_uIcc at h₁ h₂, exact Icc_subset_Icc (le_inf h₁.1 h₂.1) (sup_le h₁.2 h₂.2) } @@ -576,11 +579,15 @@ by { rw mem_Icc at ha hb, exact Icc_subset_Icc (le_inf ha.1 hb.1) (sup_le ha.2 h lemma uIcc_subset_uIcc_iff_mem : [a₁, b₁] ⊆ [a₂, b₂] ↔ a₁ ∈ [a₂, b₂] ∧ b₁ ∈ [a₂, b₂] := ⟨λ h, ⟨h left_mem_uIcc, h right_mem_uIcc⟩, λ h, uIcc_subset_uIcc h.1 h.2⟩ -lemma uIcc_subset_uIcc_iff_le' : [a₁, b₁] ⊆ [a₂, b₂] ↔ a₂ ⊓ b₂ ≤ a₁ ⊓ b₁ ∧ a₁ ⊔ b₁ ≤ a₂ ⊔ b₂ := +lemma uIcc_subset_uIcc_iff_le' : + [a₁, b₁] ⊆ [a₂, b₂] ↔ a₂ ⊓ b₂ ≤ a₁ ⊓ b₁ ∧ a₁ ⊔ b₁ ≤ a₂ ⊔ b₂ := Icc_subset_Icc_iff inf_le_sup -lemma uIcc_subset_uIcc_right (h : x ∈ [a, b]) : [x, b] ⊆ [a, b] := uIcc_subset_uIcc h right_mem_uIcc -lemma uIcc_subset_uIcc_left (h : x ∈ [a, b]) : [a, x] ⊆ [a, b] := uIcc_subset_uIcc left_mem_uIcc h +lemma uIcc_subset_uIcc_right (h : x ∈ [a, b]) : [x, b] ⊆ [a, b] := +uIcc_subset_uIcc h right_mem_uIcc + +lemma uIcc_subset_uIcc_left (h : x ∈ [a, b]) : [a, x] ⊆ [a, b] := +uIcc_subset_uIcc left_mem_uIcc h end lattice diff --git a/src/data/finsupp/interval.lean b/src/data/finsupp/interval.lean index b71f5d4d4987c..b05e89043e75b 100644 --- a/src/data/finsupp/interval.lean +++ b/src/data/finsupp/interval.lean @@ -106,6 +106,15 @@ by rw [card_Ioo_eq_card_Icc_sub_two, card_Icc] end partial_order +section lattice +variables [lattice α] [has_zero α] [locally_finite_order α] (f g : ι →₀ α) + +lemma card_uIcc [decidable_eq ι] : + (uIcc f g).card = ∏ i in f.support ∪ g.support, (uIcc (f i) (g i)).card := +by { rw ←support_inf_union_support_sup, exact card_Icc _ _ } + +end lattice + section canonically_ordered variables [canonically_ordered_add_monoid α] [locally_finite_order α] diff --git a/src/data/finsupp/order.lean b/src/data/finsupp/order.lean index 889ee49238192..b5aa762727bc9 100644 --- a/src/data/finsupp/order.lean +++ b/src/data/finsupp/order.lean @@ -86,8 +86,18 @@ instance [semilattice_sup α] : semilattice_sup (ι →₀ α) := @[simp] lemma sup_apply [semilattice_sup α] {i : ι} {f g : ι →₀ α} : (f ⊔ g) i = f i ⊔ g i := rfl -instance lattice [lattice α] : lattice (ι →₀ α) := -{ .. finsupp.semilattice_inf, .. finsupp.semilattice_sup } +instance [lattice α] : lattice (ι →₀ α) := { ..finsupp.semilattice_inf, ..finsupp.semilattice_sup } + +section lattice +variables [decidable_eq ι] [lattice α] (f g : ι →₀ α) + +lemma support_inf_union_support_sup : (f ⊓ g).support ∪ (f ⊔ g).support = f.support ∪ g.support := +coe_injective $ compl_injective $ by { ext, simp [inf_eq_and_sup_eq_iff] } + +lemma support_sup_union_support_inf : (f ⊔ g).support ∪ (f ⊓ g).support = f.support ∪ g.support := +(union_comm _ _).trans $ support_inf_union_support_sup _ _ + +end lattice end has_zero diff --git a/src/data/int/interval.lean b/src/data/int/interval.lean index 4856c9e400009..4b0c1c7b84641 100644 --- a/src/data/int/interval.lean +++ b/src/data/int/interval.lean @@ -92,18 +92,18 @@ lemma Ioc_eq_finset_map : lemma Ioo_eq_finset_map : Ioo a b = (finset.range (b - a - 1).to_nat).map (nat.cast_embedding.trans $ add_left_embedding (a + 1)) := rfl +lemma uIcc_eq_finset_map : uIcc a b = (range (max a b + 1 - min a b).to_nat).map + (nat.cast_embedding.trans $ add_left_embedding $ min a b) := rfl -@[simp] lemma card_Icc : (Icc a b).card = (b + 1 - a).to_nat := -by { change (finset.map _ _).card = _, rw [finset.card_map, finset.card_range] } +@[simp] lemma card_Icc : (Icc a b).card = (b + 1 - a).to_nat := (card_map _).trans $ card_range _ +@[simp] lemma card_Ico : (Ico a b).card = (b - a).to_nat := (card_map _).trans $ card_range _ +@[simp] lemma card_Ioc : (Ioc a b).card = (b - a).to_nat := (card_map _).trans $ card_range _ +@[simp] lemma card_Ioo : (Ioo a b).card = (b - a - 1).to_nat := (card_map _).trans $ card_range _ -@[simp] lemma card_Ico : (Ico a b).card = (b - a).to_nat := -by { change (finset.map _ _).card = _, rw [finset.card_map, finset.card_range] } - -@[simp] lemma card_Ioc : (Ioc a b).card = (b - a).to_nat := -by { change (finset.map _ _).card = _, rw [finset.card_map, finset.card_range] } - -@[simp] lemma card_Ioo : (Ioo a b).card = (b - a - 1).to_nat := -by { change (finset.map _ _).card = _, rw [finset.card_map, finset.card_range] } +@[simp] lemma card_uIcc : (uIcc a b).card = (b - a).nat_abs + 1 := +(card_map _).trans $ int.coe_nat_inj $ by rw [card_range, sup_eq_max, inf_eq_min, + int.to_nat_of_nonneg (sub_nonneg_of_le $ le_add_one min_le_max), int.coe_nat_add, int.coe_nat_abs, + add_comm, add_sub_assoc, max_sub_min_eq_abs, add_comm, int.coe_nat_one] lemma card_Icc_of_le (h : a ≤ b + 1) : ((Icc a b).card : ℤ) = b + 1 - a := by rw [card_Icc, to_nat_sub_of_le h] @@ -129,6 +129,9 @@ by rw [←card_Ioc, fintype.card_of_finset] @[simp] lemma card_fintype_Ioo : fintype.card (set.Ioo a b) = (b - a - 1).to_nat := by rw [←card_Ioo, fintype.card_of_finset] +@[simp] lemma card_fintype_uIcc : fintype.card (set.uIcc a b) = (b - a).nat_abs + 1 := +by rw [←card_uIcc, fintype.card_of_finset] + lemma card_fintype_Icc_of_le (h : a ≤ b + 1) : (fintype.card (set.Icc a b) : ℤ) = b + 1 - a := by rw [card_fintype_Icc, to_nat_sub_of_le h] diff --git a/src/data/multiset/interval.lean b/src/data/multiset/interval.lean index 33ae1fb363183..8e36ee02b2efb 100644 --- a/src/data/multiset/interval.lean +++ b/src/data/multiset/interval.lean @@ -29,42 +29,52 @@ multisets are typically used computationally. open finset dfinsupp function open_locale big_operators pointwise -variables {α : Type*} {β : α → Type*} +variables {α : Type*} namespace multiset - -variables [decidable_eq α] (f g : multiset α) +variables [decidable_eq α] (s t : multiset α) instance : locally_finite_order (multiset α) := locally_finite_order.of_Icc (multiset α) - (λ f g, (finset.Icc f.to_dfinsupp g.to_dfinsupp).map + (λ s t, (finset.Icc s.to_dfinsupp t.to_dfinsupp).map (multiset.equiv_dfinsupp.to_equiv.symm.to_embedding)) - (λ f g x, by simp) + (λ s t x, by simp) lemma Icc_eq : - finset.Icc f g = - (finset.Icc f.to_dfinsupp g.to_dfinsupp).map + finset.Icc s t = + (finset.Icc s.to_dfinsupp t.to_dfinsupp).map (multiset.equiv_dfinsupp.to_equiv.symm.to_embedding) := rfl +lemma uIcc_eq : + uIcc s t = + (uIcc s.to_dfinsupp t.to_dfinsupp).map + (multiset.equiv_dfinsupp.to_equiv.symm.to_embedding) := +(Icc_eq _ _).trans $ by simp [uIcc] + lemma card_Icc : - (finset.Icc f g).card = ∏ i in f.to_finset ∪ g.to_finset, (g.count i + 1 - f.count i) := + (finset.Icc s t).card = ∏ i in s.to_finset ∪ t.to_finset, (t.count i + 1 - s.count i) := by simp_rw [Icc_eq, finset.card_map, dfinsupp.card_Icc, nat.card_Icc, multiset.to_dfinsupp_apply, to_dfinsupp_support] lemma card_Ico : - (finset.Ico f g).card = ∏ i in f.to_finset ∪ g.to_finset, (g.count i + 1 - f.count i) - 1 := + (finset.Ico s t).card = ∏ i in s.to_finset ∪ t.to_finset, (t.count i + 1 - s.count i) - 1 := by rw [card_Ico_eq_card_Icc_sub_one, card_Icc] lemma card_Ioc : - (finset.Ioc f g).card = ∏ i in f.to_finset ∪ g.to_finset, (g.count i + 1 - f.count i) - 1 := + (finset.Ioc s t).card = ∏ i in s.to_finset ∪ t.to_finset, (t.count i + 1 - s.count i) - 1 := by rw [card_Ioc_eq_card_Icc_sub_one, card_Icc] lemma card_Ioo : - (finset.Ioo f g).card = ∏ i in f.to_finset ∪ g.to_finset, (g.count i + 1 - f.count i) - 2 := + (finset.Ioo s t).card = ∏ i in s.to_finset ∪ t.to_finset, (t.count i + 1 - s.count i) - 2 := by rw [card_Ioo_eq_card_Icc_sub_two, card_Icc] +lemma card_uIcc : + (uIcc s t).card = ∏ i in s.to_finset ∪ t.to_finset, ((t.count i - s.count i : ℤ).nat_abs + 1) := +by simp_rw [uIcc_eq, finset.card_map, dfinsupp.card_uIcc, nat.card_uIcc, multiset.to_dfinsupp_apply, + to_dfinsupp_support] + lemma card_Iic : - (finset.Iic f).card = ∏ i in f.to_finset, (f.count i + 1) := + (finset.Iic s).card = ∏ i in s.to_finset, (s.count i + 1) := by simp_rw [Iic_eq_Icc, card_Icc, bot_eq_zero, to_finset_zero, empty_union, count_zero, tsub_zero] end multiset diff --git a/src/data/nat/interval.lean b/src/data/nat/interval.lean index d1a5a2e862753..a7efc2bf8d370 100644 --- a/src/data/nat/interval.lean +++ b/src/data/nat/interval.lean @@ -65,6 +65,8 @@ lemma Icc_eq_range' : Icc a b = ⟨list.range' a (b + 1 - a), list.nodup_range' lemma Ico_eq_range' : Ico a b = ⟨list.range' a (b - a), list.nodup_range' _ _⟩ := rfl lemma Ioc_eq_range' : Ioc a b = ⟨list.range' (a + 1) (b - a), list.nodup_range' _ _⟩ := rfl lemma Ioo_eq_range' : Ioo a b = ⟨list.range' (a + 1) (b - a - 1), list.nodup_range' _ _⟩ := rfl +lemma uIcc_eq_range' : + uIcc a b = ⟨list.range' (min a b) (max a b + 1 - min a b), list.nodup_range' _ _⟩ := rfl lemma Iio_eq_range : Iio = range := by { ext b x, rw [mem_Iio, mem_range] } @@ -76,6 +78,18 @@ lemma _root_.finset.range_eq_Ico : range = Ico 0 := Ico_zero_eq_range.symm @[simp] lemma card_Ico : (Ico a b).card = b - a := list.length_range' _ _ @[simp] lemma card_Ioc : (Ioc a b).card = b - a := list.length_range' _ _ @[simp] lemma card_Ioo : (Ioo a b).card = b - a - 1 := list.length_range' _ _ + +@[simp] lemma card_uIcc : (uIcc a b).card = (b - a : ℤ).nat_abs + 1 := +begin + refine (card_Icc _ _).trans (int.coe_nat_inj _), + rw [sup_eq_max, inf_eq_min, int.coe_nat_sub], + { rw [add_comm, int.coe_nat_add, add_sub_assoc], + norm_cast, + push_cast, + rw [max_sub_min_eq_abs, add_comm] }, + { exact min_le_max.trans le_self_add } +end + @[simp] lemma card_Iic : (Iic b).card = b + 1 := by rw [Iic_eq_Icc, card_Icc, bot_eq_zero, tsub_zero] diff --git a/src/data/pi/interval.lean b/src/data/pi/interval.lean index c77e21f3ce2bd..812b52ef97122 100644 --- a/src/data/pi/interval.lean +++ b/src/data/pi/interval.lean @@ -19,14 +19,14 @@ order are locally finite and calculates the cardinality of their intervals. open finset fintype open_locale big_operators -variables {ι : Type*} {α : ι → Type*} - +variables {ι : Type*} {α : ι → Type*} [fintype ι] [decidable_eq ι] [Π i, decidable_eq (α i)] namespace pi +section partial_order +variables [Π i, partial_order (α i)] -section locally_finite -variables [decidable_eq ι] [fintype ι] [Π i, decidable_eq (α i)] - [Π i, partial_order (α i)] [Π i, locally_finite_order (α i)] +section locally_finite_order +variables [Π i, locally_finite_order (α i)] instance : locally_finite_order (Π i, α i) := locally_finite_order.of_Icc _ @@ -39,21 +39,18 @@ lemma Icc_eq : Icc a b = pi_finset (λ i, Icc (a i) (b i)) := rfl lemma card_Icc : (Icc a b).card = ∏ i, (Icc (a i) (b i)).card := card_pi_finset _ -lemma card_Ico : (Ico a b).card = (∏ i, (Icc (a i) (b i)).card) - 1 := +lemma card_Ico : (Ico a b).card = ∏ i, (Icc (a i) (b i)).card - 1 := by rw [card_Ico_eq_card_Icc_sub_one, card_Icc] -lemma card_Ioc : (Ioc a b).card = (∏ i, (Icc (a i) (b i)).card) - 1 := +lemma card_Ioc : (Ioc a b).card = ∏ i, (Icc (a i) (b i)).card - 1 := by rw [card_Ioc_eq_card_Icc_sub_one, card_Icc] -lemma card_Ioo : (Ioo a b).card = (∏ i, (Icc (a i) (b i)).card) - 2 := +lemma card_Ioo : (Ioo a b).card = ∏ i, (Icc (a i) (b i)).card - 2 := by rw [card_Ioo_eq_card_Icc_sub_two, card_Icc] -end locally_finite - -section bounded -variables [decidable_eq ι] [fintype ι] [Π i, decidable_eq (α i)] [Π i, partial_order (α i)] +end locally_finite_order -section bot +section locally_finite_order_bot variables [Π i, locally_finite_order_bot (α i)] (b : Π i, α i) instance : locally_finite_order_bot (Π i, α i) := @@ -63,12 +60,12 @@ locally_finite_order_top.of_Iic _ lemma card_Iic : (Iic b).card = ∏ i, (Iic (b i)).card := card_pi_finset _ -lemma card_Iio : (Iio b).card = (∏ i, (Iic (b i)).card) - 1 := +lemma card_Iio : (Iio b).card = ∏ i, (Iic (b i)).card - 1 := by rw [card_Iio_eq_card_Iic_sub_one, card_Iic] -end bot +end locally_finite_order_bot -section top +section locally_finite_order_top variables [Π i, locally_finite_order_top (α i)] (a : Π i, α i) instance : locally_finite_order_top (Π i, α i) := @@ -76,13 +73,19 @@ locally_finite_order_top.of_Ici _ (λ a, pi_finset $ λ i, Ici (a i)) (λ a x, by simp_rw [mem_pi_finset, mem_Ici, le_def]) -lemma card_Ici : (Ici a).card = (∏ i, (Ici (a i)).card) := card_pi_finset _ +lemma card_Ici : (Ici a).card = ∏ i, (Ici (a i)).card := card_pi_finset _ -lemma card_Ioi : (Ioi a).card = (∏ i, (Ici (a i)).card) - 1 := +lemma card_Ioi : (Ioi a).card = ∏ i, (Ici (a i)).card - 1 := by rw [card_Ioi_eq_card_Ici_sub_one, card_Ici] -end top +end locally_finite_order_top +end partial_order + +section lattice +variables [Π i, lattice (α i)] [Π i, locally_finite_order (α i)] (a b : Π i, α i) -end bounded +lemma uIcc_eq : uIcc a b = pi_finset (λ i, uIcc (a i) (b i)) := rfl +lemma card_uIcc : (uIcc a b).card = ∏ i, (uIcc (a i) (b i)).card := card_Icc _ _ +end lattice end pi diff --git a/src/data/pnat/interval.lean b/src/data/pnat/interval.lean index 6d2ecf63145a6..4b6737613cc9f 100644 --- a/src/data/pnat/interval.lean +++ b/src/data/pnat/interval.lean @@ -16,7 +16,7 @@ This file proves that `ℕ+` is a `locally_finite_order` and calculates the card intervals as finsets and fintypes. -/ -open finset pnat +open finset function pnat instance : locally_finite_order ℕ+ := subtype.locally_finite_order _ @@ -27,19 +27,23 @@ lemma Icc_eq_finset_subtype : Icc a b = (Icc (a : ℕ) b).subtype (λ (n : ℕ), lemma Ico_eq_finset_subtype : Ico a b = (Ico (a : ℕ) b).subtype (λ (n : ℕ), 0 < n) := rfl lemma Ioc_eq_finset_subtype : Ioc a b = (Ioc (a : ℕ) b).subtype (λ (n : ℕ), 0 < n) := rfl lemma Ioo_eq_finset_subtype : Ioo a b = (Ioo (a : ℕ) b).subtype (λ (n : ℕ), 0 < n) := rfl +lemma uIcc_eq_finset_subtype : uIcc a b = (uIcc (a : ℕ) b).subtype (λ (n : ℕ), 0 < n) := rfl -lemma map_subtype_embedding_Icc : (Icc a b).map (function.embedding.subtype _) = Icc (a : ℕ) b := +lemma map_subtype_embedding_Icc : (Icc a b).map (embedding.subtype _) = Icc a b := map_subtype_embedding_Icc _ _ _ (λ c _ x hx _ hc _, hc.trans_le hx) -lemma map_subtype_embedding_Ico : (Ico a b).map (function.embedding.subtype _) = Ico (a : ℕ) b := +lemma map_subtype_embedding_Ico : (Ico a b).map (embedding.subtype _) = Ico a b := map_subtype_embedding_Ico _ _ _ (λ c _ x hx _ hc _, hc.trans_le hx) -lemma map_subtype_embedding_Ioc : (Ioc a b).map (function.embedding.subtype _) = Ioc (a : ℕ) b := +lemma map_subtype_embedding_Ioc : (Ioc a b).map (embedding.subtype _) = Ioc a b := map_subtype_embedding_Ioc _ _ _ (λ c _ x hx _ hc _, hc.trans_le hx) -lemma map_subtype_embedding_Ioo : (Ioo a b).map (function.embedding.subtype _) = Ioo (a : ℕ) b := +lemma map_subtype_embedding_Ioo : (Ioo a b).map (embedding.subtype _) = Ioo a b := map_subtype_embedding_Ioo _ _ _ (λ c _ x hx _ hc _, hc.trans_le hx) +lemma map_subtype_embedding_uIcc : (uIcc a b).map (embedding.subtype _) = uIcc a b := +map_subtype_embedding_Icc _ _ + @[simp] lemma card_Icc : (Icc a b).card = b + 1 - a := by rw [←nat.card_Icc, ←map_subtype_embedding_Icc, card_map] @@ -52,6 +56,9 @@ by rw [←nat.card_Ioc, ←map_subtype_embedding_Ioc, card_map] @[simp] lemma card_Ioo : (Ioo a b).card = b - a - 1 := by rw [←nat.card_Ioo, ←map_subtype_embedding_Ioo, card_map] +@[simp] lemma card_uIcc : (uIcc a b).card = (b - a : ℤ).nat_abs + 1 := +by rw [coe_coe, coe_coe, ←nat.card_uIcc, ←map_subtype_embedding_uIcc, card_map] + @[simp] lemma card_fintype_Icc : fintype.card (set.Icc a b) = b + 1 - a := by rw [←card_Icc, fintype.card_of_finset] @@ -64,4 +71,7 @@ by rw [←card_Ioc, fintype.card_of_finset] @[simp] lemma card_fintype_Ioo : fintype.card (set.Ioo a b) = b - a - 1 := by rw [←card_Ioo, fintype.card_of_finset] +@[simp] lemma card_fintype_uIcc : fintype.card (set.uIcc a b) = (b - a : ℤ).nat_abs + 1 := +by rw [←card_uIcc, fintype.card_of_finset] + end pnat diff --git a/src/order/locally_finite.lean b/src/order/locally_finite.lean index 40c3ec825ac84..f6194e68b8b81 100644 --- a/src/order/locally_finite.lean +++ b/src/order/locally_finite.lean @@ -458,17 +458,10 @@ namespace set section preorder variables [preorder α] [locally_finite_order α] (a b : α) -instance fintype_Icc : fintype (Icc a b) := -fintype.of_finset (finset.Icc a b) (λ x, by rw [finset.mem_Icc, mem_Icc]) - -instance fintype_Ico : fintype (Ico a b) := -fintype.of_finset (finset.Ico a b) (λ x, by rw [finset.mem_Ico, mem_Ico]) - -instance fintype_Ioc : fintype (Ioc a b) := -fintype.of_finset (finset.Ioc a b) (λ x, by rw [finset.mem_Ioc, mem_Ioc]) - -instance fintype_Ioo : fintype (Ioo a b) := -fintype.of_finset (finset.Ioo a b) (λ x, by rw [finset.mem_Ioo, mem_Ioo]) +instance fintype_Icc : fintype (Icc a b) := fintype.of_finset (finset.Icc a b) $ λ x, finset.mem_Icc +instance fintype_Ico : fintype (Ico a b) := fintype.of_finset (finset.Ico a b) $ λ x, finset.mem_Ico +instance fintype_Ioc : fintype (Ioc a b) := fintype.of_finset (finset.Ioc a b) $ λ x, finset.mem_Ioc +instance fintype_Ioo : fintype (Ioo a b) := fintype.of_finset (finset.Ioo a b) $ λ x, finset.mem_Ioo lemma finite_Icc : (Icc a b).finite := (Icc a b).to_finite lemma finite_Ico : (Ico a b).finite := (Ico a b).to_finite @@ -480,11 +473,8 @@ end preorder section order_top variables [preorder α] [locally_finite_order_top α] (a : α) -instance fintype_Ici : fintype (Ici a) := -fintype.of_finset (finset.Ici a) (λ x, by rw [finset.mem_Ici, mem_Ici]) - -instance fintype_Ioi : fintype (Ioi a) := -fintype.of_finset (finset.Ioi a) (λ x, by rw [finset.mem_Ioi, mem_Ioi]) +instance fintype_Ici : fintype (Ici a) := fintype.of_finset (finset.Ici a) $ λ x, finset.mem_Ici +instance fintype_Ioi : fintype (Ioi a) := fintype.of_finset (finset.Ioi a) $ λ x, finset.mem_Ioi lemma finite_Ici : (Ici a).finite := (Ici a).to_finite lemma finite_Ioi : (Ioi a).finite := (Ioi a).to_finite @@ -494,17 +484,23 @@ end order_top section order_bot variables [preorder α] [locally_finite_order_bot α] (b : α) -instance fintype_Iic : fintype (Iic b) := -fintype.of_finset (finset.Iic b) (λ x, by rw [finset.mem_Iic, mem_Iic]) - -instance fintype_Iio : fintype (Iio b) := -fintype.of_finset (finset.Iio b) (λ x, by rw [finset.mem_Iio, mem_Iio]) +instance fintype_Iic : fintype (Iic b) := fintype.of_finset (finset.Iic b) $ λ x, finset.mem_Iic +instance fintype_Iio : fintype (Iio b) := fintype.of_finset (finset.Iio b) $ λ x, finset.mem_Iio lemma finite_Iic : (Iic b).finite := (Iic b).to_finite lemma finite_Iio : (Iio b).finite := (Iio b).to_finite end order_bot +section lattice +variables [lattice α] [locally_finite_order α] (a b : α) + +instance fintype_uIcc : fintype (uIcc a b) := +fintype.of_finset (finset.uIcc a b) $ λ x, finset.mem_uIcc + +@[simp] lemma finite_interval : (uIcc a b).finite := (uIcc _ _).to_finite + +end lattice end set /-! ### Instances -/ From 6b711d2ba5d470c040677ddda0c26b0d72283886 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sat, 15 Jul 2023 00:02:14 +0000 Subject: [PATCH 23/61] feat(data/matrix/auto): lemmas for arbitrary concrete matrices generated via auto_params (#15738) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This adds a lemma `matrix.mul_fin` that works for arbitrary dimensions of concrete matrices indexed by `fin`. It uses a similar strategy to `category_theory.reassoc_of` to have a tactic populate wildcards in the lemma statement. For example: ```lean example {α} [add_comm_monoid α] [has_mul α] (a₁₁ a₁₂ a₂₁ a₂₂ b₁₁ b₁₂ b₂₁ b₂₂ : α) : !![a₁₁, a₁₂; a₂₁, a₂₂] ⬝ !![b₁₁, b₁₂; b₂₁, b₂₂] = !![a₁₁ * b₁₁ + a₁₂ * b₂₁, a₁₁ * b₁₂ + a₁₂ * b₂₂; a₂₁ * b₁₁ + a₂₂ * b₂₁, a₂₁ * b₁₂ + a₂₂ * b₂₂] := begin rw mul_fin, end example {α} [add_comm_monoid α] [has_mul α] (a₁₁ a₁₂ b₁₁ b₁₂ b₂₁ b₂₂ : α) : !![a₁₁, a₁₂] ⬝ !![b₁₁, b₁₂; b₂₁, b₂₂] = !![a₁₁ * b₁₁ + a₁₂ * b₂₁, a₁₁ * b₁₂ + a₁₂ * b₂₂;] := begin rw mul_fin, end ``` Relevant zulip threads: * [Explicit vector unfolding](https://leanprover.zulipchat.com/#narrow/stream/113489-new-members/topic/Explicit.20vector.20unfolding) * [concrete matrix multiplication](https://leanprover.zulipchat.com/#narrow/stream/217875-Is-there-code-for-X.3F/topic/concrete.20matrix.20multiplication/near/291208624) * [A tactic for expanding matrices into coefficients](https://leanprover.zulipchat.com/#narrow/stream/239415-metaprogramming-.2F-tactics/topic/A.20tactic.20for.20expanding.20matrices.20into.20coefficients) --- scripts/lint-style.py | 2 +- src/algebra/expr.lean | 48 +++++++++ src/data/matrix/auto.lean | 196 ++++++++++++++++++++++++++++++++++++ test/matrix_reflection.lean | 21 ++++ 4 files changed, 266 insertions(+), 1 deletion(-) create mode 100644 src/algebra/expr.lean create mode 100644 src/data/matrix/auto.lean diff --git a/scripts/lint-style.py b/scripts/lint-style.py index 6fbd777019874..c644ebe292a35 100755 --- a/scripts/lint-style.py +++ b/scripts/lint-style.py @@ -133,7 +133,7 @@ def small_alpha_vrachy_check(lines, path): def instance_check(lines, path): errors = [] for line_nr, line in skip_string(skip_comments(enumerate(lines, 1))): - if re.search(r'\b(_inst_)\d+', line) is not None: + if re.search(r'\b(? vector.mmap f ⟨list.fin_range n, list.length_fin_range _⟩ + +namespace matrix + +section fin_eta + +/-- Prove a statement of the form `∀ A : matrix m n α, A = !![A 0 0, ...]`. +Returns the type of this statement and its proof. -/ +meta def fin_eta.prove (m n : ℕ) : tactic (expr × expr) := +do + u ← tactic.mk_meta_univ, + α ← tactic.mk_local' `α binder_info.implicit (expr.sort u.succ), + A ← tactic.mk_local' `A binder_info.default + (expr.const `matrix [level.zero, level.zero, u] `(fin %%`(m)) `(fin %%`(n)) α), + let entries := λ (i : fin m) (j : fin n), A `(i) `(j), + let entry_vals := pi_fin.to_pexpr (λ i, pi_fin.to_pexpr (λ j, to_pexpr $ entries i j)), + let A_eta := (``(@matrix.of (fin %%`(m)) (fin %%`(n)) _).app entry_vals), + A_eq ← tactic.to_expr ``(%%A = %%A_eta), + t ← tactic.pis [α, A] A_eq, + ((), pr) ← tactic.solve_aux t `[intros α A, exact (matrix.eta_expand_eq A).symm], + pure (t, pr) + +/-- Helper tactic used as an `auto_param` for `matrix.fin_eta` -/ +meta def fin_eta.derive : tactic unit := +do + target@`(%%A' = %%A_eta) ← tactic.target, + (expr.const `matrix ls, [`(fin %%m), `(fin %%n), α]) + ← expr.get_app_fn_args <$> tactic.infer_type A', + some (m, n) ← pure (prod.mk <$> m.to_nat <*> n.to_nat) | + fail!"Dimensions {m} {n} are not numerals", + (t,pr) ← matrix.fin_eta.prove m n, + + tactic.unify target (t.instantiate_pis [α, A']), + tactic.exact (pr α A') + +/-- This lemma expands `A` into `!![A 0 0, ...]`. -/ +theorem fin_eta {α} {m n : ℕ} + (A : matrix (fin m) (fin n) α) {«!![A 0 0, ...]» : matrix (fin m) (fin n) α} + (h : A = «!![A 0 0, ...]» . matrix.fin_eta.derive) : A = «!![A 0 0, ...]» := h + +example : true := +begin + let B : matrix (fin 20) (fin 20) ℕ := 0, + have := matrix.fin_eta B, -- 400 coefficients, but very fast + have : B = B, by rw this, + trivial, +end + +end fin_eta + +section of_mul_of_fin + +/-- Choose a name suffix for a matrix index -/ +private def name_suffix {m n : ℕ} : fin m → fin n → string := +let chars := "₀₁₂₃₄₅₆₇₈₉".data in +if h : m ≤ 10 ∧ n ≤ 10 +then (λ i j, [chars.nth_le i (i.prop.trans_le h.1), chars.nth_le j (j.prop.trans_le h.2)].as_string) +else (λ i j, "_" ++ to_string i ++ "_" ++ to_string j) + +/-- `pi_fin.to_pexpr` but for matrices -/ +meta def fin_to_pexpr {m n : ℕ} (A : matrix (fin m) (fin n) pexpr) : pexpr := +``(@matrix.of (fin %%`(m)) (fin %%`(n)) _).app $ + pi_fin.to_pexpr (λ i : fin m, pi_fin.to_pexpr (λ j : fin n, A i j)) + +/-- This statement is defeq to `of_mul_of_fin`, but syntactically worse-/ +theorem of_mul_of_fin_aux (l m n : ℕ) ⦃α⦄ [has_mul α] [add_comm_monoid α] : + «forall» $ λ A : matrix (fin l) (fin m) α, + «forall» $ λ B : matrix (fin m) (fin n) α, + A.mul B = A.mulᵣ B := +by simp_rw [forall_iff, mulᵣ_eq, eq_self_iff_true, forall_const] + +/-- Prove a statement of the form +``` +∀ α [has_mul α] [add_comm_monoid α] (a₁₁ ... aₗₘ b₁₁ ... bₘₙ : α), + !![a₁₁ ⋱ aₗₘ] ⬝ !![b₁₁ ⋱ bₘₙ] = !![⋱] +``` +Returns the type of this statement and its proof. -/ +meta def of_mul_of_fin.prove (l m n : ℕ) : tactic (expr × expr) := +do + -- create all the binders, one for each coefficient + u ← tactic.mk_meta_univ, + α ← tactic.mk_local' `α binder_info.implicit (expr.sort u.succ), + has_mul_α ← tactic.mk_app `has_mul [α] >>= tactic.mk_local' `_inst_1 binder_info.inst_implicit, + add_comm_monoid_α ← + tactic.mk_app `add_comm_monoid [α] >>= tactic.mk_local' `_inst_2 binder_info.inst_implicit, + a ← (fin.mmap $ λ i : fin l, fin.mmap $ λ j : fin m, + tactic.mk_local' ((`a).append_suffix (name_suffix i j)) binder_info.default α), + b ← (fin.mmap $ λ i : fin m, fin.mmap $ λ j : fin n, + tactic.mk_local' ((`b).append_suffix (name_suffix i j)) binder_info.default α), + let a_flat := (list.fin_range l).bind (λ i, (list.fin_range m).map $ λ j, a i j), + let b_flat := (list.fin_range m).bind (λ i, (list.fin_range n).map $ λ j, b i j), + let args := [α, has_mul_α, add_comm_monoid_α] ++ a_flat ++ b_flat, + + -- build the matrices out of the coefficients + let A := matrix.fin_to_pexpr (matrix.map a to_pexpr), + let B := matrix.fin_to_pexpr (matrix.map b to_pexpr), + -- get an instance cache holding all the instances needed for matrix multiplication. There must + -- be a better way to do this. + t ← tactic.mk_instance_cache α, + has_add_α ← tactic.mk_app `has_add [α] >>= (λ t, prod.snd <$> @tactic.solve_aux unit t (do + { tmp2 ← tactic.pose `_inst_2' none add_comm_monoid_α, + tactic.reset_instance_cache, + tactic.apply_instance })), + has_zero_α ← tactic.mk_app `has_zero [α] >>= (λ t, prod.snd <$> @tactic.solve_aux unit t (do + { tmp2 ← tactic.pose `_inst_2' none add_comm_monoid_α, + tactic.reset_instance_cache, + tactic.apply_instance })), + let t := {inst := [ + (`has_mul, has_mul_α), + (`has_add, has_add_α), + (`has_zero, has_zero_α), + (`add_comm_monoid, add_comm_monoid_α)].foldl (λ n x, n.insert x.1 x.2) t.inst, + ..t}, + + -- clever trick: create algebraic instances on `expr` so that we can use `matrix.mul` or + -- `matrix.mulᵣ` to build the expression we want to end up with. It doesn't matter which we pick, + -- but the typeclasses are easier to create for the latter. + (t, has_mul_αe) ← expr.has_mul t, + (t, has_add_αe) ← expr.has_add t, + (t, has_zero_αe) ← expr.has_zero t, + let ab := @matrix.mulᵣ _ _ _ _ has_mul_αe has_add_αe has_zero_αe a b, + let AB := matrix.fin_to_pexpr (matrix.map ab to_pexpr), + + -- State and prove the equality, noting the RHS is defeq to `mulᵣ A B`. + A_eq ← tactic.to_expr ``(@matrix.mul _ _ _ _ _ %%has_mul_α %%add_comm_monoid_α %%A %%B = %%AB), + t ← tactic.pis args A_eq, + let pr := (expr.const `matrix.of_mul_of_fin_aux [u]).mk_app [`(l), `(m), `(n)], + -- This seems to create a metavariable then assign it, which ensures `pr` carries the right type. + ((), pr) ← tactic.solve_aux t $ tactic.exact pr, + + pure (t, pr) + +open_locale matrix + + +/-- Helper tactic used as an `auto_param` for `matrix.of_mul_of_fin` -/ +meta def of_mul_of_fin.derive : tactic unit := +do + target@`(@matrix.mul (fin %%l) (fin %%m) (fin %%n) %%α %%_ %%i1 %%i2 %%A %%B = %%AB) + ← tactic.target, + some (l, m, n) ← pure (prod.mk <$> l.to_nat <*> (prod.mk <$> m.to_nat <*> n.to_nat)) | + fail!"Dimensions {l}, {m} {n} are not numerals", + (t,pr) ← of_mul_of_fin.prove l m n, + tactic.apply (pr α i1 i2) {}, + tactic.done + -- TODO: should we be extracting the coefficients manually so we can do a full invocation as + -- something like: + -- tactic.unify target (t.instantiate_pis [α, A']), + -- tactic.exact (pr α A') + +/-- This lemma assumes that `a_coeffs` and `b_coeffs` refer to expressions of the form +`![![x₀₀, x₀₁], ![x₁₀, x₁₁]]`. It then uses an `auto_param` to populate `ab_coeffs` with an +expression of the same form, containing the appropriate expressions in terms of `+`, `*`, `aᵢⱼ`, +and `bⱼₖ`. -/ +theorem of_mul_of_fin {α} [has_mul α] [add_comm_monoid α] {l m n : ℕ} + {a_coeffs : fin l → fin m → α} + {b_coeffs : fin m → fin n → α} + {ab_coeffs : fin l → fin n → α} + (h : of a_coeffs ⬝ of b_coeffs = of ab_coeffs . of_mul_of_fin.derive) : + of a_coeffs ⬝ of b_coeffs = of ab_coeffs := h + +end of_mul_of_fin + +end matrix diff --git a/test/matrix_reflection.lean b/test/matrix_reflection.lean index 331d0a5573bfc..a734b8979715a 100644 --- a/test/matrix_reflection.lean +++ b/test/matrix_reflection.lean @@ -1,3 +1,4 @@ +import data.matrix.auto import data.matrix.reflection variables {α: Type*} @@ -17,3 +18,23 @@ example [add_comm_monoid α] [has_mul α] a₂₁*b₁₁ + a₂₂*b₂₁ + a₂₃*b₃₁, a₂₁*b₁₂ + a₂₂*b₂₂ + a₂₃*b₃₂, a₂₁*b₁₃ + a₂₂*b₂₃ + a₂₃*b₃₃; a₃₁*b₁₁ + a₃₂*b₂₁ + a₃₃*b₃₁, a₃₁*b₁₂ + a₃₂*b₂₂ + a₃₃*b₃₂, a₃₁*b₁₃ + a₃₂*b₂₃ + a₃₃*b₃₃] := (matrix.mulᵣ_eq _ _).symm + + +example {α} [add_comm_monoid α] [has_mul α] (a₁₁ a₁₂ a₂₁ a₂₂ b₁₁ b₁₂ b₂₁ b₂₂ : α) : + !![a₁₁, a₁₂; + a₂₁, a₂₂] ⬝ !![b₁₁, b₁₂; + b₂₁, b₂₂] = !![a₁₁ * b₁₁ + a₁₂ * b₂₁, a₁₁ * b₁₂ + a₁₂ * b₂₂; + a₂₁ * b₁₁ + a₂₂ * b₂₁, a₂₁ * b₁₂ + a₂₂ * b₂₂] := +begin + rw of_mul_of_fin, +end + +example {α} [add_comm_monoid α] [has_mul α] (a₁₁ a₁₂ b₁₁ b₁₂ b₂₁ b₂₂ : α) : + !![a₁₁, a₁₂] ⬝ !![b₁₁, b₁₂; + b₂₁, b₂₂] = !![a₁₁ * b₁₁ + a₁₂ * b₂₁, a₁₁ * b₁₂ + a₁₂ * b₂₂;] := +begin + -- if we really need it, we can get the proof directly like this + of_mul_of_fin.prove 1 2 2 >>= function.uncurry (tactic.assertv `h), + specialize @h α _ _, + rw of_mul_of_fin +end From c1acdccd694b692db2619fff903e0e40de428169 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sat, 15 Jul 2023 01:08:20 +0000 Subject: [PATCH 24/61] fix(tactic/norm_num): Do not allow typos in simp lemmas if they are unused (#17981) Previously ```lean import tactic.norm_num example : 1 + 1 = 2 := by norm_num [nonsense] ``` was legal. More realistically, if a lemma was added to the simp set then renamed, `norm_num [old_name]` would keep working even if the lemma no longer existed. This only seemed to happen once in mathlib, in the tests. It feels a bit silly to build the lemma set only to discard it, but we already rebuild the same lemma set on every iteration of the `repeat1` --- src/tactic/norm_num.lean | 9 ++++++--- test/norm_num.lean | 2 ++ test/norm_num_ext.lean | 2 +- 3 files changed, 9 insertions(+), 4 deletions(-) diff --git a/src/tactic/norm_num.lean b/src/tactic/norm_num.lean index eaeb50dd01860..27753225ea3be 100644 --- a/src/tactic/norm_num.lean +++ b/src/tactic/norm_num.lean @@ -1360,9 +1360,12 @@ use `get_step` to get the default `norm_num` set and `derive.step` for the basic simplifications. -/ meta def tactic.norm_num (step : expr → tactic (expr × expr)) (hs : list simp_arg_type) (l : interactive.loc) : tactic unit := -repeat1 $ orelse' (tactic.norm_num1 step l) $ -interactive.simp_core {} (tactic.norm_num1 step (interactive.loc.ns [none])) - ff (simp_arg_type.except ``one_div :: hs) [] l >> skip +do + -- Build and discard the simp lemma set, to validate it. + mk_simp_set_core ff [] (simp_arg_type.except ``one_div :: hs) tt, + repeat1 $ orelse' (tactic.norm_num1 step l) $ + interactive.simp_core {} (tactic.norm_num1 step (interactive.loc.ns [none])) + ff (simp_arg_type.except ``one_div :: hs) [] l >> skip /-- Carry out similar operations as `tactic.norm_num` but on an `expr` rather than a location. Given an expression `e`, returns `(e', ⊢ e = e')`. diff --git a/test/norm_num.lean b/test/norm_num.lean index d2b98fdec5b84..97dba721cd51b 100644 --- a/test/norm_num.lean +++ b/test/norm_num.lean @@ -304,6 +304,8 @@ example : (- ((- (((66 - 86) - 36) / 94) - 3) / - - (77 / (56 - - - 79))) + 87) example : 2 ^ 13 - 1 = int.of_nat 8191 := by norm_num +example : 1 + 1 = 2 := by success_if_fail { norm_num [this_doesnt_exist] }; refl + -- `^` and `•` do not have to match `monoid.has_pow` and `add_monoid.has_smul` syntactically example {α} [ring α] : (2 ^ 3 : ℕ → α) = 8 := by norm_num example {α} [ring α] : (2 • 3 : ℕ → α) = 6 := by norm_num diff --git a/test/norm_num_ext.lean b/test/norm_num_ext.lean index 96648a919b00d..5f75b9d9e5751 100644 --- a/test/norm_num_ext.lean +++ b/test/norm_num_ext.lean @@ -296,7 +296,7 @@ open_locale big_operators example : ([1, 2, 1, 3]).sum = 7 := by norm_num [-list.sum_cons] example : (([1, 2, 1, 3] : list ℚ).map (λ i, i^2)).sum = 15 := by norm_num [-list.map] example : (list.range 10).sum = 45 := by norm_num [-list.range_succ] -example : (list.fin_range 10).sum = 45 := by norm_num [-list.fin_range_succ] +example : (list.fin_range 10).sum = 45 := by norm_num [-list.fin_range_succ_eq_map] -- Multisets: example : (1 ::ₘ 2 ::ₘ 1 ::ₘ 3 ::ₘ {}).sum = 7 := by norm_num [-multiset.sum_cons] From 2fe465deb81bcd7ccafa065bb686888a82f15372 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Sat, 15 Jul 2023 12:19:30 +0000 Subject: [PATCH 25/61] chore(*): add mathlib4 synchronization comments (#19231) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.big_operators.norm_num` * `algebraic_geometry.projective_spectrum.scheme` * `analysis.analytic.inverse` * `analysis.inner_product_space.of_norm` * `analysis.normed.group.SemiNormedGroup.completion` * `category_theory.monad.monadicity` * `data.buffer.parser.basic` * `data.buffer.parser.numeral` * `data.fintype.array` * `data.hash_map` * `geometry.manifold.algebra.left_invariant_derivation` * `geometry.manifold.cont_mdiff_mfderiv` * `geometry.manifold.vector_bundle.hom` * `linear_algebra.clifford_algebra.even_equiv` * `ring_theory.etale` * `ring_theory.kaehler` * `tactic.group` * `testing.slim_check.functions` --- src/algebra/big_operators/norm_num.lean | 3 +++ src/algebraic_geometry/projective_spectrum/scheme.lean | 3 +++ src/analysis/analytic/inverse.lean | 3 +++ src/analysis/inner_product_space/of_norm.lean | 3 +++ src/analysis/normed/group/SemiNormedGroup/completion.lean | 3 +++ src/category_theory/monad/monadicity.lean | 3 +++ src/data/buffer/parser/basic.lean | 3 +++ src/data/buffer/parser/numeral.lean | 3 +++ src/data/fintype/array.lean | 3 +++ src/data/hash_map.lean | 3 +++ src/geometry/manifold/algebra/left_invariant_derivation.lean | 3 +++ src/geometry/manifold/cont_mdiff_mfderiv.lean | 3 +++ src/geometry/manifold/vector_bundle/hom.lean | 3 +++ src/linear_algebra/clifford_algebra/even_equiv.lean | 3 +++ src/ring_theory/etale.lean | 3 +++ src/ring_theory/kaehler.lean | 3 +++ src/tactic/group.lean | 3 +++ src/testing/slim_check/functions.lean | 3 +++ 18 files changed, 54 insertions(+) diff --git a/src/algebra/big_operators/norm_num.lean b/src/algebra/big_operators/norm_num.lean index a4067b58b3425..d950f11219ca6 100644 --- a/src/algebra/big_operators/norm_num.lean +++ b/src/algebra/big_operators/norm_num.lean @@ -9,6 +9,9 @@ import tactic.norm_num /-! ### `norm_num` plugin for big operators +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This `norm_num` plugin provides support for computing sums and products of lists, multisets and finsets. diff --git a/src/algebraic_geometry/projective_spectrum/scheme.lean b/src/algebraic_geometry/projective_spectrum/scheme.lean index 4bfbcf04dcd7e..4fc52b338df17 100644 --- a/src/algebraic_geometry/projective_spectrum/scheme.lean +++ b/src/algebraic_geometry/projective_spectrum/scheme.lean @@ -10,6 +10,9 @@ import ring_theory.graded_algebra.radical /-! # Proj as a scheme +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file is to prove that `Proj` is a scheme. ## Notation diff --git a/src/analysis/analytic/inverse.lean b/src/analysis/analytic/inverse.lean index 60268c17f70fd..d2c1a11a873ea 100644 --- a/src/analysis/analytic/inverse.lean +++ b/src/analysis/analytic/inverse.lean @@ -9,6 +9,9 @@ import tactic.congrm # Inverse of analytic functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We construct the left and right inverse of a formal multilinear series with invertible linear term, we prove that they coincide and study their properties (notably convergence). diff --git a/src/analysis/inner_product_space/of_norm.lean b/src/analysis/inner_product_space/of_norm.lean index be0bf32167923..7c36dba00df14 100644 --- a/src/analysis/inner_product_space/of_norm.lean +++ b/src/analysis/inner_product_space/of_norm.lean @@ -10,6 +10,9 @@ import analysis.inner_product_space.basic /-! # Inner product space derived from a norm +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines an `inner_product_space` instance from a norm that respects the parallellogram identity. The parallelogram identity is a way to express the inner product of `x` and `y` in terms of the norms of `x`, `y`, `x + y`, `x - y`. diff --git a/src/analysis/normed/group/SemiNormedGroup/completion.lean b/src/analysis/normed/group/SemiNormedGroup/completion.lean index 27a0a2315624c..8ffe9b3cff21e 100644 --- a/src/analysis/normed/group/SemiNormedGroup/completion.lean +++ b/src/analysis/normed/group/SemiNormedGroup/completion.lean @@ -10,6 +10,9 @@ import analysis.normed.group.hom_completion /-! # Completions of normed groups +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains an API for completions of seminormed groups (basic facts about objects and morphisms). diff --git a/src/category_theory/monad/monadicity.lean b/src/category_theory/monad/monadicity.lean index 8a15ac2fb4247..3caeaa2ad27b2 100644 --- a/src/category_theory/monad/monadicity.lean +++ b/src/category_theory/monad/monadicity.lean @@ -12,6 +12,9 @@ import category_theory.monad.limits /-! # Monadicity theorems +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We prove monadicity theorems which can establish a given functor is monadic. In particular, we show three versions of Beck's monadicity theorem, and the reflexive (crude) monadicity theorem: diff --git a/src/data/buffer/parser/basic.lean b/src/data/buffer/parser/basic.lean index 4c9b11621f64b..30f1b1f17d5ef 100644 --- a/src/data/buffer/parser/basic.lean +++ b/src/data/buffer/parser/basic.lean @@ -11,6 +11,9 @@ import data.buffer.parser /-! # Parsers +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + `parser α` is the type that describes a computation that can ingest a `char_buffer` and output, if successful, a term of type `α`. This file expands on the definitions in the core library, proving that all the core library diff --git a/src/data/buffer/parser/numeral.lean b/src/data/buffer/parser/numeral.lean index af91eeba878d2..050b505ee8b8f 100644 --- a/src/data/buffer/parser/numeral.lean +++ b/src/data/buffer/parser/numeral.lean @@ -8,6 +8,9 @@ import data.buffer.parser.basic /-! # Numeral parsers +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file expands on the existing `nat : parser ℕ` to provide parsers into any type `α` that can be represented by a numeral, which relies on `α` having a 0, 1, and addition operation. There are also convenience parsers that ensure that the numeral parsed in is not larger than diff --git a/src/data/fintype/array.lean b/src/data/fintype/array.lean index 5211c8cebc957..134d9d2837d71 100644 --- a/src/data/fintype/array.lean +++ b/src/data/fintype/array.lean @@ -8,6 +8,9 @@ import logic.equiv.array /-! # `array n α` is a fintype when `α` is. + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ variables {α : Type*} diff --git a/src/data/hash_map.lean b/src/data/hash_map.lean index e9293d372477a..1c105b600c392 100644 --- a/src/data/hash_map.lean +++ b/src/data/hash_map.lean @@ -12,6 +12,9 @@ import data.pnat.defs /-! # Hash maps +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Defines a hash map data structure, representing a finite key-value map with a value type that may depend on the key type. The structure requires a `nat`-valued hash function to associate keys to buckets. diff --git a/src/geometry/manifold/algebra/left_invariant_derivation.lean b/src/geometry/manifold/algebra/left_invariant_derivation.lean index 6aab779498d14..eb5dd7aff2692 100644 --- a/src/geometry/manifold/algebra/left_invariant_derivation.lean +++ b/src/geometry/manifold/algebra/left_invariant_derivation.lean @@ -11,6 +11,9 @@ import geometry.manifold.derivation_bundle # Left invariant derivations +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define the concept of left invariant derivation for a Lie group. The concept is analogous to the more classical concept of left invariant vector fields, and it holds that the derivation associated to a vector field is left invariant iff the field is. diff --git a/src/geometry/manifold/cont_mdiff_mfderiv.lean b/src/geometry/manifold/cont_mdiff_mfderiv.lean index 9efd7835a3bd5..89abe55046496 100644 --- a/src/geometry/manifold/cont_mdiff_mfderiv.lean +++ b/src/geometry/manifold/cont_mdiff_mfderiv.lean @@ -9,6 +9,9 @@ import geometry.manifold.cont_mdiff_map /-! ### Interactions between differentiability, smoothness and manifold derivatives +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We give the relation between `mdifferentiable`, `cont_mdiff`, `mfderiv`, `tangent_map` and related notions. diff --git a/src/geometry/manifold/vector_bundle/hom.lean b/src/geometry/manifold/vector_bundle/hom.lean index 9dff42177cf64..ef6b4079edcfb 100644 --- a/src/geometry/manifold/vector_bundle/hom.lean +++ b/src/geometry/manifold/vector_bundle/hom.lean @@ -8,6 +8,9 @@ import topology.vector_bundle.hom /-! # Homs of smooth vector bundles over the same base space +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Here we show that `bundle.continuous_linear_map` is a smooth vector bundle. Note that we only do this for bundles of linear maps, not for bundles of arbitrary semilinear maps. diff --git a/src/linear_algebra/clifford_algebra/even_equiv.lean b/src/linear_algebra/clifford_algebra/even_equiv.lean index caf1ccefbe519..55e8ea60c4c02 100644 --- a/src/linear_algebra/clifford_algebra/even_equiv.lean +++ b/src/linear_algebra/clifford_algebra/even_equiv.lean @@ -9,6 +9,9 @@ import linear_algebra.quadratic_form.prod /-! # Isomorphisms with the even subalgebra of a Clifford algebra +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file provides some notable isomorphisms regarding the even subalgebra, `clifford_algebra.even`. ## Main definitions diff --git a/src/ring_theory/etale.lean b/src/ring_theory/etale.lean index 09b4302def2a3..93367714fc187 100644 --- a/src/ring_theory/etale.lean +++ b/src/ring_theory/etale.lean @@ -10,6 +10,9 @@ import ring_theory.kaehler # Formally étale morphisms +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + An `R`-algebra `A` is formally étale (resp. unramified, smooth) if for every `R`-algebra, every square-zero ideal `I : ideal B` and `f : A →ₐ[R] B ⧸ I`, there exists exactly (resp. at most, at least) one lift `A →ₐ[R] B`. diff --git a/src/ring_theory/kaehler.lean b/src/ring_theory/kaehler.lean index 14eedac487b93..ec3b0a3079dc8 100644 --- a/src/ring_theory/kaehler.lean +++ b/src/ring_theory/kaehler.lean @@ -11,6 +11,9 @@ import ring_theory.is_tensor_product /-! # The module of kaehler differentials +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main results - `kaehler_differential`: The module of kaehler differentials. For an `R`-algebra `S`, we provide diff --git a/src/tactic/group.lean b/src/tactic/group.lean index 1cfe207169a8e..e9603a022658f 100644 --- a/src/tactic/group.lean +++ b/src/tactic/group.lean @@ -10,6 +10,9 @@ import algebra.group.commutator /-! # `group` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Normalizes expressions in the language of groups. The basic idea is to use the simplifier to put everything into a product of group powers (`zpow` which takes a group element and an integer), then simplify the exponents using the `ring` tactic. The process needs to be repeated diff --git a/src/testing/slim_check/functions.lean b/src/testing/slim_check/functions.lean index b70e3757746c6..7249c61434dfe 100644 --- a/src/testing/slim_check/functions.lean +++ b/src/testing/slim_check/functions.lean @@ -14,6 +14,9 @@ import testing.slim_check.testable /-! ## `slim_check`: generators for functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines `sampleable` instances for `α → β` functions and `ℤ → ℤ` injective functions. From 31b269b60935483943542d547a6dd83a66b37dc7 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Sun, 16 Jul 2023 07:50:31 +0000 Subject: [PATCH 26/61] feat(set_theory/ordinal/natural_ops): define natural multiplication (#14324) --- src/set_theory/ordinal/arithmetic.lean | 14 + src/set_theory/ordinal/natural_ops.lean | 352 ++++++++++++++++++++++-- src/set_theory/ordinal/principal.lean | 22 +- 3 files changed, 355 insertions(+), 33 deletions(-) diff --git a/src/set_theory/ordinal/arithmetic.lean b/src/set_theory/ordinal/arithmetic.lean index a860ab829134a..daeab4abbea76 100644 --- a/src/set_theory/ordinal/arithmetic.lean +++ b/src/set_theory/ordinal/arithmetic.lean @@ -1526,6 +1526,20 @@ theorem is_normal.eq_iff_zero_and_succ {f g : ordinal.{u} → ordinal.{u}} (hf : exact H b hb end)⟩ +/-- A two-argument version of `ordinal.blsub`. + +We don't develop a full API for this, since it's only used in a handful of existence results. -/ +def blsub₂ (o₁ o₂ : ordinal) (op : Π (a < o₁) (b < o₂), ordinal) : ordinal := +lsub (λ x : o₁.out.α × o₂.out.α, + op (typein (<) x.1) (typein_lt_self _) (typein (<) x.2) (typein_lt_self _)) + +theorem lt_blsub₂ {o₁ o₂ : ordinal} (op : Π (a < o₁) (b < o₂), ordinal) {a b : ordinal} + (ha : a < o₁) (hb : b < o₂) : op a ha b hb < blsub₂ o₁ o₂ op := +begin + convert lt_lsub _ (prod.mk (enum (<) a (by rwa type_lt)) (enum (<) b (by rwa type_lt))), + simp only [typein_enum] +end + /-! ### Minimum excluded ordinals -/ /-- The minimum excluded ordinal in a family of ordinals. -/ diff --git a/src/set_theory/ordinal/natural_ops.lean b/src/set_theory/ordinal/natural_ops.lean index ff6f1e849beff..38f6a566ec06e 100644 --- a/src/set_theory/ordinal/natural_ops.lean +++ b/src/set_theory/ordinal/natural_ops.lean @@ -5,6 +5,7 @@ Authors: Violeta Hernández Palacios -/ import set_theory.ordinal.arithmetic +import tactic.abel /-! # Natural operations on ordinals @@ -38,7 +39,6 @@ between both types, we attempt to prove and state most results on `ordinal`. # Todo -- Define natural multiplication and provide a basic API. - Prove the characterizations of natural addition and multiplication in terms of the Cantor normal form. -/ @@ -49,6 +49,8 @@ open function order noncomputable theory +/-! ### Basic casts between `ordinal` and `nat_ordinal` -/ + /-- A type synonym for ordinals with natural addition and multiplication. -/ @[derive [has_zero, inhabited, has_one, linear_order, succ_order, has_well_founded]] def nat_ordinal : Type* := ordinal @@ -59,10 +61,10 @@ def nat_ordinal : Type* := ordinal /-- The identity function between `nat_ordinal` and `ordinal`. -/ @[pattern] def nat_ordinal.to_ordinal : nat_ordinal ≃o ordinal := order_iso.refl _ -open ordinal - namespace nat_ordinal +open ordinal + variables {a b c : nat_ordinal.{u}} @[simp] theorem to_ordinal_symm_eq : nat_ordinal.to_ordinal.symm = ordinal.to_nat_ordinal := rfl @@ -95,8 +97,6 @@ end nat_ordinal namespace ordinal -variables {a b c : ordinal.{u}} - @[simp] theorem to_nat_ordinal_symm_eq : to_nat_ordinal.symm = nat_ordinal.to_ordinal := rfl @[simp] theorem to_nat_ordinal_to_ordinal (a : ordinal) : a.to_nat_ordinal.to_ordinal = a := rfl @@ -106,11 +106,14 @@ variables {a b c : ordinal.{u}} @[simp] theorem to_nat_ordinal_eq_zero (a) : to_nat_ordinal a = 0 ↔ a = 0 := iff.rfl @[simp] theorem to_nat_ordinal_eq_one (a) : to_nat_ordinal a = 1 ↔ a = 1 := iff.rfl -@[simp] theorem to_nat_ordinal_max : +@[simp] theorem to_nat_ordinal_max (a b : ordinal) : to_nat_ordinal (max a b) = max a.to_nat_ordinal b.to_nat_ordinal := rfl -@[simp] theorem to_nat_ordinal_min : +@[simp] theorem to_nat_ordinal_min (a b : ordinal) : (linear_order.min a b).to_nat_ordinal = linear_order.min a.to_nat_ordinal b.to_nat_ordinal := rfl +/-! We place the definitions of `nadd` and `nmul` before actually developing their API, as this +guarantees we only need to open the `natural_ops` locale once. -/ + /-- Natural addition on ordinals `a ♯ b`, also known as the Hessenberg sum, is recursively defined as the least ordinal greater than `a' ♯ b` and `a ♯ b'` for all `a' < a` and `b' < b`. In contrast to normal ordinal addition, it is commutative. @@ -125,6 +128,30 @@ using_well_founded { dec_tac := `[solve_by_elim [psigma.lex.left, psigma.lex.rig localized "infix (name := ordinal.nadd) ` ♯ `:65 := ordinal.nadd" in natural_ops +/-- Natural multiplication on ordinals `a ⨳ b`, also known as the Hessenberg product, is recursively +defined as the least ordinal such that `a ⨳ b + a' ⨳ b'` is greater than `a' ⨳ b + a ⨳ b'` for all +`a' < a` and `b < b'`. In contrast to normal ordinal multiplication, it is commutative and +distributive (over natural addition). + +Natural multiplication can equivalently be characterized as the ordinal resulting from multiplying +the Cantor normal forms of `a` and `b` as if they were polynomials in `ω`. Addition of exponents is +done via natural addition. -/ +noncomputable def nmul : ordinal.{u} → ordinal.{u} → ordinal.{u} +| a b := Inf {c | ∀ (a' < a) (b' < b), nmul a' b ♯ nmul a b' < c ♯ nmul a' b'} +using_well_founded { dec_tac := `[solve_by_elim [psigma.lex.left, psigma.lex.right]] } + +localized "infix ` ⨳ `:70 := ordinal.nmul" in natural_ops + +end ordinal + +open_locale natural_ops + +/-! ### Natural addition -/ + +namespace ordinal + +variables {a b c : ordinal.{u}} + theorem nadd_def (a b : ordinal) : a ♯ b = max (blsub.{u u} a $ λ a' h, a' ♯ b) (blsub.{u u} b $ λ b' h, a ♯ b') := @@ -243,10 +270,10 @@ end end ordinal -open ordinal - namespace nat_ordinal +open ordinal + instance : has_add nat_ordinal := ⟨nadd⟩ instance add_covariant_class_lt : @@ -280,21 +307,19 @@ instance : add_monoid_with_one nat_ordinal := add_monoid_with_one.unary begin induction n with n hn, { refl }, - { change nadd (to_ordinal n) 1 = n + 1, - rw hn, - apply nadd_one } + { change to_ordinal n ♯ 1 = n + 1, + rw hn, exact nadd_one n } end end nat_ordinal -open nat_ordinal - -open_locale natural_ops - namespace ordinal +theorem nadd_eq_add (a b : ordinal) : a ♯ b = (a.to_nat_ordinal + b.to_nat_ordinal).to_ordinal := +rfl + @[simp] theorem to_nat_ordinal_cast_nat (n : ℕ) : to_nat_ordinal n = n := -by { rw ←to_ordinal_cast_nat n, refl } +by { rw ←nat_ordinal.to_ordinal_cast_nat n, refl } theorem lt_of_nadd_lt_nadd_left : ∀ {a b c}, a ♯ b < a ♯ c → b < c := @lt_of_add_lt_add_left nat_ordinal _ _ _ @@ -314,6 +339,16 @@ theorem nadd_le_nadd_iff_left : ∀ a {b c}, a ♯ b ≤ a ♯ c ↔ b ≤ c := theorem nadd_le_nadd_iff_right : ∀ a {b c}, b ♯ a ≤ c ♯ a ↔ b ≤ c := @_root_.add_le_add_iff_right nat_ordinal _ _ _ _ +theorem nadd_le_nadd : ∀ {a b c d}, a ≤ b → c ≤ d → a ♯ c ≤ b ♯ d := +@add_le_add nat_ordinal _ _ _ _ +theorem nadd_lt_nadd : ∀ {a b c d}, a < b → c < d → a ♯ c < b ♯ d := +@add_lt_add nat_ordinal _ _ _ _ + +theorem nadd_lt_nadd_of_lt_of_le : ∀ {a b c d}, a < b → c ≤ d → a ♯ c < b ♯ d := +@add_lt_add_of_lt_of_le nat_ordinal _ _ _ _ +theorem nadd_lt_nadd_of_le_of_lt : ∀ {a b c d}, a ≤ b → c < d → a ♯ c < b ♯ d := +@add_lt_add_of_le_of_lt nat_ordinal _ _ _ _ + theorem nadd_left_cancel : ∀ {a b c}, a ♯ b = a ♯ c → b = c := @_root_.add_left_cancel nat_ordinal _ _ theorem nadd_right_cancel : ∀ {a b c}, a ♯ b = c ♯ b → a = c := @@ -323,4 +358,287 @@ theorem nadd_left_cancel_iff : ∀ {a b c}, a ♯ b = a ♯ c ↔ b = c := theorem nadd_right_cancel_iff : ∀ {a b c}, b ♯ a = c ♯ a ↔ b = c := @add_right_cancel_iff nat_ordinal _ _ +theorem le_nadd_self {a b} : a ≤ b ♯ a := +by simpa using nadd_le_nadd_right (ordinal.zero_le b) a +theorem le_nadd_left {a b c} (h : a ≤ c) : a ≤ b ♯ c := +le_nadd_self.trans (nadd_le_nadd_left h b) +theorem le_self_nadd {a b} : a ≤ a ♯ b := +by simpa using nadd_le_nadd_left (ordinal.zero_le b) a +theorem le_nadd_right {a b c} (h : a ≤ b) : a ≤ b ♯ c := +le_self_nadd.trans (nadd_le_nadd_right h c) + +theorem nadd_left_comm : ∀ a b c, a ♯ (b ♯ c) = b ♯ (a ♯ c) := +@add_left_comm nat_ordinal _ +theorem nadd_right_comm : ∀ a b c, a ♯ b ♯ c = a ♯ c ♯ b := +@add_right_comm nat_ordinal _ + +/-! ### Natural multiplication -/ + +variables {a b c d : ordinal.{u}} + +theorem nmul_def (a b : ordinal) : + a ⨳ b = Inf {c | ∀ (a' < a) (b' < b), a' ⨳ b ♯ a ⨳ b' < c ♯ a' ⨳ b'} := +by rw nmul + +/-- The set in the definition of `nmul` is nonempty. -/ +theorem nmul_nonempty (a b : ordinal.{u}) : + {c : ordinal.{u} | ∀ (a' < a) (b' < b), a' ⨳ b ♯ a ⨳ b' < c ♯ a' ⨳ b'}.nonempty := +⟨_, λ a' ha b' hb, (lt_blsub₂.{u u u} _ ha hb).trans_le le_self_nadd⟩ + +theorem nmul_nadd_lt {a' b' : ordinal} (ha : a' < a) (hb : b' < b) : + a' ⨳ b ♯ a ⨳ b' < a ⨳ b ♯ a' ⨳ b' := +by { rw nmul_def a b, exact Inf_mem (nmul_nonempty a b) a' ha b' hb } + +theorem nmul_nadd_le {a' b' : ordinal} (ha : a' ≤ a) (hb : b' ≤ b) : + a' ⨳ b ♯ a ⨳ b' ≤ a ⨳ b ♯ a' ⨳ b' := +begin + rcases lt_or_eq_of_le ha with ha | rfl, + { rcases lt_or_eq_of_le hb with hb | rfl, + { exact (nmul_nadd_lt ha hb).le }, + { rw nadd_comm } }, + { exact le_rfl } +end + +theorem lt_nmul_iff : c < a ⨳ b ↔ ∃ (a' < a) (b' < b), c ♯ a' ⨳ b' ≤ a' ⨳ b ♯ a ⨳ b' := +begin + refine ⟨λ h, _, _⟩, + { rw nmul at h, + simpa using not_mem_of_lt_cInf h ⟨0, λ _ _, bot_le⟩ }, + { rintros ⟨a', ha, b', hb, h⟩, + have := h.trans_lt (nmul_nadd_lt ha hb), + rwa nadd_lt_nadd_iff_right at this } +end + +theorem nmul_le_iff : a ⨳ b ≤ c ↔ ∀ (a' < a) (b' < b), a' ⨳ b ♯ a ⨳ b' < c ♯ a' ⨳ b' := +by { rw ←not_iff_not, simp [lt_nmul_iff] } + +theorem nmul_comm : ∀ (a b), a ⨳ b = b ⨳ a +| a b := begin + rw [nmul, nmul], + congr, ext x, split; + intros H c hc d hd, + { rw [nadd_comm, ←nmul_comm, ←nmul_comm a, ←nmul_comm d], + exact H _ hd _ hc }, + { rw [nadd_comm, nmul_comm, nmul_comm c, nmul_comm c], + exact H _ hd _ hc } +end +using_well_founded { dec_tac := `[solve_by_elim [psigma.lex.left, psigma.lex.right]] } + +@[simp] theorem nmul_zero (a) : a ⨳ 0 = 0 := +by { rw [←ordinal.le_zero, nmul_le_iff], exact λ _ _ a ha, (ordinal.not_lt_zero a ha).elim } + +@[simp] theorem zero_nmul (a) : 0 ⨳ a = 0 := +by rw [nmul_comm, nmul_zero] + +@[simp] theorem nmul_one : ∀ a, a ⨳ 1 = a +| a := begin + rw nmul, + simp only [lt_one_iff_zero, forall_eq, nmul_zero, nadd_zero], + convert cInf_Ici, + ext b, + refine ⟨λ H, le_of_forall_lt (λ c hc, _), λ ha c hc, _⟩, + { simpa only [nmul_one] using H c hc }, + { simpa only [nmul_one] using hc.trans_le ha } +end +using_well_founded { dec_tac := `[assumption] } + +@[simp] theorem one_nmul (a) : 1 ⨳ a = a := +by rw [nmul_comm, nmul_one] + +theorem nmul_lt_nmul_of_pos_left (h₁ : a < b) (h₂ : 0 < c) : c ⨳ a < c ⨳ b := +lt_nmul_iff.2 ⟨0, h₂, a, h₁, by simp⟩ + +theorem nmul_lt_nmul_of_pos_right (h₁ : a < b) (h₂ : 0 < c) : a ⨳ c < b ⨳ c := +lt_nmul_iff.2 ⟨a, h₁, 0, h₂, by simp⟩ + +theorem nmul_le_nmul_of_nonneg_left (h₁ : a ≤ b) (h₂ : 0 ≤ c) : c ⨳ a ≤ c ⨳ b := +begin + rcases lt_or_eq_of_le h₁ with h₁|rfl; + rcases lt_or_eq_of_le h₂ with h₂|rfl, + { exact (nmul_lt_nmul_of_pos_left h₁ h₂).le }, + all_goals { simp } +end + +theorem nmul_le_nmul_of_nonneg_right (h₁ : a ≤ b) (h₂ : 0 ≤ c) : a ⨳ c ≤ b ⨳ c := +begin + rw [nmul_comm, nmul_comm b], + exact nmul_le_nmul_of_nonneg_left h₁ h₂ +end + +theorem nmul_nadd : ∀ (a b c), a ⨳ (b ♯ c) = a ⨳ b ♯ a ⨳ c +| a b c := begin + apply le_antisymm (nmul_le_iff.2 $ λ a' ha d hd, _) (nadd_le_iff.2 ⟨λ d hd, _, λ d hd, _⟩), + { rw nmul_nadd, + rcases lt_nadd_iff.1 hd with ⟨b', hb, hd⟩ | ⟨c', hc, hd⟩, + { have := nadd_lt_nadd_of_lt_of_le (nmul_nadd_lt ha hb) (nmul_nadd_le ha.le hd), + rw [nmul_nadd, nmul_nadd] at this, + simp only [nadd_assoc] at this, + rwa [nadd_left_comm, nadd_left_comm _ (a ⨳ b'), nadd_left_comm (a ⨳ b), nadd_lt_nadd_iff_left, + nadd_left_comm (a' ⨳ b), nadd_left_comm (a ⨳ b), nadd_lt_nadd_iff_left, ←nadd_assoc, + ←nadd_assoc] at this }, + { have := nadd_lt_nadd_of_le_of_lt (nmul_nadd_le ha.le hd) (nmul_nadd_lt ha hc), + rw [nmul_nadd, nmul_nadd] at this, + simp only [nadd_assoc] at this, + rwa [nadd_left_comm, nadd_comm (a ⨳ c), nadd_left_comm (a' ⨳ d), nadd_left_comm (a ⨳ c'), + nadd_left_comm (a ⨳ b), nadd_lt_nadd_iff_left, nadd_comm (a' ⨳ c), nadd_left_comm (a ⨳ d), + nadd_left_comm (a' ⨳ b), nadd_left_comm (a ⨳ b), nadd_lt_nadd_iff_left, nadd_comm (a ⨳ d), + nadd_comm (a' ⨳ d), ←nadd_assoc, ←nadd_assoc] at this } }, + { rcases lt_nmul_iff.1 hd with ⟨a', ha, b', hb, hd⟩, + have := nadd_lt_nadd_of_le_of_lt hd (nmul_nadd_lt ha (nadd_lt_nadd_right hb c)), + rw [nmul_nadd, nmul_nadd, nmul_nadd a'] at this, + simp only [nadd_assoc] at this, + rwa [nadd_left_comm (a' ⨳ b'), nadd_left_comm, nadd_lt_nadd_iff_left, nadd_left_comm, + nadd_left_comm _ (a' ⨳ b'), nadd_left_comm (a ⨳ b'), nadd_lt_nadd_iff_left, + nadd_left_comm (a' ⨳ c), nadd_left_comm, nadd_lt_nadd_iff_left, nadd_left_comm, + nadd_comm _ (a' ⨳ c), nadd_lt_nadd_iff_left] at this }, + { rcases lt_nmul_iff.1 hd with ⟨a', ha, c', hc, hd⟩, + have := nadd_lt_nadd_of_lt_of_le (nmul_nadd_lt ha (nadd_lt_nadd_left hc b)) hd, + rw [nmul_nadd, nmul_nadd, nmul_nadd a'] at this, + simp only [nadd_assoc] at this, + rwa [nadd_left_comm _ (a' ⨳ b), nadd_lt_nadd_iff_left, nadd_left_comm (a' ⨳ c'), + nadd_left_comm _ (a' ⨳ c), nadd_lt_nadd_iff_left, nadd_left_comm, + nadd_comm (a' ⨳ c'), nadd_left_comm _ (a ⨳ c'), nadd_lt_nadd_iff_left, + nadd_comm _ (a' ⨳ c'), nadd_comm _ (a' ⨳ c'), nadd_left_comm, + nadd_lt_nadd_iff_left] at this } +end +using_well_founded { dec_tac := `[solve_by_elim [psigma.lex.left, psigma.lex.right]] } + +theorem nadd_nmul (a b c) : (a ♯ b) ⨳ c = a ⨳ c ♯ b ⨳ c := +by rw [nmul_comm, nmul_nadd, nmul_comm, nmul_comm c] + +theorem nmul_nadd_lt₃ {a' b' c' : ordinal} (ha : a' < a) (hb : b' < b) (hc : c' < c) : + a' ⨳ b ⨳ c ♯ a ⨳ b' ⨳ c ♯ a ⨳ b ⨳ c' ♯ a' ⨳ b' ⨳ c' < + a ⨳ b ⨳ c ♯ a' ⨳ b' ⨳ c ♯ a' ⨳ b ⨳ c' ♯ a ⨳ b' ⨳ c' := +by simpa only [nadd_nmul, ←nadd_assoc] using nmul_nadd_lt (nmul_nadd_lt ha hb) hc + +theorem nmul_nadd_le₃ {a' b' c' : ordinal} (ha : a' ≤ a) (hb : b' ≤ b) (hc : c' ≤ c) : + a' ⨳ b ⨳ c ♯ a ⨳ b' ⨳ c ♯ a ⨳ b ⨳ c' ♯ a' ⨳ b' ⨳ c' ≤ + a ⨳ b ⨳ c ♯ a' ⨳ b' ⨳ c ♯ a' ⨳ b ⨳ c' ♯ a ⨳ b' ⨳ c' := +by simpa only [nadd_nmul, ←nadd_assoc] using nmul_nadd_le (nmul_nadd_le ha hb) hc + +theorem nmul_nadd_lt₃' {a' b' c' : ordinal} (ha : a' < a) (hb : b' < b) (hc : c' < c) : + a' ⨳ (b ⨳ c) ♯ a ⨳ (b' ⨳ c) ♯ a ⨳ (b ⨳ c') ♯ a' ⨳ (b' ⨳ c') < + a ⨳ (b ⨳ c) ♯ a' ⨳ (b' ⨳ c) ♯ a' ⨳ (b ⨳ c') ♯ a ⨳ (b' ⨳ c') := +begin + simp only [nmul_comm _ (_ ⨳ _)], + convert nmul_nadd_lt₃ hb hc ha using 1; + { simp only [nadd_eq_add, nat_ordinal.to_ordinal_to_nat_ordinal], abel } +end + +theorem nmul_nadd_le₃' {a' b' c' : ordinal} (ha : a' ≤ a) (hb : b' ≤ b) (hc : c' ≤ c) : + a' ⨳ (b ⨳ c) ♯ a ⨳ (b' ⨳ c) ♯ a ⨳ (b ⨳ c') ♯ a' ⨳ (b' ⨳ c') ≤ + a ⨳ (b ⨳ c) ♯ a' ⨳ (b' ⨳ c) ♯ a' ⨳ (b ⨳ c') ♯ a ⨳ (b' ⨳ c') := +begin + simp only [nmul_comm _ (_ ⨳ _)], + convert nmul_nadd_le₃ hb hc ha using 1; + { simp only [nadd_eq_add, nat_ordinal.to_ordinal_to_nat_ordinal], abel } +end + +theorem lt_nmul_iff₃ : d < a ⨳ b ⨳ c ↔ ∃ (a' < a) (b' < b) (c' < c), + d ♯ a' ⨳ b' ⨳ c ♯ a' ⨳ b ⨳ c' ♯ a ⨳ b' ⨳ c' ≤ + a' ⨳ b ⨳ c ♯ a ⨳ b' ⨳ c ♯ a ⨳ b ⨳ c' ♯ a' ⨳ b' ⨳ c' := +begin + refine ⟨λ h, _, _⟩, + { rcases lt_nmul_iff.1 h with ⟨e, he, c', hc, H₁⟩, + rcases lt_nmul_iff.1 he with ⟨a', ha, b', hb, H₂⟩, + refine ⟨a', ha, b', hb, c', hc, _⟩, + have := nadd_le_nadd H₁ (nmul_nadd_le H₂ hc.le), + simp only [nadd_nmul, nadd_assoc] at this, + rw [nadd_left_comm, nadd_left_comm d, nadd_left_comm, nadd_le_nadd_iff_left, + nadd_left_comm (a ⨳ b' ⨳ c), nadd_left_comm (a' ⨳ b ⨳ c), nadd_left_comm (a ⨳ b ⨳ c'), + nadd_le_nadd_iff_left, nadd_left_comm (a ⨳ b ⨳ c'), nadd_left_comm (a ⨳ b ⨳ c')] at this, + simpa only [nadd_assoc] }, + { rintro ⟨a', ha, b', hb, c', hc, h⟩, + have := h.trans_lt (nmul_nadd_lt₃ ha hb hc), + repeat { rwa nadd_lt_nadd_iff_right at this } } +end + +theorem nmul_le_iff₃ : a ⨳ b ⨳ c ≤ d ↔ ∀ (a' < a) (b' < b) (c' < c), + a' ⨳ b ⨳ c ♯ a ⨳ b' ⨳ c ♯ a ⨳ b ⨳ c' ♯ a' ⨳ b' ⨳ c' < + d ♯ a' ⨳ b' ⨳ c ♯ a' ⨳ b ⨳ c' ♯ a ⨳ b' ⨳ c' := +by { rw ←not_iff_not, simp [lt_nmul_iff₃] } + +theorem lt_nmul_iff₃' : d < a ⨳ (b ⨳ c) ↔ ∃ (a' < a) (b' < b) (c' < c), + d ♯ a' ⨳ (b' ⨳ c) ♯ a' ⨳ (b ⨳ c') ♯ a ⨳ (b' ⨳ c') ≤ + a' ⨳ (b ⨳ c) ♯ a ⨳ (b' ⨳ c) ♯ a ⨳ (b ⨳ c') ♯ a' ⨳ (b' ⨳ c') := +begin + simp only [nmul_comm _ (_ ⨳ _), lt_nmul_iff₃, nadd_eq_add, nat_ordinal.to_ordinal_to_nat_ordinal], + split; rintro ⟨b', hb, c', hc, a', ha, h⟩, + { use [a', ha, b', hb, c', hc], convert h using 1; abel }, + { use [c', hc, a', ha, b', hb], convert h using 1; abel } +end + +theorem nmul_le_iff₃' : a ⨳ (b ⨳ c) ≤ d ↔ ∀ (a' < a) (b' < b) (c' < c), + a' ⨳ (b ⨳ c) ♯ a ⨳ (b' ⨳ c) ♯ a ⨳ (b ⨳ c') ♯ a' ⨳ (b' ⨳ c') < + d ♯ a' ⨳ (b' ⨳ c) ♯ a' ⨳ (b ⨳ c') ♯ a ⨳ (b' ⨳ c') := +by { rw ←not_iff_not, simp [lt_nmul_iff₃'] } + +theorem nmul_assoc : ∀ a b c, a ⨳ b ⨳ c = a ⨳ (b ⨳ c) +| a b c := begin + apply le_antisymm, + { rw nmul_le_iff₃, + intros a' ha b' hb c' hc, + repeat { rw nmul_assoc }, + exact nmul_nadd_lt₃' ha hb hc }, + { rw nmul_le_iff₃', + intros a' ha b' hb c' hc, + repeat { rw ←nmul_assoc }, + exact nmul_nadd_lt₃ ha hb hc }, +end +using_well_founded { dec_tac := `[solve_by_elim [psigma.lex.left, psigma.lex.right]] } + +end ordinal + +open ordinal + +instance : has_mul nat_ordinal := ⟨nmul⟩ + +instance : ordered_comm_semiring nat_ordinal := +{ mul := (*), + left_distrib := nmul_nadd, + right_distrib := nadd_nmul, + zero_mul := zero_nmul, + mul_zero := nmul_zero, + mul_assoc := nmul_assoc, + one := 1, + one_mul := one_nmul, + mul_one := nmul_one, + mul_comm := nmul_comm, + zero_le_one := @zero_le_one ordinal _ _ _ _, + mul_le_mul_of_nonneg_left := λ a b c, nmul_le_nmul_of_nonneg_left, + mul_le_mul_of_nonneg_right := λ a b c, nmul_le_nmul_of_nonneg_right, + ..nat_ordinal.ordered_cancel_add_comm_monoid, + ..nat_ordinal.linear_order } + +namespace ordinal + +theorem nmul_eq_mul (a b) : a ⨳ b = (a.to_nat_ordinal * b.to_nat_ordinal).to_ordinal := rfl + +theorem nmul_nadd_one : ∀ a b, a ⨳ (b ♯ 1) = a ⨳ b ♯ a := @mul_add_one nat_ordinal _ _ _ +theorem nadd_one_nmul : ∀ a b, (a ♯ 1) ⨳ b = a ⨳ b ♯ b := @add_one_mul nat_ordinal _ _ _ +theorem nmul_succ (a b) : a ⨳ succ b = a ⨳ b ♯ a := by rw [←nadd_one, nmul_nadd_one] +theorem succ_nmul (a b) : succ a ⨳ b = a ⨳ b ♯ b := by rw [←nadd_one, nadd_one_nmul] +theorem nmul_add_one : ∀ a b, a ⨳ (b + 1) = a ⨳ b ♯ a := nmul_succ +theorem add_one_nmul : ∀ a b, (a + 1) ⨳ b = a ⨳ b ♯ b := succ_nmul + end ordinal + +namespace nat_ordinal + +open ordinal + +theorem mul_le_nmul (a b : ordinal.{u}) : a * b ≤ a ⨳ b := +begin + apply b.limit_rec_on, + { simp }, + { intros c h, + rw [mul_succ, nmul_succ], + exact (add_le_nadd _ a).trans (nadd_le_nadd_right h a) }, + { intros c hc H, + rcases eq_zero_or_pos a with rfl | ha, + { simp }, + { rw [←is_normal.blsub_eq.{u u} (mul_is_normal ha) hc, blsub_le_iff], + exact λ i hi, (H i hi).trans_lt (nmul_lt_nmul_of_pos_left hi ha) } } +end + +end nat_ordinal diff --git a/src/set_theory/ordinal/principal.lean b/src/set_theory/ordinal/principal.lean index 61774bc56030a..ce63acccbfca8 100644 --- a/src/set_theory/ordinal/principal.lean +++ b/src/set_theory/ordinal/principal.lean @@ -85,31 +85,21 @@ nfp_le $ λ n, (ho.iterate_lt hao n).le /-! ### Principal ordinals are unbounded -/ -/-- The least strict upper bound of `op` applied to all pairs of ordinals less than `o`. This is -essentially a two-argument version of `ordinal.blsub`. -/ -def blsub₂ (op : ordinal → ordinal → ordinal) (o : ordinal) : ordinal := -lsub (λ x : o.out.α × o.out.α, op (typein (<) x.1) (typein (<) x.2)) - -theorem lt_blsub₂ (op : ordinal → ordinal → ordinal) {o : ordinal} {a b : ordinal} (ha : a < o) - (hb : b < o) : op a b < blsub₂ op o := -begin - convert lt_lsub _ (prod.mk (enum (<) a (by rwa type_lt)) (enum (<) b (by rwa type_lt))), - simp only [typein_enum] -end - theorem principal_nfp_blsub₂ (op : ordinal → ordinal → ordinal) (o : ordinal) : - principal op (nfp (blsub₂.{u u} op) o) := + principal op (nfp (λ o', blsub₂.{u u u} o' o' (λ a _ b _, op a b)) o) := λ a b ha hb, begin rw lt_nfp at *, cases ha with m hm, cases hb with n hn, - cases le_total ((blsub₂.{u u} op)^[m] o) ((blsub₂.{u u} op)^[n] o) with h h, + cases le_total + ((λ o', blsub₂.{u u u} o' o' (λ a _ b _, op a b))^[m] o) + ((λ o', blsub₂.{u u u} o' o' (λ a _ b _, op a b))^[n] o) with h h, { use n + 1, rw function.iterate_succ', - exact lt_blsub₂ op (hm.trans_le h) hn }, + exact lt_blsub₂ _ (hm.trans_le h) hn }, { use m + 1, rw function.iterate_succ', - exact lt_blsub₂ op hm (hn.trans_le h) }, + exact lt_blsub₂ _ hm (hn.trans_le h) } end theorem unbounded_principal (op : ordinal → ordinal → ordinal) : From 016138c2e83fa76d338d5df7d32d0acb6c587792 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sun, 16 Jul 2023 09:47:03 +0000 Subject: [PATCH 27/61] fix(tactic/linarith): instantiate metavariables in linarith (#19233) This probably isn't exhaustive, but catches the really trivial case. This doesn't appear to need forward-porting. --- src/tactic/linarith/datatypes.lean | 4 ++-- src/tactic/linarith/frontend.lean | 6 +++--- src/tactic/linarith/parsing.lean | 4 ++-- src/tactic/linarith/preprocessing.lean | 18 +++++++++--------- src/tactic/linarith/verification.lean | 6 +++--- test/linarith.lean | 9 +++++++++ 6 files changed, 28 insertions(+), 19 deletions(-) diff --git a/src/tactic/linarith/datatypes.lean b/src/tactic/linarith/datatypes.lean index e1edff788486d..52bc606a08fb8 100644 --- a/src/tactic/linarith/datatypes.lean +++ b/src/tactic/linarith/datatypes.lean @@ -369,13 +369,13 @@ Typically `R` and `R'` will be the same, except when `c = 0`, in which case `R'` If `c = 1`, `h'` is the same as `h` -- specifically, it does *not* change the type to `1*t R 0`. -/ meta def mk_single_comp_zero_pf (c : ℕ) (h : expr) : tactic (ineq × expr) := -do tp ← infer_type h, +do tp ← infer_type h >>= instantiate_mvars, some (iq, e) ← return $ parse_into_comp_and_expr tp, if c = 0 then do e' ← mk_app ``zero_mul [e], return (ineq.eq, e') else if c = 1 then return (iq, h) else - do tp ← (prod.snd <$> (infer_type h >>= get_rel_sides)) >>= infer_type, + do tp ← (prod.snd <$> (infer_type h >>= instantiate_mvars >>= get_rel_sides)) >>= infer_type, c ← tp.of_nat c, cpos ← to_expr ``(%%c > 0), (_, ex) ← solve_aux cpos `[norm_num, done], diff --git a/src/tactic/linarith/frontend.lean b/src/tactic/linarith/frontend.lean index 5ea2c37df8738..f09b0cedc5ca5 100644 --- a/src/tactic/linarith/frontend.lean +++ b/src/tactic/linarith/frontend.lean @@ -151,7 +151,7 @@ newly introduced local constant. Otherwise returns `none`. -/ meta def apply_contr_lemma : tactic (option (expr × expr)) := -do t ← target, +do t ← target >>= instantiate_mvars, match get_contr_lemma_name_and_type t with | some (nm, tp) := do refine ((expr.const nm []) pexpr.mk_placeholder), v ← intro1, return $ some (tp, v) @@ -207,7 +207,7 @@ to only those that are comparisons over the type `restr_type`. -/ meta def filter_hyps_to_type (restr_type : expr) (hyps : list expr) : tactic (list expr) := hyps.mfilter $ λ h, do - ht ← infer_type h, + ht ← infer_type h >>= instantiate_mvars, match get_contr_lemma_name_and_type ht with | some (_, htype) := succeeds $ unify htype restr_type | none := return ff @@ -238,7 +238,7 @@ expressions. meta def tactic.linarith (reduce_semi : bool) (only_on : bool) (hyps : list pexpr) (cfg : linarith_config := {}) : tactic unit := focus1 $ -do t ← target, +do t ← target >>= instantiate_mvars, -- if the target is an equality, we run `linarith` twice, to prove ≤ and ≥. if t.is_eq.is_some then linarith_trace "target is an equality: splitting" >> diff --git a/src/tactic/linarith/parsing.lean b/src/tactic/linarith/parsing.lean index c31cdc3b9fcee..f43291cd087ea 100644 --- a/src/tactic/linarith/parsing.lean +++ b/src/tactic/linarith/parsing.lean @@ -182,8 +182,8 @@ It also returns the largest variable index that appears in comparisons in `c`. -/ meta def linear_forms_and_max_var (red : transparency) (pfs : list expr) : tactic (list comp × ℕ) := -do pftps ← pfs.mmap infer_type, - (l, _, map) ← to_comp_fold red [] pftps mk_rb_map, +do pftps ← pfs.mmap (λ e, infer_type e >>= instantiate_mvars), + (l, _, map) ← to_comp_fold red [] pftps mk_rb_map, return (l, map.size - 1) diff --git a/src/tactic/linarith/preprocessing.lean b/src/tactic/linarith/preprocessing.lean index 59775da603e99..ae700421d8eeb 100644 --- a/src/tactic/linarith/preprocessing.lean +++ b/src/tactic/linarith/preprocessing.lean @@ -70,7 +70,7 @@ private meta def rearr_comp_aux : expr → expr → tactic expr and turns it into a proof of a comparison `_ R 0`, where `R ∈ {=, ≤, <}`. -/ meta def rearr_comp (e : expr) : tactic expr := -infer_type e >>= rearr_comp_aux e +infer_type e >>= instantiate_mvars >>= rearr_comp_aux e /-- If `e` is of the form `((n : ℕ) : ℤ)`, `is_nat_int_coe e` returns `n : ℕ`. -/ meta def is_nat_int_coe : expr → option expr @@ -96,7 +96,7 @@ If `pf` is a proof of a strict inequality `(a : ℤ) < b`, and similarly if `pf` proves a negated weak inequality. -/ meta def mk_non_strict_int_pf_of_strict_int_pf (pf : expr) : tactic expr := -do tp ← infer_type pf, +do tp ← infer_type pf >>= instantiate_mvars, match tp with | `(%%a < %%b) := to_expr ``(int.add_one_le_iff.mpr %%pf) | `(%%a > %%b) := to_expr ``(int.add_one_le_iff.mpr %%pf) @@ -139,7 +139,7 @@ Removes any expressions that are not proofs of inequalities, equalities, or nega meta def filter_comparisons : preprocessor := { name := "filter terms that are not proofs of comparisons", transform := λ h, -(do tp ← infer_type h, +(do tp ← infer_type h >>= instantiate_mvars, is_prop tp >>= guardb, guardb (filter_comparisons_aux tp), return [h]) @@ -152,7 +152,7 @@ For example, a proof of `¬ a < b` will become a proof of `a ≥ b`. meta def remove_negations : preprocessor := { name := "replace negations of comparisons", transform := λ h, -do tp ← infer_type h, +do tp ← infer_type h >>= instantiate_mvars, match tp with | `(¬ %%p) := singleton <$> rem_neg h p | _ := return [h] @@ -171,9 +171,9 @@ meta def nat_to_int : global_preprocessor := -- we lock the tactic state here because a `simplify` call inside of -- `zify_proof` corrupts the tactic state when run under `io.run_tactic`. do l ← lock_tactic_state $ l.mmap $ λ h, - infer_type h >>= guardb ∘ is_nat_prop >> zify_proof [] h <|> return h, + infer_type h >>= instantiate_mvars >>= guardb ∘ is_nat_prop >> zify_proof [] h <|> return h, nonnegs ← l.mfoldl (λ (es : expr_set) h, do - (a, b) ← infer_type h >>= get_rel_sides, + (a, b) ← infer_type h >>= instantiate_mvars >>= get_rel_sides, return $ (es.insert_list (get_nat_comps a)).insert_list (get_nat_comps b)) mk_rb_set, (++) l <$> nonnegs.to_list.mmap mk_coe_nat_nonneg_prf } @@ -183,7 +183,7 @@ into a proof of `t1 ≤ t2 + 1`. -/ meta def strengthen_strict_int : preprocessor := { name := "strengthen strict inequalities over int", transform := λ h, -do tp ← infer_type h, +do tp ← infer_type h >>= instantiate_mvars, guardb (is_strict_int_prop tp) >> singleton <$> mk_non_strict_int_pf_of_strict_int_pf h <|> return [h] } @@ -213,7 +213,7 @@ it tries to scale `t` to cancel out division by numerals. meta def cancel_denoms : preprocessor := { name := "cancel denominators", transform := λ pf, -(do some (_, lhs) ← parse_into_comp_and_expr <$> infer_type pf, +(do some (_, lhs) ← parse_into_comp_and_expr <$> (infer_type pf >>= instantiate_mvars), guardb $ lhs.contains_constant (= `has_div.div), singleton <$> normalize_denominators_in_lhs pf lhs) <|> return [pf] } @@ -272,7 +272,7 @@ This produces `2^n` branches when there are `n` such hypotheses in the input. -/ meta def remove_ne_aux : list expr → tactic (list branch) := λ hs, -(do e ← hs.mfind (λ e : expr, do e ← infer_type e, guard $ e.is_ne.is_some), +(do e ← hs.mfind (λ e : expr, do e ← infer_type e >>= instantiate_mvars, guard $ e.is_ne.is_some), [(_, ng1), (_, ng2)] ← to_expr ``(or.elim (lt_or_gt_of_ne %%e)) >>= apply, let do_goal : expr → tactic (list branch) := λ g, do set_goals [g], diff --git a/src/tactic/linarith/verification.lean b/src/tactic/linarith/verification.lean index 8f6e0b1e5010c..ea5123b0dd715 100644 --- a/src/tactic/linarith/verification.lean +++ b/src/tactic/linarith/verification.lean @@ -86,11 +86,11 @@ meta def mk_lt_zero_pf : list (expr × ℕ) → tactic expr /-- If `prf` is a proof of `t R s`, `term_of_ineq_prf prf` returns `t`. -/ meta def term_of_ineq_prf (prf : expr) : tactic expr := -prod.fst <$> (infer_type prf >>= get_rel_sides) +prod.fst <$> (infer_type prf >>= instantiate_mvars >>= get_rel_sides) /-- If `prf` is a proof of `t R s`, `ineq_prf_tp prf` returns the type of `t`. -/ meta def ineq_prf_tp (prf : expr) : tactic expr := -term_of_ineq_prf prf >>= infer_type +term_of_ineq_prf prf >>= infer_type >>= instantiate_mvars /-- `mk_neg_one_lt_zero_pf tp` returns a proof of `-1 < 0`, @@ -120,7 +120,7 @@ proof, it adds a proof of `-t = 0` to the list. meta def add_neg_eq_pfs : list expr → tactic (list expr) | [] := return [] | (h::t) := - do some (iq, tp) ← parse_into_comp_and_expr <$> infer_type h, + do some (iq, tp) ← parse_into_comp_and_expr <$> (infer_type h >>= instantiate_mvars), match iq with | ineq.eq := do nep ← mk_neg_eq_zero_pf h, tl ← add_neg_eq_pfs t, return $ h::nep::tl | _ := list.cons h <$> add_neg_eq_pfs t diff --git a/test/linarith.lean b/test/linarith.lean index 1d127062a60e5..365273aecad63 100644 --- a/test/linarith.lean +++ b/test/linarith.lean @@ -1,5 +1,14 @@ import tactic.linarith +example : ∀ (y : ℕ), y ≤ 37 → y < 40 := +begin + refine λ y hy, _, + -- The type of `hy` is a (solved but not instantiated) metavariable + do { tactic.get_local `hy >>= tactic.infer_type >>= guardb ∘ expr.is_mvar }, + -- But linarith should still work + linarith +end + example {α : Type} (_inst : Π (a : Prop), decidable a) [linear_ordered_field α] {a b c : α} From 88e6ad41d9c5d6da7e2ac45cae6d07a7fc6bb429 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sun, 16 Jul 2023 09:47:04 +0000 Subject: [PATCH 28/61] fix: handle archive and counterexamples correctly when adding port comments (#19237) --- scripts/add_port_comments.py | 7 +++++++ 1 file changed, 7 insertions(+) diff --git a/scripts/add_port_comments.py b/scripts/add_port_comments.py index 5c0160b5f9b56..366942aa8640f 100644 --- a/scripts/add_port_comments.py +++ b/scripts/add_port_comments.py @@ -12,6 +12,8 @@ status = PortStatus.deserialize_old() src_path = Path(__file__).parent.parent / 'src' +archive_path = Path(__file__).parent.parent / 'archive' +counterexamples_path = Path(__file__).parent.parent / 'counterexamples' def make_comment(fstatus): return textwrap.dedent(f"""\ @@ -87,6 +89,11 @@ def add_port_status(fcontent: str, fstatus: FileStatus) -> str: return fcontent def fname_for(import_path: str) -> Path: + for root in [src_path, archive_path, counterexamples_path]: + p = root / Path(*import_path.split('.')).with_suffix('.lean') + if p.exists(): + return p + # used only for error messages, a best-guess return src_path / Path(*import_path.split('.')).with_suffix('.lean') From 08b081ea92d80e3a41f899eea36ef6d56e0f1db0 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Sun, 16 Jul 2023 12:56:33 +0000 Subject: [PATCH 29/61] chore(*): add mathlib4 synchronization comments (#19238) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.expr` * `algebraic_geometry.morphisms.finite_type` * `algebraic_geometry.morphisms.ring_hom_properties` * `arithcc` * `canonically_ordered_comm_semiring_two_mul` * `char_p_zero_ne_char_zero` * `cyclotomic_105` * `data.matrix.auto` * `direct_sum_is_internal` * `examples.mersenne_primes` * `examples.prop_encodable` * `geometry.manifold.instances.sphere` * `girard` * `homogeneous_prime_not_prime` * `imo.imo1959_q1` * `imo.imo1960_q1` * `imo.imo1962_q1` * `imo.imo1962_q4` * `imo.imo1964_q1` * `imo.imo1969_q1` * `imo.imo1972_q5` * `imo.imo1975_q1` * `imo.imo1977_q6` * `imo.imo1981_q3` * `imo.imo1987_q1` * `imo.imo1988_q6` * `imo.imo1994_q1` * `imo.imo1998_q2` * `imo.imo2001_q2` * `imo.imo2001_q6` * `imo.imo2005_q3` * `imo.imo2005_q4` * `imo.imo2006_q3` * `imo.imo2006_q5` * `imo.imo2008_q2` * `imo.imo2008_q3` * `imo.imo2008_q4` * `imo.imo2011_q3` * `imo.imo2011_q5` * `imo.imo2013_q1` * `imo.imo2013_q5` * `imo.imo2019_q1` * `imo.imo2019_q2` * `imo.imo2019_q4` * `imo.imo2020_q2` * `imo.imo2021_q1` * `linear_order_with_pos_mul_pos_eq_zero` * `map_floor` * `miu_language.basic` * `miu_language.decision_nec` * `miu_language.decision_suf` * `oxford_invariants.2021summer.week3_p1` * `phillips` * `pseudoelement` * `quadratic_form` * `seminorm_lattice_not_distrib` * `sensitivity` * `sorgenfrey_line` * `wiedijk_100_theorems.abel_ruffini` * `wiedijk_100_theorems.area_of_a_circle` * `wiedijk_100_theorems.ascending_descending_sequences` * `wiedijk_100_theorems.ballot_problem` * `wiedijk_100_theorems.birthday_problem` * `wiedijk_100_theorems.cubing_a_cube` * `wiedijk_100_theorems.friendship_graphs` * `wiedijk_100_theorems.herons_formula` * `wiedijk_100_theorems.inverse_triangle_sum` * `wiedijk_100_theorems.konigsberg` * `wiedijk_100_theorems.partition` * `wiedijk_100_theorems.perfect_numbers` * `wiedijk_100_theorems.solution_of_cubic` * `wiedijk_100_theorems.sum_of_prime_reciprocals_diverges` * `zero_divisors_in_add_monoid_algebras` --- archive/arithcc.lean | 3 +++ archive/examples/mersenne_primes.lean | 3 +++ archive/examples/prop_encodable.lean | 3 +++ archive/imo/imo1959_q1.lean | 3 +++ archive/imo/imo1960_q1.lean | 3 +++ archive/imo/imo1962_q1.lean | 3 +++ archive/imo/imo1962_q4.lean | 3 +++ archive/imo/imo1964_q1.lean | 3 +++ archive/imo/imo1969_q1.lean | 3 +++ archive/imo/imo1972_q5.lean | 3 +++ archive/imo/imo1975_q1.lean | 3 +++ archive/imo/imo1977_q6.lean | 3 +++ archive/imo/imo1981_q3.lean | 3 +++ archive/imo/imo1987_q1.lean | 3 +++ archive/imo/imo1988_q6.lean | 3 +++ archive/imo/imo1994_q1.lean | 3 +++ archive/imo/imo1998_q2.lean | 3 +++ archive/imo/imo2001_q2.lean | 3 +++ archive/imo/imo2001_q6.lean | 3 +++ archive/imo/imo2005_q3.lean | 3 +++ archive/imo/imo2005_q4.lean | 3 +++ archive/imo/imo2006_q3.lean | 3 +++ archive/imo/imo2006_q5.lean | 3 +++ archive/imo/imo2008_q2.lean | 3 +++ archive/imo/imo2008_q3.lean | 3 +++ archive/imo/imo2008_q4.lean | 3 +++ archive/imo/imo2011_q3.lean | 3 +++ archive/imo/imo2011_q5.lean | 3 +++ archive/imo/imo2013_q1.lean | 3 +++ archive/imo/imo2013_q5.lean | 3 +++ archive/imo/imo2019_q1.lean | 3 +++ archive/imo/imo2019_q2.lean | 3 +++ archive/imo/imo2019_q4.lean | 3 +++ archive/imo/imo2020_q2.lean | 3 +++ archive/imo/imo2021_q1.lean | 3 +++ archive/miu_language/basic.lean | 3 +++ archive/miu_language/decision_nec.lean | 3 +++ archive/miu_language/decision_suf.lean | 3 +++ archive/oxford_invariants/2021summer/week3_p1.lean | 3 +++ archive/sensitivity.lean | 3 +++ archive/wiedijk_100_theorems/abel_ruffini.lean | 3 +++ archive/wiedijk_100_theorems/area_of_a_circle.lean | 3 +++ .../wiedijk_100_theorems/ascending_descending_sequences.lean | 3 +++ archive/wiedijk_100_theorems/ballot_problem.lean | 3 +++ archive/wiedijk_100_theorems/birthday_problem.lean | 3 +++ archive/wiedijk_100_theorems/cubing_a_cube.lean | 3 +++ archive/wiedijk_100_theorems/friendship_graphs.lean | 3 +++ archive/wiedijk_100_theorems/herons_formula.lean | 3 +++ archive/wiedijk_100_theorems/inverse_triangle_sum.lean | 3 +++ archive/wiedijk_100_theorems/konigsberg.lean | 3 +++ archive/wiedijk_100_theorems/partition.lean | 3 +++ archive/wiedijk_100_theorems/perfect_numbers.lean | 3 +++ archive/wiedijk_100_theorems/solution_of_cubic.lean | 3 +++ .../sum_of_prime_reciprocals_diverges.lean | 3 +++ counterexamples/canonically_ordered_comm_semiring_two_mul.lean | 3 +++ counterexamples/char_p_zero_ne_char_zero.lean | 3 +++ counterexamples/cyclotomic_105.lean | 3 +++ counterexamples/direct_sum_is_internal.lean | 3 +++ counterexamples/girard.lean | 3 +++ counterexamples/homogeneous_prime_not_prime.lean | 3 +++ counterexamples/linear_order_with_pos_mul_pos_eq_zero.lean | 3 +++ counterexamples/map_floor.lean | 3 +++ counterexamples/phillips.lean | 3 +++ counterexamples/pseudoelement.lean | 3 +++ counterexamples/quadratic_form.lean | 3 +++ counterexamples/seminorm_lattice_not_distrib.lean | 3 +++ counterexamples/sorgenfrey_line.lean | 3 +++ counterexamples/zero_divisors_in_add_monoid_algebras.lean | 3 +++ src/algebra/expr.lean | 3 +++ src/algebraic_geometry/morphisms/finite_type.lean | 3 +++ src/algebraic_geometry/morphisms/ring_hom_properties.lean | 3 +++ src/data/matrix/auto.lean | 3 +++ src/geometry/manifold/instances/sphere.lean | 3 +++ 73 files changed, 219 insertions(+) diff --git a/archive/arithcc.lean b/archive/arithcc.lean index 9de0770bbd8f2..f1feb07fa6173 100644 --- a/archive/arithcc.lean +++ b/archive/arithcc.lean @@ -9,6 +9,9 @@ import tactic.basic /-! # A compiler for arithmetic expressions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A formalization of the correctness of a compiler from arithmetic expressions to machine language described by McCarthy and Painter, which is considered the first proof of compiler correctness. diff --git a/archive/examples/mersenne_primes.lean b/archive/examples/mersenne_primes.lean index 2c5a28bdbea00..af4166ff31ebc 100644 --- a/archive/examples/mersenne_primes.lean +++ b/archive/examples/mersenne_primes.lean @@ -8,6 +8,9 @@ import number_theory.lucas_lehmer /-! # Explicit Mersenne primes +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We run some Lucas-Lehmer tests to prove some Mersenne primes are prime. See the discussion at the end of [src/number_theory/lucas_lehmer.lean] diff --git a/archive/examples/prop_encodable.lean b/archive/examples/prop_encodable.lean index a877eeed41334..d2436fd06f034 100644 --- a/archive/examples/prop_encodable.lean +++ b/archive/examples/prop_encodable.lean @@ -9,6 +9,9 @@ import data.W.basic /-! # W types +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The file `data/W.lean` shows that if `α` is an an encodable fintype and for every `a : α`, `β a` is encodable, then `W β` is encodable. diff --git a/archive/imo/imo1959_q1.lean b/archive/imo/imo1959_q1.lean index bb0c5f8b63b36..fe34d1b143489 100644 --- a/archive/imo/imo1959_q1.lean +++ b/archive/imo/imo1959_q1.lean @@ -10,6 +10,9 @@ import data.nat.prime /-! # IMO 1959 Q1 +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Prove that the fraction `(21n+4)/(14n+3)` is irreducible for every natural number `n`. Since Lean doesn't have a concept of "irreducible fractions" per se, we just formalize this diff --git a/archive/imo/imo1960_q1.lean b/archive/imo/imo1960_q1.lean index fa3ee7680092f..eccc3abb782e8 100644 --- a/archive/imo/imo1960_q1.lean +++ b/archive/imo/imo1960_q1.lean @@ -9,6 +9,9 @@ import data.nat.digits /-! # IMO 1960 Q1 +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Determine all three-digit numbers $N$ having the property that $N$ is divisible by 11, and $\dfrac{N}{11}$ is equal to the sum of the squares of the digits of $N$. diff --git a/archive/imo/imo1962_q1.lean b/archive/imo/imo1962_q1.lean index 0afc5e6be6726..bd586605533fb 100644 --- a/archive/imo/imo1962_q1.lean +++ b/archive/imo/imo1962_q1.lean @@ -9,6 +9,9 @@ import data.nat.digits /-! # IMO 1962 Q1 +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Find the smallest natural number $n$ which has the following properties: (a) Its decimal representation has 6 as the last digit. diff --git a/archive/imo/imo1962_q4.lean b/archive/imo/imo1962_q4.lean index 95b725d68beca..87996b4246960 100644 --- a/archive/imo/imo1962_q4.lean +++ b/archive/imo/imo1962_q4.lean @@ -9,6 +9,9 @@ import analysis.special_functions.trigonometric.complex /-! # IMO 1962 Q4 +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Solve the equation `cos x ^ 2 + cos (2 * x) ^ 2 + cos (3 * x) ^ 2 = 1`. Since Lean does not have a concept of "simplest form", we just express what is diff --git a/archive/imo/imo1964_q1.lean b/archive/imo/imo1964_q1.lean index fd987ba4f5f10..f381845902b31 100644 --- a/archive/imo/imo1964_q1.lean +++ b/archive/imo/imo1964_q1.lean @@ -10,6 +10,9 @@ import data.nat.modeq /-! # IMO 1964 Q1 +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + (a) Find all positive integers $n$ for which $2^n-1$ is divisible by $7$. (b) Prove that there is no positive integer $n$ for which $2^n+1$ is divisible by $7$. diff --git a/archive/imo/imo1969_q1.lean b/archive/imo/imo1969_q1.lean index 3eee264d353ca..b1e47a6bef7aa 100644 --- a/archive/imo/imo1969_q1.lean +++ b/archive/imo/imo1969_q1.lean @@ -11,6 +11,9 @@ import data.set.finite /-! # IMO 1969 Q1 +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Prove that there are infinitely many natural numbers $a$ with the following property: the number $z = n^4 + a$ is not prime for any natural number $n$. -/ diff --git a/archive/imo/imo1972_q5.lean b/archive/imo/imo1972_q5.lean index c796e589190f5..8dbb28ff56f04 100644 --- a/archive/imo/imo1972_q5.lean +++ b/archive/imo/imo1972_q5.lean @@ -10,6 +10,9 @@ import analysis.normed_space.basic /-! # IMO 1972 Q5 +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Problem: `f` and `g` are real-valued functions defined on the real line. For all `x` and `y`, `f(x + y) + f(x - y) = 2f(x)g(y)`. `f` is not identically zero and `|f(x)| ≤ 1` for all `x`. Prove that `|g(x)| ≤ 1` for all `x`. diff --git a/archive/imo/imo1975_q1.lean b/archive/imo/imo1975_q1.lean index eaf260b9cdbac..b135b1806671d 100644 --- a/archive/imo/imo1975_q1.lean +++ b/archive/imo/imo1975_q1.lean @@ -11,6 +11,9 @@ import algebra.big_operators.ring /-! # IMO 1975 Q1 +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Let `x₁, x₂, ... , xₙ` and `y₁, y₂, ... , yₙ` be two sequences of real numbers, such that `x₁ ≥ x₂ ≥ ... ≥ xₙ` and `y₁ ≥ y₂ ≥ ... ≥ yₙ`. Prove that if `z₁, z₂, ... , zₙ` is any permutation of `y₁, y₂, ... , yₙ`, then `∑ (xᵢ - yᵢ)^2 ≤ ∑ (xᵢ - zᵢ)^2` diff --git a/archive/imo/imo1977_q6.lean b/archive/imo/imo1977_q6.lean index 445644be7a341..8d3d44c050769 100644 --- a/archive/imo/imo1977_q6.lean +++ b/archive/imo/imo1977_q6.lean @@ -8,6 +8,9 @@ import data.pnat.basic /-! # IMO 1977 Q6 +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Suppose `f : ℕ+ → ℕ+` satisfies `f(f(n)) < f(n + 1)` for all `n`. Prove that `f(n) = n` for all `n`. diff --git a/archive/imo/imo1981_q3.lean b/archive/imo/imo1981_q3.lean index 4b40612e5a873..10a40331f2eec 100644 --- a/archive/imo/imo1981_q3.lean +++ b/archive/imo/imo1981_q3.lean @@ -10,6 +10,9 @@ import tactic.linarith /-! # IMO 1981 Q3 +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Determine the maximum value of `m ^ 2 + n ^ 2`, where `m` and `n` are integers in `{1, 2, ..., 1981}` and `(n ^ 2 - m * n - m ^ 2) ^ 2 = 1`. diff --git a/archive/imo/imo1987_q1.lean b/archive/imo/imo1987_q1.lean index b73054ddc9101..ffae4fbd202d7 100644 --- a/archive/imo/imo1987_q1.lean +++ b/archive/imo/imo1987_q1.lean @@ -11,6 +11,9 @@ import dynamics.fixed_points.basic /-! # Formalization of IMO 1987, Q1 +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Let $p_{n, k}$ be the number of permutations of a set of cardinality `n ≥ 1` that fix exactly `k` elements. Prove that $∑_{k=0}^n k p_{n,k}=n!$. diff --git a/archive/imo/imo1988_q6.lean b/archive/imo/imo1988_q6.lean index d5cd3fa2020e3..4e2699c0f61d9 100644 --- a/archive/imo/imo1988_q6.lean +++ b/archive/imo/imo1988_q6.lean @@ -13,6 +13,9 @@ import tactic.wlog /-! # IMO 1988 Q6 and constant descent Vieta jumping +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Question 6 of IMO1988 is somewhat (in)famous. Several expert problem solvers could not tackle the question within the given time limit. The problem lead to the introduction of a new proof technique, diff --git a/archive/imo/imo1994_q1.lean b/archive/imo/imo1994_q1.lean index 289fdf3250ad2..89b584ee3ebce 100644 --- a/archive/imo/imo1994_q1.lean +++ b/archive/imo/imo1994_q1.lean @@ -14,6 +14,9 @@ import tactic.by_contra /-! # IMO 1994 Q1 +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Let `m` and `n` be two positive integers. Let `a₁, a₂, ..., aₘ` be `m` different numbers from the set `{1, 2, ..., n}` such that for any two indices `i` and `j` with `1 ≤ i ≤ j ≤ m` and `aᵢ + aⱼ ≤ n`, diff --git a/archive/imo/imo1998_q2.lean b/archive/imo/imo1998_q2.lean index ab3cb7e1ccb35..7ce14e5cf3214 100644 --- a/archive/imo/imo1998_q2.lean +++ b/archive/imo/imo1998_q2.lean @@ -11,6 +11,9 @@ import tactic.noncomm_ring /-! # IMO 1998 Q2 + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. In a competition, there are `a` contestants and `b` judges, where `b ≥ 3` is an odd integer. Each judge rates each contestant as either "pass" or "fail". Suppose `k` is a number such that, for any two judges, their ratings coincide for at most `k` contestants. Prove that `k / a ≥ (b - 1) / (2b)`. diff --git a/archive/imo/imo2001_q2.lean b/archive/imo/imo2001_q2.lean index fc2d5172f39e8..346adc3db71fd 100644 --- a/archive/imo/imo2001_q2.lean +++ b/archive/imo/imo2001_q2.lean @@ -9,6 +9,9 @@ import analysis.special_functions.pow.real /-! # IMO 2001 Q2 +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Let $a$, $b$, $c$ be positive reals. Prove that $$ \frac{a}{\sqrt{a^2 + 8bc}} + diff --git a/archive/imo/imo2001_q6.lean b/archive/imo/imo2001_q6.lean index c2fa9f829a855..4b96c07090e1c 100644 --- a/archive/imo/imo2001_q6.lean +++ b/archive/imo/imo2001_q6.lean @@ -9,6 +9,9 @@ import tactic.linear_combination /-! # IMO 2001 Q6 + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. Let $a$, $b$, $c$, $d$ be integers with $a > b > c > d > 0$. Suppose that $$ a*c + b*d = (a + b - c + d) * (-a + b + c + d). $$ diff --git a/archive/imo/imo2005_q3.lean b/archive/imo/imo2005_q3.lean index bfcb287444ded..9d4ef279453ea 100644 --- a/archive/imo/imo2005_q3.lean +++ b/archive/imo/imo2005_q3.lean @@ -8,6 +8,9 @@ import tactic.positivity /-! # IMO 2005 Q3 + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. Let `x`, `y` and `z` be positive real numbers such that `xyz ≥ 1`. Prove that: `(x^5 - x^2)/(x^5 + y^2 + z^2) + (y^5 - y^2)/(y^5 + z^2 + x^2) + (z^5 - z^2)/(z^5 + x^2 + y^2) ≥ 0` diff --git a/archive/imo/imo2005_q4.lean b/archive/imo/imo2005_q4.lean index 9caceeddf5abf..ac90ef53f0bcc 100644 --- a/archive/imo/imo2005_q4.lean +++ b/archive/imo/imo2005_q4.lean @@ -8,6 +8,9 @@ import field_theory.finite.basic /-! # IMO 2005 Q4 +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Problem: Determine all positive integers relatively prime to all the terms of the infinite sequence `a n = 2 ^ n + 3 ^ n + 6 ^ n - 1`, for `n ≥ 1`. diff --git a/archive/imo/imo2006_q3.lean b/archive/imo/imo2006_q3.lean index 36a22e34bf2d0..8e3a0862d3ad4 100644 --- a/archive/imo/imo2006_q3.lean +++ b/archive/imo/imo2006_q3.lean @@ -9,6 +9,9 @@ import analysis.special_functions.sqrt /-! # IMO 2006 Q3 +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Determine the least real number $M$ such that $$ \left| ab(a^2 - b^2) + bc(b^2 - c^2) + ca(c^2 - a^2) \right| diff --git a/archive/imo/imo2006_q5.lean b/archive/imo/imo2006_q5.lean index 682b1994adb37..ec9929b2d4c1c 100644 --- a/archive/imo/imo2006_q5.lean +++ b/archive/imo/imo2006_q5.lean @@ -10,6 +10,9 @@ import dynamics.periodic_pts /-! # IMO 2006 Q5 +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Let $P(x)$ be a polynomial of degree $n>1$ with integer coefficients, and let $k$ be a positive integer. Consider the polynomial $Q(x) = P(P(\ldots P(P(x))\ldots))$, where $P$ occurs $k$ times. Prove that there are at most $n$ integers $t$ such that $Q(t)=t$. diff --git a/archive/imo/imo2008_q2.lean b/archive/imo/imo2008_q2.lean index eb4df1664207b..43c6b652cb127 100644 --- a/archive/imo/imo2008_q2.lean +++ b/archive/imo/imo2008_q2.lean @@ -8,6 +8,9 @@ import data.set.finite /-! # IMO 2008 Q2 + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. (a) Prove that ``` x^2 / (x-1)^2 + y^2 / (y-1)^2 + z^2 / (z-1)^2 ≥ 1 diff --git a/archive/imo/imo2008_q3.lean b/archive/imo/imo2008_q3.lean index 16a937f43fce5..ce1470e753e1b 100644 --- a/archive/imo/imo2008_q3.lean +++ b/archive/imo/imo2008_q3.lean @@ -12,6 +12,9 @@ import tactic.linear_combination /-! # IMO 2008 Q3 + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. Prove that there exist infinitely many positive integers `n` such that `n^2 + 1` has a prime divisor which is greater than `2n + √(2n)`. diff --git a/archive/imo/imo2008_q4.lean b/archive/imo/imo2008_q4.lean index 70339914ed2d3..c1cae17f559fa 100644 --- a/archive/imo/imo2008_q4.lean +++ b/archive/imo/imo2008_q4.lean @@ -10,6 +10,9 @@ import tactic.linear_combination /-! # IMO 2008 Q4 + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. Find all functions `f : (0,∞) → (0,∞)` (so, `f` is a function from the positive real numbers to the positive real numbers) such that ``` diff --git a/archive/imo/imo2011_q3.lean b/archive/imo/imo2011_q3.lean index c3e5dfc5d8361..67b4845b53018 100644 --- a/archive/imo/imo2011_q3.lean +++ b/archive/imo/imo2011_q3.lean @@ -9,6 +9,9 @@ import data.real.basic /-! # IMO 2011 Q3 +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Let f : ℝ → ℝ be a function that satisfies f(x + y) ≤ y * f(x) + f(f(x)) diff --git a/archive/imo/imo2011_q5.lean b/archive/imo/imo2011_q5.lean index abf6897ad352d..5bb51b977d729 100644 --- a/archive/imo/imo2011_q5.lean +++ b/archive/imo/imo2011_q5.lean @@ -8,6 +8,9 @@ import data.int.dvd.basic /-! # IMO 2011 Q5 +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Let `f` be a function from the set of integers to the set of positive integers. Suppose that, for any two integers `m` and `n`, the difference `f(m) - f(n)` is divisible by diff --git a/archive/imo/imo2013_q1.lean b/archive/imo/imo2013_q1.lean index cf7c03ce84b00..842d2de4c36f5 100644 --- a/archive/imo/imo2013_q1.lean +++ b/archive/imo/imo2013_q1.lean @@ -13,6 +13,9 @@ import tactic.field_simp /-! # IMO 2013 Q1 +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Prove that for any pair of positive integers k and n, there exist k positive integers m₁, m₂, ..., mₖ (not necessarily different) such that diff --git a/archive/imo/imo2013_q5.lean b/archive/imo/imo2013_q5.lean index d30c32da64014..16f7cdf6c6ab6 100644 --- a/archive/imo/imo2013_q5.lean +++ b/archive/imo/imo2013_q5.lean @@ -12,6 +12,9 @@ import tactic.positivity /-! # IMO 2013 Q5 +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Let `ℚ>₀` be the positive rational numbers. Let `f : ℚ>₀ → ℝ` be a function satisfying the conditions diff --git a/archive/imo/imo2019_q1.lean b/archive/imo/imo2019_q1.lean index 4e4772456a2ee..0226474656c48 100644 --- a/archive/imo/imo2019_q1.lean +++ b/archive/imo/imo2019_q1.lean @@ -8,6 +8,9 @@ import tactic.linarith /-! # IMO 2019 Q1 +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Determine all functions `f : ℤ → ℤ` such that, for all integers `a` and `b`, `f(2a) + 2f(b) = f(f(a+b))`. diff --git a/archive/imo/imo2019_q2.lean b/archive/imo/imo2019_q2.lean index 7be26dade77e2..b93b1046604ed 100644 --- a/archive/imo/imo2019_q2.lean +++ b/archive/imo/imo2019_q2.lean @@ -9,6 +9,9 @@ import geometry.euclidean.sphere.second_inter /-! # IMO 2019 Q2 +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In triangle `ABC`, point `A₁` lies on side `BC` and point `B₁` lies on side `AC`. Let `P` and `Q` be points on segments `AA₁` and `BB₁`, respectively, such that `PQ` is parallel to `AB`. Let `P₁` be a point on line `PB₁`, such that `B₁` lies strictly between `P` and `P₁`, and diff --git a/archive/imo/imo2019_q4.lean b/archive/imo/imo2019_q4.lean index 3465674493c60..05af48d4acb6e 100644 --- a/archive/imo/imo2019_q4.lean +++ b/archive/imo/imo2019_q4.lean @@ -10,6 +10,9 @@ import data.nat.multiplicity /-! # IMO 2019 Q4 +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Find all pairs `(k, n)` of positive integers such that ``` k! = (2 ^ n - 1)(2 ^ n - 2)(2 ^ n - 4)···(2 ^ n - 2 ^ (n - 1)) diff --git a/archive/imo/imo2020_q2.lean b/archive/imo/imo2020_q2.lean index beea58a78ff17..54938a0afb1ec 100644 --- a/archive/imo/imo2020_q2.lean +++ b/archive/imo/imo2020_q2.lean @@ -9,6 +9,9 @@ import analysis.mean_inequalities /-! # IMO 2020 Q2 +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The real numbers `a`, `b`, `c`, `d` are such that `a ≥ b ≥ c ≥ d > 0` and `a + b + c + d = 1`. Prove that `(a + 2b + 3c + 4d) a^a b^b c^c d^d < 1`. diff --git a/archive/imo/imo2021_q1.lean b/archive/imo/imo2021_q1.lean index 85abc0654bd36..397823fa856e1 100644 --- a/archive/imo/imo2021_q1.lean +++ b/archive/imo/imo2021_q1.lean @@ -14,6 +14,9 @@ import tactic.ring_exp /-! # IMO 2021 Q1 +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Let `n≥100` be an integer. Ivan writes the numbers `n, n+1,..., 2n` each on different cards. He then shuffles these `n+1` cards, and divides them into two piles. Prove that at least one of the piles contains two cards such that the sum of their numbers is a perfect square. diff --git a/archive/miu_language/basic.lean b/archive/miu_language/basic.lean index f4ca4329c312d..ad139d96ffe3b 100644 --- a/archive/miu_language/basic.lean +++ b/archive/miu_language/basic.lean @@ -8,6 +8,9 @@ import tactic.linarith /-! # An MIU Decision Procedure in Lean +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The [MIU formal system](https://en.wikipedia.org/wiki/MU_puzzle) was introduced by Douglas Hofstadter in the first chapter of his 1979 book, [Gödel, Escher, Bach](https://en.wikipedia.org/wiki/G%C3%B6del,_Escher,_Bach). diff --git a/archive/miu_language/decision_nec.lean b/archive/miu_language/decision_nec.lean index d70116a20b054..c6c5b287d1d28 100644 --- a/archive/miu_language/decision_nec.lean +++ b/archive/miu_language/decision_nec.lean @@ -11,6 +11,9 @@ import tactic.ring /-! # Decision procedure: necessary condition +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We introduce a condition `decstr` and show that if a string `en` is `derivable`, then `decstr en` holds. diff --git a/archive/miu_language/decision_suf.lean b/archive/miu_language/decision_suf.lean index 257bea77baa2e..8625a706de47b 100644 --- a/archive/miu_language/decision_suf.lean +++ b/archive/miu_language/decision_suf.lean @@ -9,6 +9,9 @@ import tactic.linarith /-! # Decision procedure - sufficient condition and decidability +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We give a sufficient condition for a string to be derivable in the MIU language. Together with the necessary condition, we use this to prove that `derivable` is an instance of `decidable_pred`. diff --git a/archive/oxford_invariants/2021summer/week3_p1.lean b/archive/oxford_invariants/2021summer/week3_p1.lean index eb30161c3a81c..cf1d4fe0e3ed0 100644 --- a/archive/oxford_invariants/2021summer/week3_p1.lean +++ b/archive/oxford_invariants/2021summer/week3_p1.lean @@ -11,6 +11,9 @@ import data.rat.cast /-! # The Oxford Invariants Puzzle Challenges - Summer 2021, Week 3, Problem 1 +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Original statement Let `n ≥ 3`, `a₁, ..., aₙ` be strictly positive integers such that `aᵢ ∣ aᵢ₋₁ + aᵢ₊₁` for diff --git a/archive/sensitivity.lean b/archive/sensitivity.lean index 995ff612ae9c3..1bd54993a14c4 100644 --- a/archive/sensitivity.lean +++ b/archive/sensitivity.lean @@ -15,6 +15,9 @@ import data.real.sqrt /-! # Huang's sensitivity theorem +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A formalization of Hao Huang's sensitivity theorem: in the hypercube of dimension n ≥ 1, if one colors more than half the vertices then at least one vertex has at least √n colored neighbors. diff --git a/archive/wiedijk_100_theorems/abel_ruffini.lean b/archive/wiedijk_100_theorems/abel_ruffini.lean index 02b6bb4a7d250..16e3595a8dbf8 100644 --- a/archive/wiedijk_100_theorems/abel_ruffini.lean +++ b/archive/wiedijk_100_theorems/abel_ruffini.lean @@ -12,6 +12,9 @@ import ring_theory.eisenstein_criterion /-! # Construction of an algebraic number that is not solvable by radicals. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The main ingredients are: * `solvable_by_rad.is_solvable'` in `field_theory/abel_ruffini` : an irreducible polynomial with an `is_solvable_by_rad` root has solvable Galois group diff --git a/archive/wiedijk_100_theorems/area_of_a_circle.lean b/archive/wiedijk_100_theorems/area_of_a_circle.lean index e3c8d7a0daf6f..e22a6274e698f 100644 --- a/archive/wiedijk_100_theorems/area_of_a_circle.lean +++ b/archive/wiedijk_100_theorems/area_of_a_circle.lean @@ -11,6 +11,9 @@ import measure_theory.measure.lebesgue.integral /-! # Freek № 9: The Area of a Circle +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we show that the area of a disc with nonnegative radius `r` is `π * r^2`. The main tools our proof uses are `volume_region_between_eq_integral`, which allows us to represent the area of the disc as an integral, and `interval_integral.integral_eq_sub_of_has_deriv_at'_of_le`, the diff --git a/archive/wiedijk_100_theorems/ascending_descending_sequences.lean b/archive/wiedijk_100_theorems/ascending_descending_sequences.lean index a0d947a8089a5..14232f57243c4 100644 --- a/archive/wiedijk_100_theorems/ascending_descending_sequences.lean +++ b/archive/wiedijk_100_theorems/ascending_descending_sequences.lean @@ -8,6 +8,9 @@ import data.fintype.powerset /-! # Erdős–Szekeres theorem +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves Theorem 73 from the [100 Theorems List](https://www.cs.ru.nl/~freek/100/), also known as the Erdős–Szekeres theorem: given a sequence of more than `r * s` distinct values, there is an increasing sequence of length longer than `r` or a decreasing sequence of length diff --git a/archive/wiedijk_100_theorems/ballot_problem.lean b/archive/wiedijk_100_theorems/ballot_problem.lean index a72c2e62f322e..f539b407b87e4 100644 --- a/archive/wiedijk_100_theorems/ballot_problem.lean +++ b/archive/wiedijk_100_theorems/ballot_problem.lean @@ -8,6 +8,9 @@ import probability.cond_count /-! # Ballot problem +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves Theorem 30 from the [100 Theorems List](https://www.cs.ru.nl/~freek/100/). The ballot problem asks, if in an election, candidate A receives `p` votes whereas candidate B diff --git a/archive/wiedijk_100_theorems/birthday_problem.lean b/archive/wiedijk_100_theorems/birthday_problem.lean index d4544a61f7f70..4c0ee761bc19f 100644 --- a/archive/wiedijk_100_theorems/birthday_problem.lean +++ b/archive/wiedijk_100_theorems/birthday_problem.lean @@ -10,6 +10,9 @@ import probability.notation /-! # Birthday Problem +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves Theorem 93 from the [100 Theorems List](https://www.cs.ru.nl/~freek/100/). As opposed to the standard probabilistic statement, we instead state the birthday problem diff --git a/archive/wiedijk_100_theorems/cubing_a_cube.lean b/archive/wiedijk_100_theorems/cubing_a_cube.lean index ab4b567df3ac4..f26ce9402eb3d 100644 --- a/archive/wiedijk_100_theorems/cubing_a_cube.lean +++ b/archive/wiedijk_100_theorems/cubing_a_cube.lean @@ -8,6 +8,9 @@ import data.set.finite import data.set.intervals.disjoint /-! +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Proof that a cube (in dimension n ≥ 3) cannot be cubed: There does not exist a partition of a cube into finitely many smaller cubes (at least two) of different sizes. diff --git a/archive/wiedijk_100_theorems/friendship_graphs.lean b/archive/wiedijk_100_theorems/friendship_graphs.lean index adf5f345655a9..222ca67a633ba 100644 --- a/archive/wiedijk_100_theorems/friendship_graphs.lean +++ b/archive/wiedijk_100_theorems/friendship_graphs.lean @@ -12,6 +12,9 @@ import tactic.interval_cases /-! # The Friendship Theorem +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Definitions and Statement - A `friendship` graph is one in which any two distinct vertices have exactly one neighbor in common - A `politician`, at least in the context of this problem, is a vertex in a graph which is adjacent diff --git a/archive/wiedijk_100_theorems/herons_formula.lean b/archive/wiedijk_100_theorems/herons_formula.lean index 100102469b88c..0fb04262d0fe8 100644 --- a/archive/wiedijk_100_theorems/herons_formula.lean +++ b/archive/wiedijk_100_theorems/herons_formula.lean @@ -8,6 +8,9 @@ import geometry.euclidean.triangle /-! # Freek № 57: Heron's Formula +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves Theorem 57 from the [100 Theorems List](https://www.cs.ru.nl/~freek/100/), also known as Heron's formula, which gives the area of a triangle based only on its three sides' lengths. diff --git a/archive/wiedijk_100_theorems/inverse_triangle_sum.lean b/archive/wiedijk_100_theorems/inverse_triangle_sum.lean index 0b12999fbf3d8..b2a77cc1c4be9 100644 --- a/archive/wiedijk_100_theorems/inverse_triangle_sum.lean +++ b/archive/wiedijk_100_theorems/inverse_triangle_sum.lean @@ -9,6 +9,9 @@ import data.real.basic /-! # Sum of the Reciprocals of the Triangular Numbers +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves Theorem 42 from the [100 Theorems List](https://www.cs.ru.nl/~freek/100/). We interpret “triangular numbers” as naturals of the form $\frac{k(k+1)}{2}$ for natural `k`. diff --git a/archive/wiedijk_100_theorems/konigsberg.lean b/archive/wiedijk_100_theorems/konigsberg.lean index 683d7314664db..fe5abdbe02a38 100644 --- a/archive/wiedijk_100_theorems/konigsberg.lean +++ b/archive/wiedijk_100_theorems/konigsberg.lean @@ -9,6 +9,9 @@ import tactic.derive_fintype /-! # The Königsberg bridges problem +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We show that a graph that represents the islands and mainlands of Königsberg and seven bridges between them has no Eulerian trail. -/ diff --git a/archive/wiedijk_100_theorems/partition.lean b/archive/wiedijk_100_theorems/partition.lean index 8796e3b3c2a4d..ac3d5f6c41d05 100644 --- a/archive/wiedijk_100_theorems/partition.lean +++ b/archive/wiedijk_100_theorems/partition.lean @@ -15,6 +15,9 @@ import tactic.congrm /-! # Euler's Partition Theorem +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves Theorem 45 from the [100 Theorems List](https://www.cs.ru.nl/~freek/100/). The theorem concerns the counting of integer partitions -- ways of diff --git a/archive/wiedijk_100_theorems/perfect_numbers.lean b/archive/wiedijk_100_theorems/perfect_numbers.lean index 7d1cbd45655a7..191b9bd62c36e 100644 --- a/archive/wiedijk_100_theorems/perfect_numbers.lean +++ b/archive/wiedijk_100_theorems/perfect_numbers.lean @@ -12,6 +12,9 @@ import ring_theory.multiplicity /-! # Perfect Numbers +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves Theorem 70 from the [100 Theorems List](https://www.cs.ru.nl/~freek/100/). The theorem characterizes even perfect numbers. diff --git a/archive/wiedijk_100_theorems/solution_of_cubic.lean b/archive/wiedijk_100_theorems/solution_of_cubic.lean index 95b58db873e73..57e17a74e712f 100644 --- a/archive/wiedijk_100_theorems/solution_of_cubic.lean +++ b/archive/wiedijk_100_theorems/solution_of_cubic.lean @@ -9,6 +9,9 @@ import ring_theory.polynomial.cyclotomic.roots /-! # The Solution of a Cubic +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves Theorem 37 from the [100 Theorems List](https://www.cs.ru.nl/~freek/100/). In this file, we give the solutions to the cubic equation `a * x^3 + b * x^2 + c * x + d = 0` diff --git a/archive/wiedijk_100_theorems/sum_of_prime_reciprocals_diverges.lean b/archive/wiedijk_100_theorems/sum_of_prime_reciprocals_diverges.lean index 062c3d870d2d6..f52c458f2a87a 100644 --- a/archive/wiedijk_100_theorems/sum_of_prime_reciprocals_diverges.lean +++ b/archive/wiedijk_100_theorems/sum_of_prime_reciprocals_diverges.lean @@ -9,6 +9,9 @@ import data.nat.squarefree /-! # Divergence of the Prime Reciprocal Series +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves Theorem 81 from the [100 Theorems List](https://www.cs.ru.nl/~freek/100/). The theorem states that the sum of the reciprocals of all prime numbers diverges. The formalization follows Erdős's proof by upper and lower estimates. diff --git a/counterexamples/canonically_ordered_comm_semiring_two_mul.lean b/counterexamples/canonically_ordered_comm_semiring_two_mul.lean index 012d49af2dcee..624471d03aeb5 100644 --- a/counterexamples/canonically_ordered_comm_semiring_two_mul.lean +++ b/counterexamples/canonically_ordered_comm_semiring_two_mul.lean @@ -21,6 +21,9 @@ multiplication cannot be strengthened to **strict** monotonicity. Reference: https://leanprover.zulipchat.com/#narrow/stream/113489-new-members/topic/canonically_ordered.20pathology + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ namespace counterexample diff --git a/counterexamples/char_p_zero_ne_char_zero.lean b/counterexamples/char_p_zero_ne_char_zero.lean index a6b9cecde95cc..a567d155f59ec 100644 --- a/counterexamples/char_p_zero_ne_char_zero.lean +++ b/counterexamples/char_p_zero_ne_char_zero.lean @@ -7,6 +7,9 @@ import algebra.char_p.basic /-! # `char_p R 0` and `char_zero R` need not coincide for semirings +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + For rings, the two notions coincide. In fact, `char_p.of_char_zero` shows that `char_zero R` implies `char_p R 0` for any `char_zero` diff --git a/counterexamples/cyclotomic_105.lean b/counterexamples/cyclotomic_105.lean index cfbb5aa23bb94..9499e39f118ac 100644 --- a/counterexamples/cyclotomic_105.lean +++ b/counterexamples/cyclotomic_105.lean @@ -9,6 +9,9 @@ import ring_theory.polynomial.cyclotomic.basic /-! # Not all coefficients of cyclotomic polynomials are -1, 0, or 1 +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We show that not all coefficients of cyclotomic polynomials are equal to `0`, `-1` or `1`, in the theorem `not_forall_coeff_cyclotomic_neg_one_zero_one`. We prove this with the counterexample `coeff_cyclotomic_105 : coeff (cyclotomic 105 ℤ) 7 = -2`. diff --git a/counterexamples/direct_sum_is_internal.lean b/counterexamples/direct_sum_is_internal.lean index 9a4b75ece1807..08d293d5e7eb4 100644 --- a/counterexamples/direct_sum_is_internal.lean +++ b/counterexamples/direct_sum_is_internal.lean @@ -11,6 +11,9 @@ import tactic.fin_cases /-! # Not all complementary decompositions of a module over a semiring make up a direct sum +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This shows that while `ℤ≤0` and `ℤ≥0` are complementary `ℕ`-submodules of `ℤ`, which in turn implies as a collection they are `complete_lattice.independent` and that they span all of `ℤ`, they do not form a decomposition into a direct sum. diff --git a/counterexamples/girard.lean b/counterexamples/girard.lean index 090551161b38a..de764493e7027 100644 --- a/counterexamples/girard.lean +++ b/counterexamples/girard.lean @@ -8,6 +8,9 @@ import logic.basic /-! # Girard's paradox +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Girard's paradox is a proof that `Type : Type` entails a contradiction. We can't say this directly in Lean because `Type : Type 1` and it's not possible to give `Type` a different type via an axiom, so instead we axiomatize the behavior of the Pi type and application if the typing rule for Pi was diff --git a/counterexamples/homogeneous_prime_not_prime.lean b/counterexamples/homogeneous_prime_not_prime.lean index 1e895db644e53..41e0b98c88cdc 100644 --- a/counterexamples/homogeneous_prime_not_prime.lean +++ b/counterexamples/homogeneous_prime_not_prime.lean @@ -10,6 +10,9 @@ import tactic.derive_fintype /-! # A homogeneous prime that is homogeneously prime but not prime +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In `src/ring_theory/graded_algebra/radical.lean`, we assumed that the underline grading is indexed by a `linear_ordered_cancel_add_comm_monoid` to prove that a homogeneous ideal is prime if and only if it is homogeneously prime. This file is aimed to show that even if this assumption isn't strictly diff --git a/counterexamples/linear_order_with_pos_mul_pos_eq_zero.lean b/counterexamples/linear_order_with_pos_mul_pos_eq_zero.lean index 0fac502d4e5e3..b83f9dca5525b 100644 --- a/counterexamples/linear_order_with_pos_mul_pos_eq_zero.lean +++ b/counterexamples/linear_order_with_pos_mul_pos_eq_zero.lean @@ -16,6 +16,9 @@ The order is `0 < ε < 1`. Since `ε ^ 2 = 0`, the product of strictly positive Relevant Zulip chat: https://leanprover.zulipchat.com/#narrow/stream/116395-maths/topic/mul_pos + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ namespace counterexample diff --git a/counterexamples/map_floor.lean b/counterexamples/map_floor.lean index 2fd1b31c7cedc..105a24f6e9574 100644 --- a/counterexamples/map_floor.lean +++ b/counterexamples/map_floor.lean @@ -9,6 +9,9 @@ import data.polynomial.reverse /-! # Floors and ceils aren't preserved under ordered ring homomorphisms +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Intuitively, if `f : α → β` is an ordered ring homomorphism, then floors and ceils should be preserved by `f` because: * `f` preserves the naturals/integers in `α` and `β` because it's a ring hom. diff --git a/counterexamples/phillips.lean b/counterexamples/phillips.lean index eafafd6d169bd..9ef4b12c2300f 100644 --- a/counterexamples/phillips.lean +++ b/counterexamples/phillips.lean @@ -11,6 +11,9 @@ import topology.continuous_function.bounded /-! # A counterexample on Pettis integrability +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + There are several theories of integration for functions taking values in Banach spaces. Bochner integration, requiring approximation by simple functions, is the analogue of the one-dimensional theory. It is very well behaved, but only works for functions with second-countable range. diff --git a/counterexamples/pseudoelement.lean b/counterexamples/pseudoelement.lean index 0e2ddaf8cc23b..a05f885f332ae 100644 --- a/counterexamples/pseudoelement.lean +++ b/counterexamples/pseudoelement.lean @@ -9,6 +9,9 @@ import algebra.category.Module.biproducts /-! # Pseudoelements and pullbacks + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. Borceux claims in Proposition 1.9.5 that the pseudoelement constructed in `category_theory.abelian.pseudoelement.pseudo_pullback` is unique. We show here that this claim is false. This means in particular that we cannot have an extensionality principle for pullbacks in diff --git a/counterexamples/quadratic_form.lean b/counterexamples/quadratic_form.lean index d9f029822dce1..50214d8f685dd 100644 --- a/counterexamples/quadratic_form.lean +++ b/counterexamples/quadratic_form.lean @@ -10,6 +10,9 @@ import data.zmod.basic /-! # `quadratic_form R M` and `subtype bilin_form.is_symm` are distinct notions in characteristic 2 +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The main result of this file is `bilin_form.not_inj_on_to_quadratic_form_is_symm`. The counterexample we use is $B (x, y) (x', y') ↦ xy' + x'y$ where `x y x' y' : zmod 2`. diff --git a/counterexamples/seminorm_lattice_not_distrib.lean b/counterexamples/seminorm_lattice_not_distrib.lean index 95b4e997ea397..ca9b74e8f6934 100644 --- a/counterexamples/seminorm_lattice_not_distrib.lean +++ b/counterexamples/seminorm_lattice_not_distrib.lean @@ -7,6 +7,9 @@ import analysis.seminorm /-! # The lattice of seminorms is not distributive +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We provide an example of three seminorms over the ℝ-vector space ℝ×ℝ which don't satisfy the lattice distributivity property `(p ⊔ q1) ⊓ (p ⊔ q2) ≤ p ⊔ (q1 ⊓ q2)`. diff --git a/counterexamples/sorgenfrey_line.lean b/counterexamples/sorgenfrey_line.lean index 065ecbecd57b5..ac740f1bb34da 100644 --- a/counterexamples/sorgenfrey_line.lean +++ b/counterexamples/sorgenfrey_line.lean @@ -13,6 +13,9 @@ import data.set.intervals.monotone /-! # Sorgenfrey line +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define `sorgenfrey_line` (notation: `ℝₗ`) to be the Sorgenfrey line. It is the real line with the topology space structure generated by half-open intervals `set.Ico a b`. diff --git a/counterexamples/zero_divisors_in_add_monoid_algebras.lean b/counterexamples/zero_divisors_in_add_monoid_algebras.lean index 88e430e411b30..8b52293f1958d 100644 --- a/counterexamples/zero_divisors_in_add_monoid_algebras.lean +++ b/counterexamples/zero_divisors_in_add_monoid_algebras.lean @@ -12,6 +12,9 @@ import data.zmod.basic /-! # Examples of zero-divisors in `add_monoid_algebra`s +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains an easy source of zero-divisors in an `add_monoid_algebra`. If `k` is a field and `G` is an additive group containing a non-zero torsion element, then `add_monoid_algebra k G` contains non-zero zero-divisors: this is lemma `zero_divisors_of_torsion`. diff --git a/src/algebra/expr.lean b/src/algebra/expr.lean index 7c6a862f0abfc..89641583e33bd 100644 --- a/src/algebra/expr.lean +++ b/src/algebra/expr.lean @@ -7,6 +7,9 @@ import tactic.core /-! ### Helpers to invoke functions involving algebra at tactic time +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + It's not clear whether using `instance_cache` is a sensible choice here. In particular, we need to use these tactics below when the algebraic instances are local variables that aren't in the "real" instance cache (the one used by `tactic.reset_instance_cache`). diff --git a/src/algebraic_geometry/morphisms/finite_type.lean b/src/algebraic_geometry/morphisms/finite_type.lean index 65ab36ddc83ec..ce417476ad834 100644 --- a/src/algebraic_geometry/morphisms/finite_type.lean +++ b/src/algebraic_geometry/morphisms/finite_type.lean @@ -9,6 +9,9 @@ import ring_theory.ring_hom.finite_type /-! # Morphisms of finite type +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A morphism of schemes `f : X ⟶ Y` is locally of finite type if for each affine `U ⊆ Y` and `V ⊆ f ⁻¹' U`, The induced map `Γ(Y, U) ⟶ Γ(X, V)` is of finite type. diff --git a/src/algebraic_geometry/morphisms/ring_hom_properties.lean b/src/algebraic_geometry/morphisms/ring_hom_properties.lean index e389bdb008d2d..d30f2f44c77ec 100644 --- a/src/algebraic_geometry/morphisms/ring_hom_properties.lean +++ b/src/algebraic_geometry/morphisms/ring_hom_properties.lean @@ -10,6 +10,9 @@ import ring_theory.local_properties # Properties of morphisms from properties of ring homs. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We provide the basic framework for talking about properties of morphisms that come from properties of ring homs. For `P` a property of ring homs, we have two ways of defining a property of scheme morphisms: diff --git a/src/data/matrix/auto.lean b/src/data/matrix/auto.lean index 1a351ef13828c..0d75a92737462 100644 --- a/src/data/matrix/auto.lean +++ b/src/data/matrix/auto.lean @@ -8,6 +8,9 @@ import data.matrix.reflection /-! # Automatically generated lemmas for working with concrete matrices +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains "magic" lemmas which autogenerate to the correct size of matrix. For instance, `matrix.of_mul_of_fin` can be used as: ```lean diff --git a/src/geometry/manifold/instances/sphere.lean b/src/geometry/manifold/instances/sphere.lean index e274632af5e28..71158b59e4aa9 100644 --- a/src/geometry/manifold/instances/sphere.lean +++ b/src/geometry/manifold/instances/sphere.lean @@ -15,6 +15,9 @@ import geometry.manifold.cont_mdiff_mfderiv /-! # Manifold structure on the sphere +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines stereographic projection from the sphere in an inner product space `E`, and uses it to put a smooth manifold structure on the sphere. From bf2428c9486c407ca38b5b3fb10b87dad0bc99fa Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sun, 16 Jul 2023 19:01:39 +0000 Subject: [PATCH 30/61] feat(order/irreducible): Sup-irreducible elements (#18999) Define sup- and inf- irreducible and prime elements in a lattice. --- src/data/fin/tuple/bubble_sort_induction.lean | 3 +- src/data/fintype/card.lean | 20 +- src/order/irreducible.lean | 229 ++++++++++++++++++ 3 files changed, 239 insertions(+), 13 deletions(-) create mode 100644 src/order/irreducible.lean diff --git a/src/data/fin/tuple/bubble_sort_induction.lean b/src/data/fin/tuple/bubble_sort_induction.lean index 8b496b9e8de09..8e7a0ca355e02 100644 --- a/src/data/fin/tuple/bubble_sort_induction.lean +++ b/src/data/fin/tuple/bubble_sort_induction.lean @@ -39,8 +39,7 @@ lemma bubble_sort_induction' {n : ℕ} {α : Type*} [linear_order α] {f : fin n P (f ∘ sort f) := begin letI := @preorder.lift _ (lex (fin n → α)) _ (λ σ : equiv.perm (fin n), to_lex (f ∘ σ)), - refine @well_founded.induction_bot' _ _ _ - (@finite.preorder.well_founded_lt (equiv.perm (fin n)) _ _) + refine @well_founded.induction_bot' _ _ _ (is_well_founded.wf : well_founded (<)) (equiv.refl _) (sort f) P (λ σ, f ∘ σ) (λ σ hσ hfσ, _) hf, obtain ⟨i, j, hij₁, hij₂⟩ := antitone_pair_of_not_sorted' hσ, exact ⟨σ * equiv.swap i j, pi.lex_desc hij₁ hij₂, h σ i j hij₁ hij₂ hfσ⟩, diff --git a/src/data/fintype/card.lean b/src/data/fintype/card.lean index 43d7bf3b7f0f5..be6e5f6be0c7c 100644 --- a/src/data/fintype/card.lean +++ b/src/data/fintype/card.lean @@ -726,17 +726,15 @@ have ∀ x y, r x y → (univ.filter (λ z, r z x)).card < (univ.filter (λ z, r exact ⟨λ z hzx, trans hzx hxy, not_forall_of_exists_not ⟨x, not_imp.2 ⟨hxy, irrefl x⟩⟩⟩, subrelation.wf this (measure_wf _) -lemma preorder.well_founded_lt [preorder α] : well_founded ((<) : α → α → Prop) := -well_founded_of_trans_of_irrefl _ - -lemma preorder.well_founded_gt [preorder α] : well_founded ((>) : α → α → Prop) := -well_founded_of_trans_of_irrefl _ - -@[priority 10] instance linear_order.is_well_order_lt [linear_order α] : is_well_order α (<) := -{ wf := preorder.well_founded_lt } - -@[priority 10] instance linear_order.is_well_order_gt [linear_order α] : is_well_order α (>) := -{ wf := preorder.well_founded_gt } +@[priority 100] -- See note [lower instance priority] +instance finite.to_well_founded_lt [preorder α] : well_founded_lt α := +⟨well_founded_of_trans_of_irrefl _⟩ +@[priority 100] -- See note [lower instance priority] +instance finite.to_well_founded_gt [preorder α] : well_founded_gt α := +⟨well_founded_of_trans_of_irrefl _⟩ + +@[priority 10] instance linear_order.is_well_order_lt [linear_order α] : is_well_order α (<) := {} +@[priority 10] instance linear_order.is_well_order_gt [linear_order α] : is_well_order α (>) := {} end finite diff --git a/src/order/irreducible.lean b/src/order/irreducible.lean new file mode 100644 index 0000000000000..07bec21729d7c --- /dev/null +++ b/src/order/irreducible.lean @@ -0,0 +1,229 @@ +/- +Copyright (c) 2023 Yaël Dillies. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies +-/ +import data.finset.lattice +import data.fintype.card + +/-! +# Irreducible and prime elements in an order + +This file defines irreducible and prime elements in an order and shows that in a well-founded +lattice every element decomposes as a supremum of irreducible elements. + +An element is sup-irreducible (resp. inf-irreducible) if it isn't `⊥` and can't be written as the +supremum of any strictly smaller elements. An element is sup-prime (resp. inf-prime) if it isn't `⊥` +and is greater than the supremum of any two elements less than it. + +Primality implies irreducibility in general. The converse only holds in distributive lattices. +Both hold for all (non-minimal) elements in a linear order. + +## Main declarations + +* `sup_irred a`: Sup-irreducibility, `a` isn't minimal and `a = b ⊔ c → a = b ∨ a = c` +* `inf_irred a`: Inf-irreducibility, `a` isn't maximal and `a = b ⊓ c → a = b ∨ a = c` +* `sup_prime a`: Sup-primality, `a` isn't minimal and `a ≤ b ⊔ c → a ≤ b ∨ a ≤ c` +* `inf_irred a`: Inf-primality, `a` isn't maximal and `a ≥ b ⊓ c → a ≥ b ∨ a ≥ c` +* `exists_sup_irred_decomposition`/`exists_inf_irred_decomposition`: Decomposition into irreducibles + in a well-founded semilattice. +-/ + +open finset order_dual + +variables {ι α : Type*} + +/-! ### Irreducible and prime elements -/ + +section semilattice_sup +variables [semilattice_sup α] {a b c : α} + +/-- A sup-irreducible element is a non-bottom element which isn't the supremum of anything smaller. +-/ +def sup_irred (a : α) : Prop := ¬ is_min a ∧ ∀ ⦃b c⦄, b ⊔ c = a → b = a ∨ c = a + +/-- A sup-prime element is a non-bottom element which isn't less than the supremum of anything +smaller. -/ +def sup_prime (a : α) : Prop := ¬ is_min a ∧ ∀ ⦃b c⦄, a ≤ b ⊔ c → a ≤ b ∨ a ≤ c + +lemma sup_irred.not_is_min (ha : sup_irred a) : ¬ is_min a := ha.1 +lemma sup_prime.not_is_min (ha : sup_prime a) : ¬ is_min a := ha.1 +lemma is_min.not_sup_irred (ha : is_min a) : ¬ sup_irred a := λ h, h.1 ha +lemma is_min.not_sup_prime (ha : is_min a) : ¬ sup_prime a := λ h, h.1 ha + +@[simp] lemma not_sup_irred : ¬ sup_irred a ↔ is_min a ∨ ∃ b c, b ⊔ c = a ∧ b < a ∧ c < a := +begin + rw [sup_irred, not_and_distrib], + push_neg, + rw exists₂_congr, + simp [@eq_comm _ _ a] { contextual := tt }, +end + +@[simp] lemma not_sup_prime : ¬ sup_prime a ↔ is_min a ∨ ∃ b c, a ≤ b ⊔ c ∧ ¬ a ≤ b ∧ ¬ a ≤ c := +by { rw [sup_prime, not_and_distrib], push_neg, refl } + +protected lemma sup_prime.sup_irred : sup_prime a → sup_irred a := +and.imp_right $ λ h b c ha, by simpa [←ha] using h ha.ge + +lemma sup_prime.le_sup (ha : sup_prime a) : a ≤ b ⊔ c ↔ a ≤ b ∨ a ≤ c := +⟨λ h, ha.2 h, λ h, h.elim le_sup_of_le_left le_sup_of_le_right⟩ + +variables [order_bot α] {s : finset ι} {f : ι → α} + +@[simp] lemma not_sup_irred_bot : ¬ sup_irred (⊥ : α) := is_min_bot.not_sup_irred +@[simp] lemma not_sup_prime_bot : ¬ sup_prime (⊥ : α) := is_min_bot.not_sup_prime + +lemma sup_irred.ne_bot (ha : sup_irred a) : a ≠ ⊥ := by { rintro rfl, exact not_sup_irred_bot ha } +lemma sup_prime.ne_bot (ha : sup_prime a) : a ≠ ⊥ := by { rintro rfl, exact not_sup_prime_bot ha } + +lemma sup_irred.finset_sup_eq (ha : sup_irred a) (h : s.sup f = a) : ∃ i ∈ s, f i = a := +begin + classical, + induction s using finset.induction with i s hi ih, + { simpa [ha.ne_bot] using h.symm }, + simp only [exists_prop, exists_mem_insert] at ⊢ ih, + rw sup_insert at h, + exact (ha.2 h).imp_right ih, +end + +lemma sup_prime.le_finset_sup (ha : sup_prime a) : a ≤ s.sup f ↔ ∃ i ∈ s, a ≤ f i := +begin + classical, + induction s using finset.induction with i s hi ih, + { simp [ha.ne_bot] }, + { simp only [exists_prop, exists_mem_insert, sup_insert, ha.le_sup, ih] } +end + +variables [well_founded_lt α] + +/-- In a well-founded lattice, any element is the supremum of finitely many sup-irreducible +elements. This is the order-theoretic analogue of prime factorisation. -/ +lemma exists_sup_irred_decomposition (a : α) : + ∃ s : finset α, s.sup id = a ∧ ∀ ⦃b⦄, b ∈ s → sup_irred b := +begin + classical, + apply well_founded_lt.induction a _, + clear a, + rintro a ih, + by_cases ha : sup_irred a, + { exact ⟨{a}, by simp [ha]⟩ }, + rw not_sup_irred at ha, + obtain ha | ⟨b, c, rfl, hb, hc⟩ := ha, + { exact ⟨∅, by simp [ha.eq_bot]⟩ }, + obtain ⟨s, rfl, hs⟩ := ih _ hb, + obtain ⟨t, rfl, ht⟩ := ih _ hc, + exact ⟨s ∪ t, sup_union, forall_mem_union.2 ⟨hs, ht⟩⟩, +end + +end semilattice_sup + +section semilattice_inf +variables [semilattice_inf α] {a b c : α} + +/-- An inf-irreducible element is a non-top element which isn't the infimum of anything bigger. -/ +def inf_irred (a : α) : Prop := ¬ is_max a ∧ ∀ ⦃b c⦄, b ⊓ c = a → b = a ∨ c = a + +/-- An inf-prime element is a non-top element which isn't bigger than the infimum of anything +bigger. -/ +def inf_prime (a : α) : Prop := ¬ is_max a ∧ ∀ ⦃b c⦄, b ⊓ c ≤ a → b ≤ a ∨ c ≤ a + +@[simp] lemma is_max.not_inf_irred (ha : is_max a) : ¬ inf_irred a := λ h, h.1 ha +@[simp] lemma is_max.not_inf_prime (ha : is_max a) : ¬ inf_prime a := λ h, h.1 ha + +@[simp] lemma not_inf_irred : ¬ inf_irred a ↔ is_max a ∨ ∃ b c, b ⊓ c = a ∧ a < b ∧ a < c := +@not_sup_irred αᵒᵈ _ _ + +@[simp] lemma not_inf_prime : ¬ inf_prime a ↔ is_max a ∨ ∃ b c, b ⊓ c ≤ a ∧ ¬ b ≤ a ∧ ¬ c ≤ a := +@not_sup_prime αᵒᵈ _ _ + +protected lemma inf_prime.inf_irred : inf_prime a → inf_irred a := +and.imp_right $ λ h b c ha, by simpa [←ha] using h ha.le + +lemma inf_prime.inf_le (ha : inf_prime a) : b ⊓ c ≤ a ↔ b ≤ a ∨ c ≤ a := +⟨λ h, ha.2 h, λ h, h.elim inf_le_of_left_le inf_le_of_right_le⟩ + +variables [order_top α] {s : finset ι} {f : ι → α} + +@[simp] lemma not_inf_irred_top : ¬ inf_irred (⊤ : α) := is_max_top.not_inf_irred +@[simp] lemma not_inf_prime_top : ¬ inf_prime (⊤ : α) := is_max_top.not_inf_prime + +lemma inf_irred.ne_top (ha : inf_irred a) : a ≠ ⊤ := by { rintro rfl, exact not_inf_irred_top ha } +lemma inf_prime.ne_top (ha : inf_prime a) : a ≠ ⊤ := by { rintro rfl, exact not_inf_prime_top ha } + +lemma inf_irred.finset_inf_eq : inf_irred a → s.inf f = a → ∃ i ∈ s, f i = a := +@sup_irred.finset_sup_eq _ αᵒᵈ _ _ _ _ _ + +lemma inf_prime.finset_inf_le (ha : inf_prime a) : s.inf f ≤ a ↔ ∃ i ∈ s, f i ≤ a := +@sup_prime.le_finset_sup _ αᵒᵈ _ _ _ _ _ ha + +variables [well_founded_gt α] + +/-- In a cowell-founded lattice, any element is the infimum of finitely many inf-irreducible +elements. This is the order-theoretic analogue of prime factorisation. -/ +lemma exists_inf_irred_decomposition (a : α) : + ∃ s : finset α, s.inf id = a ∧ ∀ ⦃b⦄, b ∈ s → inf_irred b := +@exists_sup_irred_decomposition αᵒᵈ _ _ _ _ + +end semilattice_inf + +section semilattice_sup +variables [semilattice_sup α] + +@[simp] lemma inf_irred_to_dual {a : α} : inf_irred (to_dual a) ↔ sup_irred a := iff.rfl +@[simp] lemma inf_prime_to_dual {a : α} : inf_prime (to_dual a) ↔ sup_prime a := iff.rfl +@[simp] lemma sup_irred_of_dual {a : αᵒᵈ} : sup_irred (of_dual a) ↔ inf_irred a := iff.rfl +@[simp] lemma sup_prime_of_dual {a : αᵒᵈ} : sup_prime (of_dual a) ↔ inf_prime a := iff.rfl + +alias inf_irred_to_dual ↔ _ sup_irred.dual +alias inf_prime_to_dual ↔ _ sup_prime.dual +alias sup_irred_of_dual ↔ _ inf_irred.of_dual +alias sup_prime_of_dual ↔ _ inf_prime.of_dual + +end semilattice_sup + +section semilattice_inf +variables [semilattice_inf α] + +@[simp] lemma sup_irred_to_dual {a : α} : sup_irred (to_dual a) ↔ inf_irred a := iff.rfl +@[simp] lemma sup_prime_to_dual {a : α} : sup_prime (to_dual a) ↔ inf_prime a := iff.rfl +@[simp] lemma inf_irred_of_dual {a : αᵒᵈ} : inf_irred (of_dual a) ↔ sup_irred a := iff.rfl +@[simp] lemma inf_prime_of_dual {a : αᵒᵈ} : inf_prime (of_dual a) ↔ sup_prime a := iff.rfl + +alias sup_irred_to_dual ↔ _ inf_irred.dual +alias sup_prime_to_dual ↔ _ inf_prime.dual +alias inf_irred_of_dual ↔ _ sup_irred.of_dual +alias inf_prime_of_dual ↔ _ sup_prime.of_dual + +end semilattice_inf + +section distrib_lattice +variables [distrib_lattice α] {a b c : α} + +@[simp] lemma sup_prime_iff_sup_irred : sup_prime a ↔ sup_irred a := +⟨sup_prime.sup_irred, and.imp_right $ λ h b c, + by { simp_rw [←inf_eq_left, inf_sup_left], exact @h _ _ }⟩ + +@[simp] lemma inf_prime_iff_inf_irred : inf_prime a ↔ inf_irred a := +⟨inf_prime.inf_irred, and.imp_right $ λ h b c, + by { simp_rw [←sup_eq_left, sup_inf_left], exact @h _ _ }⟩ + +alias sup_prime_iff_sup_irred ↔ _ sup_irred.sup_prime +alias inf_prime_iff_inf_irred ↔ _ inf_irred.inf_prime + +attribute [protected] sup_irred.sup_prime inf_irred.inf_prime + +end distrib_lattice + +section linear_order +variables [linear_order α] {a : α} + +@[simp] lemma sup_prime_iff_not_is_min : sup_prime a ↔ ¬ is_min a := and_iff_left $ by simp +@[simp] lemma inf_prime_iff_not_is_max : inf_prime a ↔ ¬ is_max a := and_iff_left $ by simp + +@[simp] lemma sup_irred_iff_not_is_min : sup_irred a ↔ ¬ is_min a := +and_iff_left $ λ _ _, by simpa only [sup_eq_max, max_eq_iff] using or.imp and.left and.left + +@[simp] lemma inf_irred_iff_not_is_max : inf_irred a ↔ ¬ is_max a := +and_iff_left $ λ _ _, by simpa only [inf_eq_min, min_eq_iff] using or.imp and.left and.left + +end linear_order From 8ea5598db6caeddde6cb734aa179cc2408dbd345 Mon Sep 17 00:00:00 2001 From: negiizhao Date: Wed, 19 Jul 2023 06:33:41 +0000 Subject: [PATCH 31/61] chore(set_theory/ordinal/basic): golf (#18547) Co-authored-by: Eric Wieser --- src/order/initial_seg.lean | 17 +++++++++ src/set_theory/ordinal/basic.lean | 60 ++++++++++--------------------- 2 files changed, 35 insertions(+), 42 deletions(-) diff --git a/src/order/initial_seg.lean b/src/order/initial_seg.lean index 2d3ef20adcc9d..e0365f4bc7537 100644 --- a/src/order/initial_seg.lean +++ b/src/order/initial_seg.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Floris van Doorn -/ +import logic.equiv.set import order.rel_iso.set import order.well_founded @@ -311,6 +312,22 @@ def of_element {α : Type*} (r : α → α → Prop) (a : α) : subrel r {b | r @[simp] theorem of_element_top {α : Type*} (r : α → α → Prop) (a : α) : (of_element r a).top = a := rfl +/-- For any principal segment `r ≺i s`, there is a `subrel` of `s` order isomorphic to `r`. -/ +@[simps symm_apply] +noncomputable def subrel_iso (f : r ≺i s) : subrel s {b | s b f.top} ≃r r := +rel_iso.symm +{ to_equiv := ((equiv.of_injective f f.injective).trans (equiv.set_congr + (funext (λ x, propext f.down.symm)))), + map_rel_iff' := λ a₁ a₂, f.map_rel_iff } + +@[simp] theorem apply_subrel_iso (f : r ≺i s) (b : {b | s b f.top}) : + f (f.subrel_iso b) = b := +equiv.apply_of_injective_symm f.injective _ + +@[simp] theorem subrel_iso_apply (f : r ≺i s) (a : α) : + f.subrel_iso ⟨f a, f.down.mpr ⟨a, rfl⟩⟩ = a := +equiv.of_injective_symm_apply f.injective _ + /-- Restrict the codomain of a principal segment -/ def cod_restrict (p : set β) (f : r ≺i s) (H : ∀ a, f a ∈ p) (H₂ : f.top ∈ p) : r ≺i subrel s p := diff --git a/src/set_theory/ordinal/basic.lean b/src/set_theory/ordinal/basic.lean index 2e146285a1f99..c9ccd28a73e48 100644 --- a/src/set_theory/ordinal/basic.lean +++ b/src/set_theory/ordinal/basic.lean @@ -358,28 +358,29 @@ injective_of_increasing r (<) (typein r) (λ x y, (typein_lt_typein r).2) {a b} : typein r a = typein r b ↔ a = b := (typein_injective r).eq_iff +/-- Principal segment version of the `typein` function, embedding a well order into + ordinals as a principal segment. -/ +def typein.principal_seg (r : α → α → Prop) [is_well_order α r] : + r ≺i ((<) : ordinal → ordinal → Prop) := +⟨⟨⟨typein r, typein_injective r⟩, λ a b, typein_lt_typein r⟩, type r, + λ o, ⟨typein_surj r, λ ⟨a, h⟩, h ▸ typein_lt_type r a⟩⟩ + +@[simp] theorem typein.principal_seg_coe (r : α → α → Prop) [is_well_order α r] : + (typein.principal_seg r : α → ordinal) = typein r := rfl + /-! ### Enumerating elements in a well-order with ordinals. -/ /-- `enum r o h` is the `o`-th element of `α` ordered by `r`. That is, `enum` maps an initial segment of the ordinals, those less than the order type of `r`, to the elements of `α`. -/ -def enum (r : α → α → Prop) [is_well_order α r] (o) : o < type r → α := -quot.rec_on o (λ ⟨β, s, _⟩ h, (classical.choice h).top) $ -λ ⟨β, s, _⟩ ⟨γ, t, _⟩ ⟨h⟩, begin - resetI, refine funext (λ (H₂ : type t < type r), _), - have H₁ : type s < type r, {rwa type_eq.2 ⟨h⟩}, - have : ∀ {o e} (H : o < type r), @@eq.rec - (λ (o : ordinal), o < type r → α) - (λ (h : type s < type r), (classical.choice h).top) - e H = (classical.choice H₁).top, {intros, subst e}, - exact (this H₂).trans (principal_seg.top_eq h - (classical.choice H₁) (classical.choice H₂)) -end +def enum (r : α → α → Prop) [is_well_order α r] (o) (h : o < type r) : α := +(typein.principal_seg r).subrel_iso ⟨o, h⟩ theorem enum_type {α β} {r : α → α → Prop} {s : β → β → Prop} [is_well_order α r] [is_well_order β s] (f : s ≺i r) {h : type s < type r} : enum r (type s) h = f.top := -principal_seg.top_eq (rel_iso.refl _) _ _ +(typein.principal_seg r).injective $ + ((typein.principal_seg r).apply_subrel_iso _).trans (typein_top _).symm @[simp] theorem enum_typein (r : α → α → Prop) [is_well_order α r] (a : α) : enum r (typein r a) (typein_lt_type r a) = a := @@ -412,13 +413,8 @@ lemma rel_iso_enum {α β : Type u} {r : α → α → Prop} {s : β → β → rel_iso_enum' _ _ _ _ theorem lt_wf : @well_founded ordinal (<) := -⟨λ a, induction_on a $ λ α r wo, by exactI -suffices ∀ a, acc (<) (typein r a), from -⟨_, λ o h, let ⟨a, e⟩ := typein_surj r h in e ▸ this a⟩, -λ a, acc.rec_on (wo.wf.apply a) $ λ x H IH, ⟨_, λ o h, begin - rcases typein_surj r (lt_trans h (typein_lt_type r _)) with ⟨b, rfl⟩, - exact IH _ ((typein_lt_typein r).1 h) -end⟩⟩ +well_founded_iff_well_founded_subrel.mpr $ λ a, induction_on a $ λ α r wo, by exactI + rel_hom_class.well_founded (typein.principal_seg r).subrel_iso wo.wf instance : has_well_founded ordinal := ⟨(<), lt_wf⟩ @@ -428,18 +424,6 @@ lemma induction {p : ordinal.{u} → Prop} (i : ordinal.{u}) (h : ∀ j, (∀ k, k < j → p k) → p j) : p i := lt_wf.induction i h -/-- Principal segment version of the `typein` function, embedding a well order into - ordinals as a principal segment. -/ -def typein.principal_seg {α : Type u} (r : α → α → Prop) [is_well_order α r] : - @principal_seg α ordinal.{u} r (<) := -⟨rel_embedding.of_monotone (typein r) - (λ a b, (typein_lt_typein r).2), type r, λ b, - ⟨λ h, ⟨enum r _ h, typein_enum r h⟩, - λ ⟨a, e⟩, e ▸ typein_lt_type _ _⟩⟩ - -@[simp] theorem typein.principal_seg_coe (r : α → α → Prop) [is_well_order α r] : - (typein.principal_seg r : α → ordinal) = typein r := rfl - /-! ### Cardinality of ordinals -/ /-- The cardinal of an ordinal is the cardinality of any type on which a relation with that order @@ -820,21 +804,13 @@ by { rw [←enum_typein (<) a, enum_le_enum', ←lt_succ_iff], apply typein_lt_s @[simp] theorem enum_inj {r : α → α → Prop} [is_well_order α r] {o₁ o₂ : ordinal} (h₁ : o₁ < type r) (h₂ : o₂ < type r) : enum r o₁ h₁ = enum r o₂ h₂ ↔ o₁ = o₂ := -⟨λ h, begin - by_contra hne, - cases lt_or_gt_of_ne hne with hlt hlt; - apply (is_well_order.is_irrefl r).1, - { rwa [←@enum_lt_enum α r _ o₁ o₂ h₁ h₂, h] at hlt }, - { change _ < _ at hlt, rwa [←@enum_lt_enum α r _ o₂ o₁ h₂ h₁, h] at hlt } -end, λ h, by simp_rw h⟩ +(typein.principal_seg r).subrel_iso.injective.eq_iff.trans subtype.mk_eq_mk /-- A well order `r` is order isomorphic to the set of ordinals smaller than `type r`. -/ @[simps] def enum_iso (r : α → α → Prop) [is_well_order α r] : subrel (<) (< type r) ≃r r := { to_fun := λ x, enum r x.1 x.2, inv_fun := λ x, ⟨typein r x, typein_lt_type r x⟩, - left_inv := λ ⟨o, h⟩, subtype.ext_val (typein_enum _ _), - right_inv := λ h, enum_typein _ _, - map_rel_iff' := by { rintros ⟨a, _⟩ ⟨b, _⟩, apply enum_lt_enum } } + ..(typein.principal_seg r).subrel_iso } /-- The order isomorphism between ordinals less than `o` and `o.out.α`. -/ @[simps] noncomputable def enum_iso_out (o : ordinal) : set.Iio o ≃o o.out.α := From 573eea921b01c49712ac02471911df0719297349 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Sat, 22 Jul 2023 09:12:42 +0000 Subject: [PATCH 32/61] chore(*): add mathlib4 synchronization comments (#19239) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `order.irreducible` --- src/order/irreducible.lean | 3 +++ 1 file changed, 3 insertions(+) diff --git a/src/order/irreducible.lean b/src/order/irreducible.lean index 07bec21729d7c..0862574732a13 100644 --- a/src/order/irreducible.lean +++ b/src/order/irreducible.lean @@ -9,6 +9,9 @@ import data.fintype.card /-! # Irreducible and prime elements in an order +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines irreducible and prime elements in an order and shows that in a well-founded lattice every element decomposes as a supremum of irreducible elements. From 8047de4d911cdef39c2d646165eea972f7f9f539 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sat, 22 Jul 2023 18:41:27 +0000 Subject: [PATCH 33/61] feat(topology/metric_space/basic): decomposition of a "sphere" hypercube (#18875) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The main result here is `sphere x r = (⋃ i : β, function.eval i ⁻¹' sphere (x i) r) ∩ closed_ball x r`, which attempts to express that you can form the surface of a cube by taking the union of the faces in each axis. The `prod` result, `sphere (x, y) r = sphere x r ×ˢ closed_ball y r ∪ closed_ball x r ×ˢ sphere y r`, is a little easier to follow. I can imagine these being useful if we wanted to prove that the surface area of a "sphere" in `fin 3 -> R` was `24*r*r`! Co-authored-by: Yaël Dillies --- src/topology/metric_space/basic.lean | 66 +++++++++++++++++++++++++++- 1 file changed, 65 insertions(+), 1 deletion(-) diff --git a/src/topology/metric_space/basic.lean b/src/topology/metric_space/basic.lean index 4a34ea24186fa..de0baf9a61b0a 100644 --- a/src/topology/metric_space/basic.lean +++ b/src/topology/metric_space/basic.lean @@ -439,6 +439,12 @@ theorem mem_sphere' : y ∈ sphere x ε ↔ dist x y = ε := by rw [dist_comm, m theorem ne_of_mem_sphere (h : y ∈ sphere x ε) (hε : ε ≠ 0) : y ≠ x := by { contrapose! hε, symmetry, simpa [hε] using h } +theorem nonneg_of_mem_sphere (hy : y ∈ sphere x ε) : 0 ≤ ε := +dist_nonneg.trans_eq hy + +@[simp] theorem sphere_eq_empty_of_neg (hε : ε < 0) : sphere x ε = ∅ := +set.eq_empty_iff_forall_not_mem.mpr $ λ y hy, (nonneg_of_mem_sphere hy).not_lt hε + theorem sphere_eq_empty_of_subsingleton [subsingleton α] (hε : ε ≠ 0) : sphere x ε = ∅ := set.eq_empty_iff_forall_not_mem.mpr $ λ y hy, ne_of_mem_sphere hy hε (subsingleton.elim _ _) @@ -456,6 +462,10 @@ show dist x x ≤ ε, by rw dist_self; assumption @[simp] lemma closed_ball_eq_empty : closed_ball x ε = ∅ ↔ ε < 0 := by rw [← not_nonempty_iff_eq_empty, nonempty_closed_ball, not_le] +/-- Closed balls and spheres coincide when the radius is non-positive -/ +theorem closed_ball_eq_sphere_of_nonpos (hε : ε ≤ 0) : closed_ball x ε = sphere x ε := +set.ext $ λ _, (hε.trans dist_nonneg).le_iff_eq + theorem ball_subset_closed_ball : ball x ε ⊆ closed_ball x ε := assume y (hy : _ < _), le_of_lt hy @@ -1562,6 +1572,20 @@ theorem closed_ball_prod_same (x : α) (y : β) (r : ℝ) : closed_ball x r ×ˢ closed_ball y r = closed_ball (x, y) r := ext $ λ z, by simp [prod.dist_eq] +theorem sphere_prod (x : α × β) (r : ℝ) : + sphere x r = sphere x.1 r ×ˢ closed_ball x.2 r ∪ closed_ball x.1 r ×ˢ sphere x.2 r := +begin + obtain hr | rfl | hr := lt_trichotomy r 0, + { simp [hr], }, + { cases x, + simp_rw [←closed_ball_eq_sphere_of_nonpos le_rfl, union_self, closed_ball_prod_same] }, + { ext ⟨x', y'⟩, + simp_rw [set.mem_union, set.mem_prod, metric.mem_closed_ball, metric.mem_sphere, + prod.dist_eq, max_eq_iff], + refine or_congr (and_congr_right _) ((and_comm _ _).trans (and_congr_left _)), + all_goals { rintro rfl, refl } }, +end + end prod theorem uniform_continuous_dist : uniform_continuous (λp:α×α, dist p.1 p.2) := @@ -1802,11 +1826,25 @@ lemma nndist_pi_le_iff {f g : Πb, π b} {r : ℝ≥0} : nndist f g ≤ r ↔ ∀b, nndist (f b) (g b) ≤ r := by simp [nndist_pi_def] +lemma nndist_pi_lt_iff {f g : Πb, π b} {r : ℝ≥0} (hr : 0 < r) : + nndist f g < r ↔ ∀ b, nndist (f b) (g b) < r := +by simp [nndist_pi_def, finset.sup_lt_iff (show ⊥ < r, from hr)] + +lemma nndist_pi_eq_iff {f g : Π b, π b} {r : ℝ≥0} (hr : 0 < r) : + nndist f g = r ↔ (∃ i, nndist (f i) (g i) = r) ∧ ∀ b, nndist (f b) (g b) ≤ r := +begin + rw [eq_iff_le_not_lt, nndist_pi_lt_iff hr, nndist_pi_le_iff, not_forall, and_comm], + simp_rw [not_lt, and.congr_left_iff, le_antisymm_iff], + intro h, + refine exists_congr (λ b, _), + apply (and_iff_right $ h _).symm, +end + lemma dist_pi_lt_iff {f g : Πb, π b} {r : ℝ} (hr : 0 < r) : dist f g < r ↔ ∀b, dist (f b) (g b) < r := begin lift r to ℝ≥0 using hr.le, - simp [dist_pi_def, finset.sup_lt_iff (show ⊥ < r, from hr)], + exact nndist_pi_lt_iff hr, end lemma dist_pi_le_iff {f g : Πb, π b} {r : ℝ} (hr : 0 ≤ r) : @@ -1816,6 +1854,13 @@ begin exact nndist_pi_le_iff end +lemma dist_pi_eq_iff {f g : Πb, π b} {r : ℝ} (hr : 0 < r) : + dist f g = r ↔ (∃ i, dist (f i) (g i) = r) ∧ ∀ b, dist (f b) (g b) ≤ r := +begin + lift r to ℝ≥0 using hr.le, + simp_rw [←coe_nndist, nnreal.coe_eq, nndist_pi_eq_iff hr, nnreal.coe_le_coe], +end + lemma dist_pi_le_iff' [nonempty β] {f g : Π b, π b} {r : ℝ} : dist f g ≤ r ↔ ∀ b, dist (f b) (g b) ≤ r := begin @@ -1867,6 +1912,25 @@ lemma closed_ball_pi' [nonempty β] (x : Π b, π b) (r : ℝ) : closed_ball x r = set.pi univ (λ b, closed_ball (x b) r) := (le_or_lt 0 r).elim (closed_ball_pi x) $ λ hr, by simp [closed_ball_eq_empty.2 hr] +/-- A sphere in a product space is a union of spheres on each component restricted to the closed +ball. -/ +lemma sphere_pi (x : Πb, π b) {r : ℝ} (h : 0 < r ∨ nonempty β) : + sphere x r = (⋃ i : β, function.eval i ⁻¹' sphere (x i) r) ∩ closed_ball x r := +begin + obtain hr | rfl | hr := lt_trichotomy r 0, + { simp [hr], }, + { rw [closed_ball_eq_sphere_of_nonpos le_rfl, eq_comm, set.inter_eq_right_iff_subset], + letI := h.resolve_left (lt_irrefl _), + inhabit β, + refine subset_Union_of_subset default _, + intros x hx, + replace hx := hx.le, + rw [dist_pi_le_iff le_rfl] at hx, + exact le_antisymm (hx default) dist_nonneg }, + { ext, + simp [dist_pi_eq_iff hr, dist_pi_le_iff hr.le] }, +end + @[simp] lemma fin.nndist_insert_nth_insert_nth {n : ℕ} {α : fin (n + 1) → Type*} [Π i, pseudo_metric_space (α i)] (i : fin (n + 1)) (x y : α i) (f g : Π j, α (i.succ_above j)) : nndist (i.insert_nth x f) (i.insert_nth y g) = max (nndist x y) (nndist f g) := From 88fcdc3da43943f5b01925deddaa5bf0c0e85e4e Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sun, 23 Jul 2023 09:51:10 +0000 Subject: [PATCH 34/61] feat(ring_theory/tensor_product): add missing scalar tower instances (#19143) --- src/linear_algebra/tensor_product.lean | 11 ++++- src/ring_theory/tensor_product.lean | 65 ++++++++++++++++++++------ 2 files changed, 60 insertions(+), 16 deletions(-) diff --git a/src/linear_algebra/tensor_product.lean b/src/linear_algebra/tensor_product.lean index 7467afbf19bf2..2e031039e0618 100644 --- a/src/linear_algebra/tensor_product.lean +++ b/src/linear_algebra/tensor_product.lean @@ -293,7 +293,16 @@ section -- Like `R'`, `R'₂` provides a `distrib_mul_action R'₂ (M ⊗[R] N)` variables {R'₂ : Type*} [monoid R'₂] [distrib_mul_action R'₂ M] -variables [smul_comm_class R R'₂ M] [has_smul R'₂ R'] +variables [smul_comm_class R R'₂ M] + +/-- `smul_comm_class R' R'₂ M` implies `smul_comm_class R' R'₂ (M ⊗[R] N)` -/ +instance smul_comm_class_left [smul_comm_class R' R'₂ M] : smul_comm_class R' R'₂ (M ⊗[R] N) := +{ smul_comm := λ r' r'₂ x, tensor_product.induction_on x + (by simp_rw tensor_product.smul_zero) + (λ m n, by simp_rw [smul_tmul', smul_comm]) + (λ x y ihx ihy, by { simp_rw tensor_product.smul_add, rw [ihx, ihy] }),} + +variables [has_smul R'₂ R'] /-- `is_scalar_tower R'₂ R' M` implies `is_scalar_tower R'₂ R' (M ⊗[R] N)` -/ instance is_scalar_tower_left [is_scalar_tower R'₂ R' M] : diff --git a/src/ring_theory/tensor_product.lean b/src/ring_theory/tensor_product.lean index 9e3c0fdaa57d3..fe3a5edcf8001 100644 --- a/src/ring_theory/tensor_product.lean +++ b/src/ring_theory/tensor_product.lean @@ -322,16 +322,16 @@ begin { intros, simp only [linear_map.map_add, *, linear_map.add_apply], }, end -lemma mul_assoc (x y z : A ⊗[R] B) : mul (mul x y) z = mul x (mul y z) := +protected lemma mul_assoc (x y z : A ⊗[R] B) : mul (mul x y) z = mul x (mul y z) := mul_assoc' mul (by { intros, simp only [mul_apply, mul_assoc], }) x y z -lemma one_mul (x : A ⊗[R] B) : mul (1 ⊗ₜ 1) x = x := +protected lemma one_mul (x : A ⊗[R] B) : mul (1 ⊗ₜ 1) x = x := begin apply tensor_product.induction_on x; simp {contextual := tt}, end -lemma mul_one (x : A ⊗[R] B) : mul x (1 ⊗ₜ 1) = x := +protected lemma mul_one (x : A ⊗[R] B) : mul x (1 ⊗ₜ 1) = x := begin apply tensor_product.induction_on x; simp {contextual := tt}, @@ -347,9 +347,9 @@ instance : semiring (A ⊗[R] B) := add := (+), one := 1, mul := λ a b, mul a b, - one_mul := one_mul, - mul_one := mul_one, - mul_assoc := mul_assoc, + one_mul := algebra.tensor_product.one_mul, + mul_one := algebra.tensor_product.mul_one, + mul_assoc := algebra.tensor_product.mul_assoc, zero_mul := by simp, mul_zero := by simp, left_distrib := by simp, @@ -383,22 +383,57 @@ def include_left_ring_hom : A →+* A ⊗[R] B := map_one' := rfl, map_mul' := by simp } -variables {S : Type*} [comm_semiring S] [algebra S A] +variables {S : Type*} + +-- we want `is_scalar_tower_right` to take priority since it's better for unification elsewhere +@[priority 100] +instance is_scalar_tower_right + [monoid S] [distrib_mul_action S A] [is_scalar_tower S A A] [smul_comm_class R S A] : + is_scalar_tower S (A ⊗[R] B) (A ⊗[R] B) := +{ smul_assoc := λ r x y, begin + change (r • x) * y = r • (x * y), + apply tensor_product.induction_on y, + { simp [smul_zero], }, + { apply tensor_product.induction_on x, + { simp [smul_zero] }, + { intros a b a' b', + dsimp, + rw [tensor_product.smul_tmul', tensor_product.smul_tmul', tmul_mul_tmul, smul_mul_assoc], }, + { intros, simp [smul_add, add_mul, *], } }, + { intros, simp [smul_add, mul_add, *], }, + end } + +-- we want `algebra.to_smul_comm_class` to take priority since it's better for unification elsewhere +@[priority 100] +instance smul_comm_class_right + [monoid S] [distrib_mul_action S A] [smul_comm_class S A A] [smul_comm_class R S A] : + smul_comm_class S (A ⊗[R] B) (A ⊗[R] B) := +{ smul_comm := λ r x y, begin + change r • (x * y) = x * r • y, + apply tensor_product.induction_on y, + { simp [smul_zero], }, + { apply tensor_product.induction_on x, + { simp [smul_zero] }, + { intros a b a' b', + dsimp, + rw [tensor_product.smul_tmul', tensor_product.smul_tmul', tmul_mul_tmul, mul_smul_comm], }, + { intros, simp [smul_add, add_mul, *], } }, + { intros, simp [smul_add, mul_add, *], }, + end } + +variables [comm_semiring S] [algebra S A] instance left_algebra [smul_comm_class R S A] : algebra S (A ⊗[R] B) := { commutes' := λ r x, begin - apply tensor_product.induction_on x, - { simp, }, - { intros a b, dsimp, rw [algebra.commutes, _root_.mul_one, _root_.one_mul], }, - { intros y y' h h', dsimp at h h' ⊢, simp only [mul_add, add_mul, h, h'], }, + dsimp only [ring_hom.to_fun_eq_coe, ring_hom.comp_apply, include_left_ring_hom_apply], + rw [algebra_map_eq_smul_one, ←smul_tmul', ←one_def, mul_smul_comm, smul_mul_assoc, mul_one, + one_mul], end, smul_def' := λ r x, begin - apply tensor_product.induction_on x, - { simp [smul_zero], }, - { intros a b, dsimp, rw [tensor_product.smul_tmul', algebra.smul_def r a, _root_.one_mul] }, - { intros, dsimp, simp [smul_add, mul_add, *], }, + dsimp only [ring_hom.to_fun_eq_coe, ring_hom.comp_apply, include_left_ring_hom_apply], + rw [algebra_map_eq_smul_one, ←smul_tmul', smul_mul_assoc, ←one_def, one_mul], end, .. tensor_product.include_left_ring_hom.comp (algebra_map S A), .. (by apply_instance : module S (A ⊗[R] B)) }. From 4b92a463033b5587bb011657e25e4710bfca7364 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 24 Jul 2023 07:47:40 +0000 Subject: [PATCH 35/61] chore(ring_theory/kaehler): cleanup instances (#19234) The previous module instance has the wrong defeqs, and was more work to construct. The new one remains propositionally equal to the old one. --- src/ring_theory/ideal/cotangent.lean | 12 ++------ src/ring_theory/kaehler.lean | 45 +++++++++++----------------- 2 files changed, 21 insertions(+), 36 deletions(-) diff --git a/src/ring_theory/ideal/cotangent.lean b/src/ring_theory/ideal/cotangent.lean index 7743e96ac51d3..fe303936f6e4e 100644 --- a/src/ring_theory/ideal/cotangent.lean +++ b/src/ring_theory/ideal/cotangent.lean @@ -38,16 +38,10 @@ instance cotangent.module_of_tower : module S I.cotangent := submodule.quotient.module' _ instance : is_scalar_tower S S' I.cotangent := -begin - delta cotangent, - constructor, - intros s s' x, - rw [← @is_scalar_tower.algebra_map_smul S' R, ← @is_scalar_tower.algebra_map_smul S' R, - ← smul_assoc, ← is_scalar_tower.to_alg_hom_apply S S' R, map_smul], - refl -end +submodule.quotient.is_scalar_tower _ _ -instance [is_noetherian R I] : is_noetherian R I.cotangent := by { delta cotangent, apply_instance } +instance [is_noetherian R I] : is_noetherian R I.cotangent := +submodule.quotient.is_noetherian _ /-- The quotient map from `I` to `I ⧸ I ^ 2`. -/ @[simps apply (lemmas_only)] diff --git a/src/ring_theory/kaehler.lean b/src/ring_theory/kaehler.lean index ec3b0a3079dc8..2ec88c0e1d9fa 100644 --- a/src/ring_theory/kaehler.lean +++ b/src/ring_theory/kaehler.lean @@ -145,31 +145,23 @@ notation `Ω[`:100 S `⁄`:0 R `]`:0 := kaehler_differential R S instance : nonempty Ω[S⁄R] := ⟨0⟩ -instance kaehler_differential.module' {R' : Type*} [comm_ring R'] [algebra R' S] : +instance kaehler_differential.module' {R' : Type*} [comm_ring R'] [algebra R' S] + [smul_comm_class R R' S] : module R' Ω[S⁄R] := -(module.comp_hom (kaehler_differential.ideal R S).cotangent (algebra_map R' S) : _) +submodule.quotient.module' _ instance : is_scalar_tower S (S ⊗[R] S) Ω[S⁄R] := ideal.cotangent.is_scalar_tower _ instance kaehler_differential.is_scalar_tower_of_tower {R₁ R₂ : Type*} [comm_ring R₁] [comm_ring R₂] - [algebra R₁ S] [algebra R₂ S] [algebra R₁ R₂] [is_scalar_tower R₁ R₂ S] : + [algebra R₁ S] [algebra R₂ S] [has_smul R₁ R₂] + [smul_comm_class R R₁ S] [smul_comm_class R R₂ S] [is_scalar_tower R₁ R₂ S] : is_scalar_tower R₁ R₂ Ω[S⁄R] := -begin - convert restrict_scalars.is_scalar_tower R₁ R₂ Ω[S⁄R] using 1, - ext x m, - show algebra_map R₁ S x • m = algebra_map R₂ S (algebra_map R₁ R₂ x) • m, - rw ← is_scalar_tower.algebra_map_apply, -end +submodule.quotient.is_scalar_tower _ _ instance kaehler_differential.is_scalar_tower' : is_scalar_tower R (S ⊗[R] S) Ω[S⁄R] := -begin - convert restrict_scalars.is_scalar_tower R (S ⊗[R] S) Ω[S⁄R] using 1, - ext x m, - show algebra_map R S x • m = algebra_map R (S ⊗[R] S) x • m, - simp_rw [is_scalar_tower.algebra_map_apply R S (S ⊗[R] S), is_scalar_tower.algebra_map_smul] -end +submodule.quotient.is_scalar_tower _ _ /-- The quotient map `I → Ω[S⁄R]` with `I` being the kernel of `S ⊗[R] S → S`. -/ def kaehler_differential.from_ideal : kaehler_differential.ideal R S →ₗ[S ⊗[R] S] Ω[S⁄R] := @@ -190,13 +182,14 @@ rfl /-- The universal derivation into `Ω[S⁄R]`. -/ def kaehler_differential.D : derivation R S Ω[S⁄R] := { map_one_eq_zero' := begin - dsimp [kaehler_differential.D_linear_map_apply], + dsimp only [kaehler_differential.D_linear_map_apply], rw [ideal.to_cotangent_eq_zero, subtype.coe_mk, sub_self], exact zero_mem _ end, leibniz' := λ a b, begin - dsimp [kaehler_differential.D_linear_map_apply], - rw [← linear_map.map_smul_of_tower, ← linear_map.map_smul_of_tower, ← map_add, + dsimp only [kaehler_differential.D_linear_map_apply], + rw [← linear_map.map_smul_of_tower ((kaehler_differential.ideal R S).to_cotangent) a, + ← linear_map.map_smul_of_tower ((kaehler_differential.ideal R S).to_cotangent) b, ← map_add, ideal.to_cotangent_eq, pow_two], convert submodule.mul_mem_mul (kaehler_differential.one_smul_sub_smul_one_mem_ideal R a : _) (kaehler_differential.one_smul_sub_smul_one_mem_ideal R b : _) using 1, @@ -205,7 +198,7 @@ def kaehler_differential.D : derivation R S Ω[S⁄R] := submodule.coe_smul_of_tower, smul_sub, tensor_product.smul_tmul', smul_eq_mul, mul_one], ring_nf, end, - ..(kaehler_differential.D_linear_map R S) } + to_linear_map := kaehler_differential.D_linear_map R S } lemma kaehler_differential.D_apply (s : S) : kaehler_differential.D R S s = (kaehler_differential.ideal R S).to_cotangent @@ -373,9 +366,9 @@ into `(kaehler_differential.ideal R S).cotangent_ideal` -/ -- `derivation R S Ω[S⁄R] ≃ₗ[R] derivation R S (kaehler_differential.ideal R S).cotangent_ideal` -- But lean times-out if this is given explicitly. noncomputable -def kaehler_differential.End_equiv_derivation' := -@linear_equiv.comp_der R _ _ _ _ Ω[S⁄R] _ _ _ _ _ _ _ _ _ - ((kaehler_differential.ideal R S).cotangent_equiv_ideal.restrict_scalars S) +def kaehler_differential.End_equiv_derivation' : + derivation R S Ω[S⁄R] ≃ₗ[R] derivation R S _ := +linear_equiv.comp_der ((kaehler_differential.ideal R S).cotangent_equiv_ideal.restrict_scalars S) /-- (Implementation) An `equiv` version of `kaehler_differential.End_equiv_aux`. Used in `kaehler_differential.End_equiv`. -/ @@ -538,7 +531,7 @@ variables (A B : Type*) [comm_ring A] [comm_ring B] [algebra R A] [algebra S B] variables [algebra A B] [is_scalar_tower R S B] [is_scalar_tower R A B] -- The map `(A →₀ A) →ₗ[A] (B →₀ B)` -local notation `finsupp_map` := ((finsupp.map_range.linear_map (algebra.of_id A B).to_linear_map) +local notation `finsupp_map` := ((finsupp.map_range.linear_map (algebra.linear_map A B)) .comp (finsupp.lmap_domain A A (algebra_map A B))) lemma kaehler_differential.ker_total_map (h : function.surjective (algebra_map A B)) : @@ -552,7 +545,7 @@ begin set.image_univ, map_sub, map_add], simp only [linear_map.comp_apply, finsupp.map_range.linear_map_apply, finsupp.map_range_single, finsupp.lmap_domain_apply, finsupp.map_domain_single, alg_hom.to_linear_map_apply, - algebra.of_id_apply, ← is_scalar_tower.algebra_map_apply, map_one, map_add, map_mul], + algebra.linear_map_apply, ← is_scalar_tower.algebra_map_apply, map_one, map_add, map_mul], simp_rw [sup_assoc, ← (h.prod_map h).range_comp], congr' 3, rw [sup_eq_right], @@ -565,8 +558,6 @@ end presentation section exact_sequence -local attribute [irreducible] kaehler_differential - /- We have the commutative diagram A --→ B ↑ ↑ @@ -585,7 +576,7 @@ def derivation.comp_algebra_map [module A M] [module B M] [is_scalar_tower A B M leibniz' := λ a b, by simp, to_linear_map := d.to_linear_map.comp (is_scalar_tower.to_alg_hom R A B).to_linear_map } -variables (R B) +variables (R B) [smul_comm_class S A B] /-- The map `Ω[A⁄R] →ₗ[A] Ω[B⁄R]` given a square A --→ B From 148aefbd371a25f1cff33c85f20c661ce3155def Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 24 Jul 2023 08:57:01 +0000 Subject: [PATCH 36/61] docs(topology/dense_embedding): Improve explanation (#18134) An attempt at improving understandability. Closes #1539 --- src/topology/dense_embedding.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/src/topology/dense_embedding.lean b/src/topology/dense_embedding.lean index 3e768af19cd82..af0d11c6b63c5 100644 --- a/src/topology/dense_embedding.lean +++ b/src/topology/dense_embedding.lean @@ -15,8 +15,8 @@ import topology.bases This file defines three properties of functions: * `dense_range f` means `f` has dense image; -* `dense_inducing i` means `i` is also `inducing`; -* `dense_embedding e` means `e` is also an `embedding`. +* `dense_inducing i` means `i` is also `inducing`, namely it induces the topology on its codomain; +* `dense_embedding e` means `e` is further an `embedding`, namely it is injective and `inducing`. The main theorem `continuous_extend` gives a criterion for a function `f : X → Z` to a T₃ space Z to extend along a dense embedding From 4be589053caf347b899a494da75410deb55fb3ef Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 24 Jul 2023 08:57:03 +0000 Subject: [PATCH 37/61] fix(group_theory/subgroup/basic): generalize `centralizer` from `subgroup G` to `set G` (#18965) This is consistent with all the other `sub.centralizer` definitions. This generalization reveals that a lot of downstream results are rather strangely stated about `zpowers`. This does not attempt to change these, instead leaving the work for a follow up (either in a later mathlib3 PR or in mathlib4). --- src/group_theory/abelianization.lean | 11 +++++--- src/group_theory/commutator.lean | 2 +- src/group_theory/group_action/conj_act.lean | 3 +- src/group_theory/group_action/quotient.lean | 2 +- src/group_theory/subgroup/basic.lean | 31 +++++++++++---------- src/group_theory/subgroup/zpowers.lean | 10 ++++--- src/group_theory/sylow.lean | 12 ++++---- src/group_theory/transfer.lean | 2 +- 8 files changed, 41 insertions(+), 32 deletions(-) diff --git a/src/group_theory/abelianization.lean b/src/group_theory/abelianization.lean index 6c9fdf44175e5..c2915d7a36840 100644 --- a/src/group_theory/abelianization.lean +++ b/src/group_theory/abelianization.lean @@ -32,6 +32,8 @@ universes u v w -- Let G be a group. variables (G : Type u) [group G] +open subgroup (centralizer) + /-- The commutator subgroup of a group G is the normal subgroup generated by the commutators [p,q]=`p*q*p⁻¹*q⁻¹`. -/ @[derive subgroup.normal] @@ -64,12 +66,13 @@ begin end lemma commutator_centralizer_commutator_le_center : - ⁅(commutator G).centralizer, (commutator G).centralizer⁆ ≤ subgroup.center G := + ⁅centralizer (commutator G : set G), centralizer (commutator G : set G)⁆ ≤ subgroup.center G := begin - rw [←subgroup.centralizer_top, ←subgroup.commutator_eq_bot_iff_le_centralizer], - suffices : ⁅⁅⊤, (commutator G).centralizer⁆, (commutator G).centralizer⁆ = ⊥, + rw [←subgroup.centralizer_univ, ←subgroup.coe_top, + ←subgroup.commutator_eq_bot_iff_le_centralizer], + suffices : ⁅⁅⊤, centralizer (commutator G : set G)⁆, centralizer (commutator G : set G)⁆ = ⊥, { refine subgroup.commutator_commutator_eq_bot_of_rotate _ this, - rwa subgroup.commutator_comm (commutator G).centralizer }, + rwa subgroup.commutator_comm (centralizer (commutator G : set G)) }, rw [subgroup.commutator_comm, subgroup.commutator_eq_bot_iff_le_centralizer], exact set.centralizer_subset (subgroup.commutator_mono le_top le_top), end diff --git a/src/group_theory/commutator.lean b/src/group_theory/commutator.lean index 8d2ca4052460a..5c21d23e5f311 100644 --- a/src/group_theory/commutator.lean +++ b/src/group_theory/commutator.lean @@ -75,7 +75,7 @@ H₃.closure_le.trans ⟨λ h a b c d, h ⟨a, b, c, d, rfl⟩, λ h g ⟨a, b, lemma commutator_mono (h₁ : H₁ ≤ K₁) (h₂ : H₂ ≤ K₂) : ⁅H₁, H₂⁆ ≤ ⁅K₁, K₂⁆ := commutator_le.mpr (λ g₁ hg₁ g₂ hg₂, commutator_mem_commutator (h₁ hg₁) (h₂ hg₂)) -lemma commutator_eq_bot_iff_le_centralizer : ⁅H₁, H₂⁆ = ⊥ ↔ H₁ ≤ H₂.centralizer := +lemma commutator_eq_bot_iff_le_centralizer : ⁅H₁, H₂⁆ = ⊥ ↔ H₁ ≤ centralizer H₂ := begin rw [eq_bot_iff, commutator_le], refine forall_congr (λ p, forall_congr (λ hp, forall_congr (λ q, forall_congr (λ hq, _)))), diff --git a/src/group_theory/group_action/conj_act.lean b/src/group_theory/group_action/conj_act.lean index 12809e56e22d4..54cf1c4e85312 100644 --- a/src/group_theory/group_action/conj_act.lean +++ b/src/group_theory/group_action/conj_act.lean @@ -193,7 +193,8 @@ by { rw [is_conj_comm, is_conj_iff, mem_orbit_iff], refl } lemma orbit_rel_conj_act : (orbit_rel (conj_act G) G).rel = is_conj := funext₂ $ λ g h, by rw [orbit_rel_apply, mem_orbit_conj_act] -lemma stabilizer_eq_centralizer (g : G) : stabilizer (conj_act G) g = (zpowers g).centralizer := +lemma stabilizer_eq_centralizer (g : G) : + stabilizer (conj_act G) g = centralizer (zpowers (to_conj_act g) : set (conj_act G)) := le_antisymm (le_centralizer_iff.mp (zpowers_le.mpr (λ x, mul_inv_eq_iff_eq_mul.mp))) (λ x h, mul_inv_eq_of_eq_mul (h g (mem_zpowers g)).symm) diff --git a/src/group_theory/group_action/quotient.lean b/src/group_theory/group_action/quotient.lean index 0c2d4a21eb7af..d9bd002eff0e8 100644 --- a/src/group_theory/group_action/quotient.lean +++ b/src/group_theory/group_action/quotient.lean @@ -290,7 +290,7 @@ open quotient_group /-- Cosets of the centralizer of an element embed into the set of commutators. -/ noncomputable def quotient_centralizer_embedding (g : G) : - G ⧸ centralizer (zpowers (g : G)) ↪ commutator_set G := + G ⧸ centralizer (zpowers (g : G) : set G) ↪ commutator_set G := ((mul_action.orbit_equiv_quotient_stabilizer (conj_act G) g).trans (quotient_equiv_of_eq (conj_act.stabilizer_eq_centralizer g))).symm.to_embedding.trans ⟨λ x, ⟨x * g⁻¹, let ⟨_, x, rfl⟩ := x in ⟨x, g, rfl⟩⟩, λ x y, subtype.ext ∘ mul_right_cancel ∘ subtype.ext_iff.mp⟩ diff --git a/src/group_theory/subgroup/basic.lean b/src/group_theory/subgroup/basic.lean index 464a9504a5aaa..26cf2cc379965 100644 --- a/src/group_theory/subgroup/basic.lean +++ b/src/group_theory/subgroup/basic.lean @@ -1580,41 +1580,42 @@ normalizer_eq_top.mp (hmax.2 _ (hnc H (lt_top_iff_ne_top.mpr hmax.1))) end normalizer section centralizer +variables {H} /-- The `centralizer` of `H` is the subgroup of `g : G` commuting with every `h : H`. -/ @[to_additive "The `centralizer` of `H` is the additive subgroup of `g : G` commuting with every `h : H`."] -def centralizer : subgroup G := -{ carrier := set.centralizer H, +def centralizer (s : set G) : subgroup G := +{ carrier := set.centralizer s, inv_mem' := λ g, set.inv_mem_centralizer, - .. submonoid.centralizer ↑H } - -variables {H} + .. submonoid.centralizer s } -@[to_additive] lemma mem_centralizer_iff {g : G} : g ∈ H.centralizer ↔ ∀ h ∈ H, h * g = g * h := +@[to_additive] lemma mem_centralizer_iff {g : G} {s : set G} : + g ∈ centralizer s ↔ ∀ h ∈ s, h * g = g * h := iff.rfl -@[to_additive] lemma mem_centralizer_iff_commutator_eq_one {g : G} : - g ∈ H.centralizer ↔ ∀ h ∈ H, h * g * h⁻¹ * g⁻¹ = 1 := +@[to_additive] lemma mem_centralizer_iff_commutator_eq_one {g : G} {s : set G} : + g ∈ centralizer s ↔ ∀ h ∈ s, h * g * h⁻¹ * g⁻¹ = 1 := by simp only [mem_centralizer_iff, mul_inv_eq_iff_eq_mul, one_mul] -@[to_additive] lemma centralizer_top : centralizer ⊤ = center G := +@[to_additive] lemma centralizer_univ : centralizer set.univ = center G := set_like.ext' (set.centralizer_univ G) -@[to_additive] lemma le_centralizer_iff : H ≤ K.centralizer ↔ K ≤ H.centralizer := +@[to_additive] lemma le_centralizer_iff : H ≤ centralizer K ↔ K ≤ centralizer H := ⟨λ h x hx y hy, (h hy x hx).symm, λ h x hx y hy, (h hy x hx).symm⟩ @[to_additive] lemma center_le_centralizer (s) : center G ≤ centralizer s := set.center_subset_centralizer s -@[to_additive] lemma centralizer_le (h : H ≤ K) : centralizer K ≤ centralizer H := +@[to_additive] lemma centralizer_le {s t : set G} (h : s ⊆ t) : centralizer t ≤ centralizer s := submonoid.centralizer_le h -@[simp, to_additive] lemma centralizer_eq_top_iff_subset {s} : centralizer s = ⊤ ↔ s ≤ center G := +@[simp, to_additive] lemma centralizer_eq_top_iff_subset {s : set G} : + centralizer s = ⊤ ↔ s ⊆ center G := set_like.ext'_iff.trans set.centralizer_eq_top_iff_subset @[to_additive] instance subgroup.centralizer.characteristic [hH : H.characteristic] : - H.centralizer.characteristic := + (centralizer (H : set G)).characteristic := begin refine subgroup.characteristic_iff_comap_le.mpr (λ ϕ g hg h hh, ϕ.injective _), rw [map_mul, map_mul], @@ -1663,10 +1664,10 @@ end⟩⟩ (H.subgroup_of K).is_commutative := H.comap_injective_is_commutative subtype.coe_injective -@[to_additive] lemma le_centralizer_iff_is_commutative : K ≤ K.centralizer ↔ K.is_commutative := +@[to_additive] lemma le_centralizer_iff_is_commutative : K ≤ centralizer K ↔ K.is_commutative := ⟨λ h, ⟨⟨λ x y, subtype.ext (h y.2 x x.2)⟩⟩, λ h x hx y hy, congr_arg coe (h.1.1 ⟨y, hy⟩ ⟨x, hx⟩)⟩ -@[to_additive] lemma le_centralizer [h : H.is_commutative] : H ≤ H.centralizer := +@[to_additive] lemma le_centralizer [h : H.is_commutative] : H ≤ centralizer H := le_centralizer_iff_is_commutative.mpr h end subgroup diff --git a/src/group_theory/subgroup/zpowers.lean b/src/group_theory/subgroup/zpowers.lean index 6de48373ae8b7..f3456efdf70d3 100644 --- a/src/group_theory/subgroup/zpowers.lean +++ b/src/group_theory/subgroup/zpowers.lean @@ -166,17 +166,19 @@ zpowers_eq_bot.not subgroup.zpowers_eq_bot.mpr rfl @[to_additive] lemma centralizer_closure (S : set G) : - (closure S).centralizer = ⨅ g ∈ S, (zpowers g).centralizer := -le_antisymm (le_infi $ λ g, le_infi $ λ hg, centralizer_le $ zpowers_le.2 $ subset_closure hg) + centralizer (closure S : set G) = ⨅ g ∈ S, centralizer (zpowers g : set G) := +le_antisymm + (le_infi $ λ g, le_infi $ λ hg, centralizer_le $ set_like.coe_subset_coe.2 $ + zpowers_le.2 $ subset_closure hg) $ le_centralizer_iff.1 $ (closure_le _).2 $ λ g, set_like.mem_coe.2 ∘ zpowers_le.1 ∘ le_centralizer_iff.1 ∘ infi_le_of_le g ∘ infi_le _ @[to_additive] lemma center_eq_infi (S : set G) (hS : closure S = ⊤) : center G = ⨅ g ∈ S, centralizer (zpowers g) := -by rw [←centralizer_top, ←hS, centralizer_closure] +by rw [←centralizer_univ, ←coe_top, ←hS, centralizer_closure] @[to_additive] lemma center_eq_infi' (S : set G) (hS : closure S = ⊤) : - center G = ⨅ g : S, centralizer (zpowers g) := + center G = ⨅ g : S, centralizer (zpowers (g : G) : set G) := by rw [center_eq_infi S hS, ←infi_subtype''] end subgroup diff --git a/src/group_theory/sylow.lean b/src/group_theory/sylow.lean index 8b32b29b66805..1560211d6a95c 100644 --- a/src/group_theory/sylow.lean +++ b/src/group_theory/sylow.lean @@ -302,18 +302,19 @@ ext (λ g, sylow.smul_eq_iff_mem_normalizer) lemma sylow.conj_eq_normalizer_conj_of_mem_centralizer [fact p.prime] [finite (sylow p G)] (P : sylow p G) (x g : G) - (hx : x ∈ (P : subgroup G).centralizer) (hy : g⁻¹ * x * g ∈ (P : subgroup G).centralizer) : + (hx : x ∈ centralizer (P : set G)) (hy : g⁻¹ * x * g ∈ centralizer (P : set G)) : ∃ n ∈ (P : subgroup G).normalizer, g⁻¹ * x * g = n⁻¹ * x * n := begin - have h1 : ↑P ≤ (zpowers x).centralizer, + have h1 : ↑P ≤ centralizer (zpowers x : set G), { rwa [le_centralizer_iff, zpowers_le] }, - have h2 : ↑(g • P) ≤ (zpowers x).centralizer, + have h2 : ↑(g • P) ≤ centralizer (zpowers x : set G), { rw [le_centralizer_iff, zpowers_le], rintros - ⟨z, hz, rfl⟩, specialize hy z hz, rwa [←mul_assoc, ←eq_mul_inv_iff_mul_eq, mul_assoc, mul_assoc, mul_assoc, ←mul_assoc, eq_inv_mul_iff_mul_eq, ←mul_assoc, ←mul_assoc] at hy }, - obtain ⟨h, hh⟩ := exists_smul_eq (zpowers x).centralizer ((g • P).subtype h2) (P.subtype h1), + obtain ⟨h, hh⟩ := + exists_smul_eq (centralizer (zpowers x : set G)) ((g • P).subtype h2) (P.subtype h1), simp_rw [sylow.smul_subtype, smul_def, smul_smul] at hh, refine ⟨h * g, sylow.smul_eq_iff_mem_normalizer.mp (sylow.subtype_injective hh), _⟩, rw [←mul_assoc, commute.right_comm (h.prop x (mem_zpowers x)), mul_inv_rev, inv_mul_cancel_right] @@ -322,7 +323,8 @@ end lemma sylow.conj_eq_normalizer_conj_of_mem [fact p.prime] [finite (sylow p G)] (P : sylow p G) [hP : (P : subgroup G).is_commutative] (x g : G) (hx : x ∈ P) (hy : g⁻¹ * x * g ∈ P) : ∃ n ∈ (P : subgroup G).normalizer, g⁻¹ * x * g = n⁻¹ * x * n := -P.conj_eq_normalizer_conj_of_mem_centralizer x g (le_centralizer P hx) (le_centralizer P hy) +P.conj_eq_normalizer_conj_of_mem_centralizer x g + (le_centralizer (P : subgroup G) hx : _) (le_centralizer (P : subgroup G) hy : _) /-- Sylow `p`-subgroups are in bijection with cosets of the normalizer of a Sylow `p`-subgroup -/ noncomputable def sylow.equiv_quotient_normalizer [fact p.prime] [fintype (sylow p G)] diff --git a/src/group_theory/transfer.lean b/src/group_theory/transfer.lean index f54f4a3cc4f1d..1202268a0a7b7 100644 --- a/src/group_theory/transfer.lean +++ b/src/group_theory/transfer.lean @@ -172,7 +172,7 @@ rfl section burnside_transfer -variables {p : ℕ} (P : sylow p G) (hP : (P : subgroup G).normalizer ≤ (P : subgroup G).centralizer) +variables {p : ℕ} (P : sylow p G) (hP : (P : subgroup G).normalizer ≤ centralizer (P : set G)) include hP From c0c52abb75074ed8b73a948341f50521fbf43b4c Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 24 Jul 2023 11:50:08 +0000 Subject: [PATCH 38/61] feat(order/upper_lower/basic): Linear order (#19068) Upper/lower sets on a linear order themselves form a linear order. --- src/algebra/order/upper_lower.lean | 15 +++--- src/order/upper_lower/basic.lean | 53 +++++++++++++++++++-- src/topology/algebra/order/upper_lower.lean | 2 +- 3 files changed, 56 insertions(+), 14 deletions(-) diff --git a/src/algebra/order/upper_lower.lean b/src/algebra/order/upper_lower.lean index 3f36d91f00fa1..d52a2b05f96bc 100644 --- a/src/algebra/order/upper_lower.lean +++ b/src/algebra/order/upper_lower.lean @@ -34,13 +34,10 @@ section ordered_comm_group variables {α : Type*} [ordered_comm_group α] {s t : set α} {a : α} @[to_additive] lemma is_upper_set.smul (hs : is_upper_set s) : is_upper_set (a • s) := -begin - rintro _ y hxy ⟨x, hx, rfl⟩, - exact mem_smul_set_iff_inv_smul_mem.2 (hs (le_inv_mul_iff_mul_le.2 hxy) hx), -end +hs.image $ order_iso.mul_left _ @[to_additive] lemma is_lower_set.smul (hs : is_lower_set s) : is_lower_set (a • s) := -hs.of_dual.smul +hs.image $ order_iso.mul_left _ @[to_additive] lemma set.ord_connected.smul (hs : s.ord_connected) : (a • s).ord_connected := begin @@ -55,10 +52,10 @@ by { rw [←smul_eq_mul, ←bUnion_smul_set], exact is_upper_set_Union₂ (λ x by { rw mul_comm, exact hs.mul_left } @[to_additive] lemma is_lower_set.mul_left (ht : is_lower_set t) : is_lower_set (s * t) := -ht.of_dual.mul_left +ht.to_dual.mul_left @[to_additive] lemma is_lower_set.mul_right (hs : is_lower_set s) : is_lower_set (s * t) := -hs.of_dual.mul_right +hs.to_dual.mul_right @[to_additive] lemma is_upper_set.inv (hs : is_upper_set s) : is_lower_set s⁻¹ := λ x y h, hs $ inv_le_inv' h @@ -73,10 +70,10 @@ by { rw div_eq_mul_inv, exact ht.inv.mul_left } by { rw div_eq_mul_inv, exact hs.mul_right } @[to_additive] lemma is_lower_set.div_left (ht : is_lower_set t) : is_upper_set (s / t) := -ht.of_dual.div_left +ht.to_dual.div_left @[to_additive] lemma is_lower_set.div_right (hs : is_lower_set s) : is_lower_set (s / t) := -hs.of_dual.div_right +hs.to_dual.div_right namespace upper_set diff --git a/src/order/upper_lower/basic.lean b/src/order/upper_lower/basic.lean index d95fe66e45e02..db5b4c91df203 100644 --- a/src/order/upper_lower/basic.lean +++ b/src/order/upper_lower/basic.lean @@ -6,6 +6,7 @@ Authors: Yaël Dillies, Sara Rousta import data.set_like.basic import data.set.intervals.ord_connected import data.set.intervals.order_iso +import tactic.by_contra /-! # Up-sets and down-sets @@ -135,10 +136,10 @@ iff.rfl @[simp] lemma is_upper_set_preimage_to_dual_iff {s : set αᵒᵈ} : is_upper_set (to_dual ⁻¹' s) ↔ is_lower_set s := iff.rfl -alias is_lower_set_preimage_of_dual_iff ↔ _ is_upper_set.of_dual -alias is_upper_set_preimage_of_dual_iff ↔ _ is_lower_set.of_dual -alias is_lower_set_preimage_to_dual_iff ↔ _ is_upper_set.to_dual -alias is_upper_set_preimage_to_dual_iff ↔ _ is_lower_set.to_dual +alias is_lower_set_preimage_of_dual_iff ↔ _ is_upper_set.to_dual +alias is_upper_set_preimage_of_dual_iff ↔ _ is_lower_set.to_dual +alias is_lower_set_preimage_to_dual_iff ↔ _ is_upper_set.of_dual +alias is_upper_set_preimage_to_dual_iff ↔ _ is_lower_set.of_dual end has_le @@ -266,6 +267,24 @@ alias is_lower_set_iff_Iio_subset ↔ is_lower_set.Iio_subset _ end partial_order +section linear_order +variables [linear_order α] {s t : set α} + +lemma is_upper_set.total (hs : is_upper_set s) (ht : is_upper_set t) : s ⊆ t ∨ t ⊆ s := +begin + by_contra' h, + simp_rw set.not_subset at h, + obtain ⟨⟨a, has, hat⟩, b, hbt, hbs⟩ := h, + obtain hab | hba := le_total a b, + { exact hbs (hs hab has) }, + { exact hat (ht hba hbt) } +end + +lemma is_lower_set.total (hs : is_lower_set s) (ht : is_lower_set t) : s ⊆ t ∨ t ⊆ s := +hs.to_dual.total ht.to_dual + +end linear_order + /-! ### Bundled upper/lower sets -/ section has_le @@ -519,6 +538,32 @@ end lower_set end has_le +section linear_order +variables [linear_order α] + +instance upper_set.is_total_le : is_total (upper_set α) (≤) := ⟨λ s t, t.upper.total s.upper⟩ +instance lower_set.is_total_le : is_total (lower_set α) (≤) := ⟨λ s t, s.lower.total t.lower⟩ + +noncomputable instance : complete_linear_order (upper_set α) := +{ le_total := is_total.total, + decidable_le := classical.dec_rel _, + decidable_eq := classical.dec_rel _, + decidable_lt := classical.dec_rel _, + max_def := by classical; exact sup_eq_max_default, + min_def := by classical; exact inf_eq_min_default, + ..upper_set.complete_distrib_lattice } + +noncomputable instance : complete_linear_order (lower_set α) := +{ le_total := is_total.total, + decidable_le := classical.dec_rel _, + decidable_eq := classical.dec_rel _, + decidable_lt := classical.dec_rel _, + max_def := by classical; exact sup_eq_max_default, + min_def := by classical; exact inf_eq_min_default, + ..lower_set.complete_distrib_lattice } + +end linear_order + /-! #### Map -/ section diff --git a/src/topology/algebra/order/upper_lower.lean b/src/topology/algebra/order/upper_lower.lean index 5bc75c68c9be3..792ad0b431a15 100644 --- a/src/topology/algebra/order/upper_lower.lean +++ b/src/topology/algebra/order/upper_lower.lean @@ -85,7 +85,7 @@ protected lemma is_upper_set.interior (h : is_upper_set s) : is_upper_set (inter by { rw [←is_lower_set_compl, ←closure_compl], exact h.compl.closure } protected lemma is_lower_set.interior (h : is_lower_set s) : is_lower_set (interior s) := -h.of_dual.interior +h.to_dual.interior protected lemma set.ord_connected.interior (h : s.ord_connected) : (interior s).ord_connected := begin From 18ee599842a5d17f189fe572f0ed8cb1d064d772 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Tue, 25 Jul 2023 05:26:13 +0000 Subject: [PATCH 39/61] feat(algebraic_topology/dold_kan): tools for compatibilities, lemmas (#17923) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: Joël Riou <37772949+joelriou@users.noreply.github.com> --- .../dold_kan/compatibility.lean | 125 +++++++++++++++++- 1 file changed, 119 insertions(+), 6 deletions(-) diff --git a/src/algebraic_topology/dold_kan/compatibility.lean b/src/algebraic_topology/dold_kan/compatibility.lean index b5afeedb24a51..3567efda0ccc9 100644 --- a/src/algebraic_topology/dold_kan/compatibility.lean +++ b/src/algebraic_topology/dold_kan/compatibility.lean @@ -77,7 +77,7 @@ lemma equivalence₁_inverse : (equivalence₁ hF).inverse = e'.inverse ⋙ eA.i def equivalence₁_counit_iso : (e'.inverse ⋙ eA.inverse) ⋙ F ≅ 𝟭 B' := calc (e'.inverse ⋙ eA.inverse) ⋙ F - ≅ (e'.inverse ⋙ eA.inverse) ⋙ (eA.functor ⋙ e'.functor) : iso_whisker_left _ hF.symm + ≅ (e'.inverse ⋙ eA.inverse) ⋙ (eA.functor ⋙ e'.functor) : iso_whisker_left _ hF.symm ... ≅ e'.inverse ⋙ (eA.inverse ⋙ eA.functor) ⋙ e'.functor : iso.refl _ ... ≅ e'.inverse ⋙ 𝟭 _ ⋙ e'.functor : iso_whisker_left _ (iso_whisker_right eA.counit_iso _) ... ≅ e'.inverse ⋙ e'.functor : iso.refl _ @@ -97,7 +97,7 @@ def equivalence₁_unit_iso : calc 𝟭 A ≅ eA.functor ⋙ eA.inverse : eA.unit_iso ... ≅ eA.functor ⋙ 𝟭 A' ⋙ eA.inverse : iso.refl _ ... ≅ eA.functor ⋙ (e'.functor ⋙ e'.inverse) ⋙ eA.inverse : - iso_whisker_left _ (iso_whisker_right e'.unit_iso _) + iso_whisker_left _ (iso_whisker_right e'.unit_iso _) ... ≅ (eA.functor ⋙ e'.functor) ⋙ (e'.inverse ⋙ eA.inverse) : iso.refl _ ... ≅ F ⋙ (e'.inverse ⋙ eA.inverse) : iso_whisker_right hF _ @@ -124,9 +124,9 @@ lemma equivalence₂_inverse : (equivalence₂ eB hF).inverse = def equivalence₂_counit_iso : (eB.functor ⋙ e'.inverse ⋙ eA.inverse) ⋙ (F ⋙ eB.inverse) ≅ 𝟭 B := calc (eB.functor ⋙ e'.inverse ⋙ eA.inverse) ⋙ (F ⋙ eB.inverse) - ≅ eB.functor ⋙ (e'.inverse ⋙ eA.inverse ⋙ F) ⋙ eB.inverse : iso.refl _ + ≅ eB.functor ⋙ (e'.inverse ⋙ eA.inverse ⋙ F) ⋙ eB.inverse : iso.refl _ ... ≅ eB.functor ⋙ 𝟭 _ ⋙ eB.inverse : - iso_whisker_left _ (iso_whisker_right (equivalence₁_counit_iso hF) _) + iso_whisker_left _ (iso_whisker_right (equivalence₁_counit_iso hF) _) ... ≅ eB.functor ⋙ eB.inverse : iso.refl _ ... ≅ 𝟭 B : eB.unit_iso.symm @@ -146,7 +146,7 @@ def equivalence₂_unit_iso : calc 𝟭 A ≅ F ⋙ e'.inverse ⋙ eA.inverse : equivalence₁_unit_iso hF ... ≅ F ⋙ 𝟭 B' ⋙ (e'.inverse ⋙ eA.inverse) : iso.refl _ ... ≅ F ⋙ (eB.inverse ⋙ eB.functor) ⋙ e'.inverse ⋙ eA.inverse : - iso_whisker_left _ (iso_whisker_right eB.counit_iso.symm _) + iso_whisker_left _ (iso_whisker_right eB.counit_iso.symm _) ... ≅ (F ⋙ eB.inverse) ⋙ (eB.functor ⋙ e'.inverse ⋙ eA.inverse) : iso.refl _ lemma equivalence₂_unit_iso_eq : @@ -169,7 +169,7 @@ begin letI : is_equivalence G := begin refine is_equivalence.of_iso _ (is_equivalence.of_equivalence (equivalence₂ eB hF).symm), calc eB.functor ⋙ e'.inverse ⋙ eA.inverse - ≅ (eB.functor ⋙ e'.inverse) ⋙ eA.inverse : iso.refl _ + ≅ (eB.functor ⋙ e'.inverse) ⋙ eA.inverse : iso.refl _ ... ≅ (G ⋙ eA.functor) ⋙ eA.inverse : iso_whisker_right hG _ ... ≅ G ⋙ 𝟭 A : iso_whisker_left _ eA.unit_iso.symm ... ≅ G : functor.right_unitor G, @@ -179,6 +179,119 @@ end lemma equivalence_functor : (equivalence hF hG).functor = F ⋙ eB.inverse := rfl +omit hG hF + +/-- The isomorphism `eB.functor ⋙ e'.inverse ⋙ e'.functor ≅ eB.functor` deduced +from the counit isomorphism of `e'`. -/ +@[simps hom_app] +def τ₀ : eB.functor ⋙ e'.inverse ⋙ e'.functor ≅ eB.functor := +calc eB.functor ⋙ e'.inverse ⋙ e'.functor + ≅ eB.functor ⋙ 𝟭 _ : iso_whisker_left _ e'.counit_iso +... ≅ eB.functor : functor.right_unitor _ + +include hF hG + +/-- The isomorphism `eB.functor ⋙ e'.inverse ⋙ e'.functor ≅ eB.functor` deduced +from the isomorphisms `hF : eA.functor ⋙ e'.functor ≅ F`, +`hG : eB.functor ⋙ e'.inverse ≅ G ⋙ eA.functor` and the datum of +an isomorphism `η : G ⋙ F ≅ eB.functor`. -/ +@[simps hom_app] +def τ₁ (η : G ⋙ F ≅ eB.functor) : + eB.functor ⋙ e'.inverse ⋙ e'.functor ≅ eB.functor := +calc eB.functor ⋙ e'.inverse ⋙ e'.functor + ≅ (eB.functor ⋙ e'.inverse) ⋙ e'.functor : iso.refl _ +... ≅ (G ⋙ eA.functor) ⋙ e'.functor : iso_whisker_right hG _ +... ≅ G ⋙ (eA.functor ⋙ e'.functor) : by refl +... ≅ G ⋙ F : iso_whisker_left _ hF +... ≅ eB.functor : η + +variables (η : G ⋙ F ≅ eB.functor) (hη : τ₀ = τ₁ hF hG η) + +omit hF hG +include η + +/-- The counit isomorphism of `equivalence`. -/ +@[simps] +def equivalence_counit_iso : G ⋙ (F ⋙ eB.inverse) ≅ 𝟭 B := +calc G ⋙ (F ⋙ eB.inverse) ≅ (G ⋙ F) ⋙ eB.inverse : iso.refl _ +... ≅ eB.functor ⋙ eB.inverse : iso_whisker_right η _ +... ≅ 𝟭 B : eB.unit_iso.symm + +variables {η hF hG} +include hη + +lemma equivalence_counit_iso_eq : + (equivalence hF hG).counit_iso = equivalence_counit_iso η := +begin + ext1, apply nat_trans.ext, ext Y, + dsimp [equivalence, equivalence_counit_iso, is_equivalence.of_equivalence], + simp only [equivalence₂_counit_iso_eq eB hF], + erw [nat_trans.id_app, nat_trans.id_app], + dsimp [equivalence₂, equivalence₁], + simp only [assoc, comp_id, F.map_id, id_comp, + equivalence₂_counit_iso_hom_app, ← eB.inverse.map_comp_assoc, + ← τ₀_hom_app, hη, τ₁_hom_app], + erw hF.inv.naturality_assoc, + congr' 2, + dsimp, + simp only [assoc, ← e'.functor.map_comp_assoc, eA.functor.map_comp, + equivalence.fun_inv_map, iso.inv_hom_id_app_assoc, hG.inv_hom_id_app], + dsimp, + rw [comp_id, eA.functor_unit_iso_comp, e'.functor.map_id, id_comp, hF.inv_hom_id_app_assoc], +end + +omit hη η eB +include hF + +variable (hF) + +/-- The isomorphism `eA.functor ≅ F ⋙ e'.inverse` deduced from the +unit isomorphism of `e'` and the isomorphism `hF : eA.functor ⋙ e'.functor ≅ F`. -/ +@[simps] +def υ : eA.functor ≅ F ⋙ e'.inverse := +calc eA.functor ≅ eA.functor ⋙ 𝟭 A' : (functor.left_unitor _).symm +... ≅ eA.functor ⋙ (e'.functor ⋙ e'.inverse) : iso_whisker_left _ e'.unit_iso +... ≅ (eA.functor ⋙ e'.functor) ⋙ e'.inverse : iso.refl _ +... ≅ F ⋙ e'.inverse : iso_whisker_right hF _ + +variables (ε : eA.functor ≅ F ⋙ e'.inverse) (hε : υ hF = ε) + +include ε hG +omit hF + +variable (hG) + +/-- The unit isomorphism of `equivalence`. -/ +@[simps] +def equivalence_unit_iso : 𝟭 A ≅ (F ⋙ eB.inverse) ⋙ G := +calc 𝟭 A ≅ eA.functor ⋙ eA.inverse : eA.unit_iso +... ≅ (F ⋙ e'.inverse) ⋙ eA.inverse : iso_whisker_right ε _ +... ≅ F ⋙ 𝟭 B' ⋙ e'.inverse ⋙ eA.inverse : iso.refl _ +... ≅ F ⋙ (eB.inverse ⋙ eB.functor) ⋙ (e'.inverse ⋙ eA.inverse) : + iso_whisker_left _ (iso_whisker_right eB.counit_iso.symm _) +... ≅ (F ⋙ eB.inverse) ⋙ (eB.functor ⋙ e'.inverse) ⋙ eA.inverse : iso.refl _ +... ≅ (F ⋙ eB.inverse) ⋙ (G ⋙ eA.functor) ⋙ eA.inverse : + iso_whisker_left _ (iso_whisker_right hG _) +... ≅ (F ⋙ eB.inverse ⋙ G) ⋙ (eA.functor ⋙ eA.inverse) : iso.refl _ +... ≅ (F ⋙ eB.inverse ⋙ G) ⋙ 𝟭 A : iso_whisker_left _ eA.unit_iso.symm +... ≅ (F ⋙ eB.inverse) ⋙ G : iso.refl _ + +include hε +variables {ε hF hG} + +lemma equivalence_unit_iso_eq : + (equivalence hF hG).unit_iso = equivalence_unit_iso hG ε := +begin + ext1, apply nat_trans.ext, ext X, + dsimp [equivalence, iso.refl, nat_iso.hcomp, is_equivalence.inverse, + is_equivalence.of_equivalence], + erw [nat_trans.id_app, id_comp, G.map_id, comp_id, comp_id], + simp only [equivalence₂_unit_iso_eq eB hF, equivalence₂_unit_iso_hom_app], + dsimp [equivalence₂, equivalence₁], + simp only [assoc, equivalence_unit_iso_hom_app, nat_iso.cancel_nat_iso_hom_left, + ← eA.inverse.map_comp_assoc, ← hε, υ_hom_app], +end + end compatibility end dold_kan From d11f435d4e34a6cea0a1797d6b625b0c170be845 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 26 Jul 2023 13:39:42 +0000 Subject: [PATCH 40/61] feat(linear_algebra/quadratic_form): some work from 'Clifford Algebras and Spinor Norms Over a Commutative Ring' (#18447) --- docs/references.bib | 12 ++ src/linear_algebra/quadratic_form/basic.lean | 2 +- src/linear_algebra/quadratic_form/dual.lean | 151 +++++++++++++++++++ 3 files changed, 164 insertions(+), 1 deletion(-) create mode 100644 src/linear_algebra/quadratic_form/dual.lean diff --git a/docs/references.bib b/docs/references.bib index 03a341208d7bb..ece945a541bf9 100644 --- a/docs/references.bib +++ b/docs/references.bib @@ -1317,6 +1317,18 @@ @Article{ huneke2002 url = {https://doi.org/10.1080/00029890.2002.11919853} } +@InProceedings{ hyman1973, + author = "Bass, Hyman", + editor = "Bass, Hyman", + title = "Unitary algebraic K-theory", + booktitle = "Hermitian K-Theory and Geometric Applications", + year = "1973", + publisher = "Springer Berlin Heidelberg", + address = "Berlin, Heidelberg", + pages = "57--265", + isbn = "978-3-540-37773-3" +} + @Book{ iordanescu2003, author = {Radu {Iord\u{a}nescu}}, title = {{Jordan structures in geometry and physics. With an diff --git a/src/linear_algebra/quadratic_form/basic.lean b/src/linear_algebra/quadratic_form/basic.lean index cea7b02947391..663a7df5eafb5 100644 --- a/src/linear_algebra/quadratic_form/basic.lean +++ b/src/linear_algebra/quadratic_form/basic.lean @@ -706,7 +706,7 @@ variables [invertible (2 : R₁)] /-- `associated` is the linear map that sends a quadratic form over a commutative ring to its associated symmetric bilinear form. -/ -abbreviation associated : quadratic_form R₁ M →ₗ[R₁] bilin_form R₁ M := +@[reducible] def associated : quadratic_form R₁ M →ₗ[R₁] bilin_form R₁ M := associated_hom R₁ @[simp] lemma associated_lin_mul_lin (f g : M →ₗ[R₁] R₁) : diff --git a/src/linear_algebra/quadratic_form/dual.lean b/src/linear_algebra/quadratic_form/dual.lean new file mode 100644 index 0000000000000..68ed89fab74cb --- /dev/null +++ b/src/linear_algebra/quadratic_form/dual.lean @@ -0,0 +1,151 @@ +/- +Copyright (c) 2023 Eric Wieser. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Eric Wieser +-/ + +import linear_algebra.quadratic_form.isometry +import linear_algebra.quadratic_form.prod +/-! +# Quadratic form structures related to `module.dual` + +## Main definitions + +* `bilin_form.dual_prod R M`, the bilinear form on `(f, x) : module.dual R M × M` defined as + `f x`. +* `quadratic_form.dual_prod R M`, the quadratic form on `(f, x) : module.dual R M × M` defined as + `f x`. +* `quadratic_form.to_dual_prod : M × M →ₗ[R] module.dual R M × M` a form-preserving map from + `(Q.prod $ -Q)` to `quadratic_form.dual_prod R M`. Note that we do not have the morphism + version of `quadratic_form.isometry`, so for now this is stated without full bundling. + +-/ + +variables (R M N : Type*) + +namespace bilin_form + +section semiring +variables [comm_semiring R] [add_comm_monoid M] [module R M] + +/-- The symmetric bilinear form on `module.dual R M × M` defined as +`B (f, x) (g, y) = f y + g x`. -/ +@[simps] +def dual_prod : bilin_form R (module.dual R M × M) := +linear_map.to_bilin $ + (linear_map.applyₗ.comp (linear_map.snd R (module.dual R M) M)).compl₂ + (linear_map.fst R (module.dual R M) M) + + ((linear_map.applyₗ.comp (linear_map.snd R (module.dual R M) M)).compl₂ + (linear_map.fst R (module.dual R M) M)).flip + +lemma is_symm_dual_prod : (dual_prod R M).is_symm := +λ x y, add_comm _ _ + +end semiring + +section ring +variables [comm_ring R] [add_comm_group M] [module R M] + +lemma nondenerate_dual_prod : + (dual_prod R M).nondegenerate ↔ function.injective (module.dual.eval R M) := +begin + classical, + rw nondegenerate_iff_ker_eq_bot, + rw linear_map.ker_eq_bot, + let e := linear_equiv.prod_comm R _ _ + ≪≫ₗ module.dual_prod_dual_equiv_dual R (module.dual R M) M, + let h_d := e.symm.to_linear_map.comp (dual_prod R M).to_lin, + refine (function.injective.of_comp_iff e.symm.injective (dual_prod R M).to_lin).symm.trans _, + rw [←linear_equiv.coe_to_linear_map, ←linear_map.coe_comp], + change function.injective h_d ↔ _, + have : h_d = linear_map.prod_map (linear_map.id) (module.dual.eval R M), + { refine linear_map.ext (λ x, prod.ext _ _), + { ext, + dsimp [h_d, module.dual.eval, linear_equiv.prod_comm], + simp }, + { ext, + dsimp [h_d, module.dual.eval, linear_equiv.prod_comm], + simp } }, + rw [this, linear_map.coe_prod_map], + refine prod.map_injective.trans _, + exact and_iff_right function.injective_id, +end + +end ring + +end bilin_form + +namespace quadratic_form + +section semiring +variables [comm_semiring R] [add_comm_monoid M] [add_comm_monoid N] [module R M] [module R N] + +/-- The quadratic form on `module.dual R M × M` defined as `Q (f, x) = f x`. -/ +@[simps] +def dual_prod : quadratic_form R (module.dual R M × M) := +{ to_fun := λ p, p.1 p.2, + to_fun_smul := λ a p, by rw [prod.smul_fst, prod.smul_snd, linear_map.smul_apply, + linear_map.map_smul, smul_eq_mul, smul_eq_mul, mul_assoc], + exists_companion' := ⟨bilin_form.dual_prod R M, λ p q, begin + rw [bilin_form.dual_prod_apply, prod.fst_add, prod.snd_add, linear_map.add_apply, map_add, + map_add, add_right_comm _ (q.1 q.2), add_comm (q.1 p.2) (p.1 q.2), ←add_assoc, ←add_assoc], + end⟩ } + +@[simp] +lemma _root_.bilin_form.dual_prod.to_quadratic_form : + (bilin_form.dual_prod R M).to_quadratic_form = 2 • dual_prod R M := +ext $ λ a, (two_nsmul _).symm + +variables {R M N} + +/-- Any module isomorphism induces a quadratic isomorphism between the corresponding `dual_prod.` -/ +@[simps] +def dual_prod_isometry (f : M ≃ₗ[R] N) : + (dual_prod R M).isometry (dual_prod R N) := +{ to_linear_equiv := f.dual_map.symm.prod f, + map_app' := λ x, fun_like.congr_arg x.fst $ f.symm_apply_apply _ } + +/-- `quadratic_form.dual_prod` commutes (isometrically) with `quadratic_form.prod`. -/ +@[simps] +def dual_prod_prod_isometry : + (dual_prod R (M × N)).isometry ((dual_prod R M).prod (dual_prod R N)) := +{ to_linear_equiv := + ((module.dual_prod_dual_equiv_dual R M N).symm.prod (linear_equiv.refl R (M × N))) + ≪≫ₗ linear_equiv.prod_prod_prod_comm R _ _ M N, + map_app' := λ m, + (m.fst.map_add _ _).symm.trans $ fun_like.congr_arg m.fst $ prod.ext (add_zero _) (zero_add _) } + +end semiring + +section ring +variables [comm_ring R] [add_comm_group M] [module R M] + +variables {R M} + +/-- The isometry sending `(Q.prod $ -Q)` to `(quadratic_form.dual_prod R M)`. + +This is `σ` from Proposition 4.8, page 84 of +[*Hermitian K-Theory and Geometric Applications*][hyman1973]; though we swap the order of the pairs. +-/ +@[simps] +def to_dual_prod (Q : quadratic_form R M) [invertible (2 : R)] : + M × M →ₗ[R] module.dual R M × M := +linear_map.prod + (Q.associated.to_lin.comp (linear_map.fst _ _ _) + + Q.associated.to_lin.comp (linear_map.snd _ _ _)) + ((linear_map.fst _ _ _ - linear_map.snd _ _ _)) + +lemma to_dual_prod_isometry [invertible (2 : R)] (Q : quadratic_form R M) (x : M × M) : + quadratic_form.dual_prod R M (to_dual_prod Q x) = (Q.prod $ -Q) x := +begin + dsimp only [to_dual_prod, associated, associated_hom], + dsimp, + simp [polar_comm _ x.1 x.2, ←sub_add, mul_sub, sub_mul, smul_sub, submonoid.smul_def, + ←sub_eq_add_neg (Q x.1) (Q x.2)], +end + +-- TODO: show that `to_dual_prod` is an equivalence + +end ring + +end quadratic_form From fe18deda804e30c594e75a6e5fe0f7d14695289f Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 27 Jul 2023 10:40:44 +0000 Subject: [PATCH 41/61] feat(analysis/normed_space/continuous_linear_map): generalize typeclass assumptions (#19108) --- .../normed_space/continuous_linear_map.lean | 81 ++++++++++--------- 1 file changed, 41 insertions(+), 40 deletions(-) diff --git a/src/analysis/normed_space/continuous_linear_map.lean b/src/analysis/normed_space/continuous_linear_map.lean index a46af72230230..1ca9737ceac0a 100644 --- a/src/analysis/normed_space/continuous_linear_map.lean +++ b/src/analysis/normed_space/continuous_linear_map.lean @@ -35,14 +35,14 @@ open_locale nnreal variables {𝕜 𝕜₂ E F G : Type*} -variables [normed_field 𝕜] [normed_field 𝕜₂] /-! # General constructions -/ -section seminormed +section seminormed_add_comm_group +variables [ring 𝕜] [ring 𝕜₂] variables [seminormed_add_comm_group E] [seminormed_add_comm_group F] [seminormed_add_comm_group G] -variables [normed_space 𝕜 E] [normed_space 𝕜₂ F] [normed_space 𝕜 G] +variables [module 𝕜 E] [module 𝕜₂ F] [module 𝕜 G] variables {σ : 𝕜 →+* 𝕜₂} (f : E →ₛₗ[σ] F) /-- Construct a continuous linear map from a linear map and a bound on this linear map. @@ -51,13 +51,6 @@ The fact that the norm of the continuous linear map is then controlled is given def linear_map.mk_continuous (C : ℝ) (h : ∀x, ‖f x‖ ≤ C * ‖x‖) : E →SL[σ] F := ⟨f, add_monoid_hom_class.continuous_of_bound f C h⟩ -/-- Reinterpret a linear map `𝕜 →ₗ[𝕜] E` as a continuous linear map. This construction -is generalized to the case of any finite dimensional domain -in `linear_map.to_continuous_linear_map`. -/ -def linear_map.to_continuous_linear_map₁ (f : 𝕜 →ₗ[𝕜] E) : 𝕜 →L[𝕜] E := -f.mk_continuous (‖f 1‖) $ λ x, le_of_eq $ -by { conv_lhs { rw ← mul_one x }, rw [← smul_eq_mul, f.map_smul, norm_smul, mul_comm] } - /-- Construct a continuous linear map from a linear map and the existence of a bound on this linear map. If you have an explicit bound, use `linear_map.mk_continuous` instead, as a norm estimate will follow automatically in `linear_map.mk_continuous_norm_le`. -/ @@ -89,14 +82,6 @@ add_monoid_hom_class.continuous_of_bound φ C h_bound @[simp] lemma linear_map.mk_continuous_of_exists_bound_apply (h : ∃C, ∀x, ‖f x‖ ≤ C * ‖x‖) (x : E) : f.mk_continuous_of_exists_bound h x = f x := rfl -@[simp] lemma linear_map.to_continuous_linear_map₁_coe (f : 𝕜 →ₗ[𝕜] E) : - (f.to_continuous_linear_map₁ : 𝕜 →ₗ[𝕜] E) = f := -rfl - -@[simp] lemma linear_map.to_continuous_linear_map₁_apply (f : 𝕜 →ₗ[𝕜] E) (x) : - f.to_continuous_linear_map₁ x = f x := -rfl - namespace continuous_linear_map theorem antilipschitz_of_bound (f : E →SL[σ] F) {K : ℝ≥0} (h : ∀ x, ‖x‖ ≤ K * ‖f x‖) : @@ -125,11 +110,34 @@ def linear_equiv.to_continuous_linear_equiv_of_bounds (e : E ≃ₛₗ[σ] F) (C end -end seminormed +end seminormed_add_comm_group + +section seminormed_bounded + +variables [semi_normed_ring 𝕜] [ring 𝕜₂] [seminormed_add_comm_group E] +variables [module 𝕜 E] [has_bounded_smul 𝕜 E] + +/-- Reinterpret a linear map `𝕜 →ₗ[𝕜] E` as a continuous linear map. This construction +is generalized to the case of any finite dimensional domain +in `linear_map.to_continuous_linear_map`. -/ +def linear_map.to_continuous_linear_map₁ (f : 𝕜 →ₗ[𝕜] E) : 𝕜 →L[𝕜] E := +f.mk_continuous (‖f 1‖) $ λ x, +by { conv_lhs { rw ← mul_one x }, rw [← smul_eq_mul, f.map_smul, mul_comm],exact norm_smul_le _ _ } + +@[simp] lemma linear_map.to_continuous_linear_map₁_coe (f : 𝕜 →ₗ[𝕜] E) : + (f.to_continuous_linear_map₁ : 𝕜 →ₗ[𝕜] E) = f := +rfl + +@[simp] lemma linear_map.to_continuous_linear_map₁_apply (f : 𝕜 →ₗ[𝕜] E) (x) : + f.to_continuous_linear_map₁ x = f x := +rfl + +end seminormed_bounded section normed -variables [normed_add_comm_group E] [normed_add_comm_group F] [normed_space 𝕜 E] [normed_space 𝕜₂ F] +variables [ring 𝕜] [ring 𝕜₂] +variables [normed_add_comm_group E] [normed_add_comm_group F] [module 𝕜 E] [module 𝕜₂ F] variables {σ : 𝕜 →+* 𝕜₂} (f g : E →SL[σ] F) (x y z : E) theorem continuous_linear_map.uniform_embedding_of_bound {K : ℝ≥0} (hf : ∀ x, ‖x‖ ≤ K * ‖f x‖) : @@ -142,8 +150,9 @@ end normed section seminormed +variables [ring 𝕜] [ring 𝕜₂] variables [seminormed_add_comm_group E] [seminormed_add_comm_group F] -variables [normed_space 𝕜 E] [normed_space 𝕜₂ F] +variables [module 𝕜 E] [module 𝕜₂ F] variables {σ : 𝕜 →+* 𝕜₂} (f : E →ₛₗ[σ] F) /-- A (semi-)linear map which is a homothety is a continuous linear map. @@ -180,42 +189,34 @@ end seminormed /-! ## The span of a single vector -/ -section seminormed - -variables [seminormed_add_comm_group E] [normed_space 𝕜 E] - namespace continuous_linear_map +variables (𝕜) -variable (𝕜) +section seminormed +variables [normed_division_ring 𝕜] [seminormed_add_comm_group E] [module 𝕜 E] [has_bounded_smul 𝕜 E] lemma to_span_singleton_homothety (x : E) (c : 𝕜) : ‖linear_map.to_span_singleton 𝕜 E x c‖ = ‖x‖ * ‖c‖ := by {rw mul_comm, exact norm_smul _ _} -end continuous_linear_map +end seminormed -section +end continuous_linear_map namespace continuous_linear_equiv - variable (𝕜) +section seminormed +variables [normed_division_ring 𝕜] [seminormed_add_comm_group E] [module 𝕜 E] [has_bounded_smul 𝕜 E] + lemma to_span_nonzero_singleton_homothety (x : E) (h : x ≠ 0) (c : 𝕜) : ‖linear_equiv.to_span_nonzero_singleton 𝕜 E x h c‖ = ‖x‖ * ‖c‖ := continuous_linear_map.to_span_singleton_homothety _ _ _ -end continuous_linear_equiv - -end - end seminormed section normed - -variables [normed_add_comm_group E] [normed_space 𝕜 E] - -namespace continuous_linear_equiv -variable (𝕜) +variables [normed_field 𝕜] [normed_add_comm_group E] [normed_space 𝕜 E] /-- Given a nonzero element `x` of a normed space `E₁` over a field `𝕜`, the natural continuous linear equivalence from `E₁` to the span of `x`.-/ @@ -246,6 +247,6 @@ noncomputable def coord (x : E) (h : x ≠ 0) : (𝕜 ∙ x) →L[𝕜] 𝕜 := (coord 𝕜 x h) (⟨x, submodule.mem_span_singleton_self x⟩ : 𝕜 ∙ x) = 1 := linear_equiv.coord_self 𝕜 E x h -end continuous_linear_equiv - end normed + +end continuous_linear_equiv From e90e0a6119d01e3ef1703b3038b555785a998285 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 28 Jul 2023 07:23:47 +0000 Subject: [PATCH 42/61] feat(tactic/positivity): Extension for `ite` (#17650) Add `positivity_ite`, an extension for `ite`. Co-authored-by: Jireh Loreaux --- src/tactic/positivity.lean | 46 ++++++++++++++++++++++++++++++++++++++ test/positivity.lean | 12 ++++++++++ 2 files changed, 58 insertions(+) diff --git a/src/tactic/positivity.lean b/src/tactic/positivity.lean index 6544387973144..bec4bfd4a8cb8 100644 --- a/src/tactic/positivity.lean +++ b/src/tactic/positivity.lean @@ -352,6 +352,52 @@ variables {ι α R : Type*} /-! ### `positivity` extensions for particular arithmetic operations -/ +section ite +variables [has_zero α] {p : Prop} [decidable p] {a b : α} + +private lemma ite_pos [has_lt α] (ha : 0 < a) (hb : 0 < b) : 0 < ite p a b := +by by_cases p; simp [*] + +private lemma ite_nonneg [has_le α] (ha : 0 ≤ a) (hb : 0 ≤ b) : 0 ≤ ite p a b := +by by_cases p; simp [*] + +private lemma ite_nonneg_of_pos_of_nonneg [preorder α] (ha : 0 < a) (hb : 0 ≤ b) : 0 ≤ ite p a b := +ite_nonneg ha.le hb + +private lemma ite_nonneg_of_nonneg_of_pos [preorder α] (ha : 0 ≤ a) (hb : 0 < b) : 0 ≤ ite p a b := +ite_nonneg ha hb.le + +private lemma ite_ne_zero (ha : a ≠ 0) (hb : b ≠ 0) : ite p a b ≠ 0 := by by_cases p; simp [*] + +private lemma ite_ne_zero_of_pos_of_ne_zero [preorder α] (ha : 0 < a) (hb : b ≠ 0) : + ite p a b ≠ 0 := +ite_ne_zero ha.ne' hb + +private lemma ite_ne_zero_of_ne_zero_of_pos [preorder α] (ha : a ≠ 0) (hb : 0 < b) : + ite p a b ≠ 0 := +ite_ne_zero ha hb.ne' + +end ite + +/-- Extension for the `positivity` tactic: the `if then else` of two numbers is +positive/nonnegative/nonzero if both are. -/ +@[positivity] +meta def positivity_ite : expr → tactic strictness +| e@`(@ite %%typ %%p %%hp %%a %%b) := do + strictness_a ← core a, + strictness_b ← core b, + match strictness_a, strictness_b with + | positive pa, positive pb := positive <$> mk_app ``ite_pos [pa, pb] + | positive pa, nonnegative pb := nonnegative <$> mk_app ``ite_nonneg_of_pos_of_nonneg [pa, pb] + | nonnegative pa, positive pb := nonnegative <$> mk_app ``ite_nonneg_of_nonneg_of_pos [pa, pb] + | nonnegative pa, nonnegative pb := nonnegative <$> mk_app ``ite_nonneg [pa, pb] + | positive pa, nonzero pb := nonzero <$> to_expr ``(ite_ne_zero_of_pos_of_ne_zero %%pa %%pb) + | nonzero pa, positive pb := nonzero <$> to_expr ``(ite_ne_zero_of_ne_zero_of_pos %%pa %%pb) + | nonzero pa, nonzero pb := nonzero <$> to_expr ``(ite_ne_zero %%pa %%pb) + | sa@_, sb@ _ := positivity_fail e a b sa sb + end +| e := pp e >>= fail ∘ format.bracket "The expression `" "` isn't of the form `ite p a b`" + section linear_order variables [linear_order R] {a b c : R} diff --git a/test/positivity.lean b/test/positivity.lean index f4dc7425036ef..007d449a37bf3 100644 --- a/test/positivity.lean +++ b/test/positivity.lean @@ -80,6 +80,18 @@ example [has_zero α] [preorder α] {a : α} (ha : 0 < a) : 0 ≤ const ι a := example [has_zero α] [preorder α] {a : α} (ha : 0 ≤ a) : 0 ≤ const ι a := by positivity example [nonempty ι] [has_zero α] [preorder α] {a : α} (ha : 0 < a) : 0 < const ι a := by positivity +section ite +variables {p : Prop} [decidable p] {a b : ℤ} + +example (ha : 0 < a) (hb : 0 < b) : 0 < ite p a b := by positivity +example (ha : 0 < a) (hb : 0 ≤ b) : 0 ≤ ite p a b := by positivity +example (ha : 0 ≤ a) (hb : 0 < b) : 0 ≤ ite p a b := by positivity +example (ha : 0 < a) (hb : b ≠ 0) : ite p a b ≠ 0 := by positivity +example (ha : a ≠ 0) (hb : 0 < b) : ite p a b ≠ 0 := by positivity +example (ha : a ≠ 0) (hb : b ≠ 0) : ite p a b ≠ 0 := by positivity + +end ite + example {a b : ℚ} (ha : 0 < a) (hb : 0 < b) : 0 < min a b := by positivity example {a b : ℚ} (ha : 0 < a) (hb : 0 ≤ b) : 0 ≤ min a b := by positivity example {a b : ℚ} (ha : 0 ≤ a) (hb : 0 < b) : 0 ≤ min a b := by positivity From 8900d545017cd21961daa2a1734bb658ef52c618 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Fri, 28 Jul 2023 10:11:06 +0000 Subject: [PATCH 43/61] feat(set_theory/game/pgame): small sets of pre-games / games / surreals are bounded (#15260) --- src/set_theory/game/basic.lean | 10 ++++++ src/set_theory/game/pgame.lean | 53 +++++++++++++++++++++++++++++++ src/set_theory/surreal/basic.lean | 28 ++++++++++++++++ 3 files changed, 91 insertions(+) diff --git a/src/set_theory/game/basic.lean b/src/set_theory/game/basic.lean index e98541fca9c39..d1f574448fbed 100644 --- a/src/set_theory/game/basic.lean +++ b/src/set_theory/game/basic.lean @@ -128,6 +128,16 @@ instance ordered_add_comm_group : ordered_add_comm_group game := ..game.add_comm_group_with_one, ..game.partial_order } +/-- A small set `s` of games is bounded above. -/ +lemma bdd_above_of_small (s : set game.{u}) [small.{u} s] : bdd_above s := +⟨_, λ i hi, by simpa using pgame.le_iff_game_le.1 + (upper_bound_mem_upper_bounds _ (set.mem_image_of_mem quotient.out hi))⟩ + +/-- A small set `s` of games is bounded below. -/ +lemma bdd_below_of_small (s : set game.{u}) [small.{u} s] : bdd_below s := +⟨_, λ i hi, by simpa using pgame.le_iff_game_le.1 + (lower_bound_mem_lower_bounds _ (set.mem_image_of_mem quotient.out hi))⟩ + end game namespace pgame diff --git a/src/set_theory/game/pgame.lean b/src/set_theory/game/pgame.lean index 1cfa0e55c037c..b6a9580b9cbb2 100644 --- a/src/set_theory/game/pgame.lean +++ b/src/set_theory/game/pgame.lean @@ -6,6 +6,7 @@ Authors: Reid Barton, Mario Carneiro, Isabel Longbottom, Scott Morrison import data.fin.basic import data.list.basic import logic.relation +import logic.small.basic import order.game_add /-! @@ -491,6 +492,58 @@ lemma left_response_spec {x : pgame} (h : 0 ≤ x) (j : x.right_moves) : 0 ≤ (x.move_right j).move_left (left_response h j) := classical.some_spec $ (zero_le.1 h) j +/-- An explicit upper bound for a family of pre-games, whose left moves are the union of the left +moves of all the pre-games in the family. -/ +def upper_bound {ι : Type u} (f : ι → pgame.{u}) : pgame := +⟨Σ i, (f i).left_moves, pempty, λ x, move_left _ x.2, pempty.elim⟩ + +instance upper_bound_right_moves_empty {ι : Type u} (f : ι → pgame.{u}) : + is_empty (upper_bound f).right_moves := +pempty.is_empty + +theorem le_upper_bound {ι : Type u} (f : ι → pgame.{u}) (i : ι) : f i ≤ upper_bound f := +begin + rw [upper_bound, le_iff_forall_lf], + dsimp, + simp only [and_true, is_empty.forall_iff], + exact λ j, @move_left_lf (upper_bound f) ⟨i, j⟩ +end + +lemma upper_bound_mem_upper_bounds (s : set pgame.{u}) [small.{u} s] : + upper_bound (subtype.val ∘ (equiv_shrink s).symm) ∈ upper_bounds s := +λ i hi, by simpa using + le_upper_bound (subtype.val ∘ (equiv_shrink s).symm) (equiv_shrink s ⟨i, hi⟩) + +/-- A small set `s` of pre-games is bounded above. -/ +lemma bdd_above_of_small (s : set pgame.{u}) [small.{u} s] : bdd_above s := +⟨_, upper_bound_mem_upper_bounds s⟩ + +/-- An explicit lower bound for a family of pre-games, whose right moves are the union of the right +moves of all the pre-games in the family. -/ +def lower_bound {ι : Type u} (f : ι → pgame.{u}) : pgame := +⟨pempty, Σ i, (f i).right_moves, pempty.elim, λ x, move_right _ x.2⟩ + +instance lower_bound_left_moves_empty {ι : Type u} (f : ι → pgame.{u}) : + is_empty (lower_bound f).left_moves := +pempty.is_empty + +theorem lower_bound_le {ι : Type u} (f : ι → pgame.{u}) (i : ι) : lower_bound f ≤ f i := +begin + rw [lower_bound, le_iff_forall_lf], + dsimp, + simp only [is_empty.forall_iff, true_and], + exact λ j, @lf_move_right (lower_bound f) ⟨i, j⟩ +end + +lemma lower_bound_mem_lower_bounds (s : set pgame.{u}) [small.{u} s] : + lower_bound (subtype.val ∘ (equiv_shrink s).symm) ∈ lower_bounds s := +λ i hi, by simpa using + lower_bound_le (subtype.val ∘ (equiv_shrink s).symm) (equiv_shrink s ⟨i, hi⟩) + +/-- A small set `s` of pre-games is bounded below. -/ +lemma bdd_below_of_small (s : set pgame.{u}) [small.{u} s] : bdd_below s := +⟨_, lower_bound_mem_lower_bounds s⟩ + /-- The equivalence relation on pre-games. Two pre-games `x`, `y` are equivalent if `x ≤ y` and `y ≤ x`. diff --git a/src/set_theory/surreal/basic.lean b/src/set_theory/surreal/basic.lean index 6a9d8bf27327a..fe2df6135ac01 100644 --- a/src/set_theory/surreal/basic.lean +++ b/src/set_theory/surreal/basic.lean @@ -328,6 +328,34 @@ theorem zero_to_game : to_game 0 = 0 := rfl @[simp] theorem one_to_game : to_game 1 = 1 := rfl @[simp] theorem nat_to_game : ∀ n : ℕ, to_game n = n := map_nat_cast' _ one_to_game +theorem upper_bound_numeric {ι : Type u} {f : ι → pgame.{u}} (H : ∀ i, (f i).numeric) : + (upper_bound f).numeric := +numeric_of_is_empty_right_moves _ $ λ i, (H _).move_left _ + +theorem lower_bound_numeric {ι : Type u} {f : ι → pgame.{u}} (H : ∀ i, (f i).numeric) : + (lower_bound f).numeric := +numeric_of_is_empty_left_moves _ $ λ i, (H _).move_right _ + +/-- A small set `s` of surreals is bounded above. -/ +lemma bdd_above_of_small (s : set surreal.{u}) [small.{u} s] : bdd_above s := +begin + let g := subtype.val ∘ quotient.out ∘ subtype.val ∘ (equiv_shrink s).symm, + refine ⟨mk (upper_bound g) (upper_bound_numeric $ λ i, subtype.prop _), λ i hi, _⟩, + rw ←quotient.out_eq i, + show i.out.1 ≤ _, + simpa [g] using le_upper_bound g (equiv_shrink s ⟨i, hi⟩) +end + +/-- A small set `s` of surreals is bounded below. -/ +lemma bdd_below_of_small (s : set surreal.{u}) [small.{u} s] : bdd_below s := +begin + let g := subtype.val ∘ quotient.out ∘ subtype.val ∘ (equiv_shrink s).symm, + refine ⟨mk (lower_bound g) (lower_bound_numeric $ λ i, subtype.prop _), λ i hi, _⟩, + rw ←quotient.out_eq i, + show _ ≤ i.out.1, + simpa [g] using lower_bound_le g (equiv_shrink s ⟨i, hi⟩) +end + end surreal open surreal From 188a411e916e1119e502dbe35b8b475716362401 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9mi=20Bottinelli?= Date: Fri, 28 Jul 2023 10:11:08 +0000 Subject: [PATCH 44/61] feat(combinatorics/quiver/covering): Definition of coverings and unique lifting of paths (#17828) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: Antoine Labelle Co-authored-by: Rémi Bottinelli Co-authored-by: Kyle Miller Co-authored-by: antoinelab01 <66086247+antoinelab01@users.noreply.github.com> Co-authored-by: Antoine Labelle --- src/combinatorics/quiver/covering.lean | 266 +++++++++++++++++++++++++ 1 file changed, 266 insertions(+) create mode 100644 src/combinatorics/quiver/covering.lean diff --git a/src/combinatorics/quiver/covering.lean b/src/combinatorics/quiver/covering.lean new file mode 100644 index 0000000000000..18c98970d2cfd --- /dev/null +++ b/src/combinatorics/quiver/covering.lean @@ -0,0 +1,266 @@ +/- +Copyright (c) 2022 Antoine Labelle, Rémi Bottinelli. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Antoine Labelle, Rémi Bottinelli +-/ +import combinatorics.quiver.cast +import combinatorics.quiver.symmetric +import data.sigma.basic +import logic.equiv.basic + +/-! +# Covering + +This file defines coverings of quivers as prefunctors that are bijective on the +so-called stars and costars at each vertex of the domain. + +## Main definitions + +* `quiver.star u` is the type of all arrows with source `u`; +* `quiver.costar u` is the type of all arrows with target `u`; +* `prefunctor.star φ u` is the obvious function `star u → star (φ.obj u)`; +* `prefunctor.costar φ u` is the obvious function `costar u → costar (φ.obj u)`; +* `prefunctor.is_covering φ` means that `φ.star u` and `φ.costar u` are bijections for all `u`; +* `quiver.star_path u` is the type of all paths with source `u`; +* `prefunctor.star_path u` is the obvious function `star_path u → star_path (φ.obj u)`. + +## Main statements + +* `prefunctor.is_covering.path_star_bijective` states that if `φ` is a covering, + then `φ.star_path u` is a bijection for all `u`. + In other words, every path in the codomain of `φ` lifts uniquely to its domain. + +## TODO + +Clean up the namespaces by renaming `prefunctor` to `quiver.prefunctor`. + +## Tags + +Cover, covering, quiver, path, lift +-/ + +open function quiver + +universes u v w + +variables {U : Type*} [quiver.{u+1} U] + {V : Type*} [quiver.{v+1} V] (φ : U ⥤q V) + {W : Type*} [quiver.{w+1} W] (ψ : V ⥤q W) + +/-- The `quiver.star` at a vertex is the collection of arrows whose source is the vertex. +The type `quiver.star u` is defined to be `Σ (v : U), (u ⟶ v)`. -/ +@[reducible] def quiver.star (u : U) := Σ (v : U), (u ⟶ v) + +/-- Constructor for `quiver.star`. Defined to be `sigma.mk`. -/ +@[reducible] protected def quiver.star.mk {u v : U} (f : u ⟶ v) : quiver.star u := ⟨_, f⟩ + +/-- The `quiver.costar` at a vertex is the collection of arrows whose target is the vertex. +The type `quiver.costar v` is defined to be `Σ (u : U), (u ⟶ v)`. -/ +@[reducible] def quiver.costar (v : U) := Σ (u : U), (u ⟶ v) + +/-- Constructor for `quiver.costar`. Defined to be `sigma.mk`. -/ +@[reducible] protected def quiver.costar.mk {u v : U} (f : u ⟶ v) : quiver.costar v := ⟨_, f⟩ + +/-- A prefunctor induces a map of `quiver.star` at every vertex. -/ +@[simps] def prefunctor.star (u : U) : quiver.star u → quiver.star (φ.obj u) := +λ F, quiver.star.mk (φ.map F.2) + +/-- A prefunctor induces a map of `quiver.costar` at every vertex. -/ +@[simps] def prefunctor.costar (u : U) : quiver.costar u → quiver.costar (φ.obj u) := +λ F, quiver.costar.mk (φ.map F.2) + +@[simp] lemma prefunctor.star_apply {u v : U} (e : u ⟶ v) : + φ.star u (quiver.star.mk e) = quiver.star.mk (φ.map e) := rfl + +@[simp] lemma prefunctor.costar_apply {u v : U} (e : u ⟶ v) : + φ.costar v (quiver.costar.mk e) = quiver.costar.mk (φ.map e) := rfl + +lemma prefunctor.star_comp (u : U) : + (φ ⋙q ψ).star u = (ψ.star (φ.obj u)) ∘ ((φ.star) u) := rfl + +lemma prefunctor.costar_comp (u : U) : + (φ ⋙q ψ).costar u = (ψ.costar (φ.obj u)) ∘ ((φ.costar) u) := rfl + +/-- A prefunctor is a covering of quivers if it defines bijections on all stars and costars. -/ +protected structure prefunctor.is_covering : Prop := +(star_bijective : ∀ u, bijective (φ.star u)) +(costar_bijective : ∀ u, bijective (φ.costar u)) + +@[simp] lemma prefunctor.is_covering.map_injective (hφ : φ.is_covering) {u v : U} : + injective (λ (f : u ⟶ v), φ.map f) := +begin + rintro f g he, + have : φ.star u (quiver.star.mk f) = φ.star u (quiver.star.mk g) := by simpa using he, + simpa using (hφ.star_bijective u).left this, +end + +lemma prefunctor.is_covering.comp (hφ : φ.is_covering) (hψ : ψ.is_covering) : + (φ ⋙q ψ).is_covering := +⟨λ u, (hψ.star_bijective _).comp (hφ.star_bijective _), + λ u, (hψ.costar_bijective _).comp (hφ.costar_bijective _)⟩ + +lemma prefunctor.is_covering.of_comp_right (hψ : ψ.is_covering) (hφψ : (φ ⋙q ψ).is_covering) : + φ.is_covering := +⟨λ u, (bijective.of_comp_iff' (hψ.star_bijective _) _).mp (hφψ.star_bijective _), + λ u, (bijective.of_comp_iff' (hψ.costar_bijective _) _).mp (hφψ.costar_bijective _)⟩ + +lemma prefunctor.is_covering.of_comp_left (hφ : φ.is_covering) (hφψ : (φ ⋙q ψ).is_covering) + (φsur : surjective φ.obj) : ψ.is_covering := +begin + refine ⟨λ v, _, λ v, _⟩; + obtain ⟨u, rfl⟩ := φsur v, + exacts [(bijective.of_comp_iff _ (hφ.star_bijective u)).mp (hφψ.star_bijective u), + (bijective.of_comp_iff _ (hφ.costar_bijective u)).mp (hφψ.costar_bijective u)], +end + +/-- The star of the symmetrification of a quiver at a vertex `u` is equivalent to the sum of the +star and the costar at `u` in the original quiver. -/ +def quiver.symmetrify_star (u : U) : + quiver.star (symmetrify.of.obj u) ≃ quiver.star u ⊕ quiver.costar u := +equiv.sigma_sum_distrib _ _ + +/-- The costar of the symmetrification of a quiver at a vertex `u` is equivalent to the sum of the +costar and the star at `u` in the original quiver. -/ +def quiver.symmetrify_costar (u : U) : + quiver.costar (symmetrify.of.obj u) ≃ quiver.costar u ⊕ quiver.star u := +equiv.sigma_sum_distrib _ _ + +lemma prefunctor.symmetrify_star (u : U) : + φ.symmetrify.star u = (quiver.symmetrify_star _).symm ∘ + sum.map (φ.star u) (φ.costar u) ∘ quiver.symmetrify_star u := +begin + rw equiv.eq_symm_comp, + ext ⟨v, (f|g)⟩; + simp [quiver.symmetrify_star], +end + +protected lemma prefunctor.symmetrify_costar (u : U) : + φ.symmetrify.costar u = (quiver.symmetrify_costar _).symm ∘ + sum.map (φ.costar u) (φ.star u) ∘ quiver.symmetrify_costar u := +begin + rw equiv.eq_symm_comp, + ext ⟨v, (f|g)⟩; + simp [quiver.symmetrify_costar], +end + +protected lemma prefunctor.is_covering.symmetrify (hφ : φ.is_covering) : φ.symmetrify.is_covering := +begin + refine ⟨λ u, _, λ u, _⟩; + simp [φ.symmetrify_star, φ.symmetrify_costar, hφ.star_bijective u, hφ.costar_bijective u], +end + +/-- The path star at a vertex `u` is the type of all paths starting at `u`. +The type `quiver.path_star u` is defined to be `Σ v : U, path u v`. -/ +@[reducible] def quiver.path_star (u : U) := Σ v : U, path u v + +/-- Constructor for `quiver.path_star`. Defined to be `sigma.mk`. -/ +@[reducible] protected def quiver.path_star.mk {u v : U} (p : path u v) : + quiver.path_star u := ⟨_, p⟩ + +/-- A prefunctor induces a map of path stars. -/ +def prefunctor.path_star (u : U) : quiver.path_star u → quiver.path_star (φ.obj u) := +λ p, quiver.path_star.mk (φ.map_path p.2) + +@[simp] lemma prefunctor.path_star_apply {u v : U} (p : path u v) : + φ.path_star u (quiver.path_star.mk p) = quiver.path_star.mk (φ.map_path p) := rfl + +theorem prefunctor.path_star_injective (hφ : ∀ u, injective (φ.star u)) (u : U) : + injective (φ.path_star u) := +begin + dsimp [prefunctor.path_star, quiver.path_star.mk], + rintro ⟨v₁, p₁⟩, + induction p₁ with x₁ y₁ p₁ e₁ ih; + rintro ⟨y₂, p₂⟩; cases p₂ with x₂ _ p₂ e₂; + intro h; + simp only [prefunctor.path_star_apply, prefunctor.map_path_nil, + prefunctor.map_path_cons] at h, + { exfalso, + cases h with h h', + rw [←path.eq_cast_iff_heq rfl h.symm, path.cast_cons] at h', + exact (path.nil_ne_cons _ _) h', }, + { exfalso, + cases h with h h', + rw [←path.cast_eq_iff_heq rfl h, path.cast_cons] at h', + exact (path.cons_ne_nil _ _) h', }, + { cases h with hφy h', + rw [←path.cast_eq_iff_heq rfl hφy, path.cast_cons, path.cast_rfl_rfl] at h', + have hφx := path.obj_eq_of_cons_eq_cons h', + have hφp := path.heq_of_cons_eq_cons h', + have hφe := heq.trans (hom.cast_heq rfl hφy _).symm (path.hom_heq_of_cons_eq_cons h'), + have h_path_star : φ.path_star u ⟨x₁, p₁⟩ = φ.path_star u ⟨x₂, p₂⟩, + { simp only [prefunctor.path_star_apply, sigma.mk.inj_iff], exact ⟨hφx, hφp⟩, }, + cases ih h_path_star, + have h_star : φ.star x₁ ⟨y₁, e₁⟩ = φ.star x₁ ⟨y₂, e₂⟩, + { simp only [prefunctor.star_apply, sigma.mk.inj_iff], exact ⟨hφy, hφe⟩, }, + cases hφ x₁ h_star, + refl, }, +end + +theorem prefunctor.path_star_surjective (hφ : ∀ u, surjective (φ.star u)) (u : U) : + surjective (φ.path_star u) := +begin + dsimp [prefunctor.path_star, quiver.path_star.mk], + rintro ⟨v, p⟩, + induction p with v' v'' p' ev ih, + { use ⟨u, path.nil⟩, + simp only [prefunctor.map_path_nil, eq_self_iff_true, heq_iff_eq, and_self], }, + { obtain ⟨⟨u', q'⟩, h⟩ := ih, + simp only at h, + obtain ⟨rfl,rfl⟩ := h, + obtain ⟨⟨u'', eu⟩, k⟩ := hφ u' ⟨_, ev⟩, + simp at k, + obtain ⟨rfl,rfl⟩ := k, + use ⟨_, q'.cons eu⟩, + simp only [prefunctor.map_path_cons, eq_self_iff_true, heq_iff_eq, and_self], } +end + +theorem prefunctor.path_star_bijective (hφ : ∀ u, bijective (φ.star u)) (u : U) : + bijective (φ.path_star u) := +⟨φ.path_star_injective (λ u, (hφ u).1) _, φ.path_star_surjective (λ u, (hφ u).2) _⟩ + +namespace prefunctor.is_covering +variable {φ} + +protected theorem path_star_bijective (hφ : φ.is_covering) (u : U) : + bijective (φ.path_star u) := φ.path_star_bijective hφ.1 u + +end prefunctor.is_covering + +section has_involutive_reverse + +variables [has_involutive_reverse U] [has_involutive_reverse V] [prefunctor.map_reverse φ] + +/-- In a quiver with involutive inverses, the star and costar at every vertex are equivalent. +This map is induced by `quiver.reverse`. -/ +@[simps] def quiver.star_equiv_costar (u : U) : + quiver.star u ≃ quiver.costar u := +{ to_fun := λ e, ⟨e.1, reverse e.2⟩, + inv_fun := λ e, ⟨e.1, reverse e.2⟩, + left_inv := λ e, by simp [sigma.ext_iff], + right_inv := λ e, by simp [sigma.ext_iff] } + +@[simp] lemma quiver.star_equiv_costar_apply {u v : U} (e : u ⟶ v) : + quiver.star_equiv_costar u (quiver.star.mk e) = quiver.costar.mk (reverse e) := rfl + +@[simp] lemma quiver.star_equiv_costar_symm_apply {u v : U} (e : u ⟶ v) : + (quiver.star_equiv_costar v).symm (quiver.costar.mk e) = quiver.star.mk (reverse e) := rfl + +lemma prefunctor.costar_conj_star (u : U) : + φ.costar u = + quiver.star_equiv_costar (φ.obj u) ∘ φ.star u ∘ (quiver.star_equiv_costar u).symm := +by { ext ⟨v, f⟩; simp, } + +lemma prefunctor.bijective_costar_iff_bijective_star (u : U) : + bijective (φ.costar u) ↔ bijective (φ.star u) := +begin + rw [prefunctor.costar_conj_star, bijective.of_comp_iff', bijective.of_comp_iff]; + exact equiv.bijective _, +end + +lemma prefunctor.is_covering_of_bijective_star (h : ∀ u, bijective (φ.star u)) : + φ.is_covering := ⟨h, λ u, (φ.bijective_costar_iff_bijective_star u).2 (h u)⟩ + +lemma prefunctor.is_covering_of_bijective_costar (h : ∀ u, bijective (φ.costar u)) : + φ.is_covering := ⟨λ u, (φ.bijective_costar_iff_bijective_star u).1 (h u), h⟩ + +end has_involutive_reverse From 0c1d80f5a86b36c1db32e021e8d19ae7809d5b79 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Fri, 28 Jul 2023 10:11:09 +0000 Subject: [PATCH 45/61] feat(linear_algebra/orientation): add `orientation.reindex` (#19236) --- src/linear_algebra/alternating.lean | 30 ++++++++++++++++++++++++ src/linear_algebra/determinant.lean | 5 ++++ src/linear_algebra/orientation.lean | 36 ++++++++++++++++++++++++++--- 3 files changed, 68 insertions(+), 3 deletions(-) diff --git a/src/linear_algebra/alternating.lean b/src/linear_algebra/alternating.lean index 0fa1eccdefb48..4f23645d2b111 100644 --- a/src/linear_algebra/alternating.lean +++ b/src/linear_algebra/alternating.lean @@ -581,6 +581,12 @@ rfl (f + g).dom_dom_congr σ = f.dom_dom_congr σ + g.dom_dom_congr σ := rfl +@[simp] lemma dom_dom_congr_smul {S : Type*} + [monoid S] [distrib_mul_action S N] [smul_comm_class R S N] (σ : ι ≃ ι') (c : S) + (f : alternating_map R M N ι) : + (c • f).dom_dom_congr σ = c • f.dom_dom_congr σ := +rfl + /-- `alternating_map.dom_dom_congr` as an equivalence. This is declared separately because it does not work with dot notation. -/ @@ -593,6 +599,30 @@ def dom_dom_congr_equiv (σ : ι ≃ ι') : right_inv := λ m, by { ext, simp [function.comp] }, map_add' := dom_dom_congr_add σ } +section dom_dom_lcongr +variables (S : Type*) [semiring S] [module S N] [smul_comm_class R S N] + +/-- `alternating_map.dom_dom_congr` as a linear equivalence. -/ +@[simps apply symm_apply] +def dom_dom_lcongr (σ : ι ≃ ι') : alternating_map R M N ι ≃ₗ[S] alternating_map R M N ι' := +{ to_fun := dom_dom_congr σ, + inv_fun := dom_dom_congr σ.symm, + left_inv := λ f, by { ext, simp [function.comp] }, + right_inv := λ m, by { ext, simp [function.comp] }, + map_add' := dom_dom_congr_add σ, + map_smul' := dom_dom_congr_smul σ } + +@[simp] lemma dom_dom_lcongr_refl : + (dom_dom_lcongr S (equiv.refl ι) : alternating_map R M N ι ≃ₗ[S] alternating_map R M N ι) = + linear_equiv.refl _ _ := +linear_equiv.ext dom_dom_congr_refl + +@[simp] lemma dom_dom_lcongr_to_add_equiv (σ : ι ≃ ι') : + (dom_dom_lcongr S σ : alternating_map R M N ι ≃ₗ[S] alternating_map R M N ι').to_add_equiv + = dom_dom_congr_equiv σ := rfl + +end dom_dom_lcongr + /-- The results of applying `dom_dom_congr` to two maps are equal if and only if those maps are. -/ @[simp] lemma dom_dom_congr_eq_iff (σ : ι ≃ ι') (f g : alternating_map R M N ι) : f.dom_dom_congr σ = g.dom_dom_congr σ ↔ f = g := diff --git a/src/linear_algebra/determinant.lean b/src/linear_algebra/determinant.lean index 80d805eccb482..c005163173ff9 100644 --- a/src/linear_algebra/determinant.lean +++ b/src/linear_algebra/determinant.lean @@ -559,6 +559,11 @@ lemma basis.det_reindex {ι' : Type*} [fintype ι'] [decidable_eq ι'] (b.reindex e).det v = b.det (v ∘ e) := by rw [basis.det_apply, basis.to_matrix_reindex', det_reindex_alg_equiv, basis.det_apply] +lemma basis.det_reindex' {ι' : Type*} [fintype ι'] [decidable_eq ι'] + (b : basis ι R M) (e : ι ≃ ι') : + (b.reindex e).det = b.det.dom_dom_congr e := +alternating_map.ext $ λ _, basis.det_reindex _ _ _ + lemma basis.det_reindex_symm {ι' : Type*} [fintype ι'] [decidable_eq ι'] (b : basis ι R M) (v : ι → M) (e : ι' ≃ ι) : (b.reindex e.symm).det (v ∘ e) = b.det v := diff --git a/src/linear_algebra/orientation.lean b/src/linear_algebra/orientation.lean index a17ec1101eba4..107bd5f34b839 100644 --- a/src/linear_algebra/orientation.lean +++ b/src/linear_algebra/orientation.lean @@ -43,7 +43,7 @@ section ordered_comm_semiring variables (R : Type*) [strict_ordered_comm_semiring R] variables (M : Type*) [add_comm_monoid M] [module R M] variables {N : Type*} [add_comm_monoid N] [module R N] -variables (ι : Type*) +variables (ι ι' : Type*) /-- An orientation of a module, intended to be used when `ι` is a `fintype` with the same cardinality as a basis. -/ @@ -73,6 +73,27 @@ by rw [orientation.map, alternating_map.dom_lcongr_refl, module.ray.map_refl] @[simp] lemma orientation.map_symm (e : M ≃ₗ[R] N) : (orientation.map ι e).symm = orientation.map ι e.symm := rfl +section reindex +variables (R M) {ι ι'} + +/-- An equivalence between indices implies an equivalence between orientations. -/ +def orientation.reindex (e : ι ≃ ι') : orientation R M ι ≃ orientation R M ι' := +module.ray.map $ alternating_map.dom_dom_lcongr R e + +@[simp] lemma orientation.reindex_apply (e : ι ≃ ι') (v : alternating_map R M R ι) + (hv : v ≠ 0) : + orientation.reindex R M e (ray_of_ne_zero _ v hv) = ray_of_ne_zero _ (v.dom_dom_congr e) + (mt (v.dom_dom_congr_eq_zero_iff e).mp hv) := rfl + +@[simp] lemma orientation.reindex_refl : + (orientation.reindex R M $ equiv.refl ι) = equiv.refl _ := +by rw [orientation.reindex, alternating_map.dom_dom_lcongr_refl, module.ray.map_refl] + +@[simp] lemma orientation.reindex_symm (e : ι ≃ ι') : + (orientation.reindex R M e).symm = orientation.reindex R M e.symm := rfl + +end reindex + /-- A module is canonically oriented with respect to an empty index type. -/ @[priority 100] instance is_empty.oriented [nontrivial R] [is_empty ι] : module.oriented R M ι := @@ -107,9 +128,14 @@ variables {M N : Type*} [add_comm_group M] [add_comm_group N] [module R M] [modu orientation.map ι f (-x) = - orientation.map ι f x := module.ray.map_neg _ x +@[simp] protected lemma orientation.reindex_neg {ι ι' : Type*} (e : ι ≃ ι') + (x : orientation R M ι) : + orientation.reindex R M e (-x) = - orientation.reindex R M e x := +module.ray.map_neg _ x + namespace basis -variables {ι : Type*} +variables {ι ι' : Type*} /-- The value of `orientation.map` when the index type has the cardinality of a basis, in terms of `f.det`. -/ @@ -125,7 +151,7 @@ begin basis.det_self, mul_one, smul_eq_mul, mul_comm, mul_smul, linear_equiv.coe_inv_det], end -variables [fintype ι] [decidable_eq ι] +variables [fintype ι] [decidable_eq ι] [fintype ι'] [decidable_eq ι'] /-- The orientation given by a basis. -/ protected def orientation [nontrivial R] (e : basis ι R M) : orientation R M ι := @@ -135,6 +161,10 @@ lemma orientation_map [nontrivial R] (e : basis ι R M) (f : M ≃ₗ[R] N) : (e.map f).orientation = orientation.map ι f e.orientation := by simp_rw [basis.orientation, orientation.map_apply, basis.det_map'] +lemma orientation_reindex [nontrivial R] (e : basis ι R M) + (eι : ι ≃ ι') : (e.reindex eι).orientation = orientation.reindex R M eι e.orientation := +by simp_rw [basis.orientation, orientation.reindex_apply, basis.det_reindex'] + /-- The orientation given by a basis derived using `units_smul`, in terms of the product of those units. -/ lemma orientation_units_smul [nontrivial R] (e : basis ι R M) (w : ι → units R) : From 3b52265189f3fb43aa631edffce5d060fafaf82f Mon Sep 17 00:00:00 2001 From: AlexKontorovich <25316162+hrmacbeth@users.noreply.github.com> Date: Fri, 28 Jul 2023 12:48:03 +0000 Subject: [PATCH 46/61] feat(measure_theory/measure/haar_quotient): the Unfolding Trick (#18863) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit We prove the "unfolding trick": Given a subgroup `Γ` of a group `G`, the integral of a function `f` on `G` times the lift to `G` of a function `g` on the coset space `G ⧸ Γ` with respect to a right-invariant measure `μ` on `G`, is equal to the integral over the coset space of the automorphization of `f` times `g`. We also prove the following simplified version: Given a subgroup `Γ` of a group `G`, the integral of a function `f` on `G` with respect to a right-invariant measure `μ` is equal to the integral over the coset space `G ⧸ Γ` of the automorphization of `f`. A question: is it possible to deduce `ae_strongly_measurable (quotient_group.automorphize f) μ_𝓕` from `ae_strongly_measurable f μ` (as opposed to assuming it as a hypothesis in the main theorem)? It seems quite plausible... Co-authored-by: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> Co-authored-by: Alex Kontorovich <58564076+AlexKontorovich@users.noreply.github.com> Co-authored-by: AlexKontorovich <58564076+AlexKontorovich@users.noreply.github.com> --- src/group_theory/group_action/group.lean | 11 ++ .../function/strongly_measurable/basic.lean | 7 + .../group/fundamental_domain.lean | 9 + src/measure_theory/measure/haar/quotient.lean | 183 +++++++++++++++++- src/topology/algebra/infinite_sum/basic.lean | 141 ++++++++++++++ 5 files changed, 343 insertions(+), 8 deletions(-) diff --git a/src/group_theory/group_action/group.lean b/src/group_theory/group_action/group.lean index 02971812f5784..be6086db460c1 100644 --- a/src/group_theory/group_action/group.lean +++ b/src/group_theory/group_action/group.lean @@ -191,6 +191,17 @@ def distrib_mul_action.to_add_equiv (x : α) : β ≃+ β := { .. distrib_mul_action.to_add_monoid_hom β x, .. mul_action.to_perm_hom α β x } +/-- Each non-zero element of a `group_with_zero` defines an additive monoid isomorphism of an +`add_monoid` on which it acts distributively. + +This is a stronger version of `distrib_mul_action.to_add_monoid_hom`. -/ +def distrib_mul_action.to_add_equiv₀ {α : Type*} (β : Type*) [group_with_zero α] [add_monoid β] + [distrib_mul_action α β] (x : α) (hx : x ≠ 0) : β ≃+ β := +{ inv_fun := λ b, x⁻¹ • b, + left_inv := inv_smul_smul₀ hx, + right_inv := smul_inv_smul₀ hx, + .. distrib_mul_action.to_add_monoid_hom β x, } + variables (α β) /-- Each element of the group defines an additive monoid isomorphism. diff --git a/src/measure_theory/function/strongly_measurable/basic.lean b/src/measure_theory/function/strongly_measurable/basic.lean index c7220c3cb2d68..a2efb0496676c 100644 --- a/src/measure_theory/function/strongly_measurable/basic.lean +++ b/src/measure_theory/function/strongly_measurable/basic.lean @@ -1740,6 +1740,13 @@ end end ae_strongly_measurable +lemma ae_strongly_measurable_of_absolutely_continuous {α β : Type*} [measurable_space α] + [topological_space β] {μ ν : measure α} (h : ν ≪ μ) (g : α → β) + (hμ : ae_strongly_measurable g μ) : ae_strongly_measurable g ν := +begin + obtain ⟨g₁, hg₁, hg₁'⟩ := hμ, + refine ⟨g₁, hg₁, h.ae_eq hg₁'⟩, +end /-! ## Almost everywhere finitely strongly measurable functions -/ diff --git a/src/measure_theory/group/fundamental_domain.lean b/src/measure_theory/group/fundamental_domain.lean index ec47f9ce9daf4..c97dba8abe1bc 100644 --- a/src/measure_theory/group/fundamental_domain.lean +++ b/src/measure_theory/group/fundamental_domain.lean @@ -212,6 +212,11 @@ calc ∫⁻ x, f x ∂μ = ∑' g : G, ∫⁻ x in g • s, f x ∂μ : h.linteg tsum_congr $ λ g, ((measure_preserving_smul g⁻¹ μ).set_lintegral_comp_emb (measurable_embedding_const_smul _) _ _).symm + +@[to_additive] lemma lintegral_eq_tsum'' (h : is_fundamental_domain G s μ) (f : α → ℝ≥0∞) : + ∫⁻ x, f x ∂μ = ∑' g : G, ∫⁻ x in s, f (g • x) ∂μ := +(lintegral_eq_tsum' h f).trans ((equiv.inv G).tsum_eq (λ g, ∫⁻ (x : α) in s, f (g • x) ∂μ)) + @[to_additive] lemma set_lintegral_eq_tsum (h : is_fundamental_domain G s μ) (f : α → ℝ≥0∞) (t : set α) : ∫⁻ x in t, f x ∂μ = ∑' g : G, ∫⁻ x in t ∩ g • s, f x ∂μ := calc ∫⁻ x in t, f x ∂μ = ∑' g : G, ∫⁻ x in g • s, f x ∂(μ.restrict t) : @@ -357,6 +362,10 @@ calc ∫ x, f x ∂μ = ∑' g : G, ∫ x in g • s, f x ∂μ : h.integral_eq_ tsum_congr $ λ g, (measure_preserving_smul g⁻¹ μ).set_integral_image_emb (measurable_embedding_const_smul _) _ _ +@[to_additive] lemma integral_eq_tsum'' (h : is_fundamental_domain G s μ) + (f : α → E) (hf : integrable f μ) : ∫ x, f x ∂μ = ∑' g : G, ∫ x in s, f (g • x) ∂μ := +(integral_eq_tsum' h f hf).trans ((equiv.inv G).tsum_eq (λ g, ∫ (x : α) in s, f (g • x) ∂μ)) + @[to_additive] lemma set_integral_eq_tsum (h : is_fundamental_domain G s μ) {f : α → E} {t : set α} (hf : integrable_on f t μ) : ∫ x in t, f x ∂μ = ∑' g : G, ∫ x in t ∩ g • s, f x ∂μ := diff --git a/src/measure_theory/measure/haar/quotient.lean b/src/measure_theory/measure/haar/quotient.lean index 03f20400f0398..5bd1067dbc4e5 100644 --- a/src/measure_theory/measure/haar/quotient.lean +++ b/src/measure_theory/measure/haar/quotient.lean @@ -33,8 +33,10 @@ Note that a group `G` with Haar measure that is both left and right invariant is **unimodular**. -/ -open set measure_theory topological_space measure_theory.measure -open_locale pointwise nnreal +noncomputable theory + +open set measure_theory topological_space measure_theory.measure quotient_group +open_locale pointwise measure_theory topology big_operators nnreal ennreal variables {G : Type*} [group G] [measurable_space G] [topological_space G] [topological_group G] [borel_space G] @@ -116,8 +118,6 @@ lemma measure_theory.is_fundamental_domain.is_mul_left_invariant_map [subgroup.n { exact hA, }, end } -variables [t2_space (G ⧸ Γ)] [second_countable_topology (G ⧸ Γ)] (K : positive_compacts (G ⧸ Γ)) - /-- Given a normal subgroup `Γ` of a topological group `G` with Haar measure `μ`, which is also right-invariant, and a finite volume fundamental domain `𝓕`, the pushforward to the quotient group `G ⧸ Γ` of the restriction of `μ` to `𝓕` is a multiple of Haar measure on `G ⧸ Γ`. -/ @@ -125,7 +125,8 @@ variables [t2_space (G ⧸ Γ)] [second_countable_topology (G ⧸ Γ)] (K : posi `μ`, which is also right-invariant, and a finite volume fundamental domain `𝓕`, the pushforward to the quotient group `G ⧸ Γ` of the restriction of `μ` to `𝓕` is a multiple of Haar measure on `G ⧸ Γ`."] -lemma measure_theory.is_fundamental_domain.map_restrict_quotient [subgroup.normal Γ] +lemma measure_theory.is_fundamental_domain.map_restrict_quotient [t2_space (G ⧸ Γ)] + [second_countable_topology (G ⧸ Γ)] (K : positive_compacts (G ⧸ Γ)) [subgroup.normal Γ] [measure_theory.measure.is_haar_measure μ] [μ.is_mul_right_invariant] (h𝓕_finite : μ 𝓕 < ⊤) : measure.map (quotient_group.mk' Γ) (μ.restrict 𝓕) = (μ (𝓕 ∩ (quotient_group.mk' Γ) ⁻¹' K)) • (measure_theory.measure.haar_measure K) := @@ -151,12 +152,178 @@ end topological group `G` with Haar measure `μ`, which is also right-invariant, and a finite volume fundamental domain `𝓕`, the quotient map to `G ⧸ Γ` is measure-preserving between appropriate multiples of Haar measure on `G` and `G ⧸ Γ`."] -lemma measure_preserving_quotient_group.mk' [subgroup.normal Γ] - [measure_theory.measure.is_haar_measure μ] [μ.is_mul_right_invariant] - (h𝓕_finite : μ 𝓕 < ⊤) (c : ℝ≥0) (h : μ (𝓕 ∩ (quotient_group.mk' Γ) ⁻¹' K) = c) : +lemma measure_preserving_quotient_group.mk' [t2_space (G ⧸ Γ)] [second_countable_topology (G ⧸ Γ)] + (K : positive_compacts (G ⧸ Γ)) [subgroup.normal Γ] [measure_theory.measure.is_haar_measure μ] + [μ.is_mul_right_invariant] (h𝓕_finite : μ 𝓕 < ⊤) (c : ℝ≥0) + (h : μ (𝓕 ∩ (quotient_group.mk' Γ) ⁻¹' K) = c) : measure_preserving (quotient_group.mk' Γ) (μ.restrict 𝓕) (c • (measure_theory.measure.haar_measure K)) := { measurable := continuous_quotient_mk.measurable, map_eq := by rw [h𝓕.map_restrict_quotient K h𝓕_finite, h]; refl } + +section + +local notation `μ_𝓕` := measure.map (@quotient_group.mk G _ Γ) (μ.restrict 𝓕) + +/-- The `ess_sup` of a function `g` on the quotient space `G ⧸ Γ` with respect to the pushforward + of the restriction, `μ_𝓕`, of a right-invariant measure `μ` to a fundamental domain `𝓕`, is the + same as the `ess_sup` of `g`'s lift to the universal cover `G` with respect to `μ`. -/ +@[to_additive "The `ess_sup` of a function `g` on the additive quotient space `G ⧸ Γ` with respect + to the pushforward of the restriction, `μ_𝓕`, of a right-invariant measure `μ` to a fundamental + domain `𝓕`, is the same as the `ess_sup` of `g`'s lift to the universal cover `G` with respect + to `μ`."] +lemma ess_sup_comp_quotient_group_mk [μ.is_mul_right_invariant] {g : G ⧸ Γ → ℝ≥0∞} + (g_ae_measurable : ae_measurable g μ_𝓕) : + ess_sup g μ_𝓕 = ess_sup (λ (x : G), g x) μ := +begin + have hπ : measurable (quotient_group.mk : G → G ⧸ Γ) := continuous_quotient_mk.measurable, + rw ess_sup_map_measure g_ae_measurable hπ.ae_measurable, + refine h𝓕.ess_sup_measure_restrict _, + rintros ⟨γ, hγ⟩ x, + dsimp, + congr' 1, + exact quotient_group.mk_mul_of_mem x hγ, +end + +/-- Given a quotient space `G ⧸ Γ` where `Γ` is `countable`, and the restriction, + `μ_𝓕`, of a right-invariant measure `μ` on `G` to a fundamental domain `𝓕`, a set + in the quotient which has `μ_𝓕`-measure zero, also has measure zero under the + folding of `μ` under the quotient. Note that, if `Γ` is infinite, then the folded map + will take the value `∞` on any open set in the quotient! -/ +@[to_additive "Given an additive quotient space `G ⧸ Γ` where `Γ` is `countable`, and the + restriction, `μ_𝓕`, of a right-invariant measure `μ` on `G` to a fundamental domain `𝓕`, a set + in the quotient which has `μ_𝓕`-measure zero, also has measure zero under the + folding of `μ` under the quotient. Note that, if `Γ` is infinite, then the folded map + will take the value `∞` on any open set in the quotient!"] +lemma _root_.measure_theory.is_fundamental_domain.absolutely_continuous_map + [μ.is_mul_right_invariant] : + map (quotient_group.mk : G → G ⧸ Γ) μ ≪ map (quotient_group.mk : G → G ⧸ Γ) (μ.restrict 𝓕) := +begin + set π : G → G ⧸ Γ := quotient_group.mk, + have meas_π : measurable π := continuous_quotient_mk.measurable, + apply absolutely_continuous.mk, + intros s s_meas hs, + rw map_apply meas_π s_meas at hs ⊢, + rw measure.restrict_apply at hs, + apply h𝓕.measure_zero_of_invariant _ _ hs, + { intros γ, + ext g, + rw [set.mem_smul_set_iff_inv_smul_mem, mem_preimage, mem_preimage], + congrm _ ∈ s, + convert quotient_group.mk_mul_of_mem g (γ⁻¹).2, }, + exact measurable_set_preimage meas_π s_meas, +end + +local attribute [-instance] quotient.measurable_space + +/-- This is a simple version of the **Unfolding Trick**: Given a subgroup `Γ` of a group `G`, the + integral of a function `f` on `G` with respect to a right-invariant measure `μ` is equal to the + integral over the quotient `G ⧸ Γ` of the automorphization of `f`. -/ +@[to_additive "This is a simple version of the **Unfolding Trick**: Given a subgroup `Γ` of an + additive group `G`, the integral of a function `f` on `G` with respect to a right-invariant + measure `μ` is equal to the integral over the quotient `G ⧸ Γ` of the automorphization of `f`."] +lemma quotient_group.integral_eq_integral_automorphize {E : Type*} [normed_add_comm_group E] + [complete_space E] [normed_space ℝ E] [μ.is_mul_right_invariant] {f : G → E} + (hf₁ : integrable f μ) (hf₂ : ae_strongly_measurable (automorphize f) μ_𝓕) : + ∫ x : G, f x ∂μ = ∫ x : G ⧸ Γ, automorphize f x ∂μ_𝓕 := +calc ∫ x : G, f x ∂μ = ∑' γ : Γ.opposite, ∫ x in 𝓕, f (γ • x) ∂μ : h𝓕.integral_eq_tsum'' f hf₁ +... = ∫ x in 𝓕, ∑' γ : Γ.opposite, f (γ • x) ∂μ : + begin + rw integral_tsum, + { exact λ i, (hf₁.1.comp_quasi_measure_preserving + (measure_preserving_smul i μ).quasi_measure_preserving).restrict, }, + { rw ← h𝓕.lintegral_eq_tsum'' (λ x, ‖f x‖₊), + exact ne_of_lt hf₁.2, }, + end +... = ∫ x : G ⧸ Γ, automorphize f x ∂μ_𝓕 : + (integral_map continuous_quotient_mk.ae_measurable hf₂).symm + +/-- This is the **Unfolding Trick**: Given a subgroup `Γ` of a group `G`, the integral of a + function `f` on `G` times the lift to `G` of a function `g` on the quotient `G ⧸ Γ` with respect + to a right-invariant measure `μ` on `G`, is equal to the integral over the quotient of the + automorphization of `f` times `g`. -/ +lemma quotient_group.integral_mul_eq_integral_automorphize_mul {K : Type*} [normed_field K] + [complete_space K] [normed_space ℝ K] [μ.is_mul_right_invariant] {f : G → K} + (f_ℒ_1 : integrable f μ) {g : G ⧸ Γ → K} (hg : ae_strongly_measurable g μ_𝓕) + (g_ℒ_infinity : ess_sup (λ x, ↑‖g x‖₊) μ_𝓕 ≠ ∞) + (F_ae_measurable : ae_strongly_measurable (quotient_group.automorphize f) μ_𝓕) : + ∫ x : G, g (x : G ⧸ Γ) * (f x) ∂μ = ∫ x : G ⧸ Γ, g x * (quotient_group.automorphize f x) ∂μ_𝓕 := +begin + let π : G → G ⧸ Γ := quotient_group.mk, + have H₀ : quotient_group.automorphize ((g ∘ π) * f) = g * (quotient_group.automorphize f) := + quotient_group.automorphize_smul_left f g, + calc ∫ (x : G), g (π x) * f x ∂μ = + ∫ (x : G ⧸ Γ), quotient_group.automorphize ((g ∘ π) * f) x ∂μ_𝓕 : _ + ... = ∫ (x : G ⧸ Γ), g x * (quotient_group.automorphize f x) ∂μ_𝓕 : by simp [H₀], + have meas_π : measurable π := continuous_quotient_mk.measurable, + have H₁ : integrable ((g ∘ π) * f) μ, + { have : ae_strongly_measurable (λ x : G, g (x : G ⧸ Γ)) μ, + { refine (ae_strongly_measurable_of_absolutely_continuous _ _ hg).comp_measurable meas_π, + exact h𝓕.absolutely_continuous_map }, + refine integrable.ess_sup_smul f_ℒ_1 this _, + { have hg' : ae_strongly_measurable (λ x, ↑‖g x‖₊) μ_𝓕 := + (ennreal.continuous_coe.comp continuous_nnnorm).comp_ae_strongly_measurable hg, + rw [← ess_sup_comp_quotient_group_mk h𝓕 hg'.ae_measurable], + exact g_ℒ_infinity } }, + have H₂ : ae_strongly_measurable (quotient_group.automorphize ((g ∘ π) * f)) μ_𝓕, + { simp_rw [H₀], + exact hg.mul F_ae_measurable }, + apply quotient_group.integral_eq_integral_automorphize h𝓕 H₁ H₂, +end + +end + +section + +variables {G' : Type*} [add_group G'] [measurable_space G'] [topological_space G'] + [topological_add_group G'] [borel_space G'] + {μ' : measure G'} + {Γ' : add_subgroup G'} + [countable Γ'] [measurable_space (G' ⧸ Γ')] [borel_space (G' ⧸ Γ')] + {𝓕' : set G'} + +local notation `μ_𝓕` := measure.map (@quotient_add_group.mk G' _ Γ') (μ'.restrict 𝓕') + +/-- This is the **Unfolding Trick**: Given an additive subgroup `Γ'` of an additive group `G'`, the + integral of a function `f` on `G'` times the lift to `G'` of a function `g` on the quotient + `G' ⧸ Γ'` with respect to a right-invariant measure `μ` on `G'`, is equal to the integral over + the quotient of the automorphization of `f` times `g`. -/ +lemma quotient_add_group.integral_mul_eq_integral_automorphize_mul +{K : Type*} [normed_field K] + [complete_space K] [normed_space ℝ K] [μ'.is_add_right_invariant] {f : G' → K} + (f_ℒ_1 : integrable f μ') {g : G' ⧸ Γ' → K} (hg : ae_strongly_measurable g μ_𝓕) + (g_ℒ_infinity : ess_sup (λ x, ↑‖g x‖₊) μ_𝓕 ≠ ∞) + (F_ae_measurable : ae_strongly_measurable (quotient_add_group.automorphize f) μ_𝓕) + (h𝓕 : is_add_fundamental_domain Γ'.opposite 𝓕' μ') : + ∫ x : G', g (x : G' ⧸ Γ') * (f x) ∂μ' + = ∫ x : G' ⧸ Γ', g x * (quotient_add_group.automorphize f x) ∂μ_𝓕 := +begin + let π : G' → G' ⧸ Γ' := quotient_add_group.mk, + have H₀ : quotient_add_group.automorphize ((g ∘ π) * f) + = g * (quotient_add_group.automorphize f) := + quotient_add_group.automorphize_smul_left f g, + calc ∫ (x : G'), g (π x) * f x ∂μ' = + ∫ (x : G' ⧸ Γ'), quotient_add_group.automorphize ((g ∘ π) * f) x ∂μ_𝓕 : _ + ... = ∫ (x : G' ⧸ Γ'), g x * (quotient_add_group.automorphize f x) ∂μ_𝓕 : by simp [H₀], + have meas_π : measurable π := continuous_quotient_mk.measurable, + have H₁ : integrable ((g ∘ π) * f) μ', + { have : ae_strongly_measurable (λ x : G', g (x : G' ⧸ Γ')) μ', + { refine (ae_strongly_measurable_of_absolutely_continuous _ _ hg).comp_measurable meas_π, + exact h𝓕.absolutely_continuous_map }, + refine integrable.ess_sup_smul f_ℒ_1 this _, + { have hg' : ae_strongly_measurable (λ x, ↑‖g x‖₊) μ_𝓕 := + (ennreal.continuous_coe.comp continuous_nnnorm).comp_ae_strongly_measurable hg, + rw [← ess_sup_comp_quotient_add_group_mk h𝓕 hg'.ae_measurable], + exact g_ℒ_infinity } }, + have H₂ : ae_strongly_measurable (quotient_add_group.automorphize ((g ∘ π) * f)) μ_𝓕, + { simp_rw [H₀], + exact hg.mul F_ae_measurable }, + apply quotient_add_group.integral_eq_integral_automorphize h𝓕 H₁ H₂, +end + +end + +attribute [to_additive quotient_group.integral_mul_eq_integral_automorphize_mul] + quotient_add_group.integral_mul_eq_integral_automorphize_mul diff --git a/src/topology/algebra/infinite_sum/basic.lean b/src/topology/algebra/infinite_sum/basic.lean index 9d0486d2144d3..69334b3838d57 100644 --- a/src/topology/algebra/infinite_sum/basic.lean +++ b/src/topology/algebra/infinite_sum/basic.lean @@ -1147,9 +1147,46 @@ hf.map (distrib_mul_action.to_add_monoid_hom α _) $ continuous_const_smul _ lemma summable.const_smul (b : γ) (hf : summable f) : summable (λ i, b • f i) := (hf.has_sum.const_smul _).summable +/-- Infinite sums commute with scalar multiplication. Version for scalars living in a `monoid`, but + requiring a summability hypothesis. -/ lemma tsum_const_smul [t2_space α] (b : γ) (hf : summable f) : ∑' i, b • f i = b • ∑' i, f i := (hf.has_sum.const_smul _).tsum_eq +/-- Infinite sums commute with scalar multiplication. Version for scalars living in a `group`, but + not requiring any summability hypothesis. -/ +lemma tsum_const_smul' {γ : Type*} [group γ] [distrib_mul_action γ α] + [has_continuous_const_smul γ α] [t2_space α] (g : γ) : + ∑' (i : β), g • f i = g • ∑' (i : β), f i := +begin + by_cases hf : summable f, + { exact tsum_const_smul _ hf, }, + rw tsum_eq_zero_of_not_summable hf, + simp only [smul_zero], + let mul_g : α ≃+ α := distrib_mul_action.to_add_equiv α g, + apply tsum_eq_zero_of_not_summable, + change ¬ summable (mul_g ∘ f), + rwa summable.map_iff_of_equiv mul_g; apply continuous_const_smul, +end + +/-- Infinite sums commute with scalar multiplication. Version for scalars living in a + `division_ring`; no summability hypothesis. This could be made to work for a + `[group_with_zero γ]` if there was such a thing as `distrib_mul_action_with_zero`. -/ +lemma tsum_const_smul'' {γ : Type*} [division_ring γ] [module γ α] [has_continuous_const_smul γ α] + [t2_space α] (g : γ) : + ∑' (i : β), g • f i = g • ∑' (i : β), f i := +begin + by_cases hf : summable f, + { exact tsum_const_smul _ hf, }, + rw tsum_eq_zero_of_not_summable hf, + simp only [smul_zero], + by_cases hg : g = 0, + { simp [hg], }, + let mul_g : α ≃+ α := distrib_mul_action.to_add_equiv₀ α g hg, + apply tsum_eq_zero_of_not_summable, + change ¬ summable (mul_g ∘ f), + rwa summable.map_iff_of_equiv mul_g; apply continuous_const_smul, +end + end const_smul /-! ### Product and pi types -/ @@ -1256,3 +1293,107 @@ begin end end has_continuous_star + +section automorphize + +variables {M : Type*} [topological_space M] [add_comm_monoid M] [t2_space M] {R : Type*} + [division_ring R] [module R M] [has_continuous_const_smul R M] + +/-- Given a group `α` acting on a type `β`, and a function `f : β → M`, we "automorphize" `f` to a + function `β ⧸ α → M` by summing over `α` orbits, `b ↦ ∑' (a : α), f(a • b)`. -/ +@[to_additive "Given an additive group `α` acting on a type `β`, and a function `f : β → M`, + we automorphize `f` to a function `β ⧸ α → M` by summing over `α` orbits, + `b ↦ ∑' (a : α), f(a • b)`."] +def mul_action.automorphize [group α] [mul_action α β] (f : β → M) : + quotient (mul_action.orbit_rel α β) → M := +@quotient.lift _ _ (mul_action.orbit_rel α β) (λ b, ∑' (a : α), f(a • b)) +begin + rintros b₁ b₂ ⟨a, (rfl : a • b₂ = b₁)⟩, + simpa [mul_smul] using (equiv.mul_right a).tsum_eq (λ a', f (a' • b₂)), +end + +/-- Automorphization of a function into an `R`-`module` distributes, that is, commutes with the + `R`-scalar multiplication. -/ +lemma mul_action.automorphize_smul_left [group α] [mul_action α β] (f : β → M) + (g : quotient (mul_action.orbit_rel α β) → R) : + mul_action.automorphize ((g ∘ quotient.mk') • f) + = g • (mul_action.automorphize f : quotient (mul_action.orbit_rel α β) → M) := +begin + ext x, + apply quotient.induction_on' x, + intro b, + simp only [mul_action.automorphize, pi.smul_apply', function.comp_app], + set π : β → quotient (mul_action.orbit_rel α β) := quotient.mk', + have H₁ : ∀ a : α, π (a • b) = π b, + { intro a, + rw quotient.eq_rel, + fconstructor, + exact a, + simp, }, + change ∑' a : α, g (π (a • b)) • f (a • b) = g (π b) • ∑' a : α, f (a • b), + simp_rw [H₁], + exact tsum_const_smul'' _, +end + +/-- Automorphization of a function into an `R`-`module` distributes, that is, commutes with the + `R`-scalar multiplication. -/ +lemma add_action.automorphize_smul_left [add_group α] [add_action α β] (f : β → M) + (g : quotient (add_action.orbit_rel α β) → R) : + add_action.automorphize ((g ∘ quotient.mk') • f) + = g • (add_action.automorphize f : quotient (add_action.orbit_rel α β) → M) := +begin + ext x, + apply quotient.induction_on' x, + intro b, + simp only [add_action.automorphize, pi.smul_apply', function.comp_app], + set π : β → quotient (add_action.orbit_rel α β) := quotient.mk', + have H₁ : ∀ a : α, π (a +ᵥ b) = π b, + { intro a, + rw quotient.eq_rel, + fconstructor, + exact a, + simp, }, + change ∑' a : α, g (π (a +ᵥ b)) • f (a +ᵥ b) = g (π b) • ∑' a : α, f (a +ᵥ b), + simp_rw [H₁], + exact tsum_const_smul'' _, +end + +attribute [to_additive mul_action.automorphize_smul_left] add_action.automorphize_smul_left + +section + +variables {G : Type*} [group G] {Γ : subgroup G} + +/-- Given a subgroup `Γ` of a group `G`, and a function `f : G → M`, we "automorphize" `f` to a + function `G ⧸ Γ → M` by summing over `Γ` orbits, `g ↦ ∑' (γ : Γ), f(γ • g)`. -/ +@[to_additive "Given a subgroup `Γ` of an additive group `G`, and a function `f : G → M`, we + automorphize `f` to a function `G ⧸ Γ → M` by summing over `Γ` orbits, + `g ↦ ∑' (γ : Γ), f(γ • g)`."] +def quotient_group.automorphize (f : G → M) : G ⧸ Γ → M := mul_action.automorphize f + +/-- Automorphization of a function into an `R`-`module` distributes, that is, commutes with the + `R`-scalar multiplication. -/ +lemma quotient_group.automorphize_smul_left (f : G → M) (g : G ⧸ Γ → R) : + quotient_group.automorphize ((g ∘ quotient.mk') • f) + = g • (quotient_group.automorphize f : G ⧸ Γ → M) := +mul_action.automorphize_smul_left f g + +end + +section + +variables {G : Type*} [add_group G] {Γ : add_subgroup G} + +/-- Automorphization of a function into an `R`-`module` distributes, that is, commutes with the `R` + -scalar multiplication. -/ +lemma quotient_add_group.automorphize_smul_left (f : G → M) (g : G ⧸ Γ → R) : + quotient_add_group.automorphize ((g ∘ quotient.mk') • f) + = g • (quotient_add_group.automorphize f : G ⧸ Γ → M) := +add_action.automorphize_smul_left f g + +end + +attribute [to_additive quotient_group.automorphize_smul_left] + quotient_add_group.automorphize_smul_left + +end automorphize From 63721b2c3eba6c325ecf8ae8cca27155a4f6306f Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Sun, 30 Jul 2023 10:18:54 +0000 Subject: [PATCH 47/61] feat(algebraic_topology/dold_kan): The Dold-Kan equivalence for pseudoabelian categories (#17925) --- .../dold_kan/equivalence_pseudoabelian.lean | 123 ++++++++++++++++++ src/category_theory/functor/category.lean | 3 + 2 files changed, 126 insertions(+) create mode 100644 src/algebraic_topology/dold_kan/equivalence_pseudoabelian.lean diff --git a/src/algebraic_topology/dold_kan/equivalence_pseudoabelian.lean b/src/algebraic_topology/dold_kan/equivalence_pseudoabelian.lean new file mode 100644 index 0000000000000..dd766e8a5e41f --- /dev/null +++ b/src/algebraic_topology/dold_kan/equivalence_pseudoabelian.lean @@ -0,0 +1,123 @@ +/- +Copyright (c) 2022 Joël Riou. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joël Riou +-/ + +import algebraic_topology.dold_kan.equivalence_additive +import algebraic_topology.dold_kan.compatibility +import category_theory.idempotents.simplicial_object + +/-! + +# The Dold-Kan correspondence for pseudoabelian categories + +In this file, for any idempotent complete additive category `C`, +the Dold-Kan equivalence +`idempotents.dold_kan.equivalence C : simplicial_object C ≌ chain_complex C ℕ` +is obtained. It is deduced from the equivalence +`preadditive.dold_kan.equivalence` between the respective idempotent +completions of these categories using the fact that when `C` is idempotent complete, +then both `simplicial_object C` and `chain_complex C ℕ` are idempotent complete. + +The construction of `idempotents.dold_kan.equivalence` uses the tools +introduced in the file `compatibility.lean`. Doing so, the functor +`idempotents.dold_kan.N` of the equivalence is +the composition of `N₁ : simplicial_object C ⥤ karoubi (chain_complex C ℕ)` +(defined in `functor_n.lean`) and the inverse of the equivalence +`chain_complex C ℕ ≌ karoubi (chain_complex C ℕ)`. The functor +`idempotents.dold_kan.Γ` of the equivalence is by definition the functor +`Γ₀` introduced in `functor_gamma.lean`. + +-/ + +noncomputable theory + +open category_theory category_theory.category category_theory.limits category_theory.idempotents + +variables {C : Type*} [category C] [preadditive C] [is_idempotent_complete C] + [has_finite_coproducts C] + +namespace category_theory + +namespace idempotents + +namespace dold_kan + +open algebraic_topology.dold_kan + +/-- The functor `N` for the equivalence is obtained by composing +`N' : simplicial_object C ⥤ karoubi (chain_complex C ℕ)` and the inverse +of the equivalence `chain_complex C ℕ ≌ karoubi (chain_complex C ℕ)`. -/ +@[simps, nolint unused_arguments] +def N : simplicial_object C ⥤ chain_complex C ℕ := +N₁ ⋙ (to_karoubi_equivalence _).inverse + +/-- The functor `Γ` for the equivalence is `Γ'`. -/ +@[simps, nolint unused_arguments] +def Γ : chain_complex C ℕ ⥤ simplicial_object C := Γ₀ + +lemma hN₁ : (to_karoubi_equivalence (simplicial_object C)).functor ⋙ + preadditive.dold_kan.equivalence.functor = N₁ := +functor.congr_obj (functor_extension₁_comp_whiskering_left_to_karoubi _ _) N₁ + +lemma hΓ₀ : (to_karoubi_equivalence (chain_complex C ℕ)).functor ⋙ + preadditive.dold_kan.equivalence.inverse = Γ ⋙ (to_karoubi_equivalence _).functor := +functor.congr_obj (functor_extension₂_comp_whiskering_left_to_karoubi _ _) Γ₀ + +/-- The Dold-Kan equivalence for pseudoabelian categories given +by the functors `N` and `Γ`. It is obtained by applying the results in +`compatibility.lean` to the equivalence `preadditive.dold_kan.equivalence`. -/ +def equivalence : simplicial_object C ≌ chain_complex C ℕ := +compatibility.equivalence (eq_to_iso hN₁) (eq_to_iso hΓ₀) + +lemma equivalence_functor : (equivalence : simplicial_object C ≌ _).functor = N := rfl +lemma equivalence_inverse : (equivalence : simplicial_object C ≌ _).inverse = Γ := rfl + +/-- The natural isomorphism `NΓ' satisfies the compatibility that is needed +for the construction of our counit isomorphism `η` -/ +lemma hη : compatibility.τ₀ = + compatibility.τ₁ (eq_to_iso hN₁) (eq_to_iso hΓ₀) + (N₁Γ₀ : Γ ⋙ N₁ ≅ (to_karoubi_equivalence (chain_complex C ℕ)).functor) := +begin + ext K : 3, + simpa only [compatibility.τ₀_hom_app, compatibility.τ₁_hom_app, eq_to_iso.hom, + preadditive.dold_kan.equivalence_counit_iso, N₂Γ₂_to_karoubi_iso_hom, eq_to_hom_map, + eq_to_hom_trans_assoc, eq_to_hom_app] using N₂Γ₂_compatible_with_N₁Γ₀ K, +end + +/-- The counit isomorphism induced by `N₁Γ₀` -/ +@[simps] +def η : Γ ⋙ N ≅ 𝟭 (chain_complex C ℕ) := compatibility.equivalence_counit_iso + (N₁Γ₀ : (Γ : chain_complex C ℕ ⥤ _ ) ⋙ N₁ ≅ (to_karoubi_equivalence _).functor) + +lemma equivalence_counit_iso : + dold_kan.equivalence.counit_iso = (η : Γ ⋙ N ≅ 𝟭 (chain_complex C ℕ)) := +compatibility.equivalence_counit_iso_eq hη + +lemma hε : compatibility.υ (eq_to_iso hN₁) = + (Γ₂N₁ : (to_karoubi_equivalence _).functor ≅ (N₁ : simplicial_object C ⥤ _) ⋙ + preadditive.dold_kan.equivalence.inverse) := +begin + ext X : 4, + erw [nat_trans.comp_app, compatibility_Γ₂N₁_Γ₂N₂_nat_trans], + simp only [compatibility.υ_hom_app, compatibility_Γ₂N₁_Γ₂N₂, + preadditive.dold_kan.equivalence_unit_iso, Γ₂N₂, iso.symm_hom, as_iso_inv, assoc], + erw [← nat_trans.comp_app_assoc, is_iso.hom_inv_id], + dsimp, + simpa only [id_comp, eq_to_hom_app, eq_to_hom_map, eq_to_hom_trans], +end + +/-- The unit isomorphism induced by `Γ₂N₁`. -/ +def ε : 𝟭 (simplicial_object C) ≅ N ⋙ Γ := +compatibility.equivalence_unit_iso (eq_to_iso hΓ₀) Γ₂N₁ + +lemma equivalence_unit_iso : dold_kan.equivalence.unit_iso = + (ε : 𝟭 (simplicial_object C) ≅ N ⋙ Γ) := +compatibility.equivalence_unit_iso_eq hε + +end dold_kan + +end idempotents + +end category_theory diff --git a/src/category_theory/functor/category.lean b/src/category_theory/functor/category.lean index bf73ec5a9f3f1..e3bb5ae20656d 100644 --- a/src/category_theory/functor/category.lean +++ b/src/category_theory/functor/category.lean @@ -60,6 +60,9 @@ lemma congr_app {α β : F ⟶ G} (h : α = β) (X : C) : α.app X = β.app X := @[simp] lemma id_app (F : C ⥤ D) (X : C) : (𝟙 F : F ⟶ F).app X = 𝟙 (F.obj X) := rfl @[simp] lemma comp_app {F G H : C ⥤ D} (α : F ⟶ G) (β : G ⟶ H) (X : C) : (α ≫ β).app X = α.app X ≫ β.app X := rfl +lemma comp_app_assoc {F G H : C ⥤ D} (α : F ⟶ G) (β : G ⟶ H) (X : C) {X' : D} + (f : H.obj X ⟶ X') : + (α ≫ β).app X ≫ f = α.app X ≫ β.app X ≫ f := by rw [comp_app, assoc] lemma app_naturality {F G : C ⥤ (D ⥤ E)} (T : F ⟶ G) (X : C) {Y Z : D} (f : Y ⟶ Z) : ((F.obj X).map f) ≫ ((T.app X).app Z) = ((T.app X).app Y) ≫ ((G.obj X).map f) := From 3ba15165bd6927679be7c22d6091a87337e3cd0c Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 2 Aug 2023 13:20:57 +0000 Subject: [PATCH 48/61] feat(analysis/convex/proj_Icc): Extending convex functions (#18797) Constantly extending monotone/antitone functions preserves their convexity. --- src/algebra/order/module.lean | 12 +++ src/algebra/order/monoid/lemmas.lean | 17 ++++ src/algebra/order/smul.lean | 15 ++- src/analysis/convex/proj_Icc.lean | 91 +++++++++++++++++++ src/data/set/intervals/basic.lean | 20 ++++ .../set/intervals/unordered_interval.lean | 25 ++++- src/order/lattice.lean | 40 ++++++++ 7 files changed, 217 insertions(+), 3 deletions(-) create mode 100644 src/analysis/convex/proj_Icc.lean diff --git a/src/algebra/order/module.lean b/src/algebra/order/module.lean index f4b75426245d9..be66690ff1695 100644 --- a/src/algebra/order/module.lean +++ b/src/algebra/order/module.lean @@ -209,6 +209,18 @@ lemma bdd_above.smul_of_nonpos (hc : c ≤ 0) (hs : bdd_above s) : bdd_below (c end ordered_ring +section linear_ordered_ring +variables [linear_ordered_ring k] [linear_ordered_add_comm_group M] [module k M] [ordered_smul k M] + {a : k} + +lemma smul_max_of_nonpos (ha : a ≤ 0) (b₁ b₂ : M) : a • max b₁ b₂ = min (a • b₁) (a • b₂) := +(antitone_smul_left ha : antitone (_ : M → M)).map_max + +lemma smul_min_of_nonpos (ha : a ≤ 0) (b₁ b₂ : M) : a • min b₁ b₂ = max (a • b₁) (a • b₂) := +(antitone_smul_left ha : antitone (_ : M → M)).map_min + +end linear_ordered_ring + section linear_ordered_field variables [linear_ordered_field k] [ordered_add_comm_group M] [module k M] [ordered_smul k M] {s : set M} {c : k} diff --git a/src/algebra/order/monoid/lemmas.lean b/src/algebra/order/monoid/lemmas.lean index 2f2ca1f8b5c17..bbc1e04cab5f1 100644 --- a/src/algebra/order/monoid/lemmas.lean +++ b/src/algebra/order/monoid/lemmas.lean @@ -244,6 +244,23 @@ variables [linear_order α] {a b c d : α} [covariant_class α α (*) (<)] @[to_additive] lemma min_le_max_of_mul_le_mul (h : a * b ≤ c * d) : min a b ≤ max c d := by { simp_rw [min_le_iff, le_max_iff], contrapose! h, exact mul_lt_mul_of_lt_of_lt h.1.1 h.2.2 } +end linear_order + +section linear_order +variables [linear_order α] [covariant_class α α (*) (≤)] [covariant_class α α (swap (*)) (≤)] + {a b c d : α} + +@[to_additive max_add_add_le_max_add_max] lemma max_mul_mul_le_max_mul_max' : + max (a * b) (c * d) ≤ max a c * max b d := +max_le (mul_le_mul' (le_max_left _ _) $ le_max_left _ _) $ + mul_le_mul' (le_max_right _ _) $ le_max_right _ _ + +--TODO: Also missing `min_mul_min_le_min_mul_mul` +@[to_additive min_add_min_le_min_add_add] lemma min_mul_min_le_min_mul_mul' : + min a c * min b d ≤ min (a * b) (c * d) := +le_min (mul_le_mul' (min_le_left _ _) $ min_le_left _ _) $ + mul_le_mul' (min_le_right _ _) $ min_le_right _ _ + end linear_order end has_mul diff --git a/src/algebra/order/smul.lean b/src/algebra/order/smul.lean index b6cd7435b978b..734d7ca1d0213 100644 --- a/src/algebra/order/smul.lean +++ b/src/algebra/order/smul.lean @@ -170,12 +170,23 @@ ordered_smul.mk'' $ λ n hn, begin { cases (int.neg_succ_not_pos _).1 hn } end +section linear_ordered_semiring +variables [linear_ordered_semiring R] [linear_ordered_add_comm_monoid M] [smul_with_zero R M] + [ordered_smul R M] {a : R} + -- TODO: `linear_ordered_field M → ordered_smul ℚ M` -instance linear_ordered_semiring.to_ordered_smul {R : Type*} [linear_ordered_semiring R] : - ordered_smul R R := +instance linear_ordered_semiring.to_ordered_smul : ordered_smul R R := ordered_smul.mk'' $ λ c, strict_mono_mul_left_of_pos +lemma smul_max (ha : 0 ≤ a) (b₁ b₂ : M) : a • max b₁ b₂ = max (a • b₁) (a • b₂) := +(monotone_smul_left ha : monotone (_ : M → M)).map_max + +lemma smul_min (ha : 0 ≤ a) (b₁ b₂ : M) : a • min b₁ b₂ = min (a • b₁) (a • b₂) := +(monotone_smul_left ha : monotone (_ : M → M)).map_min + +end linear_ordered_semiring + section linear_ordered_semifield variables [linear_ordered_semifield 𝕜] [ordered_add_comm_monoid M] [ordered_add_comm_monoid N] [mul_action_with_zero 𝕜 M] [mul_action_with_zero 𝕜 N] diff --git a/src/analysis/convex/proj_Icc.lean b/src/analysis/convex/proj_Icc.lean new file mode 100644 index 0000000000000..3115cdffd85d2 --- /dev/null +++ b/src/analysis/convex/proj_Icc.lean @@ -0,0 +1,91 @@ +/- +Copyright (c) 2023 Yaël Dillies. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies +-/ +import analysis.convex.function +import data.set.intervals.proj_Icc + +/-! +# Convexity of extension from intervals + +This file proves that constantly extending monotone/antitone functions preserves their convexity. + +## TODO + +We could deduplicate the proofs if we had a typeclass stating that `segment 𝕜 x y = [x -[𝕜] y]` as +`𝕜ᵒᵈ` respects it if `𝕜` does, while `𝕜ᵒᵈ` isn't a `linear_ordered_field` if `𝕜` is. +-/ + +open set + +variables {𝕜 β : Type*} [linear_ordered_field 𝕜] [linear_ordered_add_comm_monoid β] [has_smul 𝕜 β] + {s : set 𝕜} {f : 𝕜 → β} {z : 𝕜} + +/-- A convex set extended towards minus infinity is convex. -/ +protected lemma convex.Ici_extend (hf : convex 𝕜 s) : + convex 𝕜 {x | Ici_extend (restrict (Ici z) (∈ s)) x} := +by { rw convex_iff_ord_connected at ⊢ hf, exact hf.restrict.Ici_extend } + +/-- A convex set extended towards infinity is convex. -/ +protected lemma convex.Iic_extend (hf : convex 𝕜 s) : + convex 𝕜 {x | Iic_extend (restrict (Iic z) (∈ s)) x} := +by { rw convex_iff_ord_connected at ⊢ hf, exact hf.restrict.Iic_extend } + +/-- A convex monotone function extended constantly towards minus infinity is convex. -/ +protected lemma convex_on.Ici_extend (hf : convex_on 𝕜 s f) (hf' : monotone_on f s) : + convex_on 𝕜 {x | Ici_extend (restrict (Ici z) (∈ s)) x} (Ici_extend $ restrict (Ici z) f) := +begin + refine ⟨hf.1.Ici_extend, λ x hx y hy a b ha hb hab, _⟩, + dsimp [Ici_extend_apply] at ⊢ hx hy, + refine (hf' (hf.1.ord_connected.uIcc_subset hx hy $ monotone.image_uIcc_subset (λ _ _, max_le_max + le_rfl) $ mem_image_of_mem _ $ convex_uIcc _ _ left_mem_uIcc right_mem_uIcc ha hb hab) + (hf.1 hx hy ha hb hab) _).trans (hf.2 hx hy ha hb hab), + rw [smul_max ha z, smul_max hb z], + refine le_trans _ max_add_add_le_max_add_max, + rw [convex.combo_self hab, smul_eq_mul, smul_eq_mul], +end + +/-- A convex antitone function extended constantly towards infinity is convex. -/ +protected lemma convex_on.Iic_extend (hf : convex_on 𝕜 s f) (hf' : antitone_on f s) : + convex_on 𝕜 {x | Iic_extend (restrict (Iic z) (∈ s)) x} (Iic_extend $ restrict (Iic z) f) := +begin + refine ⟨hf.1.Iic_extend, λ x hx y hy a b ha hb hab, _⟩, + dsimp [Iic_extend_apply] at ⊢ hx hy, + refine (hf' (hf.1 hx hy ha hb hab) (hf.1.ord_connected.uIcc_subset hx hy $ + monotone.image_uIcc_subset (λ _ _, min_le_min le_rfl) $ mem_image_of_mem _ $ + convex_uIcc _ _ left_mem_uIcc right_mem_uIcc ha hb hab) _).trans (hf.2 hx hy ha hb hab), + rw [smul_min ha z, smul_min hb z], + refine min_add_min_le_min_add_add.trans _ , + rw [convex.combo_self hab, smul_eq_mul, smul_eq_mul], +end + +/-- A concave antitone function extended constantly minus towards infinity is concave. -/ +protected lemma concave_on.Ici_extend (hf : concave_on 𝕜 s f) (hf' : antitone_on f s) : + concave_on 𝕜 {x | Ici_extend (restrict (Ici z) (∈ s)) x} (Ici_extend $ restrict (Ici z) f) := +hf.dual.Ici_extend hf'.dual_right + +/-- A concave monotone function extended constantly towards infinity is concave. -/ +protected lemma concave_on.Iic_extend (hf : concave_on 𝕜 s f) (hf' : monotone_on f s) : + concave_on 𝕜 {x | Iic_extend (restrict (Iic z) (∈ s)) x} (Iic_extend $ restrict (Iic z) f) := +hf.dual.Iic_extend hf'.dual_right + +/-- A convex monotone function extended constantly towards minus infinity is convex. -/ +protected lemma convex_on.Ici_extend_of_monotone (hf : convex_on 𝕜 univ f) (hf' : monotone f) : + convex_on 𝕜 univ (Ici_extend $ restrict (Ici z) f) := +hf.Ici_extend $ hf'.monotone_on _ + +/-- A convex antitone function extended constantly towards infinity is convex. -/ +protected lemma convex_on.Iic_extend_of_antitone (hf : convex_on 𝕜 univ f) (hf' : antitone f) : + convex_on 𝕜 univ (Iic_extend $ restrict (Iic z) f) := +hf.Iic_extend $ hf'.antitone_on _ + +/-- A concave antitone function extended constantly minus towards infinity is concave. -/ +protected lemma concave_on.Ici_extend_of_antitone (hf : concave_on 𝕜 univ f) (hf' : antitone f) : + concave_on 𝕜 univ (Ici_extend $ restrict (Ici z) f) := +hf.Ici_extend $ hf'.antitone_on _ + +/-- A concave monotone function extended constantly towards infinity is concave. -/ +protected lemma concave_on.Iic_extend_of_monotone (hf : concave_on 𝕜 univ f) (hf' : monotone f) : + concave_on 𝕜 univ (Iic_extend $ restrict (Iic z) f) := +hf.Iic_extend $ hf'.monotone_on _ diff --git a/src/data/set/intervals/basic.lean b/src/data/set/intervals/basic.lean index 1c4c0e2cc4b6d..250eba029ea68 100644 --- a/src/data/set/intervals/basic.lean +++ b/src/data/set/intervals/basic.lean @@ -1164,6 +1164,26 @@ begin le_of_lt h₂, le_of_lt h₁] }, end +section lattice +variables [lattice β] {f : α → β} + +lemma _root_.monotone_on.image_Icc_subset (hf : monotone_on f (Icc a b)) : + f '' Icc a b ⊆ Icc (f a) (f b) := +image_subset_iff.2 $ λ c hc, + ⟨hf (left_mem_Icc.2 $ hc.1.trans hc.2) hc hc.1, hf hc (right_mem_Icc.2 $ hc.1.trans hc.2) hc.2⟩ + +lemma _root_.antitone_on.image_Icc_subset (hf : antitone_on f (Icc a b)) : + f '' Icc a b ⊆ Icc (f b) (f a) := +image_subset_iff.2 $ λ c hc, + ⟨hf hc (right_mem_Icc.2 $ hc.1.trans hc.2) hc.2, hf (left_mem_Icc.2 $ hc.1.trans hc.2) hc hc.1⟩ + +lemma _root_.monotone.image_Icc_subset (hf : monotone f) : f '' Icc a b ⊆ Icc (f a) (f b) := +(hf.monotone_on _).image_Icc_subset + +lemma _root_.antitone.image_Icc_subset (hf : antitone f) : f '' Icc a b ⊆ Icc (f b) (f a) := +(hf.antitone_on _).image_Icc_subset + +end lattice end linear_order section lattice diff --git a/src/data/set/intervals/unordered_interval.lean b/src/data/set/intervals/unordered_interval.lean index 1e795686afe08..6edbb024832d0 100644 --- a/src/data/set/intervals/unordered_interval.lean +++ b/src/data/set/intervals/unordered_interval.lean @@ -129,7 +129,30 @@ by simpa only [uIcc_comm] using uIcc_injective_right a end distrib_lattice section linear_order -variables [linear_order α] [linear_order β] {f : α → β} {s : set α} {a a₁ a₂ b b₁ b₂ c d x : α} +variables [linear_order α] + +section lattice +variables [lattice β] {f : α → β} {s : set α} {a b : α} + +lemma _root_.monotone_on.image_uIcc_subset (hf : monotone_on f (uIcc a b)) : + f '' uIcc a b ⊆ uIcc (f a) (f b) := +hf.image_Icc_subset.trans $ + by rw [hf.map_sup left_mem_uIcc right_mem_uIcc, hf.map_inf left_mem_uIcc right_mem_uIcc, uIcc] + +lemma _root_.antitone_on.image_uIcc_subset (hf : antitone_on f (uIcc a b)) : + f '' uIcc a b ⊆ uIcc (f a) (f b) := +hf.image_Icc_subset.trans $ + by rw [hf.map_sup left_mem_uIcc right_mem_uIcc, hf.map_inf left_mem_uIcc right_mem_uIcc, uIcc] + +lemma _root_.monotone.image_uIcc_subset (hf : monotone f) : f '' uIcc a b ⊆ uIcc (f a) (f b) := +(hf.monotone_on _).image_uIcc_subset + +lemma _root_.antitone.image_uIcc_subset (hf : antitone f) : f '' uIcc a b ⊆ uIcc (f a) (f b) := +(hf.antitone_on _).image_uIcc_subset + +end lattice + +variables [linear_order β] {f : α → β} {s : set α} {a a₁ a₂ b b₁ b₂ c d x : α} lemma Icc_min_max : Icc (min a b) (max a b) = [a, b] := rfl diff --git a/src/order/lattice.lean b/src/order/lattice.lean index e3877be97db81..14128e9768f95 100644 --- a/src/order/lattice.lean +++ b/src/order/lattice.lean @@ -846,6 +846,7 @@ hf.dual.map_sup _ _ end monotone namespace monotone_on +variables {f : α → β} {s : set α} {x y : α} /-- Pointwise supremum of two monotone functions is a monotone function. -/ protected lemma sup [preorder α] [semilattice_sup β] {f g : α → β} {s : set α} @@ -867,6 +868,25 @@ protected lemma min [preorder α] [linear_order β] {f g : α → β} {s : set (hf : monotone_on f s) (hg : monotone_on g s) : monotone_on (λ x, min (f x) (g x)) s := hf.inf hg +lemma of_map_inf [semilattice_inf α] [semilattice_inf β] + (h : ∀ (x ∈ s) (y ∈ s), f (x ⊓ y) = f x ⊓ f y) : monotone_on f s := +λ x hx y hy hxy, inf_eq_left.1 $ by rw [←h _ hx _ hy, inf_eq_left.2 hxy] + +lemma of_map_sup [semilattice_sup α] [semilattice_sup β] + (h : ∀ (x ∈ s) (y ∈ s), f (x ⊔ y) = f x ⊔ f y) : monotone_on f s := +(@of_map_inf αᵒᵈ βᵒᵈ _ _ _ _ h).dual + +variables [linear_order α] + +lemma map_sup [semilattice_sup β] (hf : monotone_on f s) (hx : x ∈ s) (hy : y ∈ s) : + f (x ⊔ y) = f x ⊔ f y := +by cases le_total x y; have := hf _ _ h; + assumption <|> simp only [h, this, sup_of_le_left, sup_of_le_right] + +lemma map_inf [semilattice_inf β] (hf : monotone_on f s) (hx : x ∈ s) (hy : y ∈ s) : + f (x ⊓ y) = f x ⊓ f y := +hf.dual.map_sup hx hy + end monotone_on namespace antitone @@ -912,6 +932,7 @@ hf.dual_right.map_inf x y end antitone namespace antitone_on +variables {f : α → β} {s : set α} {x y : α} /-- Pointwise supremum of two antitone functions is a antitone function. -/ protected lemma sup [preorder α] [semilattice_sup β] {f g : α → β} {s : set α} @@ -933,6 +954,25 @@ protected lemma min [preorder α] [linear_order β] {f g : α → β} {s : set (hf : antitone_on f s) (hg : antitone_on g s) : antitone_on (λ x, min (f x) (g x)) s := hf.inf hg +lemma of_map_inf [semilattice_inf α] [semilattice_sup β] + (h : ∀ (x ∈ s) (y ∈ s), f (x ⊓ y) = f x ⊔ f y) : antitone_on f s := +λ x hx y hy hxy, sup_eq_left.1 $ by rw [←h _ hx _ hy, inf_eq_left.2 hxy] + +lemma of_map_sup [semilattice_sup α] [semilattice_inf β] + (h : ∀ (x ∈ s) (y ∈ s), f (x ⊔ y) = f x ⊓ f y) : antitone_on f s := +(@of_map_inf αᵒᵈ βᵒᵈ _ _ _ _ h).dual + +variables [linear_order α] + +lemma map_sup [semilattice_inf β] (hf : antitone_on f s) (hx : x ∈ s) (hy : y ∈ s) : + f (x ⊔ y) = f x ⊓ f y := +by cases le_total x y; have := hf _ _ h; assumption <|> + simp only [h, this, sup_of_le_left, sup_of_le_right, inf_of_le_left, inf_of_le_right] + +lemma map_inf [semilattice_sup β] (hf : antitone_on f s) (hx : x ∈ s) (hy : y ∈ s) : + f (x ⊓ y) = f x ⊔ f y := +hf.dual.map_sup hx hy + end antitone_on /-! From 48a058d7e39a80ed56858505719a0b2197900999 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 2 Aug 2023 16:47:14 +0000 Subject: [PATCH 49/61] feat(data/sum/interval): The lexicographic sum of two locally finite orders is locally finite (#11352) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This proves `locally_finite_order (α ⊕ₗ β)` where `α` and `β` are locally finite themselves. --- src/data/finset/sum.lean | 2 + src/data/sum/interval.lean | 188 ++++++++++++++++++++++++++++++++++++- 2 files changed, 185 insertions(+), 5 deletions(-) diff --git a/src/data/finset/sum.lean b/src/data/finset/sum.lean index 86d4f7ab8974f..a940205835735 100644 --- a/src/data/finset/sum.lean +++ b/src/data/finset/sum.lean @@ -54,6 +54,8 @@ multiset.mem_disj_sum @[simp] lemma inl_mem_disj_sum : inl a ∈ s.disj_sum t ↔ a ∈ s := inl_mem_disj_sum @[simp] lemma inr_mem_disj_sum : inr b ∈ s.disj_sum t ↔ b ∈ t := inr_mem_disj_sum +@[simp] lemma disj_sum_eq_empty : s.disj_sum t = ∅ ↔ s = ∅ ∧ t = ∅ := by simp [ext_iff] + lemma disj_sum_mono (hs : s₁ ⊆ s₂) (ht : t₁ ⊆ t₂) : s₁.disj_sum t₁ ⊆ s₂.disj_sum t₂ := val_le_iff.1 $ disj_sum_mono (val_le_iff.2 hs) (val_le_iff.2 ht) diff --git a/src/data/sum/interval.lean b/src/data/sum/interval.lean index 5bd433861ad78..e831e4564adfc 100644 --- a/src/data/sum/interval.lean +++ b/src/data/sum/interval.lean @@ -3,6 +3,7 @@ Copyright (c) 2022 Yaël Dillies. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ +import data.finset.sum import data.sum.order import order.locally_finite @@ -12,11 +13,8 @@ import order.locally_finite > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. > Any changes to this file require a corresponding PR to mathlib4. -This file provides the `locally_finite_order` instance for the disjoint sum of two orders. - -## TODO - -Do the same for the lexicographic sum of orders. +This file provides the `locally_finite_order` instance for the disjoint sum and linear sum of two +orders and calculates the cardinality of their finite intervals. -/ open function sum @@ -99,6 +97,106 @@ lemma sum_lift₂_mono (h₁ : ∀ a b, f₁ a b ⊆ g₁ a b) (h₂ : ∀ a b, | (inr a) (inr b) := map_subset_map.2 (h₂ _ _) end sum_lift₂ + +section sum_lex_lift +variables (f₁ f₁' : α₁ → β₁ → finset γ₁) (f₂ f₂' : α₂ → β₂ → finset γ₂) + (g₁ g₁' : α₁ → β₂ → finset γ₁) (g₂ g₂' : α₁ → β₂ → finset γ₂) + +/-- Lifts maps `α₁ → β₁ → finset γ₁`, `α₂ → β₂ → finset γ₂`, `α₁ → β₂ → finset γ₁`, +`α₂ → β₂ → finset γ₂` to a map `α₁ ⊕ α₂ → β₁ ⊕ β₂ → finset (γ₁ ⊕ γ₂)`. Could be generalized to +alternative monads if we can make sure to keep computability and universe polymorphism. -/ +def sum_lex_lift : Π (a : α₁ ⊕ α₂) (b : β₁ ⊕ β₂), finset (γ₁ ⊕ γ₂) +| (inl a) (inl b) := (f₁ a b).map embedding.inl +| (inl a) (inr b) := (g₁ a b).disj_sum (g₂ a b) +| (inr a) (inl b) := ∅ +| (inr a) (inr b) := (f₂ a b).map ⟨_, inr_injective⟩ + +@[simp] lemma sum_lex_lift_inl_inl (a : α₁) (b : β₁) : + sum_lex_lift f₁ f₂ g₁ g₂ (inl a) (inl b) = (f₁ a b).map embedding.inl := rfl + +@[simp] lemma sum_lex_lift_inl_inr (a : α₁) (b : β₂) : + sum_lex_lift f₁ f₂ g₁ g₂ (inl a) (inr b) = (g₁ a b).disj_sum (g₂ a b) := rfl + +@[simp] lemma sum_lex_lift_inr_inl (a : α₂) (b : β₁) : + sum_lex_lift f₁ f₂ g₁ g₂ (inr a) (inl b) = ∅ := rfl + +@[simp] lemma sum_lex_lift_inr_inr (a : α₂) (b : β₂) : + sum_lex_lift f₁ f₂ g₁ g₂ (inr a) (inr b) = (f₂ a b).map ⟨_, inr_injective⟩ := rfl + +variables {f₁ g₁ f₂ g₂ f₁' g₁' f₂' g₂'} {a : α₁ ⊕ α₂} {b : β₁ ⊕ β₂} {c : γ₁ ⊕ γ₂} + +lemma mem_sum_lex_lift : + c ∈ sum_lex_lift f₁ f₂ g₁ g₂ a b ↔ + (∃ a₁ b₁ c₁, a = inl a₁ ∧ b = inl b₁ ∧ c = inl c₁ ∧ c₁ ∈ f₁ a₁ b₁) ∨ + (∃ a₁ b₂ c₁, a = inl a₁ ∧ b = inr b₂ ∧ c = inl c₁ ∧ c₁ ∈ g₁ a₁ b₂) ∨ + (∃ a₁ b₂ c₂, a = inl a₁ ∧ b = inr b₂ ∧ c = inr c₂ ∧ c₂ ∈ g₂ a₁ b₂) ∨ + ∃ a₂ b₂ c₂, a = inr a₂ ∧ b = inr b₂ ∧ c = inr c₂ ∧ c₂ ∈ f₂ a₂ b₂ := +begin + split, + { cases a; cases b, + { rw [sum_lex_lift, mem_map], + rintro ⟨c, hc, rfl⟩, + exact or.inl ⟨a, b, c, rfl, rfl, rfl, hc⟩ }, + { refine λ h, (mem_disj_sum.1 h).elim _ _, + { rintro ⟨c, hc, rfl⟩, + refine or.inr (or.inl ⟨a, b, c, rfl, rfl, rfl, hc⟩) }, + { rintro ⟨c, hc, rfl⟩, + refine or.inr (or.inr $ or.inl ⟨a, b, c, rfl, rfl, rfl, hc⟩) } }, + { refine λ h, (not_mem_empty _ h).elim }, + { rw [sum_lex_lift, mem_map], + rintro ⟨c, hc, rfl⟩, + exact or.inr (or.inr $ or.inr $ ⟨a, b, c, rfl, rfl, rfl, hc⟩) } }, + { rintro (⟨a, b, c, rfl, rfl, rfl, hc⟩ | ⟨a, b, c, rfl, rfl, rfl, hc⟩ | + ⟨a, b, c, rfl, rfl, rfl, hc⟩ | ⟨a, b, c, rfl, rfl, rfl, hc⟩), + { exact mem_map_of_mem _ hc }, + { exact inl_mem_disj_sum.2 hc }, + { exact inr_mem_disj_sum.2 hc }, + { exact mem_map_of_mem _ hc } } +end + +lemma inl_mem_sum_lex_lift {c₁ : γ₁} : + inl c₁ ∈ sum_lex_lift f₁ f₂ g₁ g₂ a b ↔ + (∃ a₁ b₁, a = inl a₁ ∧ b = inl b₁ ∧ c₁ ∈ f₁ a₁ b₁) ∨ + ∃ a₁ b₂, a = inl a₁ ∧ b = inr b₂ ∧ c₁ ∈ g₁ a₁ b₂ := +by simp [mem_sum_lex_lift] + +lemma inr_mem_sum_lex_lift {c₂ : γ₂} : + inr c₂ ∈ sum_lex_lift f₁ f₂ g₁ g₂ a b ↔ + (∃ a₁ b₂, a = inl a₁ ∧ b = inr b₂ ∧ c₂ ∈ g₂ a₁ b₂) ∨ + ∃ a₂ b₂, a = inr a₂ ∧ b = inr b₂ ∧ c₂ ∈ f₂ a₂ b₂ := +by simp [mem_sum_lex_lift] + +lemma sum_lex_lift_mono (hf₁ : ∀ a b, f₁ a b ⊆ f₁' a b) (hf₂ : ∀ a b, f₂ a b ⊆ f₂' a b) + (hg₁ : ∀ a b, g₁ a b ⊆ g₁' a b) (hg₂ : ∀ a b, g₂ a b ⊆ g₂' a b) (a : α₁ ⊕ α₂) (b : β₁ ⊕ β₂) : + sum_lex_lift f₁ f₂ g₁ g₂ a b ⊆ sum_lex_lift f₁' f₂' g₁' g₂' a b := +begin + cases a; cases b, + exacts [map_subset_map.2 (hf₁ _ _), disj_sum_mono (hg₁ _ _) (hg₂ _ _), subset.rfl, + map_subset_map.2 (hf₂ _ _)], +end + +lemma sum_lex_lift_eq_empty : + (sum_lex_lift f₁ f₂ g₁ g₂ a b) = ∅ ↔ (∀ a₁ b₁, a = inl a₁ → b = inl b₁ → f₁ a₁ b₁ = ∅) ∧ + (∀ a₁ b₂, a = inl a₁ → b = inr b₂ → g₁ a₁ b₂ = ∅ ∧ g₂ a₁ b₂ = ∅) ∧ + ∀ a₂ b₂, a = inr a₂ → b = inr b₂ → f₂ a₂ b₂ = ∅ := +begin + refine ⟨λ h, ⟨_, _, _⟩, λ h, _⟩, + any_goals { rintro a b rfl rfl, exact map_eq_empty.1 h }, + { rintro a b rfl rfl, exact disj_sum_eq_empty.1 h }, + cases a; cases b, + { exact map_eq_empty.2 (h.1 _ _ rfl rfl) }, + { simp [h.2.1 _ _ rfl rfl] }, + { refl }, + { exact map_eq_empty.2 (h.2.2 _ _ rfl rfl) } +end + +lemma sum_lex_lift_nonempty : + (sum_lex_lift f₁ f₂ g₁ g₂ a b).nonempty ↔ (∃ a₁ b₁, a = inl a₁ ∧ b = inl b₁ ∧ (f₁ a₁ b₁).nonempty) + ∨ (∃ a₁ b₂, a = inl a₁ ∧ b = inr b₂ ∧ ((g₁ a₁ b₂).nonempty ∨ (g₂ a₁ b₂).nonempty)) + ∨ ∃ a₂ b₂, a = inr a₂ ∧ b = inr b₂ ∧ (f₂ a₂ b₂).nonempty := +by simp [nonempty_iff_ne_empty, sum_lex_lift_eq_empty, not_and_distrib] + +end sum_lex_lift end finset open finset function @@ -141,4 +239,84 @@ lemma Ioc_inr_inr : Ioc (inr b₁ : α ⊕ β) (inr b₂) = (Ioc b₁ b₂).map lemma Ioo_inr_inr : Ioo (inr b₁ : α ⊕ β) (inr b₂) = (Ioo b₁ b₂).map embedding.inr := rfl end disjoint + +/-! ### Lexicographical sum of orders -/ + +namespace lex +variables [preorder α] [preorder β] [order_top α] [order_bot β] [locally_finite_order α] + [locally_finite_order β] + +/-- Throwaway tactic. -/ +private meta def simp_lex : tactic unit := +`[refine to_lex.surjective.forall₃.2 _, rintro (a | a) (b | b) (c | c); simp only + [sum_lex_lift_inl_inl, sum_lex_lift_inl_inr, sum_lex_lift_inr_inl, sum_lex_lift_inr_inr, + inl_le_inl_iff, inl_le_inr, not_inr_le_inl, inr_le_inr_iff, inl_lt_inl_iff, inl_lt_inr, + not_inr_lt_inl, inr_lt_inr_iff, mem_Icc, mem_Ico, mem_Ioc, mem_Ioo, mem_Ici, mem_Ioi, mem_Iic, + mem_Iio, equiv.coe_to_embedding, to_lex_inj, exists_false, and_false, false_and, map_empty, + not_mem_empty, true_and, inl_mem_disj_sum, inr_mem_disj_sum, and_true, of_lex_to_lex, mem_map, + embedding.coe_fn_mk, exists_prop, exists_eq_right, embedding.inl_apply]] + +instance locally_finite_order : locally_finite_order (α ⊕ₗ β) := +{ finset_Icc := λ a b, + (sum_lex_lift Icc Icc (λ a _, Ici a) (λ _, Iic) (of_lex a) (of_lex b)).map to_lex.to_embedding, + finset_Ico := λ a b, + (sum_lex_lift Ico Ico (λ a _, Ici a) (λ _, Iio) (of_lex a) (of_lex b)).map to_lex.to_embedding, + finset_Ioc := λ a b, + (sum_lex_lift Ioc Ioc (λ a _, Ioi a) (λ _, Iic) (of_lex a) (of_lex b)).map to_lex.to_embedding, + finset_Ioo := λ a b, + (sum_lex_lift Ioo Ioo (λ a _, Ioi a) (λ _, Iio) (of_lex a) (of_lex b)).map to_lex.to_embedding, + finset_mem_Icc := by simp_lex, + finset_mem_Ico := by simp_lex, + finset_mem_Ioc := by simp_lex, + finset_mem_Ioo := by simp_lex } + +variables (a a₁ a₂ : α) (b b₁ b₂ : β) + +lemma Icc_inl_inl : + Icc (inlₗ a₁ : α ⊕ₗ β) (inlₗ a₂) = (Icc a₁ a₂).map (embedding.inl.trans to_lex.to_embedding) := +by { rw ←finset.map_map, refl } + +lemma Ico_inl_inl : + Ico (inlₗ a₁ : α ⊕ₗ β) (inlₗ a₂) = (Ico a₁ a₂).map (embedding.inl.trans to_lex.to_embedding) := +by { rw ←finset.map_map, refl } + +lemma Ioc_inl_inl : + Ioc (inlₗ a₁ : α ⊕ₗ β) (inlₗ a₂) = (Ioc a₁ a₂).map (embedding.inl.trans to_lex.to_embedding) := +by { rw ←finset.map_map, refl } + +lemma Ioo_inl_inl : + Ioo (inlₗ a₁ : α ⊕ₗ β) (inlₗ a₂) = (Ioo a₁ a₂).map (embedding.inl.trans to_lex.to_embedding) := +by { rw ←finset.map_map, refl } + +@[simp] lemma Icc_inl_inr : + Icc (inlₗ a) (inrₗ b) = ((Ici a).disj_sum (Iic b)).map to_lex.to_embedding := rfl +@[simp] lemma Ico_inl_inr : + Ico (inlₗ a) (inrₗ b) = ((Ici a).disj_sum (Iio b)).map to_lex.to_embedding := rfl +@[simp] lemma Ioc_inl_inr : + Ioc (inlₗ a) (inrₗ b) = ((Ioi a).disj_sum (Iic b)).map to_lex.to_embedding := rfl +@[simp] lemma Ioo_inl_inr : + Ioo (inlₗ a) (inrₗ b) = ((Ioi a).disj_sum (Iio b)).map to_lex.to_embedding := rfl + +@[simp] lemma Icc_inr_inl : Icc (inrₗ b) (inlₗ a) = ∅ := rfl +@[simp] lemma Ico_inr_inl : Ico (inrₗ b) (inlₗ a) = ∅ := rfl +@[simp] lemma Ioc_inr_inl : Ioc (inrₗ b) (inlₗ a) = ∅ := rfl +@[simp] lemma Ioo_inr_inl : Ioo (inrₗ b) (inlₗ a) = ∅ := rfl + +lemma Icc_inr_inr : + Icc (inrₗ b₁ : α ⊕ₗ β) (inrₗ b₂) = (Icc b₁ b₂).map (embedding.inr.trans to_lex.to_embedding) := +by { rw ←finset.map_map, refl } + +lemma Ico_inr_inr : + Ico (inrₗ b₁ : α ⊕ₗ β) (inrₗ b₂) = (Ico b₁ b₂).map (embedding.inr.trans to_lex.to_embedding) := +by { rw ←finset.map_map, refl } + +lemma Ioc_inr_inr : + Ioc (inrₗ b₁ : α ⊕ₗ β) (inrₗ b₂) = (Ioc b₁ b₂).map (embedding.inr.trans to_lex.to_embedding) := +by { rw ←finset.map_map, refl } + +lemma Ioo_inr_inr : + Ioo (inrₗ b₁ : α ⊕ₗ β) (inrₗ b₂) = (Ioo b₁ b₂).map (embedding.inr.trans to_lex.to_embedding) := +by { rw ←finset.map_map, refl } + +end lex end sum From 32a7e535287f9c73f2e4d2aef306a39190f0b504 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= <37772949+joelriou@users.noreply.github.com> Date: Sat, 5 Aug 2023 18:07:30 +0000 Subject: [PATCH 50/61] feat(algebraic_topology/dold_kan): The Dold-Kan equivalence for abelian categories (#17926) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: Joël Riou --- .../dold_kan/compatibility.lean | 4 +- .../dold_kan/decomposition.lean | 2 + .../dold_kan/degeneracies.lean | 2 + .../dold_kan/equivalence.lean | 175 ++++++++++++++++++ .../dold_kan/equivalence_additive.lean | 2 + .../dold_kan/equivalence_pseudoabelian.lean | 2 + src/algebraic_topology/dold_kan/faces.lean | 2 + .../dold_kan/functor_gamma.lean | 2 + .../dold_kan/functor_n.lean | 2 +- .../dold_kan/gamma_comp_n.lean | 2 + .../dold_kan/n_comp_gamma.lean | 2 + .../dold_kan/n_reflects_iso.lean | 2 + .../dold_kan/normalized.lean | 2 + .../dold_kan/notations.lean | 2 + src/algebraic_topology/dold_kan/p_infty.lean | 3 +- .../dold_kan/projections.lean | 2 + .../dold_kan/split_simplicial_object.lean | 3 + 17 files changed, 208 insertions(+), 3 deletions(-) create mode 100644 src/algebraic_topology/dold_kan/equivalence.lean diff --git a/src/algebraic_topology/dold_kan/compatibility.lean b/src/algebraic_topology/dold_kan/compatibility.lean index 3567efda0ccc9..516e38ecaf374 100644 --- a/src/algebraic_topology/dold_kan/compatibility.lean +++ b/src/algebraic_topology/dold_kan/compatibility.lean @@ -36,7 +36,9 @@ inverse of `eB`: but whose inverse functor is `G`. When extra assumptions are given, we shall also provide simplification lemmas for the -unit and counit isomorphisms of `equivalence`. (TODO) +unit and counit isomorphisms of `equivalence`. + +(See `equivalence.lean` for the general strategy of proof of the Dold-Kan equivalence.) -/ diff --git a/src/algebraic_topology/dold_kan/decomposition.lean b/src/algebraic_topology/dold_kan/decomposition.lean index 1d89283129af6..f1cbbb85fd8b7 100644 --- a/src/algebraic_topology/dold_kan/decomposition.lean +++ b/src/algebraic_topology/dold_kan/decomposition.lean @@ -29,6 +29,8 @@ role in the proof that the functor `N₁ : simplicial_object C ⥤ karoubi (chain_complex C ℕ))` reflects isomorphisms. +(See `equivalence.lean` for the general strategy of proof of the Dold-Kan equivalence.) + -/ open category_theory category_theory.category category_theory.preadditive opposite diff --git a/src/algebraic_topology/dold_kan/degeneracies.lean b/src/algebraic_topology/dold_kan/degeneracies.lean index c3b3de9892ba2..ab0584b061494 100644 --- a/src/algebraic_topology/dold_kan/degeneracies.lean +++ b/src/algebraic_topology/dold_kan/degeneracies.lean @@ -25,6 +25,8 @@ if `X : simplicial_object C` with `C` a preadditive category, `X.map θ.op ≫ P_infty.f n = 0`. It follows from the more precise statement vanishing statement `σ_comp_P_eq_zero` for the `P q`. +(See `equivalence.lean` for the general strategy of proof of the Dold-Kan equivalence.) + -/ open category_theory category_theory.category category_theory.limits diff --git a/src/algebraic_topology/dold_kan/equivalence.lean b/src/algebraic_topology/dold_kan/equivalence.lean new file mode 100644 index 0000000000000..dc025a7938b2b --- /dev/null +++ b/src/algebraic_topology/dold_kan/equivalence.lean @@ -0,0 +1,175 @@ +/- +Copyright (c) 2022 Joël Riou. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joël Riou +-/ + +import algebraic_topology.dold_kan.equivalence_pseudoabelian +import algebraic_topology.dold_kan.normalized + +/-! + +# The Dold-Kan correspondence + +The Dold-Kan correspondence states that for any abelian category `A`, there is +an equivalence between the category of simplicial objects in `A` and the +category of chain complexes in `A` (with degrees indexed by `ℕ` and the +homological convention that the degree is decreased by the differentials). + +In this file, we finish the construction of this equivalence by providing +`category_theory.abelian.dold_kan.equivalence` which is of type +`simplicial_object A ≌ chain_complex A ℕ` for any abelian category `A`. +The functor `simplicial_object A ⥤ chain_complex A ℕ` of this equivalence is +definitionally equal to `normalized_Moore_complex A`. + +## Overall strategy of the proof of the correspondence + +Before starting the implementation of the proof in Lean, the author noticed +that the Dold-Kan equivalence not only applies to abelian categories, but +should also hold generally for any pseudoabelian category `C` +(i.e. a category with instances `[preadditive C]` +`[has_finite_coproducts C]` and `[is_idempotent_complete C]`): this is +`category_theory.idempotents.dold_kan.equivalence`. + +When the alternating face map complex `K[X]` of a simplicial object `X` in an +abelian is studied, it is shown that it decomposes as a direct sum of the +normalized subcomplex and of the degenerate subcomplex. The crucial observation +is that in this decomposition, the projection on the normalized subcomplex can +be defined in each degree using simplicial operators. Then, the definition +of this projection `P_infty : K[X] ⟶ K[X]` can be carried out for any +`X : simplicial_object C` when `C` is a preadditive category. + +The construction of the endomorphism `P_infty` is done in the files +`homotopies.lean`, `faces.lean`, `projections.lean` and `p_infty.lean`. +Eventually, as we would also like to show that the inclusion of the normalized +Moore complex is a homotopy equivalence (cf. file `homotopy_equivalence.lean`), +this projection `P_infty` needs to be homotopic to the identity. In our +construction, we get this for free because `P_infty` is obtained by altering +the identity endomorphism by null homotopic maps. More details about this +aspect of the proof are in the file `homotopies.lean`. + +When the alternating face map complex `K[X]` is equipped with the idempotent +endomorphism `P_infty`, it becomes an object in `karoubi (chain_complex C ℕ)` +which is the idempotent completion of the category `chain_complex C ℕ`. In `functor_n.lean`, +we obtain this functor `N₁ : simplicial_object C ⥤ karoubi (chain_complex C ℕ)`, +which is formally extended as +`N₂ : karoubi (simplicial_object C) ⥤ karoubi (chain_complex C ℕ)`. (Here, some functors +have an index which is the number of occurrences of `karoubi` at the source or the +target.) + +In `functor_gamma.lean`, assuming that the category `C` is additive, +we define the functor in the other direction +`Γ₂ : karoubi (chain_complex C ℕ) ⥤ karoubi (simplicial_object C)` as the formal +extension of a functor `Γ₀ : chain_complex C ℕ ⥤ simplicial_object C` which is +defined similarly as in *Simplicial Homotopy Theory* by Goerss-Jardine. +In `degeneracies.lean`, we show that `P_infty` vanishes on the image of degeneracy +operators, which is one of the key properties that makes it possible to contruct +the isomorphism `N₂Γ₂ : Γ₂ ⋙ N₂ ≅ 𝟭 (karoubi (chain_complex C ℕ))`. + +The rest of the proof follows the strategy in the [original paper by Dold][dold1958]. We show +that the functor `N₂` reflects isomorphisms in `n_reflects_iso.lean`: this relies on a +decomposition of the identity of `X _[n]` using `P_infty.f n` and degeneracies obtained in +`decomposition.lean`. Then, in `n_comp_gamma.lean`, we construct a natural transformation +`Γ₂N₂.trans : N₂ ⋙ Γ₂ ⟶ 𝟭 (karoubi (simplicial_object C))`. It is shown that it is an +isomorphism using the fact that `N₂` reflects isomorphisms, and because we can show +that the composition `N₂ ⟶ N₂ ⋙ Γ₂ ⋙ N₂ ⟶ N₂` is the identity (see `identity_N₂`). The fact +that `N₂` is defined as a formal direct factor makes the proof easier because we only +have to compare endomorphisms of an alternating face map complex `K[X]` and we do not +have to worry with inclusions of kernel subobjects. + +In `equivalence_additive.lean`, we obtain +the equivalence `equivalence : karoubi (simplicial_object C) ≌ karoubi (chain_complex C ℕ)`. +It is in the namespace `category_theory.preadditive.dold_kan`. The functors in this +equivalence are named `N` and `Γ`: by definition, they are `N₂` and `Γ₂`. + +In `equivalence_pseudoabelian.lean`, assuming `C` is idempotent complete, +we obtain `equivalence : simplicial_object C ≌ chain_complex C ℕ` +in the namespace `category_theory.idempotents.dold_kan`. This could be roughly +obtained by composing the previous equivalence with the equivalences +`simplicial_object C ≌ karoubi (simplicial_object C)` and +`karoubi (chain_complex C ℕ) ≌ chain_complex C ℕ`. Instead, we polish this construction +in `compatibility.lean` by ensuring good definitional properties of the equivalence (e.g. +the inverse functor is definitionallly equal to +`Γ₀' : chain_complex C ℕ ⥤ simplicial_object C`) and +showing compatibilities for the unit and counit isomorphisms. + +In this file `equivalence.lean`, assuming the category `A` is abelian, we obtain +`equivalence : simplicial_object A ≌ chain_complex A ℕ` in the namespace +`category_theory.abelian.dold_kan`. This is obtained by replacing the functor +`category_theory.idempotents.dold_kan.N` of the equivalence in the pseudoabelian case +with the isomorphic functor `normalized_Moore_complex A` thanks to the isomorphism +obtained in `normalized.lean`. + +TODO: Show functoriality properties of the three equivalences above. More precisely, +for example in the case of abelian categories `A` and `B`, if `F : A ⥤ B` is an +additive functor, we can show that the functors `N` for `A` and `B` are compatible +with the functors `simplicial_object A ⥤ simplicial_object B` and +`chain_complex A ℕ ⥤ chain_complex B ℕ` induced by `F`. (Note that this does not +require that `F` is an exact functor!) + +TODO: Introduce the degenerate subcomplex `D[X]` which is generated by +degenerate simplices, show that the projector `P_infty` corresponds to +a decomposition `K[X] ≅ N[X] ⊞ D[X]`. + +TODO: dualise all of this as `cosimplicial_object A ⥤ cochain_complex A ℕ`. (It is unclear +what is the best way to do this. The exact design may be decided when it is needed.) + +## References +* [Albrecht Dold, Homology of Symmetric Products and Other Functors of Complexes][dold1958] +* [Paul G. Goerss, John F. Jardine, Simplicial Homotopy Theory][goerss-jardine-2009] + +-/ + +noncomputable theory + +open category_theory +open category_theory.category +open category_theory.idempotents + +variables {A : Type*} [category A] [abelian A] + +namespace category_theory + +namespace abelian + +namespace dold_kan + +open algebraic_topology.dold_kan + +/-- The functor `N` for the equivalence is `normalized_Moore_complex A` -/ +def N : simplicial_object A ⥤ chain_complex A ℕ := algebraic_topology.normalized_Moore_complex A + +/-- The functor `Γ` for the equivalence is the same as in the pseudoabelian case. -/ +def Γ : chain_complex A ℕ ⥤ simplicial_object A := idempotents.dold_kan.Γ + +/-- The comparison isomorphism between `normalized_Moore_complex A` and +the functor `idempotents.dold_kan.N` from the pseudoabelian case -/ +@[simps] +def comparison_N : (N : simplicial_object A ⥤ _) ≅ idempotents.dold_kan.N := +calc N ≅ N ⋙ 𝟭 _ : functor.left_unitor N +... ≅ N ⋙ ((to_karoubi_equivalence _).functor ⋙ (to_karoubi_equivalence _).inverse) : + iso_whisker_left _ (to_karoubi_equivalence _).unit_iso +... ≅ (N ⋙ (to_karoubi_equivalence _).functor) ⋙ (to_karoubi_equivalence _).inverse : + iso.refl _ +... ≅ N₁ ⋙ (to_karoubi_equivalence _).inverse : iso_whisker_right + (N₁_iso_normalized_Moore_complex_comp_to_karoubi A).symm _ +... ≅ idempotents.dold_kan.N : by refl + +/-- The Dold-Kan equivalence for abelian categories -/ +@[simps functor] +def equivalence : simplicial_object A ≌ chain_complex A ℕ := +begin + let F : simplicial_object A ⥤ _ := idempotents.dold_kan.N, + let hF : is_equivalence F := is_equivalence.of_equivalence idempotents.dold_kan.equivalence, + letI : is_equivalence (N : simplicial_object A ⥤ _ ) := + is_equivalence.of_iso comparison_N.symm hF, + exact N.as_equivalence, +end + +lemma equivalence_inverse : (equivalence : simplicial_object A ≌ _).inverse = Γ := rfl + +end dold_kan + +end abelian + +end category_theory diff --git a/src/algebraic_topology/dold_kan/equivalence_additive.lean b/src/algebraic_topology/dold_kan/equivalence_additive.lean index 886aefc7db66b..742196ac2a72b 100644 --- a/src/algebraic_topology/dold_kan/equivalence_additive.lean +++ b/src/algebraic_topology/dold_kan/equivalence_additive.lean @@ -14,6 +14,8 @@ import algebraic_topology.dold_kan.n_comp_gamma This file defines `preadditive.dold_kan.equivalence` which is the equivalence of categories `karoubi (simplicial_object C) ≌ karoubi (chain_complex C ℕ)`. +(See `equivalence.lean` for the general strategy of proof of the Dold-Kan equivalence.) + -/ noncomputable theory diff --git a/src/algebraic_topology/dold_kan/equivalence_pseudoabelian.lean b/src/algebraic_topology/dold_kan/equivalence_pseudoabelian.lean index dd766e8a5e41f..4d6fff16434a8 100644 --- a/src/algebraic_topology/dold_kan/equivalence_pseudoabelian.lean +++ b/src/algebraic_topology/dold_kan/equivalence_pseudoabelian.lean @@ -29,6 +29,8 @@ the composition of `N₁ : simplicial_object C ⥤ karoubi (chain_complex C ℕ) `idempotents.dold_kan.Γ` of the equivalence is by definition the functor `Γ₀` introduced in `functor_gamma.lean`. +(See `equivalence.lean` for the general strategy of proof of the Dold-Kan equivalence.) + -/ noncomputable theory diff --git a/src/algebraic_topology/dold_kan/faces.lean b/src/algebraic_topology/dold_kan/faces.lean index ae2b2f6e8fff0..d9e7af0e63960 100644 --- a/src/algebraic_topology/dold_kan/faces.lean +++ b/src/algebraic_topology/dold_kan/faces.lean @@ -25,6 +25,8 @@ The main lemma in this file is `higher_faces_vanish.induction`. It is based on two technical lemmas `higher_faces_vanish.comp_Hσ_eq` and `higher_faces_vanish.comp_Hσ_eq_zero`. +(See `equivalence.lean` for the general strategy of proof of the Dold-Kan equivalence.) + -/ open nat diff --git a/src/algebraic_topology/dold_kan/functor_gamma.lean b/src/algebraic_topology/dold_kan/functor_gamma.lean index 9a5675b665216..bd138e8ba7d11 100644 --- a/src/algebraic_topology/dold_kan/functor_gamma.lean +++ b/src/algebraic_topology/dold_kan/functor_gamma.lean @@ -27,6 +27,8 @@ By construction, `Γ₀.obj K` is a split simplicial object whose splitting is ` We also construct `Γ₂ : karoubi (chain_complex C ℕ) ⥤ karoubi (simplicial_object C)` which shall be an equivalence for any additive category `C`. +(See `equivalence.lean` for the general strategy of proof of the Dold-Kan equivalence.) + -/ noncomputable theory diff --git a/src/algebraic_topology/dold_kan/functor_n.lean b/src/algebraic_topology/dold_kan/functor_n.lean index 17a1db8bb19eb..74f4d58a98df4 100644 --- a/src/algebraic_topology/dold_kan/functor_n.lean +++ b/src/algebraic_topology/dold_kan/functor_n.lean @@ -32,7 +32,7 @@ defined in `equivalence_pseudoabelian.lean`. When the category `C` is abelian, a relation between `N₁` and the normalized Moore complex functor shall be obtained in `normalized.lean`. -(See `equivalence.lean` for the general strategy of proof.) +(See `equivalence.lean` for the general strategy of proof of the Dold-Kan equivalence.) -/ diff --git a/src/algebraic_topology/dold_kan/gamma_comp_n.lean b/src/algebraic_topology/dold_kan/gamma_comp_n.lean index b3d66c5c2cfcc..86e9a3e0230c0 100644 --- a/src/algebraic_topology/dold_kan/gamma_comp_n.lean +++ b/src/algebraic_topology/dold_kan/gamma_comp_n.lean @@ -16,6 +16,8 @@ The purpose of this file is to construct natural isomorphisms `N₁Γ₀ : Γ₀ ⋙ N₁ ≅ to_karoubi (chain_complex C ℕ)` and `N₂Γ₂ : Γ₂ ⋙ N₂ ≅ 𝟭 (karoubi (chain_complex C ℕ))`. +(See `equivalence.lean` for the general strategy of proof of the Dold-Kan equivalence.) + -/ noncomputable theory diff --git a/src/algebraic_topology/dold_kan/n_comp_gamma.lean b/src/algebraic_topology/dold_kan/n_comp_gamma.lean index 52bdb21046dc0..ac2c1537c3539 100644 --- a/src/algebraic_topology/dold_kan/n_comp_gamma.lean +++ b/src/algebraic_topology/dold_kan/n_comp_gamma.lean @@ -21,6 +21,8 @@ that it becomes an isomorphism after the application of the functor `N₂ : karoubi (simplicial_object C) ⥤ karoubi (chain_complex C ℕ)` which reflects isomorphisms. +(See `equivalence.lean` for the general strategy of proof of the Dold-Kan equivalence.) + -/ noncomputable theory diff --git a/src/algebraic_topology/dold_kan/n_reflects_iso.lean b/src/algebraic_topology/dold_kan/n_reflects_iso.lean index d9cd2e612a1e0..10541d6b39f98 100644 --- a/src/algebraic_topology/dold_kan/n_reflects_iso.lean +++ b/src/algebraic_topology/dold_kan/n_reflects_iso.lean @@ -21,6 +21,8 @@ In this file, it is shown that the functors `N₂ : karoubi (simplicial_object C) ⥤ karoubi (chain_complex C ℕ))` reflect isomorphisms for any preadditive category `C`. +(See `equivalence.lean` for the general strategy of proof of the Dold-Kan equivalence.) + -/ open category_theory diff --git a/src/algebraic_topology/dold_kan/normalized.lean b/src/algebraic_topology/dold_kan/normalized.lean index c182615e41b7c..d01fa53ec4cd0 100644 --- a/src/algebraic_topology/dold_kan/normalized.lean +++ b/src/algebraic_topology/dold_kan/normalized.lean @@ -27,6 +27,8 @@ the Dold-Kan equivalence `category_theory.abelian.dold_kan.equivalence : simplicial_object A ≌ chain_complex A ℕ` with a functor (definitionally) equal to `normalized_Moore_complex A`. +(See `equivalence.lean` for the general strategy of proof of the Dold-Kan equivalence.) + -/ open category_theory category_theory.category category_theory.limits diff --git a/src/algebraic_topology/dold_kan/notations.lean b/src/algebraic_topology/dold_kan/notations.lean index 67fda8f17e311..3276ab335e94a 100644 --- a/src/algebraic_topology/dold_kan/notations.lean +++ b/src/algebraic_topology/dold_kan/notations.lean @@ -17,6 +17,8 @@ This file defines the notation `K[X] : chain_complex C ℕ` for the alternating map complex of `(X : simplicial_object C)` where `C` is a preadditive category, as well as `N[X]` for the normalized subcomplex in the case `C` is an abelian category. +(See `equivalence.lean` for the general strategy of proof of the Dold-Kan equivalence.) + -/ localized "notation (name := alternating_face_map_complex) `K[`X`]` := diff --git a/src/algebraic_topology/dold_kan/p_infty.lean b/src/algebraic_topology/dold_kan/p_infty.lean index b95df628099e3..7b1c539dca8bf 100644 --- a/src/algebraic_topology/dold_kan/p_infty.lean +++ b/src/algebraic_topology/dold_kan/p_infty.lean @@ -22,7 +22,8 @@ to the limit the projections `P q` defined in `projections.lean`. This projection is a critical tool in this formalisation of the Dold-Kan correspondence, because in the case of abelian categories, `P_infty` corresponds to the projection on the normalized Moore subcomplex, with kernel the degenerate subcomplex. -(See `equivalence.lean` for the general strategy of proof.) + +(See `equivalence.lean` for the general strategy of proof of the Dold-Kan equivalence.) -/ diff --git a/src/algebraic_topology/dold_kan/projections.lean b/src/algebraic_topology/dold_kan/projections.lean index 2439974a9359a..4b9ea21fdc872 100644 --- a/src/algebraic_topology/dold_kan/projections.lean +++ b/src/algebraic_topology/dold_kan/projections.lean @@ -30,6 +30,8 @@ By passing to the limit, these endomorphisms `P q` shall be used in `p_infty.lea in order to define `P_infty : K[X] ⟶ K[X]`, see `equivalence.lean` for the general strategy of proof of the Dold-Kan equivalence. +(See `equivalence.lean` for the general strategy of proof of the Dold-Kan equivalence.) + -/ open category_theory category_theory.category category_theory.limits diff --git a/src/algebraic_topology/dold_kan/split_simplicial_object.lean b/src/algebraic_topology/dold_kan/split_simplicial_object.lean index f3b2c7fec482f..997675cde1821 100644 --- a/src/algebraic_topology/dold_kan/split_simplicial_object.lean +++ b/src/algebraic_topology/dold_kan/split_simplicial_object.lean @@ -18,6 +18,9 @@ import algebraic_topology.dold_kan.functor_n In this file we define a functor `nondeg_complex : simplicial_object.split C ⥤ chain_complex C ℕ` when `C` is a preadditive category with finite coproducts, and get an isomorphism `to_karoubi_nondeg_complex_iso_N₁ : nondeg_complex ⋙ to_karoubi _ ≅ forget C ⋙ dold_kan.N₁`. + +(See `equivalence.lean` for the general strategy of proof of the Dold-Kan equivalence.) + -/ noncomputable theory From ffde2d8a6e689149e44fd95fa862c23a57f8c780 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 28 Aug 2023 11:58:43 +0000 Subject: [PATCH 51/61] chore(order/liminf_limsup): Generalise and move lemmas (#18628) Generalise lemmas from semilattices to codirected orders. Move topology-less lemmas from `topology.algebra.order.liminf_limsup` to `order.liminf_limsup`. Also turn arguments to `bdd_above_insert` and friends implicit. --- src/data/set/finite.lean | 17 ++-- src/order/bounds/basic.lean | 77 +++++++++++++------ src/order/directed.lean | 40 ---------- src/order/liminf_limsup.lean | 46 +++++++++-- src/topology/algebra/order/liminf_limsup.lean | 35 --------- 5 files changed, 99 insertions(+), 116 deletions(-) diff --git a/src/data/set/finite.lean b/src/data/set/finite.lean index 67eb58978f762..f86cca6bfa443 100644 --- a/src/data/set/finite.lean +++ b/src/data/set/finite.lean @@ -1280,7 +1280,7 @@ end section -variables [semilattice_sup α] [nonempty α] {s : set α} +variables [preorder α] [is_directed α (≤)] [nonempty α] {s : set α} /--A finite set is bounded above.-/ protected lemma finite.bdd_above (hs : s.finite) : bdd_above s := @@ -1288,7 +1288,7 @@ finite.induction_on hs bdd_above_empty $ λ a s _ _ h, h.insert a /--A finite union of sets which are all bounded above is still bounded above.-/ lemma finite.bdd_above_bUnion {I : set β} {S : β → set α} (H : I.finite) : - (bdd_above (⋃i∈I, S i)) ↔ (∀i ∈ I, bdd_above (S i)) := + bdd_above (⋃ i ∈ I, S i) ↔ ∀ i ∈ I, bdd_above (S i) := finite.induction_on H (by simp only [bUnion_empty, bdd_above_empty, ball_empty_iff]) (λ a s ha _ hs, by simp only [bUnion_insert, ball_insert_iff, bdd_above_union, hs]) @@ -1299,22 +1299,17 @@ end section -variables [semilattice_inf α] [nonempty α] {s : set α} +variables [preorder α] [is_directed α (≥)] [nonempty α] {s : set α} /--A finite set is bounded below.-/ -protected lemma finite.bdd_below (hs : s.finite) : bdd_below s := @finite.bdd_above αᵒᵈ _ _ _ hs +protected lemma finite.bdd_below (hs : s.finite) : bdd_below s := @finite.bdd_above αᵒᵈ _ _ _ _ hs /--A finite union of sets which are all bounded below is still bounded below.-/ lemma finite.bdd_below_bUnion {I : set β} {S : β → set α} (H : I.finite) : bdd_below (⋃ i ∈ I, S i) ↔ ∀ i ∈ I, bdd_below (S i) := -@finite.bdd_above_bUnion αᵒᵈ _ _ _ _ _ H +@finite.bdd_above_bUnion αᵒᵈ _ _ _ _ _ _ H -lemma infinite_of_not_bdd_below : ¬ bdd_below s → s.infinite := -begin - contrapose!, - rw not_infinite, - apply finite.bdd_below, -end +lemma infinite_of_not_bdd_below : ¬ bdd_below s → s.infinite := mt finite.bdd_below end diff --git a/src/order/bounds/basic.lean b/src/order/bounds/basic.lean index 6cfd3c71db1d3..6add0bb14d555 100644 --- a/src/order/bounds/basic.lean +++ b/src/order/bounds/basic.lean @@ -5,6 +5,7 @@ Authors: Johannes Hölzl, Yury Kudryashov -/ import data.set.intervals.basic import data.set.n_ary +import order.directed /-! @@ -283,31 +284,30 @@ h.mono $ inter_subset_left s t lemma bdd_below.inter_of_right (h : bdd_below t) : bdd_below (s ∩ t) := h.mono $ inter_subset_right s t -/-- If `s` and `t` are bounded above sets in a `semilattice_sup`, then so is `s ∪ t`. -/ -lemma bdd_above.union [semilattice_sup γ] {s t : set γ} : +/-- In a directed order, the union of bounded above sets is bounded above. -/ +lemma bdd_above.union [is_directed α (≤)] {s t : set α} : bdd_above s → bdd_above t → bdd_above (s ∪ t) := begin - rintros ⟨bs, hs⟩ ⟨bt, ht⟩, - use bs ⊔ bt, - rw upper_bounds_union, - exact ⟨upper_bounds_mono_mem le_sup_left hs, - upper_bounds_mono_mem le_sup_right ht⟩ + rintro ⟨a, ha⟩ ⟨b, hb⟩, + obtain ⟨c, hca, hcb⟩ := exists_ge_ge a b, + rw [bdd_above, upper_bounds_union], + exact ⟨c, upper_bounds_mono_mem hca ha, upper_bounds_mono_mem hcb hb⟩, end -/-- The union of two sets is bounded above if and only if each of the sets is. -/ -lemma bdd_above_union [semilattice_sup γ] {s t : set γ} : +/-- In a directed order, the union of two sets is bounded above if and only if both sets are. -/ +lemma bdd_above_union [is_directed α (≤)] {s t : set α} : bdd_above (s ∪ t) ↔ bdd_above s ∧ bdd_above t := -⟨λ h, ⟨h.mono $ subset_union_left s t, h.mono $ subset_union_right s t⟩, - λ h, h.1.union h.2⟩ +⟨λ h, ⟨h.mono $ subset_union_left _ _, h.mono $ subset_union_right _ _⟩, λ h, h.1.union h.2⟩ -lemma bdd_below.union [semilattice_inf γ] {s t : set γ} : +/-- In a codirected order, the union of bounded below sets is bounded below. -/ +lemma bdd_below.union [is_directed α (≥)] {s t : set α} : bdd_below s → bdd_below t → bdd_below (s ∪ t) := -@bdd_above.union γᵒᵈ _ s t +@bdd_above.union αᵒᵈ _ _ _ _ -/--The union of two sets is bounded above if and only if each of the sets is.-/ -lemma bdd_below_union [semilattice_inf γ] {s t : set γ} : +/-- In a codirected order, the union of two sets is bounded below if and only if both sets are. -/ +lemma bdd_below_union [is_directed α (≥)] {s t : set α} : bdd_below (s ∪ t) ↔ bdd_below s ∧ bdd_below t := -@bdd_above_union γᵒᵈ _ s t +@bdd_above_union αᵒᵈ _ _ _ _ /-- If `a` is the least upper bound of `s` and `b` is the least upper bound of `t`, then `a ⊔ b` is the least upper bound of `s ∪ t`. -/ @@ -642,22 +642,22 @@ lemma nonempty_of_not_bdd_below [ha : nonempty α] (h : ¬bdd_below s) : s.nonem -/ /-- Adding a point to a set preserves its boundedness above. -/ -@[simp] lemma bdd_above_insert [semilattice_sup γ] (a : γ) {s : set γ} : +@[simp] lemma bdd_above_insert [is_directed α (≤)] {s : set α} {a : α} : bdd_above (insert a s) ↔ bdd_above s := by simp only [insert_eq, bdd_above_union, bdd_above_singleton, true_and] -lemma bdd_above.insert [semilattice_sup γ] (a : γ) {s : set γ} (hs : bdd_above s) : - bdd_above (insert a s) := -(bdd_above_insert a).2 hs +protected lemma bdd_above.insert [is_directed α (≤)] {s : set α} (a : α) : + bdd_above s → bdd_above (insert a s) := +bdd_above_insert.2 /--Adding a point to a set preserves its boundedness below.-/ -@[simp] lemma bdd_below_insert [semilattice_inf γ] (a : γ) {s : set γ} : +@[simp] lemma bdd_below_insert [is_directed α (≥)] {s : set α} {a : α} : bdd_below (insert a s) ↔ bdd_below s := by simp only [insert_eq, bdd_below_union, bdd_below_singleton, true_and] -lemma bdd_below.insert [semilattice_inf γ] (a : γ) {s : set γ} (hs : bdd_below s) : - bdd_below (insert a s) := -(bdd_below_insert a).2 hs +lemma bdd_below.insert [is_directed α (≥)] {s : set α} (a : α) : + bdd_below s → bdd_below (insert a s) := +bdd_below_insert.2 lemma is_lub.insert [semilattice_sup γ] (a) {b} {s : set γ} (hs : is_lub s b) : is_lub (insert a s) (a ⊔ b) := @@ -1191,3 +1191,32 @@ end lemma is_glb_prod [preorder α] [preorder β] {s : set (α × β)} (p : α × β) : is_glb s p ↔ is_glb (prod.fst '' s) p.1 ∧ is_glb (prod.snd '' s) p.2 := @is_lub_prod αᵒᵈ βᵒᵈ _ _ _ _ + +section scott_continuous +variables [preorder α] [preorder β] {f : α → β} {a : α} + +/-- A function between preorders is said to be Scott continuous if it preserves `is_lub` on directed +sets. It can be shown that a function is Scott continuous if and only if it is continuous wrt the +Scott topology. + +The dual notion + +```lean +∀ ⦃d : set α⦄, d.nonempty → directed_on (≥) d → ∀ ⦃a⦄, is_glb d a → is_glb (f '' d) (f a) +``` + +does not appear to play a significant role in the literature, so is omitted here. +-/ +def scott_continuous (f : α → β) : Prop := +∀ ⦃d : set α⦄, d.nonempty → directed_on (≤) d → ∀ ⦃a⦄, is_lub d a → is_lub (f '' d) (f a) + +protected lemma scott_continuous.monotone (h : scott_continuous f) : monotone f := +begin + refine λ a b hab, (h (insert_nonempty _ _) (directed_on_pair le_refl hab) _).1 + (mem_image_of_mem _ $ mem_insert _ _), + rw [is_lub, upper_bounds_insert, upper_bounds_singleton, + inter_eq_self_of_subset_right (Ici_subset_Ici.2 hab)], + exact is_least_Ici, +end + +end scott_continuous diff --git a/src/order/directed.lean b/src/order/directed.lean index 83c44b4bf2b62..0456ab1e72b13 100644 --- a/src/order/directed.lean +++ b/src/order/directed.lean @@ -6,7 +6,6 @@ Authors: Johannes Hölzl import data.set.image import order.lattice import order.max -import order.bounds.basic /-! # Directed indexed families and sets @@ -259,42 +258,3 @@ instance order_top.to_is_directed_le [has_le α] [order_top α] : is_directed α @[priority 100] -- see Note [lower instance priority] instance order_bot.to_is_directed_ge [has_le α] [order_bot α] : is_directed α (≥) := ⟨λ a b, ⟨⊥, bot_le, bot_le⟩⟩ - -section scott_continuous - -variables [preorder α] {a : α} - -/-- -A function between preorders is said to be Scott continuous if it preserves `is_lub` on directed -sets. It can be shown that a function is Scott continuous if and only if it is continuous wrt the -Scott topology. - -The dual notion - -```lean -∀ ⦃d : set α⦄, d.nonempty → directed_on (≥) d → ∀ ⦃a⦄, is_glb d a → is_glb (f '' d) (f a) -``` - -does not appear to play a significant role in the literature, so is omitted here. --/ -def scott_continuous [preorder β] (f : α → β) : Prop := -∀ ⦃d : set α⦄, d.nonempty → directed_on (≤) d → ∀ ⦃a⦄, is_lub d a → is_lub (f '' d) (f a) - -protected lemma scott_continuous.monotone [preorder β] {f : α → β} - (h : scott_continuous f) : - monotone f := -begin - intros a b hab, - have e1 : is_lub (f '' {a, b}) (f b), - { apply h, - { exact set.insert_nonempty _ _ }, - { exact directed_on_pair le_refl hab }, - { rw [is_lub, upper_bounds_insert, upper_bounds_singleton, - set.inter_eq_self_of_subset_right (set.Ici_subset_Ici.mpr hab)], - exact is_least_Ici } }, - apply e1.1, - rw set.image_pair, - exact set.mem_insert _ _, -end - -end scott_continuous diff --git a/src/order/liminf_limsup.lean b/src/order/liminf_limsup.lean index adafe31cd2008..ececbb65f3515 100644 --- a/src/order/liminf_limsup.lean +++ b/src/order/liminf_limsup.lean @@ -122,7 +122,7 @@ lemma not_is_bounded_under_of_tendsto_at_bot [preorder β] [no_min_order β] {f ¬ is_bounded_under (≥) l f := @not_is_bounded_under_of_tendsto_at_top α βᵒᵈ _ _ _ _ _ hf -lemma is_bounded_under.bdd_above_range_of_cofinite [semilattice_sup β] {f : α → β} +lemma is_bounded_under.bdd_above_range_of_cofinite [preorder β] [is_directed β (≤)] {f : α → β} (hf : is_bounded_under (≤) cofinite f) : bdd_above (range f) := begin rcases hf with ⟨b, hb⟩, @@ -131,17 +131,17 @@ begin exact ⟨⟨b, ball_image_iff.2 $ λ x, id⟩, (hb.image f).bdd_above⟩ end -lemma is_bounded_under.bdd_below_range_of_cofinite [semilattice_inf β] {f : α → β} +lemma is_bounded_under.bdd_below_range_of_cofinite [preorder β] [is_directed β (≥)] {f : α → β} (hf : is_bounded_under (≥) cofinite f) : bdd_below (range f) := -@is_bounded_under.bdd_above_range_of_cofinite α βᵒᵈ _ _ hf +@is_bounded_under.bdd_above_range_of_cofinite α βᵒᵈ _ _ _ hf -lemma is_bounded_under.bdd_above_range [semilattice_sup β] {f : ℕ → β} +lemma is_bounded_under.bdd_above_range [preorder β] [is_directed β (≤)] {f : ℕ → β} (hf : is_bounded_under (≤) at_top f) : bdd_above (range f) := by { rw ← nat.cofinite_eq_at_top at hf, exact hf.bdd_above_range_of_cofinite } -lemma is_bounded_under.bdd_below_range [semilattice_inf β] {f : ℕ → β} +lemma is_bounded_under.bdd_below_range [preorder β] [is_directed β (≥)] {f : ℕ → β} (hf : is_bounded_under (≥) at_top f) : bdd_below (range f) := -@is_bounded_under.bdd_above_range βᵒᵈ _ _ hf +@is_bounded_under.bdd_above_range βᵒᵈ _ _ _ hf /-- `is_cobounded (≺) f` states that the filter `f` does not tend to infinity w.r.t. `≺`. This is also called frequently bounded. Will be usually instantiated with `≤` or `≥`. @@ -198,6 +198,31 @@ lemma is_cobounded.mono (h : f ≤ g) : f.is_cobounded r → g.is_cobounded r end relation +section nonempty +variables [preorder α] [nonempty α] {f : filter β} {u : β → α} + +lemma is_bounded_le_at_bot : (at_bot : filter α).is_bounded (≤) := +‹nonempty α›.elim $ λ a, ⟨a, eventually_le_at_bot _⟩ + +lemma is_bounded_ge_at_top : (at_top : filter α).is_bounded (≥) := +‹nonempty α›.elim $ λ a, ⟨a, eventually_ge_at_top _⟩ + +lemma tendsto.is_bounded_under_le_at_bot (h : tendsto u f at_bot) : f.is_bounded_under (≤) u := +is_bounded_le_at_bot.mono h + +lemma tendsto.is_bounded_under_ge_at_top (h : tendsto u f at_top) : f.is_bounded_under (≥) u := +is_bounded_ge_at_top.mono h + +lemma bdd_above_range_of_tendsto_at_top_at_bot [is_directed α (≤)] {u : ℕ → α} + (hx : tendsto u at_top at_bot) : bdd_above (set.range u) := +hx.is_bounded_under_le_at_bot.bdd_above_range + +lemma bdd_below_range_of_tendsto_at_top_at_top [is_directed α (≥)] {u : ℕ → α} + (hx : tendsto u at_top at_top) : bdd_below (set.range u) := +hx.is_bounded_under_ge_at_top.bdd_below_range + +end nonempty + lemma is_cobounded_le_of_bot [preorder α] [order_bot α] {f : filter α} : f.is_cobounded (≤) := ⟨⊥, assume a h, bot_le⟩ @@ -955,6 +980,15 @@ lemma frequently_lt_of_liminf_lt {α β} [conditionally_complete_linear_order β ∃ᶠ x in f, u x < b := @frequently_lt_of_lt_limsup _ βᵒᵈ _ f u b hu h +variables [conditionally_complete_linear_order α] {f : filter α} {b : α} + +lemma lt_mem_sets_of_Limsup_lt (h : f.is_bounded (≤)) (l : f.Limsup < b) : ∀ᶠ a in f, a < b := +let ⟨c, (h : ∀ᶠ a in f, a ≤ c), hcb⟩ := exists_lt_of_cInf_lt h l in +mem_of_superset h $ λ a, hcb.trans_le' + +lemma gt_mem_sets_of_Liminf_gt : f.is_bounded (≥) → b < f.Liminf → ∀ᶠ a in f, b < a := +@lt_mem_sets_of_Limsup_lt αᵒᵈ _ _ _ + end conditionally_complete_linear_order end filter diff --git a/src/topology/algebra/order/liminf_limsup.lean b/src/topology/algebra/order/liminf_limsup.lean index 4e819a7cdbfb6..c16bae4dda725 100644 --- a/src/topology/algebra/order/liminf_limsup.lean +++ b/src/topology/algebra/order/liminf_limsup.lean @@ -51,19 +51,6 @@ lemma filter.tendsto.is_cobounded_under_ge {f : filter β} {u : β → α} {a : [ne_bot f] (h : tendsto u f (𝓝 a)) : f.is_cobounded_under (≥) u := h.is_bounded_under_le.is_cobounded_flip -lemma is_bounded_le_at_bot (α : Type*) [hα : nonempty α] [preorder α] : - (at_bot : filter α).is_bounded (≤) := -is_bounded_iff.2 ⟨set.Iic hα.some, mem_at_bot _, hα.some, λ x hx, hx⟩ - -lemma filter.tendsto.is_bounded_under_le_at_bot {α : Type*} [nonempty α] [preorder α] - {f : filter β} {u : β → α} (h : tendsto u f at_bot) : - f.is_bounded_under (≤) u := -(is_bounded_le_at_bot α).mono h - -lemma bdd_above_range_of_tendsto_at_top_at_bot {α : Type*} [nonempty α] [semilattice_sup α] - {u : ℕ → α} (hx : tendsto u at_top at_bot) : bdd_above (set.range u) := -(filter.tendsto.is_bounded_under_le_at_bot hx).bdd_above_range - end order_closed_topology section order_closed_topology @@ -90,33 +77,11 @@ lemma filter.tendsto.is_cobounded_under_le {f : filter β} {u : β → α} {a : [ne_bot f] (h : tendsto u f (𝓝 a)) : f.is_cobounded_under (≤) u := h.is_bounded_under_ge.is_cobounded_flip -lemma is_bounded_ge_at_top (α : Type*) [hα : nonempty α] [preorder α] : - (at_top : filter α).is_bounded (≥) := -is_bounded_le_at_bot αᵒᵈ - -lemma filter.tendsto.is_bounded_under_ge_at_top {α : Type*} [nonempty α] [preorder α] - {f : filter β} {u : β → α} (h : tendsto u f at_top) : - f.is_bounded_under (≥) u := -(is_bounded_ge_at_top α).mono h - -lemma bdd_below_range_of_tendsto_at_top_at_top {α : Type*} [nonempty α] [semilattice_inf α] - {u : ℕ → α} (hx : tendsto u at_top at_top) : bdd_below (set.range u) := -(filter.tendsto.is_bounded_under_ge_at_top hx).bdd_below_range - end order_closed_topology section conditionally_complete_linear_order variables [conditionally_complete_linear_order α] -theorem lt_mem_sets_of_Limsup_lt {f : filter α} {b} (h : f.is_bounded (≤)) (l : f.Limsup < b) : - ∀ᶠ a in f, a < b := -let ⟨c, (h : ∀ᶠ a in f, a ≤ c), hcb⟩ := exists_lt_of_cInf_lt h l in -mem_of_superset h $ assume a hac, lt_of_le_of_lt hac hcb - -theorem gt_mem_sets_of_Liminf_gt : ∀ {f : filter α} {b}, f.is_bounded (≥) → b < f.Liminf → - ∀ᶠ a in f, b < a := -@lt_mem_sets_of_Limsup_lt αᵒᵈ _ - variables [topological_space α] [order_topology α] /-- If the liminf and the limsup of a filter coincide, then this filter converges to From 442a83d738cb208d3600056c489be16900ba701d Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 31 Aug 2023 08:51:42 +0000 Subject: [PATCH 52/61] feat(data/finset/lattice): `sup'`/`inf'` lemmas (#18989) Match (most of) the lemmas between `finset.sup`/`finset.inf` and `finset.sup'`/`finset.inf'`. Also golf two proofs using `eq_of_forall_ge_iff` to make sure both APIs prove their lemmas in as closely as possible a way. Also define `finset.nontrivial` to match `set.nontrivial`. --- src/data/dfinsupp/multiset.lean | 5 +- src/data/finset/basic.lean | 80 +++++++------ src/data/finset/lattice.lean | 107 +++++++++++++++--- src/data/finset/locally_finite.lean | 6 +- .../gauss_eisenstein_lemmas.lean | 2 +- 5 files changed, 140 insertions(+), 60 deletions(-) diff --git a/src/data/dfinsupp/multiset.lean b/src/data/dfinsupp/multiset.lean index de2ed3c8f1031..9b20c21615ef9 100644 --- a/src/data/dfinsupp/multiset.lean +++ b/src/data/dfinsupp/multiset.lean @@ -53,9 +53,8 @@ def to_dfinsupp : multiset α →+ Π₀ a : α, ℕ := @[simp] lemma to_dfinsupp_apply (s : multiset α) (a : α) : s.to_dfinsupp a = s.count a := rfl -@[simp] lemma to_dfinsupp_support (s : multiset α) : - s.to_dfinsupp.support = s.to_finset := -(finset.filter_eq_self _).mpr (λ x hx, count_ne_zero.mpr $ multiset.mem_to_finset.1 hx) +@[simp] lemma to_dfinsupp_support (s : multiset α) : s.to_dfinsupp.support = s.to_finset := +finset.filter_true_of_mem $ λ x hx, count_ne_zero.mpr $ multiset.mem_to_finset.1 hx @[simp] lemma to_dfinsupp_replicate (a : α) (n : ℕ) : to_dfinsupp (multiset.replicate n a) = dfinsupp.single a n := diff --git a/src/data/finset/basic.lean b/src/data/finset/basic.lean index ae0db3da0b9b8..9bb6c39eb6a06 100644 --- a/src/data/finset/basic.lean +++ b/src/data/finset/basic.lean @@ -536,14 +536,27 @@ by rw [←coe_ssubset, coe_singleton, set.ssubset_singleton_iff, coe_eq_empty] lemma eq_empty_of_ssubset_singleton {s : finset α} {x : α} (hs : s ⊂ {x}) : s = ∅ := ssubset_singleton_iff.1 hs -lemma eq_singleton_or_nontrivial (ha : a ∈ s) : s = {a} ∨ (s : set α).nontrivial := +/-- A finset is nontrivial if it has at least two elements. -/ +@[reducible] protected def nontrivial (s : finset α) : Prop := (s : set α).nontrivial + +@[simp] lemma not_nontrivial_empty : ¬ (∅ : finset α).nontrivial := by simp [finset.nontrivial] + +@[simp] lemma not_nontrivial_singleton : ¬ ({a} : finset α).nontrivial := +by simp [finset.nontrivial] + +lemma nontrivial.ne_singleton (hs : s.nontrivial) : s ≠ {a} := +by { rintro rfl, exact not_nontrivial_singleton hs } + +lemma eq_singleton_or_nontrivial (ha : a ∈ s) : s = {a} ∨ s.nontrivial := by { rw ←coe_eq_singleton, exact set.eq_singleton_or_nontrivial ha } -lemma nonempty.exists_eq_singleton_or_nontrivial : - s.nonempty → (∃ a, s = {a}) ∨ (s : set α).nontrivial := +lemma nontrivial_iff_ne_singleton (ha : a ∈ s) : s.nontrivial ↔ s ≠ {a} := +⟨nontrivial.ne_singleton, (eq_singleton_or_nontrivial ha).resolve_left⟩ + +lemma nonempty.exists_eq_singleton_or_nontrivial : s.nonempty → (∃ a, s = {a}) ∨ s.nontrivial := λ ⟨a, ha⟩, (eq_singleton_or_nontrivial ha).imp_left $ exists.intro a -instance [nonempty α] : nontrivial (finset α) := +instance nontrivial' [nonempty α] : nontrivial (finset α) := ‹nonempty α›.elim $ λ a, ⟨⟨{a}, ∅, singleton_ne_empty _⟩⟩ instance [is_empty α] : unique (finset α) := @@ -577,6 +590,8 @@ lemma forall_of_forall_cons {p : α → Prop} {h : a ∉ s} (H : ∀ x, x ∈ co @[simp] lemma mk_cons {s : multiset α} (h : (a ::ₘ s).nodup) : (⟨a ::ₘ s, h⟩ : finset α) = cons a ⟨s, (nodup_cons.1 h).2⟩ (nodup_cons.1 h).1 := rfl +@[simp] lemma cons_empty (a : α) : cons a ∅ (not_mem_empty _) = {a} := rfl + @[simp] lemma nonempty_cons (h : a ∉ s) : (cons a s h).nonempty := ⟨a, mem_cons.2 $ or.inl rfl⟩ @[simp] lemma nonempty_mk {m : multiset α} {hm} : (⟨m, hm⟩ : finset α).nonempty ↔ m ≠ 0 := @@ -1379,7 +1394,7 @@ lemma inter_sdiff (s t u : finset α) : s ∩ (t \ u) = s ∩ t \ u := by { ext @[simp] lemma sdiff_inter_self (s₁ s₂ : finset α) : (s₂ \ s₁) ∩ s₁ = ∅ := inf_sdiff_self_left -@[simp] lemma sdiff_self (s₁ : finset α) : s₁ \ s₁ = ∅ := sdiff_self +@[simp] protected lemma sdiff_self (s₁ : finset α) : s₁ \ s₁ = ∅ := sdiff_self lemma sdiff_inter_distrib_right (s t u : finset α) : s \ (t ∩ u) = (s \ t) ∪ (s \ u) := sdiff_inf @@ -1531,7 +1546,7 @@ by rw [←sdiff_singleton_eq_erase, sdiff_sdiff_eq_sdiff_union (singleton_subset union_comm] lemma sdiff_erase_self (ha : a ∈ s) : s \ s.erase a = {a} := -by rw [sdiff_erase ha, sdiff_self, insert_emptyc_eq] +by rw [sdiff_erase ha, finset.sdiff_self, insert_emptyc_eq] lemma sdiff_sdiff_self_left (s t : finset α) : s \ (s \ t) = s ∩ t := sdiff_sdiff_right_self @@ -1588,10 +1603,10 @@ theorem sizeof_lt_sizeof_of_mem [has_sizeof α] {x : α} {s : finset α} (hx : x @[simp] theorem attach_empty : attach (∅ : finset α) = ∅ := rfl -@[simp] lemma attach_nonempty_iff (s : finset α) : s.attach.nonempty ↔ s.nonempty := +@[simp] lemma attach_nonempty_iff {s : finset α} : s.attach.nonempty ↔ s.nonempty := by simp [finset.nonempty] -@[simp] lemma attach_eq_empty_iff (s : finset α) : s.attach = ∅ ↔ s = ∅ := +@[simp] lemma attach_eq_empty_iff {s : finset α} : s.attach = ∅ ↔ s = ∅ := by simpa [eq_empty_iff_forall_not_mem] /-! ### piecewise -/ @@ -1747,7 +1762,7 @@ end decidable_pi_exists /-! ### filter -/ section filter -variables (p q : α → Prop) [decidable_pred p] [decidable_pred q] +variables (p q : α → Prop) [decidable_pred p] [decidable_pred q] {s : finset α} /-- `filter p s` is the set of elements of `s` that satisfy `p`. -/ def filter (s : finset α) : finset α := ⟨_, s.2.filter p⟩ @@ -1772,38 +1787,34 @@ variable (p) theorem filter_filter (s : finset α) : (s.filter p).filter q = s.filter (λa, p a ∧ q a) := ext $ assume a, by simp only [mem_filter, and_comm, and.left_comm] -lemma filter_true {s : finset α} [h : decidable_pred (λ _, true)] : - @finset.filter α (λ _, true) h s = s := -by ext; simp +lemma filter_comm (s : finset α) : (s.filter p).filter q = (s.filter q).filter p := +by simp_rw [filter_filter, and_comm] + +/- We can simplify an application of filter where the decidability is inferred in "the wrong way" -/ +@[simp] lemma filter_congr_decidable (s : finset α) (p : α → Prop) (h : decidable_pred p) + [decidable_pred p] : @filter α p h s = s.filter p := +by congr -@[simp] theorem filter_false {h} (s : finset α) : @filter α (λa, false) h s = ∅ := -ext $ assume a, by simp only [mem_filter, and_false]; refl +lemma filter_true {h} (s : finset α) : @filter α (λ a, true) h s = s := by ext; simp +lemma filter_false {h} (s : finset α) : @filter α (λ a, false) h s = ∅ := by ext; simp variables {p q} -lemma filter_eq_self (s : finset α) : - s.filter p = s ↔ ∀ x ∈ s, p x := -by simp [finset.ext_iff] +lemma filter_eq_self : s.filter p = s ↔ ∀ ⦃x⦄, x ∈ s → p x := by simp [finset.ext_iff] +lemma filter_eq_empty_iff : s.filter p = ∅ ↔ ∀ ⦃x⦄, x ∈ s → ¬ p x := by simp [finset.ext_iff] + +lemma filter_nonempty_iff {s : finset α} : (s.filter p).nonempty ↔ ∃ a ∈ s, p a := +by simp only [nonempty_iff_ne_empty, ne.def, filter_eq_empty_iff, not_not, not_forall] /-- If all elements of a `finset` satisfy the predicate `p`, `s.filter p` is `s`. -/ -@[simp] lemma filter_true_of_mem {s : finset α} (h : ∀ x ∈ s, p x) : s.filter p = s := -(filter_eq_self s).mpr h +@[simp] lemma filter_true_of_mem (h : ∀ x ∈ s, p x) : s.filter p = s := filter_eq_self.2 h /-- If all elements of a `finset` fail to satisfy the predicate `p`, `s.filter p` is `∅`. -/ -lemma filter_false_of_mem {s : finset α} (h : ∀ x ∈ s, ¬ p x) : s.filter p = ∅ := -eq_empty_of_forall_not_mem (by simpa) +@[simp] lemma filter_false_of_mem (h : ∀ x ∈ s, ¬ p x) : s.filter p = ∅ := filter_eq_empty_iff.2 h -lemma filter_eq_empty_iff (s : finset α) : - (s.filter p = ∅) ↔ ∀ x ∈ s, ¬ p x := -begin - refine ⟨_, filter_false_of_mem⟩, - intros hs, - injection hs with hs', - rwa filter_eq_nil at hs' -end - -lemma filter_nonempty_iff {s : finset α} : (s.filter p).nonempty ↔ ∃ a ∈ s, p a := -by simp only [nonempty_iff_ne_empty, ne.def, filter_eq_empty_iff, not_not, not_forall] +@[simp] lemma filter_const (p : Prop) [decidable p] (s : finset α) : + s.filter (λ a, p) = if p then s else ∅ := +by split_ifs; simp [*] lemma filter_congr {s : finset α} (H : ∀ x ∈ s, p x ↔ q x) : filter p s = filter q s := eq_of_veq $ filter_congr H @@ -1944,11 +1955,6 @@ begin { intro x, simp, intros hx hx₂, refine ⟨or.resolve_left (h hx) hx₂, hx₂⟩ } end -/- We can simplify an application of filter where the decidability is inferred in "the wrong way" -/ -@[simp] lemma filter_congr_decidable {α} (s : finset α) (p : α → Prop) (h : decidable_pred p) - [decidable_pred p] : @filter α p h s = s.filter p := -by congr - section classical open_locale classical /-- The following instance allows us to write `{x ∈ s | p x}` for `finset.filter p s`. diff --git a/src/data/finset/lattice.lean b/src/data/finset/lattice.lean index 208cbac617b55..60a235f53593d 100644 --- a/src/data/finset/lattice.lean +++ b/src/data/finset/lattice.lean @@ -55,10 +55,6 @@ fold_map @[simp] lemma sup_singleton {b : β} : ({b} : finset β).sup f = f b := sup_singleton -lemma sup_union [decidable_eq β] : (s₁ ∪ s₂).sup f = s₁.sup f ⊔ s₂.sup f := -finset.induction_on s₁ (by rw [empty_union, sup_empty, bot_sup_eq]) $ λ a s has ih, -by rw [insert_union, sup_insert, sup_insert, ih, sup_assoc] - lemma sup_sup : s.sup (f ⊔ g) = s.sup f ⊔ s.sup g := begin refine finset.cons_induction_on s _ (λ b t _ h, _), @@ -91,6 +87,9 @@ lemma le_sup {b : β} (hb : b ∈ s) : f b ≤ s.sup f := finset.sup_le_iff.1 le lemma le_sup_of_le {b : β} (hb : b ∈ s) (h : a ≤ f b) : a ≤ s.sup f := h.trans $ le_sup hb +lemma sup_union [decidable_eq β] : (s₁ ∪ s₂).sup f = s₁.sup f ⊔ s₂.sup f := +eq_of_forall_ge_iff $ λ c, by simp [or_imp_distrib, forall_and_distrib] + @[simp] lemma sup_bUnion [decidable_eq β] (s : finset γ) (t : γ → finset β) : (s.bUnion t).sup f = s.sup (λ x, (t x).sup f) := eq_of_forall_ge_iff $ λ c, by simp [@forall_swap _ β] @@ -117,11 +116,7 @@ lemma sup_mono (h : s₁ ⊆ s₂) : s₁.sup f ≤ s₂.sup f := finset.sup_le protected lemma sup_comm (s : finset β) (t : finset γ) (f : β → γ → α) : s.sup (λ b, t.sup (f b)) = t.sup (λ c, s.sup (λ b, f b c)) := -begin - refine eq_of_forall_ge_iff (λ a, _), - simp_rw finset.sup_le_iff, - exact ⟨λ h c hc b hb, h b hb c hc, λ h b hb c hc, h c hc b hb⟩, -end +eq_of_forall_ge_iff $ λ a, by simpa using forall₂_swap @[simp] lemma sup_attach (s : finset β) (f : β → α) : s.attach.sup (λ x, f x) = s.sup f := (s.attach.sup_map (function.embedding.subtype _) f).symm.trans $ congr_arg _ attach_map_val @@ -129,10 +124,7 @@ end /-- See also `finset.product_bUnion`. -/ lemma sup_product_left (s : finset β) (t : finset γ) (f : β × γ → α) : (s ×ˢ t).sup f = s.sup (λ i, t.sup $ λ i', f ⟨i, i'⟩) := -begin - simp only [le_antisymm_iff, finset.sup_le_iff, mem_product, and_imp, prod.forall], - exact ⟨λ b c hb hc, (le_sup hb).trans' $ le_sup hc, λ b hb c hc, le_sup $ mem_product.2 ⟨hb, hc⟩⟩, -end +eq_of_forall_ge_iff $ λ a, by simp [@forall_swap _ γ] lemma sup_product_right (s : finset β) (t : finset γ) (f : β × γ → α) : (s ×ˢ t).sup f = t.sup (λ i', s.sup $ λ i, f ⟨i, i'⟩) := @@ -288,9 +280,6 @@ fold_map @[simp] lemma inf_singleton {b : β} : ({b} : finset β).inf f = f b := inf_singleton -lemma inf_union [decidable_eq β] : (s₁ ∪ s₂).inf f = s₁.inf f ⊓ s₂.inf f := -@sup_union αᵒᵈ _ _ _ _ _ _ _ - lemma inf_inf : s.inf (f ⊓ g) = s.inf f ⊓ s.inf g := @sup_sup αᵒᵈ _ _ _ _ _ _ @@ -301,6 +290,9 @@ by subst hs; exact finset.fold_congr hfg (f : F) (s : finset ι) (g : ι → α) : f (s.inf g) = s.inf (f ∘ g) := finset.cons_induction_on s (map_top f) $ λ i s _ h, by rw [inf_cons, inf_cons, map_inf, h] +lemma inf_union [decidable_eq β] : (s₁ ∪ s₂).inf f = s₁.inf f ⊓ s₂.inf f := +@sup_union αᵒᵈ _ _ _ _ _ _ _ + @[simp] lemma inf_bUnion [decidable_eq β] (s : finset γ) (t : γ → finset β) : (s.bUnion t).inf f = s.inf (λ x, (t x).inf f) := @sup_bUnion αᵒᵈ _ _ _ _ _ _ _ _ @@ -628,11 +620,28 @@ end @[simp] lemma sup'_le_iff {a : α} : s.sup' H f ≤ a ↔ ∀ b ∈ s, f b ≤ a := iff.intro (λ h b hb, trans (le_sup' f hb) h) (sup'_le H f) +lemma sup'_union [decidable_eq β] {s₁ s₂ : finset β} (h₁ : s₁.nonempty) (h₂ : s₂.nonempty) + (f : β → α) : + (s₁ ∪ s₂).sup' (h₁.mono $ subset_union_left _ _) f = s₁.sup' h₁ f ⊔ s₂.sup' h₂ f := +eq_of_forall_ge_iff $ λ a, by simp [or_imp_distrib, forall_and_distrib] + lemma sup'_bUnion [decidable_eq β] {s : finset γ} (Hs : s.nonempty) {t : γ → finset β} (Ht : ∀ b, (t b).nonempty) : (s.bUnion t).sup' (Hs.bUnion (λ b _, Ht b)) f = s.sup' Hs (λ b, (t b).sup' (Ht b) f) := eq_of_forall_ge_iff $ λ c, by simp [@forall_swap _ β] +protected lemma sup'_comm {t : finset γ} (hs : s.nonempty) (ht : t.nonempty) (f : β → γ → α) : + s.sup' hs (λ b, t.sup' ht (f b)) = t.sup' ht (λ c, s.sup' hs $ λ b, f b c) := +eq_of_forall_ge_iff $ λ a, by simpa using forall₂_swap + +lemma sup'_product_left {t : finset γ} (hs : s.nonempty) (ht : t.nonempty) (f : β × γ → α) : + (s ×ˢ t).sup' (hs.product ht) f = s.sup' hs (λ i, t.sup' ht $ λ i', f ⟨i, i'⟩) := +eq_of_forall_ge_iff $ λ a, by simp [@forall_swap _ γ] + +lemma sup'_product_right {t : finset γ} (hs : s.nonempty) (ht : t.nonempty) (f : β × γ → α) : + (s ×ˢ t).sup' (hs.product ht) f = t.sup' ht (λ i', s.sup' hs $ λ i, f ⟨i, i'⟩) := +by rw [sup'_product_left, finset.sup'_comm] + lemma comp_sup'_eq_sup'_comp [semilattice_sup γ] {s : finset β} (H : s.nonempty) {f : β → α} (g : α → γ) (g_sup : ∀ x y, g (x ⊔ y) = g x ⊔ g y) : g (s.sup' H f) = s.sup' H (g ∘ f) := @@ -675,6 +684,15 @@ begin simp only [sup'_le_iff, h₂] { contextual := tt } end +@[simp] lemma _root_.map_finset_sup' [semilattice_sup β] [sup_hom_class F α β] (f : F) + {s : finset ι} (hs) (g : ι → α) : f (s.sup' hs g) = s.sup' hs (f ∘ g) := +by refine hs.cons_induction _ _; intros; simp [*] + +@[simp] lemma sup'_image [decidable_eq β] {s : finset γ} {f : γ → β} (hs : (s.image f).nonempty) + (g : β → α) (hs': s.nonempty := (nonempty.image_iff _).1 hs) : + (s.image f).sup' hs g = s.sup' hs' (g ∘ f) := +by { rw ←with_bot.coe_eq_coe, simp only [coe_sup', sup_image, with_bot.coe_sup] } + @[simp] lemma sup'_map {s : finset γ} {f : γ ↪ β} (g : β → α) (hs : (s.map f).nonempty) (hs': s.nonempty := finset.map_nonempty.mp hs) : (s.map f).sup' hs g = s.sup' hs' (g ∘ f) := @@ -719,11 +737,28 @@ lemma inf'_le_of_le (hb : b ∈ s) (h : f b ≤ a) : s.inf' ⟨b, hb⟩ f ≤ a @[simp] lemma le_inf'_iff : a ≤ s.inf' H f ↔ ∀ b ∈ s, a ≤ f b := @sup'_le_iff αᵒᵈ _ _ _ H f _ +lemma inf'_union [decidable_eq β] {s₁ s₂ : finset β} (h₁ : s₁.nonempty) (h₂ : s₂.nonempty) + (f : β → α) : + (s₁ ∪ s₂).inf' (h₁.mono $ subset_union_left _ _) f = s₁.inf' h₁ f ⊓ s₂.inf' h₂ f := +@sup'_union αᵒᵈ _ _ _ _ _ h₁ h₂ _ + lemma inf'_bUnion [decidable_eq β] {s : finset γ} (Hs : s.nonempty) {t : γ → finset β} (Ht : ∀ b, (t b).nonempty) : (s.bUnion t).inf' (Hs.bUnion (λ b _, Ht b)) f = s.inf' Hs (λ b, (t b).inf' (Ht b) f) := @sup'_bUnion αᵒᵈ _ _ _ _ _ _ Hs _ Ht +protected lemma inf'_comm {t : finset γ} (hs : s.nonempty) (ht : t.nonempty) (f : β → γ → α) : + s.inf' hs (λ b, t.inf' ht (f b)) = t.inf' ht (λ c, s.inf' hs $ λ b, f b c) := +@finset.sup'_comm αᵒᵈ _ _ _ _ _ hs ht _ + +lemma inf'_product_left {t : finset γ} (hs : s.nonempty) (ht : t.nonempty) (f : β × γ → α) : + (s ×ˢ t).inf' (hs.product ht) f = s.inf' hs (λ i, t.inf' ht $ λ i', f ⟨i, i'⟩) := +@sup'_product_left αᵒᵈ _ _ _ _ _ hs ht _ + +lemma inf'_product_right {t : finset γ} (hs : s.nonempty) (ht : t.nonempty) (f : β × γ → α) : + (s ×ˢ t).inf' (hs.product ht) f = t.inf' ht (λ i', s.inf' hs $ λ i, f ⟨i, i'⟩) := +@sup'_product_right αᵒᵈ _ _ _ _ _ hs ht _ + lemma comp_inf'_eq_inf'_comp [semilattice_inf γ] {s : finset β} (H : s.nonempty) {f : β → α} (g : α → γ) (g_inf : ∀ x y, g (x ⊓ y) = g x ⊓ g y) : g (s.inf' H f) = s.inf' H (g ∘ f) := @@ -741,6 +776,15 @@ lemma inf'_mem (s : set α) (w : ∀ x y ∈ s, x ⊓ y ∈ s) s.inf' H f = t.inf' (h₁ ▸ H) g := @sup'_congr αᵒᵈ _ _ _ H _ _ _ h₁ h₂ +@[simp] lemma _root_.map_finset_inf' [semilattice_inf β] [inf_hom_class F α β] (f : F) + {s : finset ι} (hs) (g : ι → α) : f (s.inf' hs g) = s.inf' hs (f ∘ g) := +by refine hs.cons_induction _ _; intros; simp [*] + +@[simp] lemma inf'_image [decidable_eq β] {s : finset γ} {f : γ → β} (hs : (s.image f).nonempty) + (g : β → α) (hs': s.nonempty := (nonempty.image_iff _).1 hs) : + (s.image f).inf' hs g = s.inf' hs' (g ∘ f) := +@sup'_image αᵒᵈ _ _ _ _ _ _ hs _ hs' + @[simp] lemma inf'_map {s : finset γ} {f : γ ↪ β} (g : β → α) (hs : (s.map f).nonempty) (hs': s.nonempty := finset.map_nonempty.mp hs) : (s.map f).inf' hs g = s.inf' hs' (g ∘ f) := @@ -829,6 +873,37 @@ end inf' @[simp] lemma of_dual_inf' [semilattice_sup α] {s : finset ι} (hs : s.nonempty) (f : ι → αᵒᵈ) : of_dual (s.inf' hs f) = s.sup' hs (of_dual ∘ f) := rfl +section distrib_lattice +variables [distrib_lattice α] {s : finset ι} {t : finset κ} (hs : s.nonempty) (ht : t.nonempty) + {f : ι → α} {g : κ → α} {a : α} + +lemma sup'_inf_distrib_left (f : ι → α) (a : α) : a ⊓ s.sup' hs f = s.sup' hs (λ i, a ⊓ f i) := +begin + refine hs.cons_induction (λ i, _) (λ i s hi hs ih, _) , + { simp }, + { simp_rw [sup'_cons hs, inf_sup_left], + rw ih } +end + +lemma sup'_inf_distrib_right (f : ι → α) (a : α) : s.sup' hs f ⊓ a = s.sup' hs (λ i, f i ⊓ a) := +by { rw [inf_comm, sup'_inf_distrib_left], simp_rw inf_comm } + +lemma sup'_inf_sup' (f : ι → α) (g : κ → α) : + s.sup' hs f ⊓ t.sup' ht g = (s ×ˢ t).sup' (hs.product ht) (λ i, f i.1 ⊓ g i.2) := +by simp_rw [finset.sup'_inf_distrib_right, finset.sup'_inf_distrib_left, sup'_product_left] + +lemma inf'_sup_distrib_left (f : ι → α) (a : α) : a ⊔ s.inf' hs f = s.inf' hs (λ i, a ⊔ f i) := +@sup'_inf_distrib_left αᵒᵈ _ _ _ hs _ _ + +lemma inf'_sup_distrib_right (f : ι → α) (a : α) : s.inf' hs f ⊔ a = s.inf' hs (λ i, f i ⊔ a) := +@sup'_inf_distrib_right αᵒᵈ _ _ _ hs _ _ + +lemma inf'_sup_inf' (f : ι → α) (g : κ → α) : + s.inf' hs f ⊔ t.inf' ht g = (s ×ˢ t).inf' (hs.product ht) (λ i, f i.1 ⊔ g i.2) := +@sup'_inf_sup' αᵒᵈ _ _ _ _ _ hs ht _ _ + +end distrib_lattice + section linear_order variables [linear_order α] {s : finset ι} (H : s.nonempty) {f : ι → α} {a : α} diff --git a/src/data/finset/locally_finite.lean b/src/data/finset/locally_finite.lean index f8b6e5b429cc9..7864e60893459 100644 --- a/src/data/finset/locally_finite.lean +++ b/src/data/finset/locally_finite.lean @@ -199,16 +199,16 @@ end lemma Icc_filter_lt_of_lt_right {a b c : α} [decidable_pred (< c)] (h : b < c) : (Icc a b).filter (< c) = Icc a b := -(finset.filter_eq_self _).2 (λ x hx, lt_of_le_of_lt (mem_Icc.1 hx).2 h) +filter_true_of_mem $ λ x hx, (mem_Icc.1 hx).2.trans_lt h lemma Ioc_filter_lt_of_lt_right {a b c : α} [decidable_pred (< c)] (h : b < c) : (Ioc a b).filter (< c) = Ioc a b := -(finset.filter_eq_self _).2 (λ x hx, lt_of_le_of_lt (mem_Ioc.1 hx).2 h) +filter_true_of_mem $ λ x hx, (mem_Ioc.1 hx).2.trans_lt h lemma Iic_filter_lt_of_lt_right {α} [preorder α] [locally_finite_order_bot α] {a c : α} [decidable_pred (< c)] (h : a < c) : (Iic a).filter (< c) = Iic a := -(finset.filter_eq_self _).2 (λ x hx, lt_of_le_of_lt (mem_Iic.1 hx) h) +filter_true_of_mem $ λ x hx, (mem_Iic.1 hx).trans_lt h variables (a b) [fintype α] diff --git a/src/number_theory/legendre_symbol/gauss_eisenstein_lemmas.lean b/src/number_theory/legendre_symbol/gauss_eisenstein_lemmas.lean index 344697e529726..171dcc9523d6f 100644 --- a/src/number_theory/legendre_symbol/gauss_eisenstein_lemmas.lean +++ b/src/number_theory/legendre_symbol/gauss_eisenstein_lemmas.lean @@ -90,7 +90,7 @@ calc (a ^ (p / 2) * (p / 2)! : zmod p) = (λ _ _ _ _ _ _, id) (λ b h _, ⟨b, by simp [-not_le, *] at *⟩) (by intros; split_ifs at *; simp * at *), - by rw [prod_mul_distrib, this]; simp + by rw [prod_mul_distrib, this, prod_const] ... = (-1)^((Ico 1 (p / 2).succ).filter (λ x : ℕ, ¬(a * x : zmod p).val ≤ p / 2)).card * (p / 2)! : by rw [← prod_nat_cast, finset.prod_eq_multiset_prod, From 8818fdefc78642a7e6afcd20be5c184f3c7d9699 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 12 Sep 2023 13:12:04 +0000 Subject: [PATCH 53/61] feat(combinatorics/set_family/ahlswede_zhang): Ahlswede-Zhang identity, part I (#18612) The Ahlswede-Zhang identity is a sharpening of the [Lubell-Yamamoto-Meshalkin identity](https://leanprover-community.github.io/mathlib_docs/combinatorics/set_family/lym.html#finset.sum_card_slice_div_choose_le_one), by expliciting the correction term. This PR defines `finset.truncated_sup`/`finset.truncated_inf`, whose cardinalities show up in the correction term. Co-authored-by: Vladimir Ivanov --- .../set_family/ahlswede_zhang.lean | 264 ++++++++++++++++++ src/data/finset/sups.lean | 44 +++ .../gauss_eisenstein_lemmas.lean | 3 +- 3 files changed, 310 insertions(+), 1 deletion(-) create mode 100644 src/combinatorics/set_family/ahlswede_zhang.lean diff --git a/src/combinatorics/set_family/ahlswede_zhang.lean b/src/combinatorics/set_family/ahlswede_zhang.lean new file mode 100644 index 0000000000000..276084c473ce1 --- /dev/null +++ b/src/combinatorics/set_family/ahlswede_zhang.lean @@ -0,0 +1,264 @@ +/- +Copyright (c) 2023 Yaël Dillies, Vladimir Ivanov. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies, Vladimir Ivanov +-/ +import data.finset.sups + +/-! +# The Ahlswede-Zhang identity + +This file proves the Ahlswede-Zhang identity, which is a nontrivial relation between the size of the +"truncated unions" of a set family. It sharpens the Lubell-Yamamoto-Meshalkin inequality +`finset.sum_card_slice_div_choose_le_one`, by making explicit the correction term. + +For a set family `𝒜`, the Ahlswede-Zhang identity states that the sum of +`|⋂ B ∈ 𝒜, B ⊆ A, B|/(|A| * n.choose |A|)` is exactly `1`. + +## Main declarations + +* `finset.truncated_sup`: `s.truncated_sup a` is the supremum of all `b ≤ a` in `𝒜` if there are + some, or `⊤` if there are none. +* `finset.truncated_inf` `s.truncated_inf a` is the infimum of all `b ≥ a` in `𝒜` if there are + some, or `⊥` if there are none. + +## References + +* [R. Ahlswede, Z. Zhang, *An identity in combinatorial extremal theory*](https://doi.org/10.1016/0001-8708(90)90023-G) +* [D. T. Tru, *An AZ-style identity and Bollobás deficiency*](https://doi.org/10.1016/j.jcta.2007.03.005) +-/ + +open_locale finset_family + +namespace finset +variables {α β : Type*} + +/-! ### Truncated supremum, truncated infimum -/ + +section semilattice_sup +variables [semilattice_sup α] [order_top α] [@decidable_rel α (≤)] + [semilattice_sup β] [bounded_order β] [@decidable_rel β (≤)] {s t : finset α} {a b : α} + +private lemma sup_aux : a ∈ lower_closure (s : set α) → (s.filter $ λ b, a ≤ b).nonempty := +λ ⟨b, hb, hab⟩, ⟨b, mem_filter.2 ⟨hb, hab⟩⟩ + +/-- The infimum of the elements of `s` less than `a` if there are some, otherwise `⊤`. -/ +def truncated_sup (s : finset α) (a : α) : α := +if h : a ∈ lower_closure (s : set α) then (s.filter $ λ b, a ≤ b).sup' (sup_aux h) id else ⊤ + +lemma truncated_sup_of_mem (h : a ∈ lower_closure (s : set α)) : + truncated_sup s a = (s.filter $ λ b, a ≤ b).sup' (sup_aux h) id := dif_pos h + +lemma truncated_sup_of_not_mem (h : a ∉ lower_closure (s : set α)) : truncated_sup s a = ⊤ := +dif_neg h + +@[simp] lemma truncated_sup_empty (a : α) : truncated_sup ∅ a = ⊤ := +truncated_sup_of_not_mem $ by simp + +lemma le_truncated_sup : a ≤ truncated_sup s a := +begin + rw truncated_sup, + split_ifs, + { obtain ⟨ℬ, hb, h⟩ := h, + exact h.trans (le_sup' _ $ mem_filter.2 ⟨hb, h⟩) }, + { exact le_top } +end + +lemma map_truncated_sup (e : α ≃o β) (s : finset α) (a : α) : + e (truncated_sup s a) = truncated_sup (s.map e.to_equiv.to_embedding) (e a) := +begin + have : e a ∈ lower_closure (s.map e.to_equiv.to_embedding : set β) + ↔ a ∈ lower_closure (s : set α), + { simp }, + simp_rw [truncated_sup, apply_dite e, map_finset_sup', map_top, this], + congr' with h, + simp only [filter_map, function.comp, equiv.coe_to_embedding, rel_iso.coe_fn_to_equiv, + order_iso.le_iff_le, id.def], + rw sup'_map, -- TODO: Why can't `simp` use `finset.sup'_map`? + simp only [equiv.coe_to_embedding, rel_iso.coe_fn_to_equiv], +end + +variables [decidable_eq α] + +private lemma lower_aux : + a ∈ lower_closure (↑(s ∪ t) : set α) ↔ + a ∈ lower_closure (s : set α) ∨ a ∈ lower_closure (t : set α) := +by rw [coe_union, lower_closure_union, lower_set.mem_sup_iff] + +lemma truncated_sup_union (hs : a ∈ lower_closure (s : set α)) + (ht : a ∈ lower_closure (t : set α)) : + truncated_sup (s ∪ t) a = truncated_sup s a ⊔ truncated_sup t a := +by simpa only [truncated_sup_of_mem, hs, ht, lower_aux.2 (or.inl hs), filter_union] + using sup'_union _ _ _ + +lemma truncated_sup_union_left (hs : a ∈ lower_closure (s : set α)) + (ht : a ∉ lower_closure (t : set α)) : + truncated_sup (s ∪ t) a = truncated_sup s a := +begin + simp only [mem_lower_closure, mem_coe, exists_prop, not_exists, not_and] at ht, + simp only [truncated_sup_of_mem, hs, filter_union, filter_false_of_mem ht, union_empty, + lower_aux.2 (or.inl hs), ht], +end + +lemma truncated_sup_union_right (hs : a ∉ lower_closure (s : set α)) + (ht : a ∈ lower_closure (t : set α)) : + truncated_sup (s ∪ t) a = truncated_sup t a := +by rw [union_comm, truncated_sup_union_left ht hs] + +lemma truncated_sup_union_of_not_mem (hs : a ∉ lower_closure (s : set α)) + (ht : a ∉ lower_closure (t : set α)) : + truncated_sup (s ∪ t) a = ⊤ := +truncated_sup_of_not_mem $ λ h, (lower_aux.1 h).elim hs ht + +end semilattice_sup + +section semilattice_inf +variables [semilattice_inf α] [bounded_order α] [@decidable_rel α (≤)] + [semilattice_inf β] [bounded_order β] [@decidable_rel β (≤)] {s t : finset α} {a : α} + +private lemma inf_aux : a ∈ upper_closure (s : set α) → (s.filter $ λ b, b ≤ a).nonempty := +λ ⟨b, hb, hab⟩, ⟨b, mem_filter.2 ⟨hb, hab⟩⟩ + +/-- The infimum of the elements of `s` less than `a` if there are some, otherwise `⊥`. -/ +def truncated_inf (s : finset α) (a : α) : α := +if h : a ∈ upper_closure (s : set α) then (s.filter $ λ b, b ≤ a).inf' (inf_aux h) id else ⊥ + +lemma truncated_inf_of_mem (h : a ∈ upper_closure (s : set α)) : + truncated_inf s a = (s.filter $ λ b, b ≤ a).inf' (inf_aux h) id := dif_pos h + +lemma truncated_inf_of_not_mem (h : a ∉ upper_closure (s : set α)) : truncated_inf s a = ⊥ := +dif_neg h + +lemma truncated_inf_le (s : finset α) (a : α) : truncated_inf s a ≤ a := +begin + unfold truncated_inf, + split_ifs, + { obtain ⟨ℬ, hb, h⟩ := h, + exact (inf'_le _ $ mem_filter.2 ⟨hb, h⟩).trans h }, + { exact bot_le } +end + +@[simp] lemma truncated_inf_empty (a : α) : truncated_inf ∅ a = ⊥ := +truncated_inf_of_not_mem $ by simp + +lemma map_truncated_inf (e : α ≃o β) (s : finset α) (a : α) : + e (truncated_inf s a) = truncated_inf (s.map e.to_equiv.to_embedding) (e a) := +begin + have : e a ∈ upper_closure (s.map e.to_equiv.to_embedding : set β) + ↔ a ∈ upper_closure (s : set α), + { simp }, + simp_rw [truncated_inf, apply_dite e, map_finset_inf', map_bot, this], + congr' with h, + simp only [filter_map, function.comp, equiv.coe_to_embedding, rel_iso.coe_fn_to_equiv, + order_iso.le_iff_le, id.def], + rw inf'_map, -- TODO: Why can't `simp` use `finset.inf'_map`? + simp only [equiv.coe_to_embedding, rel_iso.coe_fn_to_equiv], +end + +variables [decidable_eq α] + +private lemma upper_aux : + a ∈ upper_closure (↑(s ∪ t) : set α) ↔ + a ∈ upper_closure (s : set α) ∨ a ∈ upper_closure (t : set α) := +by rw [coe_union, upper_closure_union, upper_set.mem_inf_iff] + +lemma truncated_inf_union (hs : a ∈ upper_closure (s : set α)) + (ht : a ∈ upper_closure (t : set α)) : + truncated_inf (s ∪ t) a = truncated_inf s a ⊓ truncated_inf t a := +by simpa only [truncated_inf_of_mem, hs, ht, upper_aux.2 (or.inl hs), filter_union] + using inf'_union _ _ _ + +lemma truncated_inf_union_left (hs : a ∈ upper_closure (s : set α)) + (ht : a ∉ upper_closure (t : set α)) : + truncated_inf (s ∪ t) a = truncated_inf s a := +begin + simp only [mem_upper_closure, mem_coe, exists_prop, not_exists, not_and] at ht, + simp only [truncated_inf_of_mem, hs, filter_union, filter_false_of_mem ht, union_empty, + upper_aux.2 (or.inl hs), ht], +end + +lemma truncated_inf_union_right (hs : a ∉ upper_closure (s : set α)) + (ht : a ∈ upper_closure (t : set α)) : + truncated_inf (s ∪ t) a = truncated_inf t a := +by rw [union_comm, truncated_inf_union_left ht hs] + +lemma truncated_inf_union_of_not_mem (hs : a ∉ upper_closure (s : set α)) + (ht : a ∉ upper_closure (t : set α)) : + truncated_inf (s ∪ t) a = ⊥ := +truncated_inf_of_not_mem $ by { rw [coe_union, upper_closure_union], exact λ h, h.elim hs ht } + +end semilattice_inf + +section distrib_lattice +variables [distrib_lattice α] [bounded_order α] [decidable_eq α] [@decidable_rel α (≤)] + {s t : finset α} {a : α} + +private lemma infs_aux + : a ∈ lower_closure (↑(s ⊼ t) : set α) ↔ a ∈ lower_closure (s : set α) ⊓ lower_closure t := +by rw [coe_infs, lower_closure_infs, lower_set.mem_inf_iff] + +private lemma sups_aux : + a ∈ upper_closure (↑(s ⊻ t) : set α) ↔ a ∈ upper_closure (s : set α) ⊔ upper_closure t := +by rw [coe_sups, upper_closure_sups, upper_set.mem_sup_iff] + +lemma truncated_sup_infs (hs : a ∈ lower_closure (s : set α)) (ht : a ∈ lower_closure (t : set α)) : + truncated_sup (s ⊼ t) a = truncated_sup s a ⊓ truncated_sup t a := +begin + simp only [truncated_sup_of_mem, hs, ht, infs_aux.2 ⟨hs, ht⟩, sup'_inf_sup', filter_infs_ge], + simp_rw ←image_inf_product, + rw sup'_image, + refl, +end + +lemma truncated_inf_sups (hs : a ∈ upper_closure (s : set α)) (ht : a ∈ upper_closure (t : set α)) : + truncated_inf (s ⊻ t) a = truncated_inf s a ⊔ truncated_inf t a := +begin + simp only [truncated_inf_of_mem, hs, ht, sups_aux.2 ⟨hs, ht⟩, inf'_sup_inf', filter_sups_le], + simp_rw ←image_sup_product, + rw inf'_image, + refl, +end + +lemma truncated_sup_infs_of_not_mem (ha : a ∉ lower_closure (s : set α) ⊓ lower_closure t) : + truncated_sup (s ⊼ t) a = ⊤ := +truncated_sup_of_not_mem $ by rwa [coe_infs, lower_closure_infs] + +lemma truncated_inf_sups_of_not_mem (ha : a ∉ upper_closure (s : set α) ⊔ upper_closure t) : + truncated_inf (s ⊻ t) a = ⊥ := +truncated_inf_of_not_mem $ by rwa [coe_sups, upper_closure_sups] + +end distrib_lattice + +section boolean_algebra +variables [boolean_algebra α] [@decidable_rel α (≤)] {s : finset α} {a : α} + +@[simp] lemma compl_truncated_sup (s : finset α) (a : α) : + (truncated_sup s a)ᶜ = truncated_inf (s.map ⟨compl, compl_injective⟩) aᶜ := +map_truncated_sup (order_iso.compl α) _ _ + +@[simp] lemma compl_truncated_inf (s : finset α) (a : α) : + (truncated_inf s a)ᶜ = truncated_sup (s.map ⟨compl, compl_injective⟩) aᶜ := +map_truncated_inf (order_iso.compl α) _ _ + +end boolean_algebra + +variables [decidable_eq α] [fintype α] + +lemma card_truncated_sup_union_add_card_truncated_sup_infs (𝒜 ℬ : finset (finset α)) + (s : finset α) : + (truncated_sup (𝒜 ∪ ℬ) s).card + (truncated_sup (𝒜 ⊼ ℬ) s).card = + (truncated_sup 𝒜 s).card + (truncated_sup ℬ s).card := +begin + by_cases h𝒜 : s ∈ lower_closure (𝒜 : set $ finset α); + by_cases hℬ : s ∈ lower_closure (ℬ : set $ finset α), + { rw [truncated_sup_union h𝒜 hℬ, truncated_sup_infs h𝒜 hℬ], + exact card_union_add_card_inter _ _ }, + { rw [truncated_sup_union_left h𝒜 hℬ, truncated_sup_of_not_mem hℬ, + truncated_sup_infs_of_not_mem (λ h, hℬ h.2)] }, + { rw [truncated_sup_union_right h𝒜 hℬ, truncated_sup_of_not_mem h𝒜, + truncated_sup_infs_of_not_mem (λ h, h𝒜 h.1), add_comm] }, + { rw [truncated_sup_of_not_mem h𝒜, truncated_sup_of_not_mem hℬ, + truncated_sup_union_of_not_mem h𝒜 hℬ, truncated_sup_infs_of_not_mem (λ h, h𝒜 h.1)] } +end + +end finset diff --git a/src/data/finset/sups.lean b/src/data/finset/sups.lean index 0c2c69a056427..b94a8777cedc4 100644 --- a/src/data/finset/sups.lean +++ b/src/data/finset/sups.lean @@ -36,6 +36,20 @@ We define the following notation in locale `finset_family`: open function open_locale set_family +-- TODO: Is there a better spot for those two instances? +namespace finset +variables {α : Type*} [preorder α] {s t : set α} {a : α} + +instance decidable_pred_mem_upper_closure (s : finset α) [@decidable_rel α (≤)] : + decidable_pred (∈ upper_closure (s : set α)) := +λ _, finset.decidable_dexists_finset + +instance decidable_pred_mem_lower_closure (s : finset α) [@decidable_rel α (≤)] : + decidable_pred (∈ lower_closure (s : set α)) := +λ _, finset.decidable_dexists_finset + +end finset + variables {α : Type*} [decidable_eq α] namespace finset @@ -116,6 +130,21 @@ lemma sups_right_comm : (s ⊻ t) ⊻ u = (s ⊻ u) ⊻ t := image₂_right_comm lemma sups_sups_sups_comm : (s ⊻ t) ⊻ (u ⊻ v) = (s ⊻ u) ⊻ (t ⊻ v) := image₂_image₂_image₂_comm sup_sup_sup_comm +variables [@decidable_rel α (≤)] + +lemma filter_sups_le (s t : finset α) (a : α) : + (s ⊻ t).filter (λ b, b ≤ a) = s.filter (λ b, b ≤ a) ⊻ t.filter (λ b, b ≤ a) := +begin + ext b, + simp only [mem_filter, mem_sups], + split, + { rintro ⟨⟨b, hb, c, hc, rfl⟩, ha⟩, + rw sup_le_iff at ha, + exact ⟨_, ⟨hb, ha.1⟩, _, ⟨hc, ha.2⟩, rfl⟩ }, + { rintro ⟨b, hb, c, hc, _, rfl⟩, + exact ⟨⟨_, hb.1, _, hc.1, rfl⟩, sup_le hb.2 hc.2⟩ } +end + end sups section infs @@ -195,6 +224,21 @@ lemma infs_right_comm : (s ⊼ t) ⊼ u = (s ⊼ u) ⊼ t := image₂_right_comm lemma infs_infs_infs_comm : (s ⊼ t) ⊼ (u ⊼ v) = (s ⊼ u) ⊼ (t ⊼ v) := image₂_image₂_image₂_comm inf_inf_inf_comm +variables [@decidable_rel α (≤)] + +lemma filter_infs_ge (s t : finset α) (a : α) : + (s ⊼ t).filter (λ b, a ≤ b) = s.filter (λ b, a ≤ b) ⊼ t.filter (λ b, a ≤ b) := +begin + ext b, + simp only [mem_filter, mem_infs], + split, + { rintro ⟨⟨b, hb, c, hc, rfl⟩, ha⟩, + rw le_inf_iff at ha, + exact ⟨_, ⟨hb, ha.1⟩, _, ⟨hc, ha.2⟩, rfl⟩ }, + { rintro ⟨b, hb, c, hc, _, rfl⟩, + exact ⟨⟨_, hb.1, _, hc.1, rfl⟩, le_inf hb.2 hc.2⟩ } +end + end infs open_locale finset_family diff --git a/src/number_theory/legendre_symbol/gauss_eisenstein_lemmas.lean b/src/number_theory/legendre_symbol/gauss_eisenstein_lemmas.lean index 171dcc9523d6f..293f6f86796bb 100644 --- a/src/number_theory/legendre_symbol/gauss_eisenstein_lemmas.lean +++ b/src/number_theory/legendre_symbol/gauss_eisenstein_lemmas.lean @@ -114,13 +114,14 @@ lemma gauss_lemma {p : ℕ} [fact p.prime] {a : ℤ} (hp : p ≠ 2) (ha0 : (a : (λ x : ℕ, p / 2 < (a * x : zmod p).val)).card := begin haveI hp' : fact (p % 2 = 1) := ⟨nat.prime.mod_two_eq_one_iff_ne_two.mpr hp⟩, + haveI : fact (2 < p) := ⟨hp.lt_of_le' (fact.out p.prime).two_le⟩, have : (legendre_sym p a : zmod p) = (((-1)^((Ico 1 (p / 2).succ).filter (λ x : ℕ, p / 2 < (a * x : zmod p).val)).card : ℤ) : zmod p) := by { rw [legendre_sym.eq_pow, gauss_lemma_aux p ha0]; simp }, cases legendre_sym.eq_one_or_neg_one p ha0; cases neg_one_pow_eq_or ℤ ((Ico 1 (p / 2).succ).filter (λ x : ℕ, p / 2 < (a * x : zmod p).val)).card; - simp [*, ne_neg_self p one_ne_zero, (ne_neg_self p one_ne_zero).symm] at * + simp only [*, neg_one_ne_one, neg_one_ne_one.symm, algebra_map.coe_one, int.cast_neg] at *, end private lemma eisenstein_lemma_aux₁ (p : ℕ) [fact p.prime] [hp2 : fact (p % 2 = 1)] From 001ffdc42920050657fd45bd2b8bfbec8eaaeb29 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 12 Sep 2023 23:46:39 +0000 Subject: [PATCH 54/61] feat(probability/independence): Independence of singletons (#18506) Characterisation of independence in terms of measure distributing over finite intersections, and lemmas connecting the different concepts of independence. Also add supporting `measurable_space` and `set.preimage` lemmas --- src/data/set/basic.lean | 3 + src/data/set/image.lean | 10 +++ src/measure_theory/measurable_space.lean | 93 ++++++++++++++++++++ src/measure_theory/measurable_space_def.lean | 6 +- src/probability/independence/basic.lean | 39 +++++++- 5 files changed, 146 insertions(+), 5 deletions(-) diff --git a/src/data/set/basic.lean b/src/data/set/basic.lean index 9ac1590fab649..065bb077cb029 100644 --- a/src/data/set/basic.lean +++ b/src/data/set/basic.lean @@ -1906,3 +1906,6 @@ lemma subset_right_of_subset_union (h : s ⊆ t ∪ u) (hab : disjoint s t) : s hab.left_le_of_le_sup_left h end disjoint + +@[simp] lemma Prop.compl_singleton (p : Prop) : ({p}ᶜ : set Prop) = {¬ p} := +ext $ λ q, by simpa [@iff.comm q] using not_iff diff --git a/src/data/set/image.lean b/src/data/set/image.lean index 4e942d43a7266..9b00e12e92842 100644 --- a/src/data/set/image.lean +++ b/src/data/set/image.lean @@ -71,6 +71,9 @@ theorem subset_preimage_univ {s : set α} : s ⊆ f ⁻¹' univ := subset_univ _ @[simp] theorem preimage_diff (f : α → β) (s t : set β) : f ⁻¹' (s \ t) = f ⁻¹' s \ f ⁻¹' t := rfl +@[simp] lemma preimage_symm_diff (f : α → β) (s t : set β) : + f ⁻¹' (s ∆ t) = (f ⁻¹' s) ∆ (f ⁻¹' t) := rfl + @[simp] theorem preimage_ite (f : α → β) (s t₁ t₂ : set β) : f ⁻¹' (s.ite t₁ t₂) = (f ⁻¹' s).ite (f ⁻¹' t₁) (f ⁻¹' t₂) := rfl @@ -120,6 +123,10 @@ lemma nonempty_of_nonempty_preimage {s : set β} {f : α → β} (hf : (f ⁻¹' s.nonempty := let ⟨x, hx⟩ := hf in ⟨f x, hx⟩ +@[simp] lemma preimage_singleton_true (p : α → Prop) : p ⁻¹' {true} = {a | p a} := by { ext, simp } +@[simp] lemma preimage_singleton_false (p : α → Prop) : p ⁻¹' {false} = {a | ¬ p a} := +by { ext, simp } + lemma preimage_subtype_coe_eq_compl {α : Type*} {s u v : set α} (hsuv : s ⊆ u ∪ v) (H : s ∩ (u ∩ v) = ∅) : (coe : s → α) ⁻¹' u = (coe ⁻¹' v)ᶜ := begin @@ -1003,6 +1010,9 @@ lemma left_inverse.preimage_preimage {g : β → α} (h : left_inverse g f) (s : f ⁻¹' (g ⁻¹' s) = s := by rw [← preimage_comp, h.comp_eq_id, preimage_id] +protected lemma involutive.preimage {f : α → α} (hf : involutive f) : involutive (preimage f) := +hf.right_inverse.preimage_preimage + end function namespace equiv_like diff --git a/src/measure_theory/measurable_space.lean b/src/measure_theory/measurable_space.lean index 98a586b1589a6..74b43d1b7bce7 100644 --- a/src/measure_theory/measurable_space.lean +++ b/src/measure_theory/measurable_space.lean @@ -7,6 +7,7 @@ Authors: Johannes Hölzl, Mario Carneiro import data.prod.tprod import group_theory.coset import logic.equiv.fin +import logic.lemmas import measure_theory.measurable_space_def import order.filter.small_sets import order.liminf_limsup @@ -136,6 +137,12 @@ lemma le_map_comap : m ≤ (m.comap g).map g := (gc_comap_map g).le_u_l _ end functors +@[simp] lemma map_const {m} (b : β) : measurable_space.map (λ a : α, b) m = ⊤ := +eq_top_iff.2 $ by { rintro s hs, by_cases b ∈ s; change measurable_set (preimage _ _); simp [*] } + +@[simp] lemma comap_const {m} (b : β) : measurable_space.comap (λ a : α, b) m = ⊥ := +eq_bot_iff.2 $ by { rintro _ ⟨s, -, rfl⟩, by_cases b ∈ s; simp [*] } + lemma comap_generate_from {f : α → β} {s : set (set β)} : (generate_from s).comap f = generate_from (preimage f '' s) := le_antisymm @@ -291,6 +298,7 @@ section constructions instance : measurable_space empty := ⊤ instance : measurable_space punit := ⊤ -- this also works for `unit` instance : measurable_space bool := ⊤ +instance Prop.measurable_space : measurable_space Prop := ⊤ instance : measurable_space ℕ := ⊤ instance : measurable_space ℤ := ⊤ instance : measurable_space ℚ := ⊤ @@ -298,6 +306,7 @@ instance : measurable_space ℚ := ⊤ instance : measurable_singleton_class empty := ⟨λ _, trivial⟩ instance : measurable_singleton_class punit := ⟨λ _, trivial⟩ instance : measurable_singleton_class bool := ⟨λ _, trivial⟩ +instance Prop.measurable_singleton_class : measurable_singleton_class Prop := ⟨λ _, trivial⟩ instance : measurable_singleton_class ℕ := ⟨λ _, trivial⟩ instance : measurable_singleton_class ℤ := ⟨λ _, trivial⟩ instance : measurable_singleton_class ℚ := ⟨λ _, trivial⟩ @@ -340,6 +349,15 @@ begin exact h, end +lemma measurable_to_prop {f : α → Prop} (h : measurable_set (f⁻¹' {true})) : measurable f := +begin + refine measurable_to_countable' (λ x, _), + by_cases hx : x, + { simpa [hx] using h }, + { simpa only [hx, ←preimage_compl, Prop.compl_singleton, not_true, preimage_singleton_false] + using h.compl } +end + lemma measurable_find_greatest' {p : α → ℕ → Prop} [∀ x, decidable_pred (p x)] {N : ℕ} (hN : ∀ k ≤ N, measurable_set {x | nat.find_greatest (p x) N = k}) : measurable (λ x, nat.find_greatest (p x) N) := @@ -860,8 +878,38 @@ end sum instance {α} {β : α → Type*} [m : Πa, measurable_space (β a)] : measurable_space (sigma β) := ⨅a, (m a).map (sigma.mk a) +section prop +variables {p : α → Prop} + +variables [measurable_space α] + +@[simp] lemma measurable_set_set_of : measurable_set {a | p a} ↔ measurable p := +⟨λ h, measurable_to_prop $ by simpa only [preimage_singleton_true], λ h, + by simpa using h (measurable_set_singleton true)⟩ + +@[simp] lemma measurable_mem : measurable (∈ s) ↔ measurable_set s := measurable_set_set_of.symm + +alias measurable_set_set_of ↔ _ measurable.set_of +alias measurable_mem ↔ _ measurable_set.mem + +end prop end constructions +namespace measurable_space + +/-- The sigma-algebra generated by a single set `s` is `{∅, s, sᶜ, univ}`. -/ +@[simp] lemma generate_from_singleton (s : set α) : + generate_from {s} = measurable_space.comap (∈ s) ⊤ := +begin + classical, + letI : measurable_space α := generate_from {s}, + refine le_antisymm (generate_from_le $ λ t ht, ⟨{true}, trivial, by simp [ht.symm]⟩) _, + rintro _ ⟨u, -, rfl⟩, + exact (show measurable_set s, from generate_measurable.basic _ $ mem_singleton s).mem trivial, +end + +end measurable_space + /-- A map `f : α → β` is called a *measurable embedding* if it is injective, measurable, and sends measurable sets to measurable sets. The latter assumption can be replaced with “`f` has measurable inverse `g : range f → α`”, see `measurable_embedding.measurable_range_splitting`, @@ -1052,6 +1100,9 @@ e.to_equiv.image_eq_preimage s measurable_set (e '' s) ↔ measurable_set s := by rw [image_eq_preimage, measurable_set_preimage] +@[simp] lemma map_eq (e : α ≃ᵐ β) : measurable_space.map e ‹_› = ‹_› := +e.measurable.le_map.antisymm' $ λ s, e.measurable_set_preimage.1 + /-- A measurable equivalence is a measurable embedding. -/ protected lemma measurable_embedding (e : α ≃ᵐ β) : measurable_embedding e := { injective := e.injective, @@ -1277,12 +1328,41 @@ def sum_compl {s : set α} [decidable_pred s] (hs : measurable_set s) : s ⊕ (s measurable_to_fun := by {apply measurable.sum_elim; exact measurable_subtype_coe}, measurable_inv_fun := measurable.dite measurable_inl measurable_inr hs } +/-- Convert a measurable involutive function `f` to a measurable permutation with +`to_fun = inv_fun = f`. See also `function.involutive.to_perm`. -/ +@[simps to_equiv] def of_involutive (f : α → α) (hf : involutive f) (hf' : measurable f) : α ≃ᵐ α := +{ measurable_to_fun := hf', + measurable_inv_fun := hf', + ..hf.to_perm _ } + +@[simp] lemma of_involutive_apply (f : α → α) (hf : involutive f) (hf' : measurable f) (a : α) : + of_involutive f hf hf' a = f a := rfl + +@[simp] lemma of_involutive_symm (f : α → α) (hf : involutive f) (hf' : measurable f) : + (of_involutive f hf hf').symm = of_involutive f hf hf' := rfl + end measurable_equiv namespace measurable_embedding variables [measurable_space α] [measurable_space β] [measurable_space γ] {f : α → β} {g : β → α} +@[simp] lemma comap_eq (hf : measurable_embedding f) : measurable_space.comap f ‹_› = ‹_› := +hf.measurable.comap_le.antisymm $ λ s h, + ⟨_, hf.measurable_set_image' h, hf.injective.preimage_image _⟩ + +lemma iff_comap_eq : + measurable_embedding f ↔ + injective f ∧ measurable_space.comap f ‹_› = ‹_› ∧ measurable_set (range f) := +⟨λ hf, ⟨hf.injective, hf.comap_eq, hf.measurable_set_range⟩, λ hf, + { injective := hf.1, + measurable := by { rw ←hf.2.1, exact comap_measurable f }, + measurable_set_image' := begin + rw ←hf.2.1, + rintro _ ⟨s, hs, rfl⟩, + simpa only [image_preimage_eq_inter_range] using hs.inter hf.2.2, + end }⟩ + /-- A set is equivalent to its image under a function `f` as measurable spaces, if `f` is a measurable embedding -/ noncomputable def equiv_image (s : set α) (hf : measurable_embedding f) : @@ -1377,6 +1457,19 @@ end end measurable_embedding +lemma measurable_space.comap_compl {m' : measurable_space β} [boolean_algebra β] + (h : measurable (compl : β → β)) (f : α → β) : + measurable_space.comap (λ a, (f a)ᶜ) infer_instance = measurable_space.comap f infer_instance := +begin + rw ←measurable_space.comap_comp, + congr', + exact (measurable_equiv.of_involutive _ compl_involutive h).measurable_embedding.comap_eq, +end + +@[simp] lemma measurable_space.comap_not (p : α → Prop) : + measurable_space.comap (λ a, ¬ p a) infer_instance = measurable_space.comap p infer_instance := +measurable_space.comap_compl (λ _ _, trivial) _ + section countably_generated namespace measurable_space diff --git a/src/measure_theory/measurable_space_def.lean b/src/measure_theory/measurable_space_def.lean index 642af17ecd600..6fb980dffc3f7 100644 --- a/src/measure_theory/measurable_space_def.lean +++ b/src/measure_theory/measurable_space_def.lean @@ -196,7 +196,7 @@ by { cases i, exacts [h₂, h₁] } measurable_set (disjointed f n) := disjointed_rec (λ t i ht, measurable_set.diff ht $ h _) (h n) -@[simp] lemma measurable_set.const (p : Prop) : measurable_set {a : α | p} := +lemma measurable_set.const (p : Prop) : measurable_set {a : α | p} := by { by_cases p; simp [h, measurable_set.empty]; apply measurable_set.univ } /-- Every set has a measurable superset. Declare this as local instance as needed. -/ @@ -377,10 +377,10 @@ begin { exact measurable_set_generate_from ht, }, end -@[simp] lemma generate_from_singleton_empty : generate_from {∅} = (⊥ : measurable_space α) := +lemma generate_from_singleton_empty : generate_from {∅} = (⊥ : measurable_space α) := by { rw [eq_bot_iff, generate_from_le_iff], simp, } -@[simp] lemma generate_from_singleton_univ : generate_from {set.univ} = (⊥ : measurable_space α) := +lemma generate_from_singleton_univ : generate_from {set.univ} = (⊥ : measurable_space α) := by { rw [eq_bot_iff, generate_from_le_iff], simp, } lemma measurable_set_bot_iff {s : set α} : @measurable_set α ⊥ s ↔ (s = ∅ ∨ s = univ) := diff --git a/src/probability/independence/basic.lean b/src/probability/independence/basic.lean index 7223f9b23c01e..03218d5783112 100644 --- a/src/probability/independence/basic.lean +++ b/src/probability/independence/basic.lean @@ -64,7 +64,7 @@ when defining `μ` in the example above, the measurable space used is the last o Part A, Chapter 4. -/ -open measure_theory measurable_space +open measure_theory measurable_space set open_locale big_operators measure_theory ennreal namespace probability_theory @@ -577,7 +577,8 @@ We prove the following equivalences on `indep_set`, for measurable sets `s, t`. * `indep_set s t μ ↔ indep_sets {s} {t} μ`. -/ -variables {s t : set Ω} (S T : set (set Ω)) +variables {s t : set Ω} (S T : set (set Ω)) {π : ι → set (set Ω)} {f : ι → set Ω} + {m0 : measurable_space Ω} {μ : measure Ω} lemma indep_set_iff_indep_sets_singleton {m0 : measurable_space Ω} (hs_meas : measurable_set s) (ht_meas : measurable_set t) @@ -624,6 +625,40 @@ lemma indep_iff_forall_indep_set (m₁ m₂ : measurable_space Ω) {m0 : measura λ h s t hs ht, h s t hs ht s t (measurable_set_generate_from (set.mem_singleton s)) (measurable_set_generate_from (set.mem_singleton t))⟩ +lemma Indep_sets.meas_Inter [fintype ι] (h : Indep_sets π μ) (hf : ∀ i, f i ∈ π i) : + μ (⋂ i, f i) = ∏ i, μ (f i) := +by simp [← h _ (λ i _, hf _)] + +lemma Indep_comap_mem_iff : Indep (λ i, measurable_space.comap (∈ f i) ⊤) μ ↔ Indep_set f μ := +by { simp_rw ←generate_from_singleton, refl } + +alias Indep_comap_mem_iff ↔ _ Indep_set.Indep_comap_mem + +lemma Indep_sets_singleton_iff : + Indep_sets (λ i, {f i}) μ ↔ ∀ t, μ (⋂ i ∈ t, f i) = ∏ i in t, μ (f i) := +forall_congr $ λ t, + ⟨λ h, h $ λ _ _, mem_singleton _, + λ h f hf, begin + refine eq.trans _ (h.trans $ finset.prod_congr rfl $ λ i hi, congr_arg _ $ (hf i hi).symm), + rw Inter₂_congr hf, + end⟩ + +variables [is_probability_measure μ] + +lemma Indep_set_iff_Indep_sets_singleton (hf : ∀ i, measurable_set (f i)) : + Indep_set f μ ↔ Indep_sets (λ i, {f i}) μ := +⟨Indep.Indep_sets $ λ _, rfl, Indep_sets.Indep _ + (λ i, generate_from_le $ by { rintro t (rfl : t = _), exact hf _}) _ + (λ _, is_pi_system.singleton _) $ λ _, rfl⟩ + +lemma Indep_set_iff_measure_Inter_eq_prod (hf : ∀ i, measurable_set (f i)) : + Indep_set f μ ↔ ∀ s, μ (⋂ i ∈ s, f i) = ∏ i in s, μ (f i) := +(Indep_set_iff_Indep_sets_singleton hf).trans Indep_sets_singleton_iff + +lemma Indep_sets.Indep_set_of_mem (hfπ : ∀ i, f i ∈ π i) (hf : ∀ i, measurable_set (f i)) + (hπ : Indep_sets π μ) : Indep_set f μ := +(Indep_set_iff_measure_Inter_eq_prod hf).2 $ λ t, hπ _ $ λ i _, hfπ _ + end indep_set section indep_fun From ce64cd319bb6b3e82f31c2d38e79080d377be451 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 18 Sep 2023 15:31:36 +0000 Subject: [PATCH 55/61] feat(topology/algebra/order/liminf_limsup): Eventual boundedness of neighborhoods (#18629) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Generalise `bounded_le_nhds`/`bounded_ge_nhds` using two ad hoc typeclasses. The goal here is to circumvent the fact that the product of order topologies is *not* an order topology, and apply those lemmas to `ℝⁿ`. --- src/order/filter/pi.lean | 9 ++ src/topology/algebra/order/liminf_limsup.lean | 144 ++++++++++++++---- 2 files changed, 123 insertions(+), 30 deletions(-) diff --git a/src/order/filter/pi.lean b/src/order/filter/pi.lean index d069ed0dcd512..1d9167060ab6f 100644 --- a/src/order/filter/pi.lean +++ b/src/order/filter/pi.lean @@ -28,6 +28,7 @@ open_locale classical filter namespace filter variables {ι : Type*} {α : ι → Type*} {f f₁ f₂ : Π i, filter (α i)} {s : Π i, set (α i)} + {p : Π i, α i → Prop} section pi @@ -93,6 +94,14 @@ end I.pi s ∈ pi f ↔ ∀ i ∈ I, s i ∈ f i := ⟨λ h i hi, mem_of_pi_mem_pi h hi, pi_mem_pi hI⟩ +lemma eventually.eval_pi {i : ι} (hf : ∀ᶠ (x : α i) in f i, p i x) : + ∀ᶠ (x : Π (i : ι), α i) in pi f, p i (x i) := +(tendsto_eval_pi _ _).eventually hf + +lemma eventually_pi [finite ι] (hf : ∀ i, ∀ᶠ x in f i, p i x) : + ∀ᶠ (x : Π i, α i) in pi f, ∀ i, p i (x i) := +eventually_all.2 $ λ i, (hf _).eval_pi + lemma has_basis_pi {ι' : ι → Type} {s : Π i, ι' i → set (α i)} {p : Π i, ι' i → Prop} (h : ∀ i, (f i).has_basis (p i) (s i)) : (pi f).has_basis (λ If : set ι × Π i, ι' i, If.1.finite ∧ ∀ i ∈ If.1, p i (If.2 i)) diff --git a/src/topology/algebra/order/liminf_limsup.lean b/src/topology/algebra/order/liminf_limsup.lean index c16bae4dda725..74946971a14e3 100644 --- a/src/topology/algebra/order/liminf_limsup.lean +++ b/src/topology/algebra/order/liminf_limsup.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2017 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Johannes Hölzl, Mario Carneiro, Yury Kudryashov +Authors: Johannes Hölzl, Mario Carneiro, Yury Kudryashov, Yaël Dillies -/ import algebra.big_operators.intervals import algebra.big_operators.order @@ -16,68 +16,152 @@ import topology.order.basic > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. > Any changes to this file require a corresponding PR to mathlib4. + +## Main declarations + +* `bounded_le_nhds_class`: Typeclass stating that neighborhoods are eventually bounded above. +* `bounded_ge_nhds_class`: Typeclass stating that neighborhoods are eventually bounded below. + +## Implementation notes + +The same lemmas are true in `ℝ`, `ℝ × ℝ`, `ι → ℝ`, `euclidean_space ι ℝ`. To avoid code +duplication, we provide an ad hoc axiomatisation of the properties we need. -/ open filter topological_space open_locale topology classical universes u v -variables {α : Type u} {β : Type v} +variables {ι α β R S : Type*} {π : ι → Type*} -section liminf_limsup +/-- Ad hoc typeclass stating that neighborhoods are eventually bounded above. -/ +class bounded_le_nhds_class (α : Type*) [preorder α] [topological_space α] : Prop := +(is_bounded_le_nhds (a : α) : (𝓝 a).is_bounded (≤)) + +/-- Ad hoc typeclass stating that neighborhoods are eventually bounded below. -/ +class bounded_ge_nhds_class (α : Type*) [preorder α] [topological_space α] : Prop := +(is_bounded_ge_nhds (a : α) : (𝓝 a).is_bounded (≥)) + +section preorder +variables [preorder α] [preorder β] [topological_space α] [topological_space β] -section order_closed_topology -variables [semilattice_sup α] [topological_space α] [order_topology α] +section bounded_le_nhds_class +variables [bounded_le_nhds_class α] [bounded_le_nhds_class β] {f : filter ι} {u : ι → α} {a : α} lemma is_bounded_le_nhds (a : α) : (𝓝 a).is_bounded (≤) := -(is_top_or_exists_gt a).elim (λ h, ⟨a, eventually_of_forall h⟩) (λ ⟨b, hb⟩, ⟨b, ge_mem_nhds hb⟩) +bounded_le_nhds_class.is_bounded_le_nhds _ -lemma filter.tendsto.is_bounded_under_le {f : filter β} {u : β → α} {a : α} - (h : tendsto u f (𝓝 a)) : f.is_bounded_under (≤) u := +lemma filter.tendsto.is_bounded_under_le (h : tendsto u f (𝓝 a)) : + f.is_bounded_under (≤) u := (is_bounded_le_nhds a).mono h -lemma filter.tendsto.bdd_above_range_of_cofinite {u : β → α} {a : α} +lemma filter.tendsto.bdd_above_range_of_cofinite [is_directed α (≤)] (h : tendsto u cofinite (𝓝 a)) : bdd_above (set.range u) := h.is_bounded_under_le.bdd_above_range_of_cofinite -lemma filter.tendsto.bdd_above_range {u : ℕ → α} {a : α} - (h : tendsto u at_top (𝓝 a)) : bdd_above (set.range u) := +lemma filter.tendsto.bdd_above_range [is_directed α (≤)] {u : ℕ → α} (h : tendsto u at_top (𝓝 a)) : + bdd_above (set.range u) := h.is_bounded_under_le.bdd_above_range lemma is_cobounded_ge_nhds (a : α) : (𝓝 a).is_cobounded (≥) := (is_bounded_le_nhds a).is_cobounded_flip -lemma filter.tendsto.is_cobounded_under_ge {f : filter β} {u : β → α} {a : α} - [ne_bot f] (h : tendsto u f (𝓝 a)) : f.is_cobounded_under (≥) u := +lemma filter.tendsto.is_cobounded_under_ge [ne_bot f] (h : tendsto u f (𝓝 a)) : + f.is_cobounded_under (≥) u := h.is_bounded_under_le.is_cobounded_flip -end order_closed_topology +instance : bounded_ge_nhds_class αᵒᵈ := ⟨@is_bounded_le_nhds α _ _ _⟩ + +instance : bounded_le_nhds_class (α × β) := +begin + refine ⟨λ x, _⟩, + obtain ⟨a, ha⟩ := is_bounded_le_nhds x.1, + obtain ⟨b, hb⟩ := is_bounded_le_nhds x.2, + rw [←@prod.mk.eta _ _ x, nhds_prod_eq], + exact ⟨(a, b), ha.prod_mk hb⟩, +end + +instance [finite ι] [Π i, preorder (π i)] [Π i, topological_space (π i)] + [Π i, bounded_le_nhds_class (π i)] : bounded_le_nhds_class (Π i, π i) := +begin + refine ⟨λ x, _⟩, + rw nhds_pi, + choose f hf using λ i, is_bounded_le_nhds (x i), + exact ⟨f, eventually_pi hf⟩, +end + +end bounded_le_nhds_class -section order_closed_topology -variables [semilattice_inf α] [topological_space α] [order_topology α] +section bounded_ge_nhds_class +variables [bounded_ge_nhds_class α] [bounded_ge_nhds_class β] {f : filter ι} {u : ι → α} {a : α} -lemma is_bounded_ge_nhds (a : α) : (𝓝 a).is_bounded (≥) := @is_bounded_le_nhds αᵒᵈ _ _ _ a +lemma is_bounded_ge_nhds (a : α) : (𝓝 a).is_bounded (≥) := +bounded_ge_nhds_class.is_bounded_ge_nhds _ -lemma filter.tendsto.is_bounded_under_ge {f : filter β} {u : β → α} {a : α} - (h : tendsto u f (𝓝 a)) : f.is_bounded_under (≥) u := +lemma filter.tendsto.is_bounded_under_ge (h : tendsto u f (𝓝 a)) : + f.is_bounded_under (≥) u := (is_bounded_ge_nhds a).mono h -lemma filter.tendsto.bdd_below_range_of_cofinite {u : β → α} {a : α} +lemma filter.tendsto.bdd_below_range_of_cofinite [is_directed α (≥)] (h : tendsto u cofinite (𝓝 a)) : bdd_below (set.range u) := h.is_bounded_under_ge.bdd_below_range_of_cofinite -lemma filter.tendsto.bdd_below_range {u : ℕ → α} {a : α} - (h : tendsto u at_top (𝓝 a)) : bdd_below (set.range u) := +lemma filter.tendsto.bdd_below_range [is_directed α (≥)] {u : ℕ → α} (h : tendsto u at_top (𝓝 a)) : + bdd_below (set.range u) := h.is_bounded_under_ge.bdd_below_range lemma is_cobounded_le_nhds (a : α) : (𝓝 a).is_cobounded (≤) := (is_bounded_ge_nhds a).is_cobounded_flip -lemma filter.tendsto.is_cobounded_under_le {f : filter β} {u : β → α} {a : α} - [ne_bot f] (h : tendsto u f (𝓝 a)) : f.is_cobounded_under (≤) u := +lemma filter.tendsto.is_cobounded_under_le [ne_bot f] (h : tendsto u f (𝓝 a)) : + f.is_cobounded_under (≤) u := h.is_bounded_under_ge.is_cobounded_flip -end order_closed_topology +instance : bounded_le_nhds_class αᵒᵈ := ⟨@is_bounded_ge_nhds α _ _ _⟩ + +instance : bounded_ge_nhds_class (α × β) := +begin + refine ⟨λ x, _⟩, + obtain ⟨a, ha⟩ := is_bounded_ge_nhds x.1, + obtain ⟨b, hb⟩ := is_bounded_ge_nhds x.2, + rw [←@prod.mk.eta _ _ x, nhds_prod_eq], + exact ⟨(a, b), ha.prod_mk hb⟩, +end + +instance [finite ι] [Π i, preorder (π i)] [Π i, topological_space (π i)] + [Π i, bounded_ge_nhds_class (π i)] : bounded_ge_nhds_class (Π i, π i) := +begin + refine ⟨λ x, _⟩, + rw nhds_pi, + choose f hf using λ i, is_bounded_ge_nhds (x i), + exact ⟨f, eventually_pi hf⟩, +end + +end bounded_ge_nhds_class + +@[priority 100] -- See note [lower instance priority] +instance order_top.to_bounded_le_nhds_class [order_top α] : bounded_le_nhds_class α := +⟨λ a, is_bounded_le_of_top⟩ + +@[priority 100] -- See note [lower instance priority] +instance order_bot.to_bounded_ge_nhds_class [order_bot α] : bounded_ge_nhds_class α := +⟨λ a, is_bounded_ge_of_bot⟩ + +@[priority 100] -- See note [lower instance priority] +instance order_topology.to_bounded_le_nhds_class [is_directed α (≤)] [order_topology α] : + bounded_le_nhds_class α := +⟨λ a, (is_top_or_exists_gt a).elim (λ h, ⟨a, eventually_of_forall h⟩) $ Exists.imp $ λ b, + ge_mem_nhds⟩ + +@[priority 100] -- See note [lower instance priority] +instance order_topology.to_bounded_ge_nhds_class [is_directed α (≥)] [order_topology α] : + bounded_ge_nhds_class α := +⟨λ a, (is_bot_or_exists_lt a).elim (λ h, ⟨a, eventually_of_forall h⟩) $ Exists.imp $ λ b, + le_mem_nhds⟩ + +end preorder + +section liminf_limsup section conditionally_complete_linear_order variables [conditionally_complete_linear_order α] @@ -224,7 +308,7 @@ end liminf_limsup section monotone -variables {ι R S : Type*} {F : filter ι} [ne_bot F] +variables {F : filter ι} [ne_bot F] [complete_linear_order R] [topological_space R] [order_topology R] [complete_linear_order S] [topological_space S] [order_topology S] @@ -334,7 +418,7 @@ open_locale topology open filter set -variables {ι : Type*} {R : Type*} [complete_linear_order R] [topological_space R] [order_topology R] +variables [complete_linear_order R] [topological_space R] [order_topology R] lemma infi_eq_of_forall_le_of_tendsto {x : R} {as : ι → R} (x_le : ∀ i, x ≤ as i) {F : filter ι} [filter.ne_bot F] (as_lim : filter.tendsto as F (𝓝 x)) : @@ -349,7 +433,7 @@ lemma supr_eq_of_forall_le_of_tendsto {x : R} {as : ι → R} (⨆ i, as i) = x := @infi_eq_of_forall_le_of_tendsto ι (order_dual R) _ _ _ x as le_x F _ as_lim -lemma Union_Ici_eq_Ioi_of_lt_of_tendsto {ι : Type*} (x : R) {as : ι → R} (x_lt : ∀ i, x < as i) +lemma Union_Ici_eq_Ioi_of_lt_of_tendsto (x : R) {as : ι → R} (x_lt : ∀ i, x < as i) {F : filter ι} [filter.ne_bot F] (as_lim : filter.tendsto as F (𝓝 x)) : (⋃ (i : ι), Ici (as i)) = Ioi x := begin @@ -361,10 +445,10 @@ begin exact Union_Ici_eq_Ioi_infi obs, end -lemma Union_Iic_eq_Iio_of_lt_of_tendsto {ι : Type*} (x : R) {as : ι → R} (lt_x : ∀ i, as i < x) +lemma Union_Iic_eq_Iio_of_lt_of_tendsto (x : R) {as : ι → R} (lt_x : ∀ i, as i < x) {F : filter ι} [filter.ne_bot F] (as_lim : filter.tendsto as F (𝓝 x)) : (⋃ (i : ι), Iic (as i)) = Iio x := -@Union_Ici_eq_Ioi_of_lt_of_tendsto (order_dual R) _ _ _ ι x as lt_x F _ as_lim +@Union_Ici_eq_Ioi_of_lt_of_tendsto ι Rᵒᵈ _ _ _ _ _ lt_x F _ as_lim end infi_and_supr From c8f305514e0d47dfaa710f5a52f0d21b588e6328 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 9 Oct 2023 14:37:58 +0000 Subject: [PATCH 56/61] feat(topology/metric_space): diameter of pointwise zero and addition (#19028) --- src/analysis/normed/group/pointwise.lean | 13 ++++++++++ src/topology/metric_space/antilipschitz.lean | 25 +++++++++++++++++++- src/topology/metric_space/basic.lean | 11 ++++++++- src/topology/metric_space/emetric_space.lean | 4 +++- src/topology/metric_space/lipschitz.lean | 25 ++++++++++++++++++++ 5 files changed, 75 insertions(+), 3 deletions(-) diff --git a/src/analysis/normed/group/pointwise.lean b/src/analysis/normed/group/pointwise.lean index 672dd575ae904..340dc562045b2 100644 --- a/src/analysis/normed/group/pointwise.lean +++ b/src/analysis/normed/group/pointwise.lean @@ -5,6 +5,7 @@ Authors: Sébastien Gouëzel, Yaël Dillies -/ import analysis.normed.group.basic import topology.metric_space.hausdorff_distance +import topology.metric_space.isometric_smul /-! # Properties of pointwise addition of sets in normed groups @@ -24,6 +25,7 @@ variables {E : Type*} section seminormed_group variables [seminormed_group E] {ε δ : ℝ} {s t : set E} {x y : E} +-- note: we can't use `lipschitz_on_with.bounded_image2` here without adding `[isometric_smul E E]` @[to_additive] lemma metric.bounded.mul (hs : bounded s) (ht : bounded t) : bounded (s * t) := begin obtain ⟨Rs, hRs⟩ : ∃ R, ∀ x ∈ s, ‖x‖ ≤ R := hs.exists_norm_le', @@ -33,6 +35,10 @@ begin exact norm_mul_le_of_le (hRs x hx) (hRt y hy), end +@[to_additive] lemma metric.bounded.of_mul (hst : bounded (s * t)) : + bounded s ∨ bounded t := +antilipschitz_with.bounded_of_image2_left _ (λ x, (isometry_mul_right x).antilipschitz) hst + @[to_additive] lemma metric.bounded.inv : bounded s → bounded s⁻¹ := by { simp_rw [bounded_iff_forall_norm_le', ←image_inv, ball_image_iff, norm_inv'], exact id } @@ -55,6 +61,13 @@ eq_of_forall_le_iff $ λ r, by simp_rw [le_inf_edist, ←image_inv, ball_image_i lemma inf_edist_inv_inv (x : E) (s : set E) : inf_edist x⁻¹ s⁻¹ = inf_edist x s := by rw [inf_edist_inv, inv_inv] +@[to_additive] lemma ediam_mul_le (x y : set E) : + emetric.diam (x * y) ≤ emetric.diam x + emetric.diam y := +(lipschitz_on_with.ediam_image2_le (*) _ _ + (λ _ _, (isometry_mul_right _).lipschitz.lipschitz_on_with _) + (λ _ _, (isometry_mul_left _).lipschitz.lipschitz_on_with _)).trans_eq $ + by simp only [ennreal.coe_one, one_mul] + end emetric variables (ε δ s t x y) diff --git a/src/topology/metric_space/antilipschitz.lean b/src/topology/metric_space/antilipschitz.lean index c08fbf9d2ef79..7974ab7cbf9d0 100644 --- a/src/topology/metric_space/antilipschitz.lean +++ b/src/topology/metric_space/antilipschitz.lean @@ -192,7 +192,8 @@ namespace antilipschitz_with open metric -variables [pseudo_metric_space α] [pseudo_metric_space β] {K : ℝ≥0} {f : α → β} +variables [pseudo_metric_space α] [pseudo_metric_space β] [pseudo_metric_space γ] +variables {K : ℝ≥0} {f : α → β} lemma bounded_preimage (hf : antilipschitz_with K f) {s : set β} (hs : bounded s) : @@ -219,6 +220,28 @@ begin exact (hf.image_preimage _).symm end +lemma bounded_of_image2_left (f : α → β → γ) {K₁ : ℝ≥0} + (hf : ∀ b, antilipschitz_with K₁ (λ a, f a b)) + {s : set α} {t : set β} (hst : bounded (set.image2 f s t)) : + bounded s ∨ bounded t := +begin + contrapose! hst, + obtain ⟨b, hb⟩ : t.nonempty := nonempty_of_unbounded hst.2, + have : ¬bounded (set.image2 f s {b}), + { intro h, + apply hst.1, + rw set.image2_singleton_right at h, + replace h := (hf b).bounded_preimage h, + refine h.mono (subset_preimage_image _ _) }, + exact mt (bounded.mono (image2_subset subset.rfl (singleton_subset_iff.mpr hb))) this, +end + +lemma bounded_of_image2_right {f : α → β → γ} {K₂ : ℝ≥0} + (hf : ∀ a, antilipschitz_with K₂ (f a)) + {s : set α} {t : set β} (hst : bounded (set.image2 f s t)) : + bounded s ∨ bounded t := +or.symm $ bounded_of_image2_left (flip f) hf $ image2_swap f s t ▸ hst + end antilipschitz_with lemma lipschitz_with.to_right_inverse [pseudo_emetric_space α] [pseudo_emetric_space β] {K : ℝ≥0} diff --git a/src/topology/metric_space/basic.lean b/src/topology/metric_space/basic.lean index de0baf9a61b0a..76e23c4a7f391 100644 --- a/src/topology/metric_space/basic.lean +++ b/src/topology/metric_space/basic.lean @@ -52,7 +52,7 @@ metric, pseudo_metric, dist open set filter topological_space bornology -open_locale uniformity topology big_operators filter nnreal ennreal +open_locale uniformity topology big_operators filter nnreal ennreal pointwise universes u v w variables {α : Type u} {β : Type v} {X ι : Type*} @@ -2171,6 +2171,13 @@ end @[simp] lemma bounded_empty : bounded (∅ : set α) := ⟨0, by simp⟩ +lemma nonempty_of_unbounded (h : ¬ bounded s) : s.nonempty := +begin + rw nonempty_iff_ne_empty, + rintro rfl, + exact h bounded_empty +end + lemma bounded_iff_mem_bounded : bounded s ↔ ∀ x ∈ s, bounded s := ⟨λ h _ _, h, λ H, s.eq_empty_or_nonempty.elim @@ -2453,6 +2460,8 @@ diam_subsingleton subsingleton_empty @[simp] lemma diam_singleton : diam ({x} : set α) = 0 := diam_subsingleton subsingleton_singleton +@[simp, to_additive] lemma diam_one [has_one α] : diam (1 : set α) = 0 := diam_singleton + -- Does not work as a simp-lemma, since {x, y} reduces to (insert y {x}) lemma diam_pair : diam ({x, y} : set α) = dist x y := by simp only [diam, emetric.diam_pair, dist_edist] diff --git a/src/topology/metric_space/emetric_space.lean b/src/topology/metric_space/emetric_space.lean index 983cb4e4990ce..c30cfc396c9a9 100644 --- a/src/topology/metric_space/emetric_space.lean +++ b/src/topology/metric_space/emetric_space.lean @@ -32,7 +32,7 @@ to `emetric_space` at the end. open set filter classical -open_locale uniformity topology big_operators filter nnreal ennreal +open_locale uniformity topology big_operators filter nnreal ennreal pointwise universes u v w variables {α : Type u} {β : Type v} {X : Type*} @@ -793,6 +793,8 @@ diam_subsingleton subsingleton_empty @[simp] lemma diam_singleton : diam ({x} : set α) = 0 := diam_subsingleton subsingleton_singleton +@[simp, to_additive] lemma diam_one [has_one α] : diam (1 : set α) = 0 := diam_singleton + lemma diam_Union_mem_option {ι : Type*} (o : option ι) (s : ι → set α) : diam (⋃ i ∈ o, s i) = ⨆ i ∈ o, diam (s i) := by cases o; simp diff --git a/src/topology/metric_space/lipschitz.lean b/src/topology/metric_space/lipschitz.lean index 08d7ec3473268..235c668ff913e 100644 --- a/src/topology/metric_space/lipschitz.lean +++ b/src/topology/metric_space/lipschitz.lean @@ -465,6 +465,20 @@ protected lemma comp {g : β → γ} {t : set β} {Kg : ℝ≥0} (hg : lipschitz lipschitz_on_with (Kg * K) (g ∘ f) s := lipschitz_on_with_iff_restrict.mpr $ hg.to_restrict.comp (hf.to_restrict_maps_to hmaps) +lemma ediam_image2_le (f : α → β → γ) {K₁ K₂ : ℝ≥0} + (s : set α) (t : set β) + (hf₁ : ∀ b ∈ t, lipschitz_on_with K₁ (λ a, f a b) s) + (hf₂ : ∀ a ∈ s, lipschitz_on_with K₂ (f a) t) : + emetric.diam (set.image2 f s t) ≤ ↑K₁ * emetric.diam s + ↑K₂ * emetric.diam t := +begin + apply emetric.diam_le, + rintros _ ⟨a₁, b₁, ha₁, hb₁, rfl⟩ _ ⟨a₂, b₂, ha₂, hb₂, rfl⟩, + refine (edist_triangle _ (f a₂ b₁) _).trans _, + exact add_le_add + ((hf₁ b₁ hb₁ ha₁ ha₂).trans $ ennreal.mul_left_mono $ emetric.edist_le_diam_of_mem ha₁ ha₂) + ((hf₂ a₂ ha₂ hb₁ hb₂).trans $ ennreal.mul_left_mono $ emetric.edist_le_diam_of_mem hb₁ hb₂), +end + end emetric section metric @@ -512,6 +526,17 @@ protected lemma iff_le_add_mul {f : α → ℝ} {K : ℝ≥0} : lipschitz_on_with K f s ↔ ∀ (x ∈ s) (y ∈ s), f x ≤ f y + K * dist x y := ⟨lipschitz_on_with.le_add_mul, lipschitz_on_with.of_le_add_mul K⟩ +lemma bounded_image2 (f : α → β → γ) {K₁ K₂ : ℝ≥0} + {s : set α} {t : set β} (hs : metric.bounded s) (ht : metric.bounded t) + (hf₁ : ∀ b ∈ t, lipschitz_on_with K₁ (λ a, f a b) s) + (hf₂ : ∀ a ∈ s, lipschitz_on_with K₂ (f a) t) : + metric.bounded (set.image2 f s t) := +metric.bounded_iff_ediam_ne_top.2 $ ne_top_of_le_ne_top + (ennreal.add_ne_top.mpr ⟨ + ennreal.mul_ne_top ennreal.coe_ne_top hs.ediam_ne_top, + ennreal.mul_ne_top ennreal.coe_ne_top ht.ediam_ne_top⟩) + (ediam_image2_le _ _ _ hf₁ hf₂) + end metric end lipschitz_on_with From 3ac76ec59aac51aceffc1e4888f870497dc0ab64 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Tue, 10 Oct 2023 22:55:10 +0000 Subject: [PATCH 57/61] doc: Add a warning mentioning Lean 4 to the readme (#19243) Also correct links to point to the lean3 webpages; this means that users clicking them end up on pages which also have scary banners telling them not to use Lean 3. --- README.md | 33 ++++++++++++++++++++------------- 1 file changed, 20 insertions(+), 13 deletions(-) diff --git a/README.md b/README.md index 88b3c18f3ccbe..53d8c3defebf5 100644 --- a/README.md +++ b/README.md @@ -1,28 +1,32 @@ -# Lean mathlib +# Lean 3's mathlib + +> [!WARNING] +> Lean 3 and Mathlib 3 are no longer actively maintained. +> It is strongly recommended that you use [mathlib4](https://github.com/leanprover-community/mathlib4) for Lean 4 instead. ![](https://github.com/leanprover-community/mathlib/workflows/continuous%20integration/badge.svg?branch=master) [![Bors enabled](https://bors.tech/images/badge_small.svg)](https://app.bors.tech/repositories/24316) [![project chat](https://img.shields.io/badge/zulip-join_chat-brightgreen.svg)](https://leanprover.zulipchat.com) [![Gitpod Ready-to-Code](https://img.shields.io/badge/Gitpod-ready--to--code-blue?logo=gitpod)](https://gitpod.io/#https://github.com/leanprover-community/mathlib) -[Mathlib](https://leanprover-community.github.io) is a user maintained library for the [Lean theorem prover](https://leanprover.github.io). +[Mathlib](https://leanprover-community.github.io) is a user maintained library for the [Lean 3 theorem prover](https://github.com/leanprover-community/lean). It contains both programming infrastructure and mathematics, as well as tactics that use the former and allow to develop the latter. ## Installation -You can find detailed instructions to install Lean, mathlib, and supporting tools on [our website](https://leanprover-community.github.io/get_started.html). +You can find detailed instructions to install Lean 3, mathlib 3, and supporting tools on [our website](https://leanprover-community.github.io/lean3/get_started.html). ## Experimenting -Got everything installed? Why not start with the [tutorial project](https://leanprover-community.github.io/install/project.html)? +Got everything installed? Why not start with the [tutorial project](https://leanprover-community.github.io/lean3/install/project.html)? -For more pointers, see [Learning Lean](https://leanprover-community.github.io/learn.html). +For more pointers, see [Learning Lean](https://leanprover-community.github.io/lean3/learn.html). ## Documentation Besides the installation guides above and [Lean's general -documentation](https://leanprover.github.io/documentation/), the documentation +documentation](https://leanprover.github.io/lean3/documentation/), the documentation of mathlib consists of: - [The mathlib docs](https://leanprover-community.github.io/mathlib_docs): documentation [generated @@ -33,10 +37,10 @@ of mathlib consists of: - [hole commands](https://leanprover-community.github.io/mathlib_docs/hole_commands.html), and - [attributes](https://leanprover-community.github.io/mathlib_docs/attributes.html). - A description of [currently covered theories](https://leanprover-community.github.io/theories.html), - as well as an [overview](https://leanprover-community.github.io/mathlib-overview.html) for mathematicians. + as well as an [overview](https://leanprover-community.github.io/lean3/mathlib-overview.html) for mathematicians. - A couple of [tutorial Lean files](docs/tutorial/) -- Some [extra Lean documentation](https://leanprover-community.github.io/learn.html) not specific to mathlib (see "Miscellaneous topics") -- Documentation for people who would like to [contribute to mathlib](https://leanprover-community.github.io/contribute/index.html) +- Some [extra Lean documentation](https://leanprover-community.github.io/lean3/learn.html) not specific to mathlib (see "Miscellaneous topics") +- Documentation for people who would like to [contribute to mathlib3](https://leanprover-community.github.io/lean3/contribute/index.html) Much of the discussion surrounding mathlib occurs in a [Zulip chat room](https://leanprover.zulipchat.com/). Since this @@ -49,8 +53,11 @@ welcomed. ## Contributing +> [!WARNING] +> Contributions are no longer accepted to mathlib 3; contribute to mathlib 4 instead! + The complete documentation for contributing to ``mathlib`` is located -[on the community guide contribute to mathlib](https://leanprover-community.github.io/contribute/index.html) +[on the community guide contribute to mathlib](https://leanprover-community.github.io/lean3/contribute/index.html) The process is different from other projects where one should not fork the repository. Instead write permission for non-master branches should be requested on [Zulip](https://leanprover.zulipchat.com) @@ -60,9 +67,9 @@ by introducing yourself, providing your GitHub handle and what contribution you Mathlib has the following guidelines and conventions that must be followed - - The [style guide](https://leanprover-community.github.io/contribute/style.html) - - A guide on the [naming convention](https://leanprover-community.github.io/contribute/naming.html) - - The [documentation style](https://leanprover-community.github.io/contribute/doc.html) + - The [style guide](https://leanprover-community.github.io/lean3/contribute/style.html) + - A guide on the [naming convention](https://leanprover-community.github.io/lean3/contribute/naming.html) + - The [documentation style](https://leanprover-community.github.io/lean3/contribute/doc.html) - The [commit naming conventions](https://github.com/leanprover-community/lean/blob/master/doc/commit_convention.md) Note: the title of a PR should follow the commit naming convention. From 19c869efa56bbb8b500f2724c0b77261edbfa28c Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Wed, 11 Oct 2023 09:24:47 +0200 Subject: [PATCH 58/61] move old README.md to OLD_README.md --- OLD_README.md | 115 ++++++++++++++++++++++++++++++++++++++++++++++++++ README.md | 110 +---------------------------------------------- 2 files changed, 116 insertions(+), 109 deletions(-) create mode 100644 OLD_README.md diff --git a/OLD_README.md b/OLD_README.md new file mode 100644 index 0000000000000..53d8c3defebf5 --- /dev/null +++ b/OLD_README.md @@ -0,0 +1,115 @@ +# Lean 3's mathlib + +> [!WARNING] +> Lean 3 and Mathlib 3 are no longer actively maintained. +> It is strongly recommended that you use [mathlib4](https://github.com/leanprover-community/mathlib4) for Lean 4 instead. + +![](https://github.com/leanprover-community/mathlib/workflows/continuous%20integration/badge.svg?branch=master) +[![Bors enabled](https://bors.tech/images/badge_small.svg)](https://app.bors.tech/repositories/24316) +[![project chat](https://img.shields.io/badge/zulip-join_chat-brightgreen.svg)](https://leanprover.zulipchat.com) +[![Gitpod Ready-to-Code](https://img.shields.io/badge/Gitpod-ready--to--code-blue?logo=gitpod)](https://gitpod.io/#https://github.com/leanprover-community/mathlib) + +[Mathlib](https://leanprover-community.github.io) is a user maintained library for the [Lean 3 theorem prover](https://github.com/leanprover-community/lean). +It contains both programming infrastructure and mathematics, +as well as tactics that use the former and allow to develop the latter. + +## Installation + +You can find detailed instructions to install Lean 3, mathlib 3, and supporting tools on [our website](https://leanprover-community.github.io/lean3/get_started.html). + +## Experimenting + +Got everything installed? Why not start with the [tutorial project](https://leanprover-community.github.io/lean3/install/project.html)? + +For more pointers, see [Learning Lean](https://leanprover-community.github.io/lean3/learn.html). + +## Documentation + +Besides the installation guides above and [Lean's general +documentation](https://leanprover.github.io/lean3/documentation/), the documentation +of mathlib consists of: + +- [The mathlib docs](https://leanprover-community.github.io/mathlib_docs): documentation [generated + automatically](https://github.com/leanprover-community/doc-gen) from the source `.lean` files. + In addition to the pages generated for each file in the library, the docs also include pages on: + - [tactics](https://leanprover-community.github.io/mathlib_docs/tactics.html), + - [commands](https://leanprover-community.github.io/mathlib_docs/commands.html), + - [hole commands](https://leanprover-community.github.io/mathlib_docs/hole_commands.html), and + - [attributes](https://leanprover-community.github.io/mathlib_docs/attributes.html). +- A description of [currently covered theories](https://leanprover-community.github.io/theories.html), + as well as an [overview](https://leanprover-community.github.io/lean3/mathlib-overview.html) for mathematicians. +- A couple of [tutorial Lean files](docs/tutorial/) +- Some [extra Lean documentation](https://leanprover-community.github.io/lean3/learn.html) not specific to mathlib (see "Miscellaneous topics") +- Documentation for people who would like to [contribute to mathlib3](https://leanprover-community.github.io/lean3/contribute/index.html) + +Much of the discussion surrounding mathlib occurs in a +[Zulip chat room](https://leanprover.zulipchat.com/). Since this +chatroom is only visible to registered users, we provide an +[openly accessible archive](https://leanprover-community.github.io/archive/) +of the public discussions. This is useful for quick reference; for a +better browsing interface, and to participate in the discussions, we strongly +suggest joining the chat. Questions from users at all levels of expertise are +welcomed. + +## Contributing + +> [!WARNING] +> Contributions are no longer accepted to mathlib 3; contribute to mathlib 4 instead! + +The complete documentation for contributing to ``mathlib`` is located +[on the community guide contribute to mathlib](https://leanprover-community.github.io/lean3/contribute/index.html) + +The process is different from other projects where one should not fork the repository. +Instead write permission for non-master branches should be requested on [Zulip](https://leanprover.zulipchat.com) +by introducing yourself, providing your GitHub handle and what contribution you are planning on doing. + +### Guidelines + +Mathlib has the following guidelines and conventions that must be followed + + - The [style guide](https://leanprover-community.github.io/lean3/contribute/style.html) + - A guide on the [naming convention](https://leanprover-community.github.io/lean3/contribute/naming.html) + - The [documentation style](https://leanprover-community.github.io/lean3/contribute/doc.html) + - The [commit naming conventions](https://github.com/leanprover-community/lean/blob/master/doc/commit_convention.md) + +Note: the title of a PR should follow the commit naming convention. + +### Using ``leanproject`` to contribute + +Running the ``leanproject get -b mathlib:shiny_lemma`` command will create a new worktree ``mathlib_shiny_lemma`` +with a local branch called ``shiny_lemma`` which has a copy of mathlib to work on. + +``leanproject build`` will check that nothing broke. +Be warned that this will take some time if a fundamental file was changed. + +## Maintainers: + +* Anne Baanen (@Vierkantor): algebra, number theory, tactics +* Reid Barton (@rwbarton): category theory, topology +* Riccardo Brasca (@riccardobrasca): algebra, number theory, algebraic geometry, category theory +* Mario Carneiro (@digama0): lean formalization, tactics, type theory, proof engineering +* Bryan Gin-ge Chen (@bryangingechen): documentation, infrastructure +* Johan Commelin (@jcommelin): algebra, number theory, category theory, algebraic geometry +* Rémy Degenne (@RemyDegenne): probability, measure theory, analysis +* Floris van Doorn (@fpvandoorn): measure theory, model theory, tactics +* Frédéric Dupuis (@dupuisf): linear algebra, functional analysis +* Gabriel Ebner (@gebner): tactics, infrastructure, core, formal languages +* Sébastien Gouëzel (@sgouezel): topology, calculus, geometry, analysis, measure theory +* Markus Himmel (@TwoFX): category theory +* Chris Hughes (@ChrisHughes24): algebra +* Yury G. Kudryashov (@urkud): analysis, topology, measure theory +* Robert Y. Lewis (@robertylewis): tactics, documentation +* Heather Macbeth (@hrmacbeth): geometry, analysis +* Patrick Massot (@patrickmassot): documentation, topology, geometry +* Bhavik Mehta (@b-mehta): category theory, combinatorics +* Kyle Miller (@kmill): combinatorics, documentation +* Scott Morrison (@semorrison): category theory, tactics +* Oliver Nash (@ocfnash): algebra, geometry, topology +* Adam Topaz (@adamtopaz): algebra, category theory, algebraic geometry +* Eric Wieser (@eric-wieser): algebra, infrastructure + +## Emeritus maintainers: + +* Jeremy Avigad (@avigad): analysis +* Johannes Hölzl (@johoelzl): measure theory, topology +* Simon Hudon (@cipher1024): tactics diff --git a/README.md b/README.md index 53d8c3defebf5..f850c525cc3b0 100644 --- a/README.md +++ b/README.md @@ -4,112 +4,4 @@ > Lean 3 and Mathlib 3 are no longer actively maintained. > It is strongly recommended that you use [mathlib4](https://github.com/leanprover-community/mathlib4) for Lean 4 instead. -![](https://github.com/leanprover-community/mathlib/workflows/continuous%20integration/badge.svg?branch=master) -[![Bors enabled](https://bors.tech/images/badge_small.svg)](https://app.bors.tech/repositories/24316) -[![project chat](https://img.shields.io/badge/zulip-join_chat-brightgreen.svg)](https://leanprover.zulipchat.com) -[![Gitpod Ready-to-Code](https://img.shields.io/badge/Gitpod-ready--to--code-blue?logo=gitpod)](https://gitpod.io/#https://github.com/leanprover-community/mathlib) - -[Mathlib](https://leanprover-community.github.io) is a user maintained library for the [Lean 3 theorem prover](https://github.com/leanprover-community/lean). -It contains both programming infrastructure and mathematics, -as well as tactics that use the former and allow to develop the latter. - -## Installation - -You can find detailed instructions to install Lean 3, mathlib 3, and supporting tools on [our website](https://leanprover-community.github.io/lean3/get_started.html). - -## Experimenting - -Got everything installed? Why not start with the [tutorial project](https://leanprover-community.github.io/lean3/install/project.html)? - -For more pointers, see [Learning Lean](https://leanprover-community.github.io/lean3/learn.html). - -## Documentation - -Besides the installation guides above and [Lean's general -documentation](https://leanprover.github.io/lean3/documentation/), the documentation -of mathlib consists of: - -- [The mathlib docs](https://leanprover-community.github.io/mathlib_docs): documentation [generated - automatically](https://github.com/leanprover-community/doc-gen) from the source `.lean` files. - In addition to the pages generated for each file in the library, the docs also include pages on: - - [tactics](https://leanprover-community.github.io/mathlib_docs/tactics.html), - - [commands](https://leanprover-community.github.io/mathlib_docs/commands.html), - - [hole commands](https://leanprover-community.github.io/mathlib_docs/hole_commands.html), and - - [attributes](https://leanprover-community.github.io/mathlib_docs/attributes.html). -- A description of [currently covered theories](https://leanprover-community.github.io/theories.html), - as well as an [overview](https://leanprover-community.github.io/lean3/mathlib-overview.html) for mathematicians. -- A couple of [tutorial Lean files](docs/tutorial/) -- Some [extra Lean documentation](https://leanprover-community.github.io/lean3/learn.html) not specific to mathlib (see "Miscellaneous topics") -- Documentation for people who would like to [contribute to mathlib3](https://leanprover-community.github.io/lean3/contribute/index.html) - -Much of the discussion surrounding mathlib occurs in a -[Zulip chat room](https://leanprover.zulipchat.com/). Since this -chatroom is only visible to registered users, we provide an -[openly accessible archive](https://leanprover-community.github.io/archive/) -of the public discussions. This is useful for quick reference; for a -better browsing interface, and to participate in the discussions, we strongly -suggest joining the chat. Questions from users at all levels of expertise are -welcomed. - -## Contributing - -> [!WARNING] -> Contributions are no longer accepted to mathlib 3; contribute to mathlib 4 instead! - -The complete documentation for contributing to ``mathlib`` is located -[on the community guide contribute to mathlib](https://leanprover-community.github.io/lean3/contribute/index.html) - -The process is different from other projects where one should not fork the repository. -Instead write permission for non-master branches should be requested on [Zulip](https://leanprover.zulipchat.com) -by introducing yourself, providing your GitHub handle and what contribution you are planning on doing. - -### Guidelines - -Mathlib has the following guidelines and conventions that must be followed - - - The [style guide](https://leanprover-community.github.io/lean3/contribute/style.html) - - A guide on the [naming convention](https://leanprover-community.github.io/lean3/contribute/naming.html) - - The [documentation style](https://leanprover-community.github.io/lean3/contribute/doc.html) - - The [commit naming conventions](https://github.com/leanprover-community/lean/blob/master/doc/commit_convention.md) - -Note: the title of a PR should follow the commit naming convention. - -### Using ``leanproject`` to contribute - -Running the ``leanproject get -b mathlib:shiny_lemma`` command will create a new worktree ``mathlib_shiny_lemma`` -with a local branch called ``shiny_lemma`` which has a copy of mathlib to work on. - -``leanproject build`` will check that nothing broke. -Be warned that this will take some time if a fundamental file was changed. - -## Maintainers: - -* Anne Baanen (@Vierkantor): algebra, number theory, tactics -* Reid Barton (@rwbarton): category theory, topology -* Riccardo Brasca (@riccardobrasca): algebra, number theory, algebraic geometry, category theory -* Mario Carneiro (@digama0): lean formalization, tactics, type theory, proof engineering -* Bryan Gin-ge Chen (@bryangingechen): documentation, infrastructure -* Johan Commelin (@jcommelin): algebra, number theory, category theory, algebraic geometry -* Rémy Degenne (@RemyDegenne): probability, measure theory, analysis -* Floris van Doorn (@fpvandoorn): measure theory, model theory, tactics -* Frédéric Dupuis (@dupuisf): linear algebra, functional analysis -* Gabriel Ebner (@gebner): tactics, infrastructure, core, formal languages -* Sébastien Gouëzel (@sgouezel): topology, calculus, geometry, analysis, measure theory -* Markus Himmel (@TwoFX): category theory -* Chris Hughes (@ChrisHughes24): algebra -* Yury G. Kudryashov (@urkud): analysis, topology, measure theory -* Robert Y. Lewis (@robertylewis): tactics, documentation -* Heather Macbeth (@hrmacbeth): geometry, analysis -* Patrick Massot (@patrickmassot): documentation, topology, geometry -* Bhavik Mehta (@b-mehta): category theory, combinatorics -* Kyle Miller (@kmill): combinatorics, documentation -* Scott Morrison (@semorrison): category theory, tactics -* Oliver Nash (@ocfnash): algebra, geometry, topology -* Adam Topaz (@adamtopaz): algebra, category theory, algebraic geometry -* Eric Wieser (@eric-wieser): algebra, infrastructure - -## Emeritus maintainers: - -* Jeremy Avigad (@avigad): analysis -* Johannes Hölzl (@johoelzl): measure theory, topology -* Simon Hudon (@cipher1024): tactics +(If you need to read the old `README.md`, please see `OLD_README.md`.) From b1abe23ae96fef89ad30d9f4362c307f72a55010 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 16 Oct 2023 09:24:06 +0000 Subject: [PATCH 59/61] =?UTF-8?q?feat(measure=5Ftheory/order/upper=5Flower?= =?UTF-8?q?):=20Order-connected=20sets=20in=20`=E2=84=9D=E2=81=BF`=20are?= =?UTF-8?q?=20measurable=20(#16976)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Prove that the frontier of an order-connected set in `ℝⁿ` (with the `∞`-metric, but it doesn't actually matter) has measure zero. As a corollary, antichains in `ℝⁿ` have measure zero. Co-authored-by: @JasonKYi --- src/analysis/normed/order/upper_lower.lean | 161 +++++++++++++++++++- src/data/real/nnreal.lean | 2 + src/measure_theory/order/upper_lower.lean | 131 ++++++++++++++++ src/order/bounds/basic.lean | 2 + src/topology/algebra/order/upper_lower.lean | 29 +++- 5 files changed, 320 insertions(+), 5 deletions(-) create mode 100644 src/measure_theory/order/upper_lower.lean diff --git a/src/analysis/normed/order/upper_lower.lean b/src/analysis/normed/order/upper_lower.lean index c834d2d671ddc..a081b70aa9623 100644 --- a/src/analysis/normed/order/upper_lower.lean +++ b/src/analysis/normed/order/upper_lower.lean @@ -23,10 +23,11 @@ are measurable. -/ open function metric set +open_locale pointwise variables {α ι : Type*} -section metric_space +section normed_ordered_group variables [normed_ordered_group α] {s : set α} @[to_additive is_upper_set.thickening] @@ -49,12 +50,22 @@ protected lemma is_lower_set.cthickening' (hs : is_lower_set s) (ε : ℝ) : is_lower_set (cthickening ε s) := by { rw cthickening_eq_Inter_thickening'', exact is_lower_set_Inter₂ (λ δ hδ, hs.thickening' _) } -end metric_space +@[to_additive upper_closure_interior_subset] +lemma upper_closure_interior_subset' (s : set α) : + (upper_closure (interior s) : set α) ⊆ interior (upper_closure s) := +upper_closure_min (interior_mono subset_upper_closure) (upper_closure s).upper.interior + +@[to_additive lower_closure_interior_subset] +lemma lower_closure_interior_subset' (s : set α) : + (upper_closure (interior s) : set α) ⊆ interior (upper_closure s) := +upper_closure_min (interior_mono subset_upper_closure) (upper_closure s).upper.interior + +end normed_ordered_group /-! ### `ℝⁿ` -/ section finite -variables [finite ι] {s : set (ι → ℝ)} {x y : ι → ℝ} {δ : ℝ} +variables [finite ι] {s : set (ι → ℝ)} {x y : ι → ℝ} lemma is_upper_set.mem_interior_of_forall_lt (hs : is_upper_set s) (hx : x ∈ closure s) (h : ∀ i, x i < y i) : @@ -99,7 +110,78 @@ end end finite section fintype -variables [fintype ι] {s : set (ι → ℝ)} {x y : ι → ℝ} {δ : ℝ} +variables [fintype ι] {s t : set (ι → ℝ)} {a₁ a₂ b₁ b₂ x y : ι → ℝ} {δ : ℝ} + +-- TODO: Generalise those lemmas so that they also apply to `ℝ` and `euclidean_space ι ℝ` +lemma dist_inf_sup (x y : ι → ℝ) : dist (x ⊓ y) (x ⊔ y) = dist x y := +begin + refine congr_arg coe (finset.sup_congr rfl $ λ i _, _), + simp only [real.nndist_eq', sup_eq_max, inf_eq_min, max_sub_min_eq_abs, pi.inf_apply, + pi.sup_apply, real.nnabs_of_nonneg, abs_nonneg, real.to_nnreal_abs], +end + +lemma dist_mono_left : monotone_on (λ x, dist x y) (Ici y) := +begin + refine λ y₁ hy₁ y₂ hy₂ hy, nnreal.coe_le_coe.2 (finset.sup_mono_fun $ λ i _, _), + rw [real.nndist_eq, real.nnabs_of_nonneg (sub_nonneg_of_le (‹y ≤ _› i : y i ≤ y₁ i)), + real.nndist_eq, real.nnabs_of_nonneg (sub_nonneg_of_le (‹y ≤ _› i : y i ≤ y₂ i))], + exact real.to_nnreal_mono (sub_le_sub_right (hy _) _), +end + +lemma dist_mono_right : monotone_on (dist x) (Ici x) := +by simpa only [dist_comm] using dist_mono_left + +lemma dist_anti_left : antitone_on (λ x, dist x y) (Iic y) := +begin + refine λ y₁ hy₁ y₂ hy₂ hy, nnreal.coe_le_coe.2 (finset.sup_mono_fun $ λ i _, _), + rw [real.nndist_eq', real.nnabs_of_nonneg (sub_nonneg_of_le (‹_ ≤ y› i : y₂ i ≤ y i)), + real.nndist_eq', real.nnabs_of_nonneg (sub_nonneg_of_le (‹_ ≤ y› i : y₁ i ≤ y i))], + exact real.to_nnreal_mono (sub_le_sub_left (hy _) _), +end + +lemma dist_anti_right : antitone_on (dist x) (Iic x) := +by simpa only [dist_comm] using dist_anti_left + +lemma dist_le_dist_of_le (ha : a₂ ≤ a₁) (h₁ : a₁ ≤ b₁) (hb : b₁ ≤ b₂) : dist a₁ b₁ ≤ dist a₂ b₂ := +(dist_mono_right h₁ (h₁.trans hb) hb).trans $ + dist_anti_left (ha.trans $ h₁.trans hb) (h₁.trans hb) ha + +protected lemma metric.bounded.bdd_below : bounded s → bdd_below s := +begin + rintro ⟨r, hr⟩, + obtain rfl | ⟨x, hx⟩ := s.eq_empty_or_nonempty, + { exact bdd_below_empty }, + { exact ⟨x - const _ r, λ y hy i, sub_le_comm.1 + (abs_sub_le_iff.1 $ (dist_le_pi_dist _ _ _).trans $ hr _ hx _ hy).1⟩ } +end + +protected lemma metric.bounded.bdd_above : bounded s → bdd_above s := +begin + rintro ⟨r, hr⟩, + obtain rfl | ⟨x, hx⟩ := s.eq_empty_or_nonempty, + { exact bdd_above_empty }, + { exact ⟨x + const _ r, λ y hy i, sub_le_iff_le_add'.1 $ + (abs_sub_le_iff.1 $ (dist_le_pi_dist _ _ _).trans $ hr _ hx _ hy).2⟩ } +end + +protected lemma bdd_below.bounded : bdd_below s → bdd_above s → bounded s := +begin + rintro ⟨a, ha⟩ ⟨b, hb⟩, + refine ⟨dist a b, λ x hx y hy, _⟩, + rw ←dist_inf_sup, + exact dist_le_dist_of_le (le_inf (ha hx) $ ha hy) inf_le_sup (sup_le (hb hx) $ hb hy), +end + +protected lemma bdd_above.bounded : bdd_above s → bdd_below s → bounded s := flip bdd_below.bounded + +lemma bounded_iff_bdd_below_bdd_above : bounded s ↔ bdd_below s ∧ bdd_above s := +⟨λ h, ⟨h.bdd_below, h.bdd_above⟩, λ h, h.1.bounded h.2⟩ + +lemma bdd_below.bounded_inter (hs : bdd_below s) (ht : bdd_above t) : bounded (s ∩ t) := +(hs.mono $ inter_subset_left _ _).bounded $ ht.mono $ inter_subset_right _ _ + +lemma bdd_above.bounded_inter (hs : bdd_above s) (ht : bdd_below t) : bounded (s ∩ t) := +(hs.mono $ inter_subset_left _ _).bounded $ ht.mono $ inter_subset_right _ _ lemma is_upper_set.exists_subset_ball (hs : is_upper_set s) (hx : x ∈ closure s) (hδ : 0 < δ) : ∃ y, closed_ball y (δ/4) ⊆ closed_ball x δ ∧ closed_ball y (δ/4) ⊆ interior s := @@ -140,3 +222,74 @@ begin end end fintype + +section finite +variables [finite ι] {s t : set (ι → ℝ)} {a₁ a₂ b₁ b₂ x y : ι → ℝ} {δ : ℝ} + +lemma is_antichain.interior_eq_empty [nonempty ι] (hs : is_antichain (≤) s) : interior s = ∅ := +begin + casesI nonempty_fintype ι, + refine eq_empty_of_forall_not_mem (λ x hx, _), + have hx' := interior_subset hx, + rw [mem_interior_iff_mem_nhds, metric.mem_nhds_iff] at hx, + obtain ⟨ε, hε, hx⟩ := hx, + refine hs.not_lt hx' (hx _) (lt_add_of_pos_right _ (by positivity : 0 < const ι (ε / 2))), + simpa [const, @pi_norm_const ι ℝ _ _ _ (ε / 2), abs_of_nonneg hε.lt.le], +end + +/-! +#### Note + +The closure and frontier of an antichain might not be antichains. Take for example the union +of the open segments from `(0, 2)` to `(1, 1)` and from `(2, 1)` to `(3, 0)`. `(1, 1)` and `(2, 1)` +are comparable and both in the closure/frontier. +-/ + +protected lemma is_closed.upper_closure (hs : is_closed s) (hs' : bdd_below s) : + is_closed (upper_closure s : set (ι → ℝ)) := +begin + casesI nonempty_fintype ι, + refine is_seq_closed.is_closed (λ f x hf hx, _), + choose g hg hgf using hf, + obtain ⟨a, ha⟩ := hx.bdd_above_range, + obtain ⟨b, hb, φ, hφ, hbf⟩ := tendsto_subseq_of_bounded (hs'.bounded_inter + bdd_above_Iic) (λ n, ⟨hg n, (hgf _).trans $ ha $ mem_range_self _⟩), + exact ⟨b, closure_minimal (inter_subset_left _ _) hs hb, + le_of_tendsto_of_tendsto' hbf (hx.comp hφ.tendsto_at_top) $ λ _, hgf _⟩, +end + +protected lemma is_closed.lower_closure (hs : is_closed s) (hs' : bdd_above s) : + is_closed (lower_closure s : set (ι → ℝ)) := +begin + casesI nonempty_fintype ι, + refine is_seq_closed.is_closed (λ f x hf hx, _), + choose g hg hfg using hf, + haveI : bounded_ge_nhds_class ℝ := by apply_instance, + obtain ⟨a, ha⟩ := hx.bdd_below_range, + obtain ⟨b, hb, φ, hφ, hbf⟩ := tendsto_subseq_of_bounded (hs'.bounded_inter + bdd_below_Ici) (λ n, ⟨hg n, (ha $ mem_range_self _).trans $ hfg _⟩), + exact ⟨b, closure_minimal (inter_subset_left _ _) hs hb, + le_of_tendsto_of_tendsto' (hx.comp hφ.tendsto_at_top) hbf $ λ _, hfg _⟩, +end + +protected lemma is_clopen.upper_closure (hs : is_clopen s) (hs' : bdd_below s) : + is_clopen (upper_closure s : set (ι → ℝ)) := +⟨hs.1.upper_closure, hs.2.upper_closure hs'⟩ + +protected lemma is_clopen.lower_closure (hs : is_clopen s) (hs' : bdd_above s) : + is_clopen (lower_closure s : set (ι → ℝ)) := +⟨hs.1.lower_closure, hs.2.lower_closure hs'⟩ + +lemma closure_upper_closure_comm (hs : bdd_below s) : + closure (upper_closure s : set (ι → ℝ)) = upper_closure (closure s) := +(closure_minimal (upper_closure_anti subset_closure) $ + is_closed_closure.upper_closure hs.closure).antisymm $ + upper_closure_min (closure_mono subset_upper_closure) (upper_closure s).upper.closure + +lemma closure_lower_closure_comm (hs : bdd_above s) : + closure (lower_closure s : set (ι → ℝ)) = lower_closure (closure s) := +(closure_minimal (lower_closure_mono subset_closure) $ + is_closed_closure.lower_closure hs.closure).antisymm $ + lower_closure_min (closure_mono subset_lower_closure) (lower_closure s).lower.closure + +end finite diff --git a/src/data/real/nnreal.lean b/src/data/real/nnreal.lean index cca419cf48609..47dba7306a2d2 100644 --- a/src/data/real/nnreal.lean +++ b/src/data/real/nnreal.lean @@ -822,6 +822,8 @@ lemma nnabs_coe (x : ℝ≥0) : nnabs x = x := by simp lemma coe_to_nnreal_le (x : ℝ) : (to_nnreal x : ℝ) ≤ |x| := max_le (le_abs_self _) (abs_nonneg _) +@[simp] lemma to_nnreal_abs (x : ℝ) : |x|.to_nnreal = x.nnabs := nnreal.coe_injective $ by simp + lemma cast_nat_abs_eq_nnabs_cast (n : ℤ) : (n.nat_abs : ℝ≥0) = nnabs n := by { ext, rw [nnreal.coe_nat_cast, int.cast_nat_abs, real.coe_nnabs] } diff --git a/src/measure_theory/order/upper_lower.lean b/src/measure_theory/order/upper_lower.lean new file mode 100644 index 0000000000000..d6bc49a4cfffa --- /dev/null +++ b/src/measure_theory/order/upper_lower.lean @@ -0,0 +1,131 @@ +/- +Copyright (c) 2022 Yaël Dillies, Kexing Ying. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies, Kexing Ying +-/ +import analysis.normed.order.upper_lower +import logic.lemmas +import measure_theory.covering.besicovitch_vector_space + +/-! +# Order-connected sets are null-measurable + +This file proves that order-connected sets in `ℝⁿ` under the pointwise order are null-measurable. +Recall that `x ≤ y` iff `∀ i, x i ≤ y i`, and `s` is order-connected iff +`∀ x y ∈ s, ∀ z, x ≤ z → z ≤ y → z ∈ s`. + +## Main declarations + +* `set.ord_connected.null_frontier`: The frontier of an order-connected set in `ℝⁿ` has measure `0`. + +## Notes + +We prove null-measurability in `ℝⁿ` with the `∞`-metric, but this transfers directly to `ℝⁿ` with +the Euclidean metric because they have the same measurable sets. + +Null-measurability can't be strengthened to measurability because any antichain (and in particular +any subset of the antidiagonal `{(x, y) | x + y = 0}`) is order-connected. + +## TODO + +Generalize so that it also applies to `ℝ × ℝ`, for example. +-/ + +open filter measure_theory metric set +open_locale topology + +variables {ι : Type*} [fintype ι] {s : set (ι → ℝ)} {x y : ι → ℝ} {δ : ℝ} + +/-- If we can fit a small ball inside a set `s` intersected with any neighborhood of `x`, then the +density of `s` near `x` is not `0`. Along with `aux₁`, this proves that `x` is a Lebesgue point of +`s`. This will be used to prove that the frontier of an order-connected set is null. -/ +private lemma aux₀ + (h : ∀ δ, 0 < δ → ∃ y, closed_ball y (δ/4) ⊆ closed_ball x δ ∧ closed_ball y (δ/4) ⊆ interior s) : + ¬ tendsto (λ r, volume (closure s ∩ closed_ball x r) / volume (closed_ball x r)) (𝓝[>] 0) + (𝓝 0) := +begin + choose f hf₀ hf₁ using h, + intros H, + obtain ⟨ε, hε, hε', hε₀⟩ := exists_seq_strict_anti_tendsto_nhds_within (0 : ℝ), + refine not_eventually.2 (frequently_of_forall $ λ _, lt_irrefl $ + ennreal.of_real $ 4⁻¹ ^ fintype.card ι) + ((tendsto.eventually_lt (H.comp hε₀) tendsto_const_nhds _).mono $ λ n, lt_of_le_of_lt _), + swap, + refine (ennreal.div_le_div_right (volume.mono $ subset_inter + ((hf₁ _ $ hε' n).trans interior_subset_closure) $ hf₀ _ $ hε' n) _).trans_eq' _, + dsimp, + have := hε' n, + rw [real.volume_pi_closed_ball, real.volume_pi_closed_ball, ←ennreal.of_real_div_of_pos, ←div_pow, + mul_div_mul_left _ _ (two_ne_zero' ℝ), div_right_comm, div_self, one_div], + all_goals { positivity }, +end + +/-- If we can fit a small ball inside a set `sᶜ` intersected with any neighborhood of `x`, then the +density of `s` near `x` is not `1`. Along with `aux₀`, this proves that `x` is a Lebesgue point of +`s`. This will be used to prove that the frontier of an order-connected set is null. -/ +private lemma aux₁ + (h : ∀ δ, 0 < δ → + ∃ y, closed_ball y (δ/4) ⊆ closed_ball x δ ∧ closed_ball y (δ/4) ⊆ interior sᶜ) : + ¬ tendsto (λ r, volume (closure s ∩ closed_ball x r) / volume (closed_ball x r)) (𝓝[>] 0) + (𝓝 1) := +begin + choose f hf₀ hf₁ using h, + intros H, + obtain ⟨ε, hε, hε', hε₀⟩ := exists_seq_strict_anti_tendsto_nhds_within (0 : ℝ), + refine not_eventually.2 (frequently_of_forall $ λ _, lt_irrefl $ + 1 - ennreal.of_real (4⁻¹ ^ fintype.card ι)) + ((tendsto.eventually_lt tendsto_const_nhds (H.comp hε₀) $ + ennreal.sub_lt_self ennreal.one_ne_top one_ne_zero _).mono $ λ n, lt_of_le_of_lt' _), + swap, + refine (ennreal.div_le_div_right (volume.mono _) _).trans_eq _, + { exact closed_ball x (ε n) \ closed_ball (f (ε n) $ hε' n) (ε n / 4) }, + { rw diff_eq_compl_inter, + refine inter_subset_inter_left _ _, + rw [subset_compl_comm, ←interior_compl], + exact hf₁ _ _ }, + dsimp, + have := hε' n, + rw [measure_diff (hf₀ _ _) _ ((real.volume_pi_closed_ball _ _).trans_ne ennreal.of_real_ne_top), + real.volume_pi_closed_ball, real.volume_pi_closed_ball, ennreal.sub_div (λ _ _, _), + ennreal.div_self _ ennreal.of_real_ne_top, ←ennreal.of_real_div_of_pos, ←div_pow, + mul_div_mul_left _ _ (two_ne_zero' ℝ), div_right_comm, div_self, one_div], + all_goals { positivity <|> measurability }, +end + +lemma is_upper_set.null_frontier (hs : is_upper_set s) : volume (frontier s) = 0 := +begin + refine eq_bot_mono (volume.mono $ λ x hx, _) + (besicovitch.ae_tendsto_measure_inter_div_of_measurable_set _ is_closed_closure.measurable_set), + { exact s }, + by_cases x ∈ closure s; simp [h], + { exact aux₁ (λ _, hs.compl.exists_subset_ball $ frontier_subset_closure $ + by rwa frontier_compl) }, + { exact aux₀ (λ _, hs.exists_subset_ball $ frontier_subset_closure hx) } +end + +lemma is_lower_set.null_frontier (hs : is_lower_set s) : volume (frontier s) = 0 := +begin + refine eq_bot_mono (volume.mono $ λ x hx, _) + (besicovitch.ae_tendsto_measure_inter_div_of_measurable_set _ is_closed_closure.measurable_set), + { exact s }, + by_cases x ∈ closure s; simp [h], + { exact aux₁ (λ _, hs.compl.exists_subset_ball $ frontier_subset_closure $ + by rwa frontier_compl) }, + { exact aux₀ (λ _, hs.exists_subset_ball $ frontier_subset_closure hx) } +end + +lemma set.ord_connected.null_frontier (hs : s.ord_connected) : volume (frontier s) = 0 := +begin + rw ← hs.upper_closure_inter_lower_closure, + refine le_bot_iff.1 ((volume.mono $ (frontier_inter_subset _ _).trans $ union_subset_union + (inter_subset_left _ _) $ inter_subset_right _ _).trans $ (measure_union_le _ _).trans_eq _), + rw [(upper_set.upper _).null_frontier, (lower_set.lower _).null_frontier, zero_add, bot_eq_zero], +end + +protected lemma set.ord_connected.null_measurable_set (hs : s.ord_connected) : + null_measurable_set s := +null_measurable_set_of_null_frontier hs.null_frontier + +lemma is_antichain.volume_eq_zero [nonempty ι] (hs : is_antichain (≤) s) : volume s = 0 := +le_bot_iff.1 $ (volume.mono $ by { rw [←closure_diff_interior, hs.interior_eq_empty, diff_empty], + exact subset_closure }).trans_eq hs.ord_connected.null_frontier diff --git a/src/order/bounds/basic.lean b/src/order/bounds/basic.lean index 6add0bb14d555..ab986bd7f2fd4 100644 --- a/src/order/bounds/basic.lean +++ b/src/order/bounds/basic.lean @@ -62,6 +62,8 @@ def is_glb (s : set α) : α → Prop := is_greatest (lower_bounds s) lemma mem_upper_bounds : a ∈ upper_bounds s ↔ ∀ x ∈ s, x ≤ a := iff.rfl lemma mem_lower_bounds : a ∈ lower_bounds s ↔ ∀ x ∈ s, a ≤ x := iff.rfl +lemma mem_upper_bounds_iff_subset_Iic : a ∈ upper_bounds s ↔ s ⊆ Iic a := iff.rfl +lemma mem_lower_bounds_iff_subset_Ici : a ∈ lower_bounds s ↔ s ⊆ Ici a := iff.rfl lemma bdd_above_def : bdd_above s ↔ ∃ x, ∀ y ∈ s, y ≤ x := iff.rfl lemma bdd_below_def : bdd_below s ↔ ∃ x, ∀ y ∈ s, x ≤ y := iff.rfl diff --git a/src/topology/algebra/order/upper_lower.lean b/src/topology/algebra/order/upper_lower.lean index 792ad0b431a15..c1286acac2cdc 100644 --- a/src/topology/algebra/order/upper_lower.lean +++ b/src/topology/algebra/order/upper_lower.lean @@ -5,6 +5,7 @@ Authors: Yaël Dillies -/ import algebra.order.upper_lower import topology.algebra.group.basic +import topology.order.basic /-! # Topological facts about upper/lower/order-connected sets @@ -46,7 +47,33 @@ instance ordered_comm_group.to_has_upper_lower_closure [ordered_comm_group α] is_open_upper_closure := λ s hs, by { rw [←mul_one s, ←mul_upper_closure], exact hs.mul_right }, is_open_lower_closure := λ s hs, by { rw [←mul_one s, ←mul_lower_closure], exact hs.mul_right } } -variables [preorder α] [has_upper_lower_closure α] {s : set α} +variables [preorder α] + +section order_closed_topology +variables [order_closed_topology α] {s : set α} + +@[simp] lemma upper_bounds_closure (s : set α) : + upper_bounds (closure s : set α) = upper_bounds s := +ext $ λ a, by simp_rw [mem_upper_bounds_iff_subset_Iic, is_closed_Iic.closure_subset_iff] + +@[simp] lemma lower_bounds_closure (s : set α) : + lower_bounds (closure s : set α) = lower_bounds s := +ext $ λ a, by simp_rw [mem_lower_bounds_iff_subset_Ici, is_closed_Ici.closure_subset_iff] + +@[simp] lemma bdd_above_closure : bdd_above (closure s) ↔ bdd_above s := +by simp_rw [bdd_above, upper_bounds_closure] + +@[simp] lemma bdd_below_closure : bdd_below (closure s) ↔ bdd_below s := +by simp_rw [bdd_below, lower_bounds_closure] + +alias bdd_above_closure ↔ bdd_above.of_closure bdd_above.closure +alias bdd_below_closure ↔ bdd_below.of_closure bdd_below.closure + +attribute [protected] bdd_above.closure bdd_below.closure + +end order_closed_topology + +variables [has_upper_lower_closure α] {s : set α} protected lemma is_upper_set.closure : is_upper_set s → is_upper_set (closure s) := has_upper_lower_closure.is_upper_set_closure _ From 3365b20c2ffa7c35e47e5209b89ba9abdddf3ffe Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sun, 22 Oct 2023 19:51:38 +0000 Subject: [PATCH 60/61] feat(combinatorics/simple_graph): More clique lemmas (#19203) More lemmas about `is_clique`, `is_n_clique`, `edge_set`. Also define `clique_free_on`, a local version of `clique_free`. --- src/combinatorics/simple_graph/basic.lean | 122 +++++++++- src/combinatorics/simple_graph/clique.lean | 211 ++++++++++++++++-- .../simple_graph/triangle/basic.lean | 3 +- src/data/finset/card.lean | 7 + src/data/finset/preimage.lean | 4 + src/logic/basic.lean | 9 + src/logic/relation.lean | 23 +- 7 files changed, 349 insertions(+), 30 deletions(-) diff --git a/src/combinatorics/simple_graph/basic.lean b/src/combinatorics/simple_graph/basic.lean index 5c66632b5e4b5..0b1d706d9e560 100644 --- a/src/combinatorics/simple_graph/basic.lean +++ b/src/combinatorics/simple_graph/basic.lean @@ -3,6 +3,7 @@ Copyright (c) 2020 Aaron Anderson, Jalex Stark, Kyle Miller. All rights reserved Released under Apache 2.0 license as described in the file LICENSE. Authors: Aaron Anderson, Jalex Stark, Kyle Miller, Alena Gusakov, Hunter Monroe -/ +import data.fun_like.fintype import data.rel import data.set.finite import data.sym.sym2 @@ -288,6 +289,13 @@ instance : complete_boolean_algebra (simple_graph V) := @[simps] instance (V : Type u) : inhabited (simple_graph V) := ⟨⊥⟩ +instance [subsingleton V] : unique (simple_graph V) := +{ default := ⊥, + uniq := λ G, by ext a b; simp [subsingleton.elim a b] } + +instance [nontrivial V] : nontrivial (simple_graph V) := +⟨⟨⊥, ⊤, λ h, not_subsingleton V ⟨by simpa [ext_iff, function.funext_iff] using h⟩⟩⟩ + section decidable variables (V) (H : simple_graph V) [decidable_rel G.adj] [decidable_rel H.adj] @@ -372,6 +380,17 @@ by { ext ⟨x, y⟩, refl } @[simp] lemma edge_set_sdiff : (G₁ \ G₂).edge_set = G₁.edge_set \ G₂.edge_set := by { ext ⟨x, y⟩, refl } +variables {G G₁ G₂} + +@[simp] lemma disjoint_edge_set : disjoint G₁.edge_set G₂.edge_set ↔ disjoint G₁ G₂ := +by rw [set.disjoint_iff, disjoint_iff_inf_le, ←edge_set_inf, ←edge_set_bot, ←set.le_iff_subset, + order_embedding.le_iff_le] + +@[simp] lemma edge_set_eq_empty : G.edge_set = ∅ ↔ G = ⊥ := by rwa [←edge_set_bot, edge_set_inj] + +@[simp] lemma edge_set_nonempty : G.edge_set.nonempty ↔ G ≠ ⊥ := +by rw [set.nonempty_iff_ne_empty, edge_set_eq_empty.ne] + /-- This lemma, combined with `edge_set_sdiff` and `edge_set_from_edge_set`, allows proving `(G \ from_edge_set s).edge_set = G.edge_set \ s` by `simp`. @@ -406,6 +425,8 @@ end lemma adj_iff_exists_edge_coe : G.adj a b ↔ ∃ (e : G.edge_set), ↑e = ⟦(a, b)⟧ := by simp only [mem_edge_set, exists_prop, set_coe.exists, exists_eq_right, subtype.coe_mk] +variables (G G₁ G₂) + lemma edge_other_ne {e : sym2 V} (he : e ∈ G.edge_set) {v : V} (h : v ∈ e) : h.other ≠ v := begin erw [← sym2.other_spec h, sym2.eq_swap] at he, @@ -482,6 +503,16 @@ begin exact λ vws _, h vws, end +@[simp] lemma disjoint_from_edge_set : disjoint G (from_edge_set s) ↔ disjoint G.edge_set s := +begin + conv_rhs { rw ←set.diff_union_inter s {e | e.is_diag} }, + rw [←disjoint_edge_set, edge_set_from_edge_set, set.disjoint_union_right, and_iff_left], + exact set.disjoint_left.2 (λ e he he', not_is_diag_of_mem_edge_set _ he he'.2), +end + +@[simp] lemma from_edge_set_disjoint : disjoint (from_edge_set s) G ↔ disjoint s G.edge_set := +by rw [disjoint.comm, disjoint_from_edge_set, disjoint.comm] + instance [decidable_eq V] [fintype s] : fintype (from_edge_set s).edge_set := by { rw edge_set_from_edge_set s, apply_instance } @@ -680,6 +711,9 @@ end edge_finset @[simp] lemma mem_neighbor_set (v w : V) : w ∈ G.neighbor_set v ↔ G.adj v w := iff.rfl +@[simp] lemma not_mem_neighbor_set_self : a ∉ G.neighbor_set a := +(mem_neighbor_set _ _ _).not.2 $ G.loopless _ + @[simp] lemma mem_incidence_set (v w : V) : ⟦(v, w)⟧ ∈ G.incidence_set v ↔ G.adj v w := by simp [incidence_set] @@ -821,6 +855,9 @@ lemma delete_edges_eq_sdiff_from_edge_set (s : set (sym2 V)) : G.delete_edges s = G \ from_edge_set s := by { ext, exact ⟨λ h, ⟨h.1, not_and_of_not_left _ h.2⟩, λ h, ⟨h.1, not_and'.mp h.2 h.ne⟩⟩ } +@[simp] lemma delete_edges_eq {s : set (sym2 V)} : G.delete_edges s = G ↔ disjoint G.edge_set s := +by rw [delete_edges_eq_sdiff_from_edge_set, sdiff_eq_left, disjoint_from_edge_set] + lemma compl_eq_delete_edges : Gᶜ = (⊤ : simple_graph V).delete_edges G.edge_set := by { ext, simp } @@ -871,18 +908,17 @@ get a graph with the property `p`. -/ def delete_far (p : simple_graph V → Prop) (r : 𝕜) : Prop := ∀ ⦃s⦄, s ⊆ G.edge_finset → p (G.delete_edges s) → r ≤ s.card -open_locale classical - variables {G} lemma delete_far_iff : - G.delete_far p r ↔ ∀ ⦃H⦄, H ≤ G → p H → r ≤ G.edge_finset.card - H.edge_finset.card := + G.delete_far p r ↔ ∀ ⦃H : simple_graph _⦄ [decidable_rel H.adj], + by exactI H ≤ G → p H → r ≤ G.edge_finset.card - H.edge_finset.card := begin - refine ⟨λ h H hHG hH, _, λ h s hs hG, _⟩, + refine ⟨λ h H _ hHG hH, _, λ h s hs hG, _⟩, { have := h (sdiff_subset G.edge_finset H.edge_finset), simp only [delete_edges_sdiff_eq_of_le _ hHG, edge_finset_mono hHG, card_sdiff, card_le_of_subset, coe_sdiff, coe_edge_finset, nat.cast_sub] at this, - exact this hH }, + convert this hH }, { simpa [card_sdiff hs, edge_finset_delete_edges, -set.to_finset_card, nat.cast_sub, card_le_of_subset hs] using h (G.delete_edges_le s) hG } end @@ -906,9 +942,22 @@ protected def map (f : V ↪ W) (G : simple_graph V) : simple_graph W := @[simp] lemma map_adj (f : V ↪ W) (G : simple_graph V) (u v : W) : (G.map f).adj u v ↔ ∃ (u' v' : V), G.adj u' v' ∧ f u' = u ∧ f v' = v := iff.rfl +lemma map_adj_apply {G : simple_graph V} {f : V ↪ W} {a b : V} : + (G.map f).adj (f a) (f b) ↔ G.adj a b := by simp + lemma map_monotone (f : V ↪ W) : monotone (simple_graph.map f) := by { rintros G G' h _ _ ⟨u, v, ha, rfl, rfl⟩, exact ⟨_, _, h ha, rfl, rfl⟩ } +@[simp] lemma map_id : G.map (function.embedding.refl _) = G := +ext _ _ $ relation.map_id_id _ + +@[simp] lemma map_map (f : V ↪ W) (g : W ↪ X) : (G.map f).map g = G.map (f.trans g) := +ext _ _ $ relation.map_map _ _ _ _ _ + +instance decidable_map (f : V ↪ W) (G : simple_graph V) [decidable_rel (relation.map G.adj f f)] : + decidable_rel (G.map f).adj := +‹decidable_rel _› + /-- Given a function, there is a contravariant induced map on graphs by pulling back the adjacency relation. This is one of the ways of creating induced graphs. See `simple_graph.induce` for a wrapper. @@ -917,6 +966,24 @@ This is surjective when `f` is injective (see `simple_graph.comap_surjective`).- @[simps] protected def comap (f : V → W) (G : simple_graph W) : simple_graph V := { adj := λ u v, G.adj (f u) (f v) } +@[simp] lemma comap_id {G : simple_graph V} : G.comap id = G := ext _ _ rfl + +@[simp] lemma comap_comap {G : simple_graph X} (f : V → W) (g : W → X) : + (G.comap g).comap f = G.comap (g ∘ f) := rfl + +instance decidable_comap (f : V → W) (G : simple_graph W) [decidable_rel G.adj] : + decidable_rel (simple_graph.comap f G).adj := +λ _ _, ‹decidable_rel G.adj› _ _ + +lemma comap_symm (G : simple_graph V) (e : V ≃ W) : + G.comap e.symm.to_embedding = G.map e.to_embedding := +by { ext, simp only [equiv.apply_eq_iff_eq_symm_apply, comap_adj, map_adj, equiv.to_embedding_apply, + exists_eq_right_right, exists_eq_right] } + +lemma map_symm (G : simple_graph W) (e : V ≃ W) : + G.map e.symm.to_embedding = G.comap e.to_embedding := +by rw [←comap_symm, e.symm_symm] + lemma comap_monotone (f : V ↪ W) : monotone (simple_graph.comap f) := by { intros G G' h _ _ ha, exact h ha } @@ -939,6 +1006,23 @@ lemma map_le_iff_le_comap (f : V ↪ W) (G : simple_graph V) (G' : simple_graph lemma map_comap_le (f : V ↪ W) (G : simple_graph W) : (G.comap f).map f ≤ G := by { rw map_le_iff_le_comap, exact le_refl _ } +/-- Equivalent types have equivalent simple graphs. -/ +@[simps apply] protected def _root_.equiv.simple_graph (e : V ≃ W) : + simple_graph V ≃ simple_graph W := +{ to_fun := simple_graph.comap e.symm, + inv_fun := simple_graph.comap e, + left_inv := λ _, by simp, + right_inv := λ _, by simp } + +@[simp] lemma _root_.equiv.simple_graph_refl : (equiv.refl V).simple_graph = equiv.refl _ := +by { ext, refl } + +@[simp] lemma _root_.equiv.simple_graph_trans (e₁ : V ≃ W) (e₂ : W ≃ X) : + (e₁.trans e₂).simple_graph = e₁.simple_graph.trans e₂.simple_graph := rfl + +@[simp] lemma _root_.equiv.symm_simple_graph (e : V ≃ W) : + e.simple_graph.symm = e.symm.simple_graph := rfl + /-! ## Induced graphs -/ /- Given a set `s` of vertices, we can restrict a graph to those vertices by restricting its @@ -1296,10 +1380,23 @@ infix ` ↪g ` : 50 := embedding infix ` ≃g ` : 50 := iso namespace hom -variables {G G'} (f : G →g G') +variables {G G'} {H : simple_graph W} (f : G →g G') /-- The identity homomorphism from a graph to itself. -/ -abbreviation id : G →g G := rel_hom.id _ +protected abbreviation id : G →g G := rel_hom.id _ + +@[simp, norm_cast] lemma coe_id : ⇑(hom.id : G →g G) = _root_.id := rfl + +instance [subsingleton (V → W)] : subsingleton (G →g H) := fun_like.coe_injective.subsingleton + +instance [is_empty V] : unique (G →g H) := +{ default := ⟨is_empty_elim, is_empty_elim⟩, + uniq := λ _, subsingleton.elim _ _ } + +noncomputable instance [fintype V] [fintype W] : fintype (G →g H) := +by classical; exact fun_like.fintype _ + +instance [finite V] [finite W] : finite (G →g H) := fun_like.finite _ lemma map_adj {v w : V} (h : G.adj v w) : G'.adj (f v) (f w) := f.map_rel' h @@ -1358,10 +1455,15 @@ abbreviation comp (f' : G' →g G'') (f : G →g G') : G →g G'' := f'.comp f @[simp] lemma coe_comp (f' : G' →g G'') (f : G →g G') : ⇑(f'.comp f) = f' ∘ f := rfl +/-- The graph homomorphism from a smaller graph to a bigger one. -/ +def of_le {H : simple_graph V} (h : G ≤ H) : G →g H := ⟨id, h⟩ + +@[simp, norm_cast] lemma coe_of_le {H : simple_graph V} (h : G ≤ H) : ⇑(of_le h) = id := rfl + end hom namespace embedding -variables {G G'} (f : G ↪g G') +variables {G G'} {H : simple_graph W} (f : G ↪g G') /-- The identity embedding from a graph to itself. -/ abbreviation refl : G ↪g G := rel_embedding.refl _ @@ -1369,7 +1471,9 @@ abbreviation refl : G ↪g G := rel_embedding.refl _ /-- An embedding of graphs gives rise to a homomorphism of graphs. -/ abbreviation to_hom : G →g G' := f.to_rel_hom -lemma map_adj_iff {v w : V} : G'.adj (f v) (f w) ↔ G.adj v w := f.map_rel_iff +@[simp] lemma coe_to_hom (f : G ↪g H) : ⇑f.to_hom = f := rfl + +@[simp] lemma map_adj_iff {v w : V} : G'.adj (f v) (f w) ↔ G.adj v w := f.map_rel_iff lemma map_mem_edge_set_iff {e : sym2 V} : e.map f ∈ G'.edge_set ↔ e ∈ G.edge_set := quotient.ind (λ ⟨v, w⟩, f.map_adj_iff) e diff --git a/src/combinatorics/simple_graph/clique.lean b/src/combinatorics/simple_graph/clique.lean index 996aeb3dd5c72..8d2b10d255ca5 100644 --- a/src/combinatorics/simple_graph/clique.lean +++ b/src/combinatorics/simple_graph/clique.lean @@ -5,6 +5,7 @@ Authors: Yaël Dillies, Bhavik Mehta -/ import combinatorics.simple_graph.basic import data.finset.pairwise +import data.finset.preimage /-! # Graph cliques @@ -25,13 +26,13 @@ adjacent. ## TODO * Clique numbers -* Do we need `clique_set`, a version of `clique_finset` for infinite graphs? +* Dualise all the API to get independent sets -/ -open finset fintype +open finset fintype function namespace simple_graph -variables {α : Type*} (G H : simple_graph α) +variables {α β : Type*} (G H : simple_graph α) /-! ### Cliques -/ @@ -61,13 +62,33 @@ end instance [decidable_eq α] [decidable_rel G.adj] {s : finset α} : decidable (G.is_clique s) := decidable_of_iff' _ G.is_clique_iff -variables {G H} +variables {G H} {a b : α} + +@[simp] lemma is_clique_empty : G.is_clique ∅ := set.pairwise_empty _ +@[simp] lemma is_clique_singleton (a : α) : G.is_clique {a} := set.pairwise_singleton _ _ + +lemma is_clique_pair : G.is_clique {a, b} ↔ a ≠ b → G.adj a b := +set.pairwise_pair_of_symmetric G.symm + +@[simp] lemma is_clique_insert : + G.is_clique (insert a s) ↔ G.is_clique s ∧ ∀ b ∈ s, a ≠ b → G.adj a b := +set.pairwise_insert_of_symmetric G.symm -lemma is_clique.mono (h : G ≤ H) : G.is_clique s → H.is_clique s := -by { simp_rw is_clique_iff, exact set.pairwise.mono' h } +lemma is_clique_insert_of_not_mem (ha : a ∉ s) : + G.is_clique (insert a s) ↔ G.is_clique s ∧ ∀ b ∈ s, G.adj a b := +set.pairwise_insert_of_symmetric_of_not_mem G.symm ha -lemma is_clique.subset (h : t ⊆ s) : G.is_clique s → G.is_clique t := -by { simp_rw is_clique_iff, exact set.pairwise.mono h } +lemma is_clique.insert (hs : G.is_clique s) (h : ∀ b ∈ s, a ≠ b → G.adj a b) : + G.is_clique (insert a s) := +hs.insert_of_symmetric G.symm h + +lemma is_clique.mono (h : G ≤ H) : G.is_clique s → H.is_clique s := set.pairwise.mono' h +lemma is_clique.subset (h : t ⊆ s) : G.is_clique s → G.is_clique t := set.pairwise.mono h + +protected lemma is_clique.map {G : simple_graph α} {s : set α} (h : G.is_clique s) {f : α ↪ β} : + (G.map f).is_clique (f '' s) := +by { rintro _ ⟨a, ha, rfl⟩ _ ⟨b, hb, rfl⟩ hab, + exact ⟨a, b, h ha hb $ ne_of_apply_ne _ hab, rfl, rfl⟩ } @[simp] lemma is_clique_bot_iff : (⊥ : simple_graph α).is_clique s ↔ (s : set α).subsingleton := set.pairwise_bot_iff @@ -93,11 +114,19 @@ instance [decidable_eq α] [decidable_rel G.adj] {n : ℕ} {s : finset α} : decidable (G.is_n_clique n s) := decidable_of_iff' _ G.is_n_clique_iff -variables {G H} +variables {G H} {a b c : α} + +@[simp] lemma is_n_clique_empty : G.is_n_clique n ∅ ↔ n = 0 := by simp [is_n_clique_iff, eq_comm] +@[simp] lemma is_n_clique_singleton : G.is_n_clique n {a} ↔ n = 1 := +by simp [is_n_clique_iff, eq_comm] lemma is_n_clique.mono (h : G ≤ H) : G.is_n_clique n s → H.is_n_clique n s := by { simp_rw is_n_clique_iff, exact and.imp_left (is_clique.mono h) } +protected lemma is_n_clique.map (h : G.is_n_clique n s) {f : α ↪ β} : + (G.map f).is_n_clique n (s.map f) := +⟨by { rw coe_map, exact h.1.map}, (card_map _).trans h.2⟩ + @[simp] lemma is_n_clique_bot_iff : (⊥ : simple_graph α).is_n_clique n s ↔ n ≤ 1 ∧ s.card = n := begin rw [is_n_clique_iff, is_clique_bot_iff], @@ -106,7 +135,22 @@ begin exact card_le_one.symm, end -variables [decidable_eq α] {a b c : α} +@[simp] lemma is_n_clique_zero : G.is_n_clique 0 s ↔ s = ∅ := +by { simp only [is_n_clique_iff, finset.card_eq_zero, and_iff_right_iff_imp], rintro rfl, simp } + +@[simp] lemma is_n_clique_one : G.is_n_clique 1 s ↔ ∃ a, s = {a} := +by { simp only [is_n_clique_iff, card_eq_one, and_iff_right_iff_imp], rintro ⟨a, rfl⟩, simp } + +variables [decidable_eq α] + +lemma is_n_clique.insert (hs : G.is_n_clique n s) (h : ∀ b ∈ s, G.adj a b) : + G.is_n_clique (n + 1) (insert a s) := +begin + split, + { push_cast, + exact hs.1.insert (λ b hb _, h _ hb) }, + { rw [card_insert_of_not_mem (λ ha, (h _ ha).ne rfl), hs.2] } +end lemma is_3_clique_triple_iff : G.is_n_clique 3 {a, b, c} ↔ G.adj a b ∧ G.adj a c ∧ G.adj b c := begin @@ -193,7 +237,7 @@ begin use (iso.complete_graph (fintype.equiv_fin α)).symm.to_embedding.trans f, end -lemma clique_free_bot (h : 2 ≤ n) : (⊥ : simple_graph α).clique_free n := +@[simp] lemma clique_free_bot (h : 2 ≤ n) : (⊥ : simple_graph α).clique_free n := begin rintro t ht, rw is_n_clique_bot_iff at ht, @@ -219,8 +263,77 @@ begin simpa using fintype.card_le_of_embedding h.some.to_embedding, end +@[simp] lemma clique_free_two : G.clique_free 2 ↔ G = ⊥ := +begin + classical, + split, + { simp_rw [←edge_set_eq_empty, set.eq_empty_iff_forall_not_mem, sym2.forall, mem_edge_set], + exact λ h a b hab, h _ ⟨by simpa [hab.ne], card_doubleton hab.ne⟩ }, + { rintro rfl, + exact clique_free_bot le_rfl } +end + end clique_free +section clique_free_on +variables {s s₁ s₂ : set α} {t : finset α} {a b : α} {m n : ℕ} + +/-- `G.clique_free_on s n` means that `G` has no `n`-cliques contained in `s`. -/ +def clique_free_on (G : simple_graph α) (s : set α) (n : ℕ) : Prop := +∀ ⦃t⦄, ↑t ⊆ s → ¬G.is_n_clique n t + +lemma clique_free_on.subset (hs : s₁ ⊆ s₂) (h₂ : G.clique_free_on s₂ n) : G.clique_free_on s₁ n := +λ t hts, h₂ $ hts.trans hs + +lemma clique_free_on.mono (hmn : m ≤ n) (hG : G.clique_free_on s m) : G.clique_free_on s n := +begin + rintro t hts ht, + obtain ⟨u, hut, hu⟩ := t.exists_smaller_set _ (hmn.trans ht.card_eq.ge), + exact hG ((coe_subset.2 hut).trans hts) ⟨ht.clique.subset hut, hu⟩, +end + +lemma clique_free_on.anti (hGH : G ≤ H) (hH : H.clique_free_on s n) : G.clique_free_on s n := +λ t hts ht, hH hts $ ht.mono hGH + +@[simp] lemma clique_free_on_empty : G.clique_free_on ∅ n ↔ n ≠ 0 := +by simp [clique_free_on, set.subset_empty_iff] + +@[simp] lemma clique_free_on_singleton : G.clique_free_on {a} n ↔ 1 < n := +by obtain _ | _ | n := n; simp [clique_free_on, set.subset_singleton_iff_eq] + +@[simp] lemma clique_free_on_univ : G.clique_free_on set.univ n ↔ G.clique_free n := +by simp [clique_free, clique_free_on] + +protected lemma clique_free.clique_free_on (hG : G.clique_free n) : G.clique_free_on s n := +λ t _, hG _ + +lemma clique_free_on_of_card_lt {s : finset α} (h : s.card < n) : G.clique_free_on s n := +λ t hts ht, h.not_le $ ht.2.symm.trans_le $ card_mono hts + +--TOOD: Restate using `simple_graph.indep_set` once we have it +@[simp] lemma clique_free_on_two : G.clique_free_on s 2 ↔ s.pairwise G.adjᶜ := +begin + classical, + refine ⟨λ h a ha b hb _ hab, h _ ⟨by simpa [hab.ne], card_doubleton hab.ne⟩, _⟩, + { push_cast, + exact set.insert_subset.2 ⟨ha, set.singleton_subset_iff.2 hb⟩ }, + simp only [clique_free_on, is_n_clique_iff, card_eq_two, coe_subset, not_and, not_exists], + rintro h t hst ht a b hab rfl, + simp only [coe_insert, coe_singleton, set.insert_subset, set.singleton_subset_iff] at hst, + refine h hst.1 hst.2 hab (ht _ _ hab); simp, +end + +lemma clique_free_on.of_succ (hs : G.clique_free_on s (n + 1)) (ha : a ∈ s) : + G.clique_free_on (s ∩ G.neighbor_set a) n := +begin + classical, + refine λ t hts ht, hs _ (ht.insert $ λ b hb, (hts hb).2), + push_cast, + exact set.insert_subset.2 ⟨ha, hts.trans $ set.inter_subset_left _ _⟩, +end + +end clique_free_on + /-! ### Set of cliques -/ section clique_set @@ -229,22 +342,60 @@ variables (G) {n : ℕ} {a b c : α} {s : finset α} /-- The `n`-cliques in a graph as a set. -/ def clique_set (n : ℕ) : set (finset α) := {s | G.is_n_clique n s} -lemma mem_clique_set_iff : s ∈ G.clique_set n ↔ G.is_n_clique n s := iff.rfl +@[simp] lemma mem_clique_set_iff : s ∈ G.clique_set n ↔ G.is_n_clique n s := iff.rfl @[simp] lemma clique_set_eq_empty_iff : G.clique_set n = ∅ ↔ G.clique_free n := by simp_rw [clique_free, set.eq_empty_iff_forall_not_mem, mem_clique_set_iff] -alias clique_set_eq_empty_iff ↔ _ clique_free.clique_set - -attribute [protected] clique_free.clique_set - variables {G H} +protected lemma clique_free.clique_set : G.clique_free n → G.clique_set n = ∅ := +G.clique_set_eq_empty_iff.2 + @[mono] lemma clique_set_mono (h : G ≤ H) : G.clique_set n ⊆ H.clique_set n := λ _, is_n_clique.mono h lemma clique_set_mono' (h : G ≤ H) : G.clique_set ≤ H.clique_set := λ _, clique_set_mono h +@[simp] lemma clique_set_zero (G : simple_graph α) : G.clique_set 0 = {∅} := +set.ext $ λ s, by simp + +@[simp] lemma clique_set_one (G : simple_graph α) : G.clique_set 1 = set.range singleton := +set.ext $ λ s, by simp [eq_comm] + +@[simp] lemma clique_set_bot (hn : 1 < n) : (⊥ : simple_graph α).clique_set n = ∅ := +(clique_free_bot hn).clique_set + +@[simp] lemma clique_set_map (hn : n ≠ 1) (G : simple_graph α) (f : α ↪ β) : + (G.map f).clique_set n = map f '' G.clique_set n := +begin + ext s, + split, + { rintro ⟨hs, rfl⟩, + have hs' : (s.preimage f $ f.injective.inj_on _).map f = s, + { classical, + rw [map_eq_image, image_preimage, filter_true_of_mem], + rintro a ha, + obtain ⟨b, hb, hba⟩ := exists_mem_ne (hn.lt_of_le' $ finset.card_pos.2 ⟨a, ha⟩) a, + obtain ⟨c, _, _, hc, _⟩ := hs ha hb hba.symm, + exact ⟨c, hc⟩ }, + refine ⟨s.preimage f $ f.injective.inj_on _, ⟨_, by rw [←card_map f, hs']⟩, hs'⟩, + rw coe_preimage, + exact λ a ha b hb hab, map_adj_apply.1 (hs ha hb $ f.injective.ne hab) }, + { rintro ⟨s, hs, rfl⟩, + exact is_n_clique.map hs } +end + +@[simp] lemma clique_set_map_of_equiv (G : simple_graph α) (e : α ≃ β) (n : ℕ) : + (G.map e.to_embedding).clique_set n = map e.to_embedding '' G.clique_set n := +begin + obtain rfl | hn := eq_or_ne n 1, + { ext, + simp [e.exists_congr_left] }, + { exact clique_set_map hn _ _ } +end + + end clique_set /-! ### Finset of cliques -/ @@ -255,10 +406,11 @@ variables (G) [fintype α] [decidable_eq α] [decidable_rel G.adj] {n : ℕ} {a /-- The `n`-cliques in a graph as a finset. -/ def clique_finset (n : ℕ) : finset (finset α) := univ.filter $ G.is_n_clique n -lemma mem_clique_finset_iff : s ∈ G.clique_finset n ↔ G.is_n_clique n s := +@[simp] lemma mem_clique_finset_iff : s ∈ G.clique_finset n ↔ G.is_n_clique n s := mem_filter.trans $ and_iff_right $ mem_univ _ -@[simp] lemma coe_clique_finset (n : ℕ) : (G.clique_finset n : set (finset α)) = G.clique_set n := +@[simp, norm_cast] lemma coe_clique_finset (n : ℕ) : + (G.clique_finset n : set (finset α)) = G.clique_set n := set.ext $ λ _, mem_clique_finset_iff _ @[simp] lemma clique_finset_eq_empty_iff : G.clique_finset n = ∅ ↔ G.clique_free n := @@ -268,10 +420,31 @@ alias clique_finset_eq_empty_iff ↔ _ _root_.simple_graph.clique_free.clique_fi attribute [protected] clique_free.clique_finset -variables {G} [decidable_rel H.adj] +variables {G} + +lemma card_clique_finset_le : (G.clique_finset n).card ≤ (card α).choose n := +begin + rw [←card_univ, ←card_powerset_len], + refine card_mono (λ s, _), + simpa [mem_powerset_len_univ_iff] using is_n_clique.card_eq, +end + +variables [decidable_rel H.adj] @[mono] lemma clique_finset_mono (h : G ≤ H) : G.clique_finset n ⊆ H.clique_finset n := monotone_filter_right _ $ λ _, is_n_clique.mono h +variables [fintype β] [decidable_eq β] (G) + +@[simp] lemma clique_finset_map (f : α ↪ β) (hn : n ≠ 1) : + (G.map f).clique_finset n = (G.clique_finset n).map ⟨map f, finset.map_injective _⟩ := +coe_injective $ + by simp_rw [coe_clique_finset, clique_set_map hn, coe_map, coe_clique_finset, embedding.coe_fn_mk] + +@[simp] lemma clique_finset_map_of_equiv (e : α ≃ β) (n : ℕ) : + (G.map e.to_embedding).clique_finset n = + (G.clique_finset n).map ⟨map e.to_embedding, finset.map_injective _⟩ := +coe_injective $ by push_cast; exact clique_set_map_of_equiv _ _ _ + end clique_finset end simple_graph diff --git a/src/combinatorics/simple_graph/triangle/basic.lean b/src/combinatorics/simple_graph/triangle/basic.lean index 639207f75591c..38c350e394e93 100644 --- a/src/combinatorics/simple_graph/triangle/basic.lean +++ b/src/combinatorics/simple_graph/triangle/basic.lean @@ -42,7 +42,8 @@ G.delete_far (λ H, H.clique_free 3) $ ε * (card α^2 : ℕ) lemma far_from_triangle_free_iff : G.far_from_triangle_free ε ↔ - ∀ ⦃H⦄, H ≤ G → H.clique_free 3 → ε * (card α^2 : ℕ) ≤ G.edge_finset.card - H.edge_finset.card := + ∀ ⦃H : simple_graph _⦄ [decidable_rel H.adj], by exactI + H ≤ G → H.clique_free 3 → ε * (card α^2 : ℕ) ≤ G.edge_finset.card - H.edge_finset.card := delete_far_iff alias far_from_triangle_free_iff ↔ far_from_triangle_free.le_card_sub_card _ diff --git a/src/data/finset/card.lean b/src/data/finset/card.lean index 98cf0399b0e58..f023ef54c1104 100644 --- a/src/data/finset/card.lean +++ b/src/data/finset/card.lean @@ -434,6 +434,13 @@ begin exact card_le_of_subset hx } end +lemma exists_mem_ne (hs : 1 < s.card) (a : α) : ∃ b ∈ s, b ≠ a := +begin + by_contra', + haveI : nonempty α := ⟨a⟩, + exact hs.not_le (card_le_one_iff_subset_singleton.2 ⟨a, subset_singleton_iff'.2 this⟩), +end + /-- A `finset` of a subsingleton type has cardinality at most one. -/ lemma card_le_one_of_subsingleton [subsingleton α] (s : finset α) : s.card ≤ 1 := finset.card_le_one_iff.2 $ λ _ _ _ _, subsingleton.elim _ _ diff --git a/src/data/finset/preimage.lean b/src/data/finset/preimage.lean index ba9639d06700d..afe69fac91482 100644 --- a/src/data/finset/preimage.lean +++ b/src/data/finset/preimage.lean @@ -61,6 +61,10 @@ finset.coe_injective (by simp) preimage sᶜ f (hf.inj_on _) = (preimage s f (hf.inj_on _))ᶜ := finset.coe_injective (by simp) +@[simp] lemma preimage_map (f : α ↪ β) (s : finset α) : + (s.map f).preimage f (f.injective.inj_on _) = s := +coe_injective $ by simp only [coe_preimage, coe_map, set.preimage_image_eq _ f.injective] + lemma monotone_preimage {f : α → β} (h : injective f) : monotone (λ s, preimage s f (h.inj_on _)) := λ s t hst x hx, mem_preimage.2 (hst $ mem_preimage.1 hx) diff --git a/src/logic/basic.lean b/src/logic/basic.lean index 71baa00527e9d..26ef5b00ad237 100644 --- a/src/logic/basic.lean +++ b/src/logic/basic.lean @@ -423,6 +423,15 @@ lemma iff.not (h : a ↔ b) : ¬ a ↔ ¬ b := not_congr h lemma iff.not_left (h : a ↔ ¬ b) : ¬ a ↔ b := h.not.trans not_not lemma iff.not_right (h : ¬ a ↔ b) : a ↔ ¬ b := not_not.symm.trans h.not +protected lemma iff.ne {α β : Sort*} {a b : α} {c d : β} : (a = b ↔ c = d) → (a ≠ b ↔ c ≠ d) := +iff.not + +lemma iff.ne_left {α β : Sort*} {a b : α} {c d : β} : (a = b ↔ c ≠ d) → (a ≠ b ↔ c = d) := +iff.not_left + +lemma iff.ne_right {α β : Sort*} {a b : α} {c d : β} : (a ≠ b ↔ c = d) → (a = b ↔ c ≠ d) := +iff.not_right + /-! ### Declarations about `xor` -/ @[simp] theorem xor_true : xor true = not := funext $ λ a, by simp [xor] diff --git a/src/logic/relation.lean b/src/logic/relation.lean index 8467caa176c36..466ad575131e0 100644 --- a/src/logic/relation.lean +++ b/src/logic/relation.lean @@ -42,7 +42,7 @@ the bundled version, see `rel`. open function -variables {α β γ δ : Type*} +variables {α β γ δ ε κ : Type*} section ne_imp @@ -186,6 +186,27 @@ related by `r`. protected def map (r : α → β → Prop) (f : α → γ) (g : β → δ) : γ → δ → Prop := λ c d, ∃ a b, r a b ∧ f a = c ∧ g b = d +section map +variables {r : α → β → Prop} {f : α → γ} {g : β → δ} {c : γ} {d : δ} + +lemma map_apply : relation.map r f g c d ↔ ∃ a b, r a b ∧ f a = c ∧ g b = d := iff.rfl + +@[simp] lemma map_id_id (r : α → β → Prop) : relation.map r id id = r := by simp [relation.map] + +@[simp] lemma map_map (r : α → β → Prop) (f₁ : α → γ) (g₁ : β → δ) (f₂ : γ → ε) (g₂ : δ → κ) : + relation.map (relation.map r f₁ g₁) f₂ g₂ = relation.map r (f₂ ∘ f₁) (g₂ ∘ g₁) := +begin + ext a b, + simp only [map_apply, function.comp_app, ←exists_and_distrib_right, @exists₂_comm γ], + refine exists₂_congr (λ a b, _), + simp [and_assoc], +end + +instance [decidable (∃ a b, r a b ∧ f a = c ∧ g b = d)] : decidable (relation.map r f g c d) := +‹decidable _› + +end map + variables {r : α → α → Prop} {a b c d : α} /-- `refl_trans_gen r`: reflexive transitive closure of `r` -/ From 65a1391a0106c9204fe45bc73a039f056558cb83 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 30 Oct 2023 15:07:49 +0000 Subject: [PATCH 61/61] feat(data/{list,multiset,finset}/*): `attach` and `filter` lemmas (#18087) Left commutativity and cardinality of `list.filter`/`multiset.filter`/`finset.filter`. Interaction of `count`/`countp` and `attach`. --- src/algebra/big_operators/basic.lean | 24 ++++++++----- src/algebra/big_operators/order.lean | 2 +- src/data/finset/card.lean | 1 + src/data/finset/image.lean | 16 +++++++++ src/data/list/basic.lean | 25 +++++++++++++ src/data/list/count.lean | 6 ++++ src/data/list/perm.lean | 2 +- src/data/multiset/basic.lean | 50 ++++++++++++++++++++------ src/data/multiset/lattice.lean | 2 +- src/data/set/finite.lean | 3 +- src/number_theory/kummer_dedekind.lean | 2 +- 11 files changed, 108 insertions(+), 25 deletions(-) diff --git a/src/algebra/big_operators/basic.lean b/src/algebra/big_operators/basic.lean index 85ce482038a50..ac6250a471925 100644 --- a/src/algebra/big_operators/basic.lean +++ b/src/algebra/big_operators/basic.lean @@ -106,6 +106,9 @@ variables {s s₁ s₂ : finset α} {a : α} {f g : α → β} @[to_additive] lemma prod_eq_multiset_prod [comm_monoid β] (s : finset α) (f : α → β) : ∏ x in s, f x = (s.1.map f).prod := rfl +@[simp, to_additive] lemma prod_map_val [comm_monoid β] (s : finset α) (f : α → β) : + (s.1.map f).prod = ∏ a in s, f a := rfl + @[to_additive] theorem prod_eq_fold [comm_monoid β] (s : finset α) (f : α → β) : ∏ x in s, f x = s.fold (*) 1 f := @@ -1418,16 +1421,19 @@ begin apply sum_congr rfl h₁ end +lemma nat_cast_card_filter [add_comm_monoid_with_one β] (p) [decidable_pred p] (s : finset α) : + ((filter p s).card : β) = ∑ a in s, if p a then 1 else 0 := +by simp only [add_zero, sum_const, nsmul_eq_mul, eq_self_iff_true, sum_const_zero, sum_ite, + nsmul_one] + +lemma card_filter (p) [decidable_pred p] (s : finset α) : + (filter p s).card = ∑ a in s, ite (p a) 1 0 := +nat_cast_card_filter _ _ + @[simp] -lemma sum_boole {s : finset α} {p : α → Prop} [non_assoc_semiring β] {hp : decidable_pred p} : - (∑ x in s, if p x then (1 : β) else (0 : β)) = (s.filter p).card := -by simp only [add_zero, - mul_one, - finset.sum_const, - nsmul_eq_mul, - eq_self_iff_true, - finset.sum_const_zero, - finset.sum_ite] +lemma sum_boole {s : finset α} {p : α → Prop} [add_comm_monoid_with_one β] {hp : decidable_pred p} : + (∑ x in s, if p x then 1 else 0 : β) = (s.filter p).card := +(nat_cast_card_filter _ _).symm lemma _root_.commute.sum_right [non_unital_non_assoc_semiring β] (s : finset α) (f : α → β) (b : β) (h : ∀ i ∈ s, commute b (f i)) : diff --git a/src/algebra/big_operators/order.lean b/src/algebra/big_operators/order.lean index 2573be61bcdb6..cda6b956aaa63 100644 --- a/src/algebra/big_operators/order.lean +++ b/src/algebra/big_operators/order.lean @@ -174,7 +174,7 @@ lemma prod_le_pow_card (s : finset ι) (f : ι → N) (n : N) (h : ∀ x ∈ s, begin refine (multiset.prod_le_pow_card (s.val.map f) n _).trans _, { simpa using h }, - { simpa } + { simp } end @[to_additive card_nsmul_le_sum] diff --git a/src/data/finset/card.lean b/src/data/finset/card.lean index f023ef54c1104..ae354aa222531 100644 --- a/src/data/finset/card.lean +++ b/src/data/finset/card.lean @@ -42,6 +42,7 @@ variables {s t : finset α} {a b : α} def card (s : finset α) : ℕ := s.1.card lemma card_def (s : finset α) : s.card = s.1.card := rfl +@[simp] lemma card_val (s : finset α) : s.1.card = s.card := rfl @[simp] lemma card_mk {m nodup} : (⟨m, nodup⟩ : finset α).card = m.card := rfl diff --git a/src/data/finset/image.lean b/src/data/finset/image.lean index e41fa0925240e..74a678a27cabc 100644 --- a/src/data/finset/image.lean +++ b/src/data/finset/image.lean @@ -127,6 +127,22 @@ lemma filter_map {p : β → Prop} [decidable_pred p] : (s.map f).filter p = (s.filter (p ∘ f)).map f := eq_of_veq (map_filter _ _ _) +lemma map_filter' (p : α → Prop) [decidable_pred p] (f : α ↪ β) (s : finset α) + [decidable_pred (λ b, ∃ a, p a ∧ f a = b)] : + (s.filter p).map f = (s.map f).filter (λ b, ∃ a, p a ∧ f a = b) := +by simp [(∘), filter_map, f.injective.eq_iff] + +lemma filter_attach' [decidable_eq α] (s : finset α) (p : s → Prop) [decidable_pred p] : + s.attach.filter p = + (s.filter $ λ x, ∃ h, p ⟨x, h⟩).attach.map ⟨subtype.map id $ filter_subset _ _, + subtype.map_injective _ injective_id⟩ := +eq_of_veq $ multiset.filter_attach' _ _ + +@[simp] lemma filter_attach (p : α → Prop) [decidable_pred p] (s : finset α) : + (s.attach.filter (λ x, p ↑x)) = + (s.filter p).attach.map ((embedding.refl _).subtype_map mem_of_mem_filter) := +eq_of_veq $ multiset.filter_attach _ _ + lemma map_filter {f : α ≃ β} {p : α → Prop} [decidable_pred p] : (s.filter p).map f.to_embedding = (s.map f.to_embedding).filter (p ∘ f.symm) := by simp only [filter_map, function.comp, equiv.to_embedding_apply, equiv.symm_apply_apply] diff --git a/src/data/list/basic.lean b/src/data/list/basic.lean index 939c431e6d038..ce40ca44d44af 100644 --- a/src/data/list/basic.lean +++ b/src/data/list/basic.lean @@ -2716,6 +2716,8 @@ end split_at_on with the same elements but in the type `{x // x ∈ l}`. -/ def attach (l : list α) : list {x // x ∈ l} := pmap subtype.mk l (λ a, id) +@[simp] lemma attach_nil : ([] : list α).attach = [] := rfl + theorem sizeof_lt_sizeof_of_mem [has_sizeof α] {x : α} {l : list α} (hx : x ∈ l) : sizeof x < sizeof l := begin @@ -3250,12 +3252,35 @@ theorem map_filter (f : β → α) (l : list β) : filter p (map f l) = map f (filter (p ∘ f) l) := by rw [← filter_map_eq_map, filter_filter_map, filter_map_filter]; refl +lemma map_filter' {f : α → β} (hf : injective f) (l : list α) + [decidable_pred (λ b, ∃ a, p a ∧ f a = b)] : + (l.filter p).map f = (l.map f).filter (λ b, ∃ a, p a ∧ f a = b) := +by simp [(∘), map_filter, hf.eq_iff] + +lemma filter_attach' (l : list α) (p : {a // a ∈ l} → Prop) [decidable_eq α] [decidable_pred p] : + l.attach.filter p = (l.filter $ λ x, ∃ h, p ⟨x, h⟩).attach.map + (subtype.map id $ λ x hx, let ⟨h, _⟩ := of_mem_filter hx in h) := +begin + classical, + refine map_injective_iff.2 subtype.coe_injective _, + simp [(∘), map_filter' _ subtype.coe_injective], +end + +@[simp] lemma filter_attach (l : list α) (p : α → Prop) [decidable_pred p] : + l.attach.filter (λ x, p ↑x) = (l.filter p).attach.map (subtype.map id $ λ _, mem_of_mem_filter) := +map_injective_iff.2 subtype.coe_injective $ by + simp_rw [map_map, (∘), subtype.map, subtype.coe_mk, id.def, ←map_filter, attach_map_coe] + @[simp] theorem filter_filter (q) [decidable_pred q] : ∀ l, filter p (filter q l) = filter (λ a, p a ∧ q a) l | [] := rfl | (a :: l) := by by_cases hp : p a; by_cases hq : q a; simp only [hp, hq, filter, if_true, if_false, true_and, false_and, filter_filter l, eq_self_iff_true] +lemma filter_comm (q) [decidable_pred q] (l : list α) : + filter p (filter q l) = filter q (filter p l) := +by simp [and_comm] + @[simp] lemma filter_true {h : decidable_pred (λ a : α, true)} (l : list α) : @filter α (λ _, true) h l = l := by convert filter_eq_self.2 (λ _ _, trivial) diff --git a/src/data/list/count.lean b/src/data/list/count.lean index 5660878757f14..25d72187efa05 100644 --- a/src/data/list/count.lean +++ b/src/data/list/count.lean @@ -84,6 +84,9 @@ by simp only [countp_eq_length_filter, filter_filter] | [] := rfl | (a::l) := by rw [map_cons, countp_cons, countp_cons, countp_map] +@[simp] lemma countp_attach (l : list α) : l.attach.countp (λ a, p ↑a) = l.countp p := +by rw [←countp_map, attach_map_coe] + variables {p q} lemma countp_mono_left (h : ∀ x ∈ l, p x → q x) : countp p l ≤ countp q l := @@ -197,6 +200,9 @@ lemma count_bind {α β} [decidable_eq β] (l : list α) (f : α → list β) (x count x (l.bind f) = sum (map (count x ∘ f) l) := by rw [list.bind, count_join, map_map] +@[simp] lemma count_attach (a : {x // x ∈ l}) : l.attach.count a = l.count a := +eq.trans (countp_congr $ λ _ _, subtype.ext_iff) $ countp_attach _ _ + @[simp] lemma count_map_of_injective {α β} [decidable_eq α] [decidable_eq β] (l : list α) (f : α → β) (hf : function.injective f) (x : α) : count (f x) (map f l) = count x l := diff --git a/src/data/list/perm.lean b/src/data/list/perm.lean index 8c1667485d02c..22c3396410970 100644 --- a/src/data/list/perm.lean +++ b/src/data/list/perm.lean @@ -396,7 +396,7 @@ theorem subperm.countp_le (p : α → Prop) [decidable_pred p] | ⟨l, p', s⟩ := p'.countp_eq p ▸ s.countp_le p theorem perm.countp_congr (s : l₁ ~ l₂) {p p' : α → Prop} [decidable_pred p] [decidable_pred p'] - (hp : ∀ x ∈ l₁, p x = p' x) : l₁.countp p = l₂.countp p' := + (hp : ∀ x ∈ l₁, p x ↔ p' x) : l₁.countp p = l₂.countp p' := begin rw ← s.countp_eq p', clear s, diff --git a/src/data/multiset/basic.lean b/src/data/multiset/basic.lean index f58ff719fad96..99873f8d0abe7 100644 --- a/src/data/multiset/basic.lean +++ b/src/data/multiset/basic.lean @@ -16,7 +16,7 @@ These are implemented as the quotient of a list by permutations. We define the global infix notation `::ₘ` for `multiset.cons`. -/ -open list subtype nat +open function list nat subtype variables {α : Type*} {β : Type*} {γ : Type*} @@ -1543,6 +1543,10 @@ le_antisymm (le_inter filter p (filter q s) = filter (λ a, p a ∧ q a) s := quot.induction_on s $ λ l, congr_arg coe $ filter_filter p q l +lemma filter_comm (q) [decidable_pred q] (s : multiset α) : + filter p (filter q s) = filter q (filter p s) := +by simp [and_comm] + theorem filter_add_filter (q) [decidable_pred q] (s : multiset α) : filter p s + filter q s = filter (λ a, p a ∨ q a) s + filter (λ a, p a ∧ q a) s := multiset.induction_on s rfl $ λ a s IH, @@ -1556,6 +1560,11 @@ theorem map_filter (f : β → α) (s : multiset β) : filter p (map f s) = map f (filter (p ∘ f) s) := quot.induction_on s (λ l, by simp [map_filter]) +lemma map_filter' {f : α → β} (hf : injective f) (s : multiset α) + [decidable_pred (λ b, ∃ a, p a ∧ f a = b)] : + (s.filter p).map f = (s.map f).filter (λ b, ∃ a, p a ∧ f a = b) := +by simp [(∘), map_filter, hf.eq_iff] + /-! ### Simultaneously filter and map elements of a multiset -/ /-- `filter_map f s` is a combination filter/map operation on `s`. @@ -1704,6 +1713,18 @@ begin card_singleton, add_comm] }, end +@[simp] lemma countp_attach (s : multiset α) : s.attach.countp (λ a, p ↑a) = s.countp p := +quotient.induction_on s $ λ l, begin + simp only [quot_mk_to_coe, coe_countp], + rw [quot_mk_to_coe, coe_attach, coe_countp], + exact list.countp_attach _ _, +end + +@[simp] lemma filter_attach (m : multiset α) (p : α → Prop) [decidable_pred p] : + (m.attach.filter (λ x, p ↑x)) = + (m.filter p).attach.map (subtype.map id $ λ _, multiset.mem_of_mem_filter) := +quotient.induction_on m $ λ l, congr_arg coe (list.filter_attach l p) + variable {p} theorem countp_pos {s} : 0 < countp p s ↔ ∃ a ∈ s, p a := @@ -1720,7 +1741,7 @@ countp_pos.2 ⟨_, h, pa⟩ theorem countp_congr {s s' : multiset α} (hs : s = s') {p p' : α → Prop} [decidable_pred p] [decidable_pred p'] - (hp : ∀ x ∈ s, p x = p' x) : s.countp p = s'.countp p' := + (hp : ∀ x ∈ s, p x ↔ p' x) : s.countp p = s'.countp p' := quot.induction_on₂ s s' (λ l l' hs hp, begin simp only [quot_mk_to_coe'', coe_eq_coe] at hs, exact hs.countp_congr hp, @@ -1731,7 +1752,7 @@ end /-! ### Multiplicity of an element -/ section -variable [decidable_eq α] +variables [decidable_eq α] {s : multiset α} /-- `count a s` is the multiplicity of `a` in `s`. -/ def count (a : α) : multiset α → ℕ := countp (eq a) @@ -1778,6 +1799,9 @@ def count_add_monoid_hom (a : α) : multiset α →+ ℕ := countp_add_monoid_ho @[simp] theorem count_nsmul (a : α) (n s) : count a (n • s) = n * count a s := by induction n; simp [*, succ_nsmul', succ_mul, zero_nsmul] +@[simp] lemma count_attach (a : {x // x ∈ s}) : s.attach.count a = s.count a := +eq.trans (countp_congr rfl $ λ _ _, subtype.ext_iff) $ countp_attach _ _ + theorem count_pos {a : α} {s : multiset α} : 0 < count a s ↔ a ∈ s := by simp [count, countp_pos] @@ -1901,13 +1925,6 @@ begin contradiction } end -@[simp] -lemma attach_count_eq_count_coe (m : multiset α) (a) : m.attach.count a = m.count (a : α) := -calc m.attach.count a - = (m.attach.map (coe : _ → α)).count (a : α) : - (multiset.count_map_eq_count' _ _ subtype.coe_injective _).symm -... = m.count (a : α) : congr_arg _ m.attach_map_coe - lemma filter_eq' (s : multiset α) (b : α) : s.filter (= b) = replicate (count b s) b := quotient.induction_on s $ λ l, congr_arg coe $ filter_eq' l b @@ -2174,6 +2191,19 @@ theorem map_injective {f : α → β} (hf : function.injective f) : function.injective (multiset.map f) := assume x y, (map_eq_map hf).1 +lemma filter_attach' (s : multiset α) (p : {a // a ∈ s} → Prop) [decidable_eq α] + [decidable_pred p] : + s.attach.filter p = + (s.filter $ λ x, ∃ h, p ⟨x, h⟩).attach.map (subtype.map id $ λ x hx, + let ⟨h, _⟩ := of_mem_filter hx in h) := +begin + classical, + refine multiset.map_injective subtype.coe_injective _, + simp only [function.comp, map_filter' _ subtype.coe_injective, subtype.exists, coe_mk, + exists_and_distrib_right, exists_eq_right, attach_map_coe, map_map, map_coe, id.def], + rw attach_map_coe, +end + end map section quot diff --git a/src/data/multiset/lattice.lean b/src/data/multiset/lattice.lean index 7b279882080d6..5e575d279b753 100644 --- a/src/data/multiset/lattice.lean +++ b/src/data/multiset/lattice.lean @@ -39,7 +39,7 @@ sup_bot_eq @[simp] lemma sup_add (s₁ s₂ : multiset α) : (s₁ + s₂).sup = s₁.sup ⊔ s₂.sup := eq.trans (by simp [sup]) (fold_add _ _ _ _ _) -lemma sup_le {s : multiset α} {a : α} : s.sup ≤ a ↔ (∀b ∈ s, b ≤ a) := +@[simp] lemma sup_le {s : multiset α} {a : α} : s.sup ≤ a ↔ (∀b ∈ s, b ≤ a) := multiset.induction_on s (by simp) (by simp [or_imp_distrib, forall_and_distrib] {contextual := tt}) diff --git a/src/data/set/finite.lean b/src/data/set/finite.lean index f86cca6bfa443..c073dce02f496 100644 --- a/src/data/set/finite.lean +++ b/src/data/set/finite.lean @@ -927,8 +927,7 @@ eq.trans (by congr) empty_card theorem card_fintype_insert_of_not_mem {a : α} (s : set α) [fintype s] (h : a ∉ s) : @fintype.card _ (fintype_insert_of_not_mem s h) = fintype.card s + 1 := -by rw [fintype_insert_of_not_mem, fintype.card_of_finset]; - simp [finset.card, to_finset]; refl +by simp [fintype_insert_of_not_mem, fintype.card_of_finset] @[simp] theorem card_insert {a : α} (s : set α) [fintype s] (h : a ∉ s) {d : fintype.{u} (insert a s : set α)} : diff --git a/src/number_theory/kummer_dedekind.lean b/src/number_theory/kummer_dedekind.lean index 28af342219f88..381da69360e7e 100644 --- a/src/number_theory/kummer_dedekind.lean +++ b/src/number_theory/kummer_dedekind.lean @@ -294,7 +294,7 @@ begin rw [multiset.count_map_eq_count' (λ f, ((normalized_factors_map_equiv_normalized_factors_min_poly_mk hI hI' hx hx').symm f : ideal S)), - multiset.attach_count_eq_count_coe], + multiset.count_attach], { exact subtype.coe_injective.comp (equiv.injective _) }, { exact (normalized_factors_map_equiv_normalized_factors_min_poly_mk hI hI' hx hx' _).prop}, { exact irreducible_of_normalized_factor _