├── .github └── workflows │ └── blueprint.yml ├── .gitignore ├── .vscode └── settings.json ├── LICENSE ├── README.md ├── SphereEversion.lean ├── SphereEversion ├── Global │ ├── Gromov.lean │ ├── Immersion.lean │ ├── Localisation.lean │ ├── LocalisationData.lean │ ├── LocalizedConstruction.lean │ ├── OneJetBundle.lean │ ├── OneJetSec.lean │ ├── ParametricityForFree.lean │ ├── Relation.lean │ ├── SmoothEmbedding.lean │ └── TwistOneJetSec.lean ├── Indexing.lean ├── InductiveConstructions.lean ├── Local │ ├── AmpleRelation.lean │ ├── Corrugation.lean │ ├── DualPair.lean │ ├── HPrinciple.lean │ ├── OneJet.lean │ ├── ParametricHPrinciple.lean │ ├── Relation.lean │ └── SphereEversion.lean ├── Loops │ ├── Basic.lean │ ├── DeltaMollifier.lean │ ├── Exists.lean │ ├── Reparametrization.lean │ └── Surrounding.lean ├── Main.lean ├── Notations.lean └── ToMathlib │ ├── Algebra │ └── Ring │ │ └── Periodic.lean │ ├── Analysis │ ├── Calculus.lean │ ├── Calculus │ │ └── AddTorsor │ │ │ └── AffineMap.lean │ ├── ContDiff.lean │ ├── Convex │ │ └── Basic.lean │ ├── CutOff.lean │ ├── InnerProductSpace │ │ ├── CrossProduct.lean │ │ ├── Dual.lean │ │ ├── Projection.lean │ │ └── Rotation.lean │ ├── Normed │ │ └── Module │ │ │ └── FiniteDimension.lean │ └── NormedSpace │ │ ├── Misc.lean │ │ └── OperatorNorm │ │ └── Prod.lean │ ├── Data │ └── Nat │ │ └── Basic.lean │ ├── Equivariant.lean │ ├── ExistsOfConvex.lean │ ├── Geometry │ └── Manifold │ │ ├── Algebra │ │ └── SmoothGerm.lean │ │ ├── Immersion.lean │ │ ├── IsManifold │ │ └── ExtChartAt.lean │ │ ├── Metrizable.lean │ │ ├── MiscManifold.lean │ │ └── VectorBundle │ │ └── Misc.lean │ ├── LinearAlgebra │ ├── Basic.lean │ └── FiniteDimensional.lean │ ├── MeasureTheory │ ├── BorelSpace.lean │ └── ParametricIntervalIntegral.lean │ ├── Order │ └── Filter │ │ └── Basic.lean │ ├── Partition.lean │ ├── SmoothBarycentric.lean │ ├── Topology │ ├── Algebra │ │ └── Module.lean │ ├── Misc.lean │ ├── Paracompact.lean │ ├── Path.lean │ └── Separation │ │ └── Hausdorff.lean │ └── Unused │ ├── EventuallyConstant.lean │ ├── Fin.lean │ ├── GeometryManifoldMisc.lean │ ├── IntervalIntegral.lean │ └── LinearAlgebra │ └── Multilinear.lean ├── blueprint ├── .gitignore ├── mathjax.sed ├── src │ ├── appendix.tex │ ├── content.tex │ ├── fix_mathjax.sh │ ├── global_convex_integration.tex │ ├── introduction.tex │ ├── latexmkrc │ ├── local_convex_integration.tex │ ├── local_to_global.tex │ ├── loops.tex │ ├── macros_common.tex │ ├── macros_print.tex │ ├── macros_web.tex │ ├── plastex.cfg │ ├── print.tex │ ├── stylecours.css │ └── web.tex └── tasks.py ├── docs_src ├── index.md ├── pygments.css ├── style.css └── template.html ├── graphs └── graphs.py ├── lake-manifest.json ├── lakefile.toml ├── lean-toolchain ├── noshake.json ├── scripts ├── build_docs.sh └── mk_all.sh └── tasks.py /.github/workflows/blueprint.yml: -------------------------------------------------------------------------------- 1 | name: Compile blueprint 2 | 3 | on: 4 | push: 5 | branches: 6 | - master 7 | pull_request: 8 | branches: 9 | - master 10 | workflow_dispatch: # Allow manual triggering of the workflow from the GitHub Actions interface 11 | 12 | # Sets permissions of the GITHUB_TOKEN to allow deployment to GitHub Pages and modify PR labels 13 | permissions: 14 | contents: read # Read access to repository contents 15 | pages: write # Write access to GitHub Pages 16 | id-token: write # Write access to ID tokens 17 | 18 | jobs: 19 | build_project: 20 | runs-on: ubuntu-latest 21 | name: Build project 22 | steps: 23 | - name: Checkout project 24 | uses: actions/checkout@v4 25 | with: 26 | fetch-depth: 0 27 | 28 | - name: Install elan 29 | run: curl https://raw.githubusercontent.com/leanprover/elan/master/elan-init.sh -sSf | sh -s -- -y 30 | 31 | - name: Get cache 32 | run: ~/.elan/bin/lake exe cache get || true 33 | 34 | - name: Build project 35 | run: ~/.elan/bin/lake build SphereEversion 36 | 37 | - name: Cache API docs 38 | uses: actions/cache@v4 39 | with: 40 | path: | 41 | docbuild/.lake/build/doc/Aesop 42 | docbuild/.lake/build/doc/Batteries 43 | docbuild/.lake/build/doc/find 44 | docbuild/.lake/build/doc/Init 45 | docbuild/.lake/build/doc/Lake 46 | docbuild/.lake/build/doc/Lean 47 | docbuild/.lake/build/doc/Mathlib 48 | docbuild/.lake/build/doc/Std 49 | key: Docs-${{ hashFiles('lake-manifest.json') }} 50 | 51 | - name: Build blueprint and copy to `docs/blueprint` 52 | uses: xu-cheng/texlive-action@v2 53 | with: 54 | docker_image: ghcr.io/xu-cheng/texlive-full:20231201 55 | run: | 56 | export PIP_BREAK_SYSTEM_PACKAGES=1 57 | apk update 58 | apk add --update make py3-pip git pkgconfig graphviz graphviz-dev gcc musl-dev 59 | git config --global --add safe.directory $GITHUB_WORKSPACE 60 | git config --global --add safe.directory `pwd` 61 | python3 -m venv env 62 | source env/bin/activate 63 | pip install --upgrade pip requests wheel 64 | pip install pygraphviz --global-option=build_ext --global-option="-L/usr/lib/graphviz/" --global-option="-R/usr/lib/graphviz/" 65 | pip install leanblueprint 66 | leanblueprint pdf 67 | mkdir docs 68 | cp blueprint/print/print.pdf docs/blueprint.pdf 69 | leanblueprint web 70 | cp -r blueprint/web docs/blueprint 71 | 72 | - name: Check declarations 73 | run: ~/.elan/bin/lake exe checkdecls blueprint/lean_decls 74 | 75 | - name: Build project API documentation 76 | run: scripts/build_docs.sh 77 | 78 | - name: Install pandoc 79 | if: github.event_name == 'push' 80 | run: | 81 | sudo apt update 82 | sudo apt install pandoc 83 | 84 | - name: Build webpage 85 | if: github.event_name == 'push' 86 | run: | 87 | cd docs_src 88 | pandoc -t html --mathjax -f markdown+tex_math_dollars+raw_tex index.md --template template.html -o ../docs/index.html 89 | cp *.css ../docs 90 | 91 | - name: Upload docs & blueprint artifact 92 | if: github.event_name == 'push' 93 | uses: actions/upload-pages-artifact@v3 94 | with: 95 | path: docs/ 96 | 97 | - name: Deploy to GitHub Pages 98 | if: github.event_name == 'push' 99 | id: deployment 100 | uses: actions/deploy-pages@v4 101 | 102 | # - name: Make sure the cache works 103 | # run: mv docs/docs .lake/build/doc 104 | -------------------------------------------------------------------------------- /.gitignore: -------------------------------------------------------------------------------- 1 | decls.pickle 2 | decls.yaml 3 | /build 4 | .lake/ 5 | .python-version 6 | .vscode 7 | lakefile.olean 8 | 9 | ## Core latex/pdflatex auxiliary files: 10 | *.aux 11 | *.lof 12 | *.log 13 | *.lot 14 | *.fls 15 | *.out 16 | *.toc 17 | *.fmt 18 | *.fot 19 | *.cb 20 | *.cb2 21 | .*.lb 22 | -------------------------------------------------------------------------------- /.vscode/settings.json: -------------------------------------------------------------------------------- 1 | { 2 | "editor.insertSpaces": true, 3 | "editor.tabSize": 2, 4 | "editor.rulers" : [100], 5 | "files.encoding": "utf8", 6 | "files.eol": "\n", 7 | "files.insertFinalNewline": true, 8 | "files.trimFinalNewlines": true, 9 | "files.trimTrailingWhitespace": true, 10 | "search.usePCRE2": true 11 | } 12 | -------------------------------------------------------------------------------- /README.md: -------------------------------------------------------------------------------- 1 | # The sphere eversion project 2 | 3 | This project formalizes the proof of a theorem implying the existence of sphere 4 | eversions. It was carried out by Patrick Massot, Floris van Doorn and Oliver 5 | Nash, with crucial help from the wider Lean community. The proof of the main theorem was completed on November 12th 2022. 6 | Details can be found on the [project website.](https://leanprover-community.github.io/sphere-eversion/) 7 | 8 | This project originally used Lean 3 but was ported to Lean 4 with crucial help from Yury Kudryashov. 9 | 10 | # Build the Lean files 11 | 12 | To build the Lean files of this project, you need to have a working version of Lean. 13 | See [the installation instructions](https://leanprover-community.github.io/get_started.html) (under Regular install). 14 | 15 | To build the project, run `lake exe cache get` and then `lake build`. 16 | 17 | # Build the blueprint 18 | 19 | To build the web version of the blueprint, you need a working LaTeX installation. 20 | Furthermore, you need some packages: 21 | ``` 22 | sudo apt install graphviz libgraphviz-dev 23 | pip3 install invoke pandoc 24 | cd .. # go to folder where you are happy clone git repos 25 | git clone git@github.com:plastex/plastex 26 | pip3 install ./plastex 27 | git clone git@github.com:PatrickMassot/leanblueprint 28 | pip3 install ./leanblueprint 29 | cd sphere-eversion 30 | ``` 31 | 32 | To actually build the blueprint, run 33 | ``` 34 | lake exe cache get 35 | lake build 36 | inv all 37 | ``` 38 | 39 | To view the web-version of the blueprint locally, run `inv serve` and navigate to 40 | `http://localhost:8000/` in your favorite browser. 41 | -------------------------------------------------------------------------------- /SphereEversion.lean: -------------------------------------------------------------------------------- 1 | import SphereEversion.Global.Gromov 2 | import SphereEversion.Global.Immersion 3 | import SphereEversion.Global.Localisation 4 | import SphereEversion.Global.LocalisationData 5 | import SphereEversion.Global.LocalizedConstruction 6 | import SphereEversion.Global.OneJetBundle 7 | import SphereEversion.Global.OneJetSec 8 | import SphereEversion.Global.ParametricityForFree 9 | import SphereEversion.Global.Relation 10 | import SphereEversion.Global.SmoothEmbedding 11 | import SphereEversion.Global.TwistOneJetSec 12 | import SphereEversion.Indexing 13 | import SphereEversion.InductiveConstructions 14 | import SphereEversion.Local.AmpleRelation 15 | import SphereEversion.Local.Corrugation 16 | import SphereEversion.Local.DualPair 17 | import SphereEversion.Local.HPrinciple 18 | import SphereEversion.Local.OneJet 19 | import SphereEversion.Local.ParametricHPrinciple 20 | import SphereEversion.Local.Relation 21 | import SphereEversion.Local.SphereEversion 22 | import SphereEversion.Loops.Basic 23 | import SphereEversion.Loops.DeltaMollifier 24 | import SphereEversion.Loops.Exists 25 | import SphereEversion.Loops.Reparametrization 26 | import SphereEversion.Loops.Surrounding 27 | import SphereEversion.Main 28 | import SphereEversion.Notations 29 | import SphereEversion.ToMathlib.Algebra.Ring.Periodic 30 | import SphereEversion.ToMathlib.Analysis.Calculus 31 | import SphereEversion.ToMathlib.Analysis.Calculus.AddTorsor.AffineMap 32 | import SphereEversion.ToMathlib.Analysis.ContDiff 33 | import SphereEversion.ToMathlib.Analysis.Convex.Basic 34 | import SphereEversion.ToMathlib.Analysis.CutOff 35 | import SphereEversion.ToMathlib.Analysis.InnerProductSpace.CrossProduct 36 | import SphereEversion.ToMathlib.Analysis.InnerProductSpace.Dual 37 | import SphereEversion.ToMathlib.Analysis.InnerProductSpace.Projection 38 | import SphereEversion.ToMathlib.Analysis.InnerProductSpace.Rotation 39 | import SphereEversion.ToMathlib.Analysis.Normed.Module.FiniteDimension 40 | import SphereEversion.ToMathlib.Analysis.NormedSpace.Misc 41 | import SphereEversion.ToMathlib.Analysis.NormedSpace.OperatorNorm.Prod 42 | import SphereEversion.ToMathlib.Data.Nat.Basic 43 | import SphereEversion.ToMathlib.Equivariant 44 | import SphereEversion.ToMathlib.ExistsOfConvex 45 | import SphereEversion.ToMathlib.Geometry.Manifold.Algebra.SmoothGerm 46 | import SphereEversion.ToMathlib.Geometry.Manifold.Immersion 47 | import SphereEversion.ToMathlib.Geometry.Manifold.IsManifold.ExtChartAt 48 | import SphereEversion.ToMathlib.Geometry.Manifold.Metrizable 49 | import SphereEversion.ToMathlib.Geometry.Manifold.MiscManifold 50 | import SphereEversion.ToMathlib.Geometry.Manifold.VectorBundle.Misc 51 | import SphereEversion.ToMathlib.LinearAlgebra.Basic 52 | import SphereEversion.ToMathlib.LinearAlgebra.FiniteDimensional 53 | import SphereEversion.ToMathlib.MeasureTheory.BorelSpace 54 | import SphereEversion.ToMathlib.MeasureTheory.ParametricIntervalIntegral 55 | import SphereEversion.ToMathlib.Order.Filter.Basic 56 | import SphereEversion.ToMathlib.Partition 57 | import SphereEversion.ToMathlib.SmoothBarycentric 58 | import SphereEversion.ToMathlib.Topology.Algebra.Module 59 | import SphereEversion.ToMathlib.Topology.Misc 60 | import SphereEversion.ToMathlib.Topology.Paracompact 61 | import SphereEversion.ToMathlib.Topology.Path 62 | import SphereEversion.ToMathlib.Topology.Separation.Hausdorff 63 | import SphereEversion.ToMathlib.Unused.Fin 64 | -------------------------------------------------------------------------------- /SphereEversion/Global/LocalisationData.lean: -------------------------------------------------------------------------------- 1 | import Mathlib.Topology.MetricSpace.PartitionOfUnity 2 | import SphereEversion.Global.SmoothEmbedding 3 | 4 | noncomputable section 5 | 6 | open scoped Manifold ContDiff 7 | 8 | open Set Metric 9 | 10 | section 11 | 12 | variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] 13 | {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] 14 | {H : Type*} [TopologicalSpace H] (I : ModelWithCorners 𝕜 E H) 15 | {M : Type*} [TopologicalSpace M] [ChartedSpace H M] [IsManifold I ∞ M] 16 | {E' : Type*} [NormedAddCommGroup E'] [NormedSpace 𝕜 E'] 17 | {H' : Type*} [TopologicalSpace H'] (I' : ModelWithCorners 𝕜 E' H') 18 | {M' : Type*} [TopologicalSpace M'] [ChartedSpace H' M'] [IsManifold I' ∞ M'] 19 | 20 | /-- Definition `def:localisation_data`. -/ 21 | structure LocalisationData (f : M → M') where 22 | cont : Continuous f 23 | ι' : Type* 24 | N : ℕ 25 | φ : IndexType N → OpenSmoothEmbedding 𝓘(𝕜, E) E I M 26 | ψ : ι' → OpenSmoothEmbedding 𝓘(𝕜, E') E' I' M' 27 | j : IndexType N → ι' 28 | h₁ : (⋃ i, φ i '' ball (0 : E) 1) = univ 29 | h₂ : (⋃ i', ψ i' '' ball (0 : E') 1) = univ 30 | h₃ : ∀ i, range (f ∘ φ i) ⊆ ψ (j i) '' ball (0 : E') 1 31 | h₄ : LocallyFinite fun i' ↦ range (ψ i') 32 | lf_φ : LocallyFinite fun i ↦ range (φ i) 33 | 34 | namespace LocalisationData 35 | 36 | -- Adaptation note(version 4.10-rc1): previously, specifying 𝕜, E, E', H and H' was not needed 37 | variable {f : M → M'} {I I'} 38 | (ld : LocalisationData (𝕜 := 𝕜) (E := E) (E' := E') (H := H) (H' := H') I I' f) 39 | 40 | abbrev ψj : IndexType ld.N → OpenSmoothEmbedding 𝓘(𝕜, E') E' I' M' := 41 | ld.ψ ∘ ld.j 42 | 43 | /-- The type indexing the source charts of the given localisation data. -/ 44 | def ι (L : LocalisationData I I' f) := 45 | IndexType L.N 46 | 47 | omit [IsManifold I ∞ M] [IsManifold I' ∞ M'] in 48 | theorem iUnion_succ' {β : Type*} (s : ld.ι → Set β) (i : IndexType ld.N) : 49 | (⋃ j ≤ i, s j) = (⋃ j < i, s j) ∪ s i := by 50 | simp only [(fun _ ↦ le_iff_lt_or_eq : ∀ j, j ≤ i ↔ j < i ∨ j = i)] 51 | erw [biUnion_union, biUnion_singleton] 52 | rfl 53 | 54 | open Filter 55 | 56 | end LocalisationData 57 | 58 | end 59 | 60 | section 61 | 62 | open ModelWithCorners 63 | 64 | variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] [FiniteDimensional ℝ E] 65 | {M : Type*} [TopologicalSpace M] [SigmaCompactSpace M] [LocallyCompactSpace M] [T2Space M] 66 | {H : Type*} [TopologicalSpace H] (I : ModelWithCorners ℝ E H) [Boundaryless I] [Nonempty M] 67 | [ChartedSpace H M] [IsManifold I ∞ M] 68 | (E' : Type*) [NormedAddCommGroup E'] [NormedSpace ℝ E'] [FiniteDimensional ℝ E'] 69 | {H' : Type*} [TopologicalSpace H'] (I' : ModelWithCorners ℝ E' H') [Boundaryless I'] 70 | {M' : Type*} [MetricSpace M'] [SigmaCompactSpace M'] [LocallyCompactSpace M'] 71 | [Nonempty M'] [ChartedSpace H' M'] [IsManifold I' ∞ M'] 72 | 73 | variable (M') 74 | 75 | theorem nice_atlas_target : 76 | ∃ n, 77 | ∃ ψ : IndexType n → OpenSmoothEmbedding 𝓘(ℝ, E') E' I' M', 78 | (LocallyFinite fun i' ↦ range (ψ i')) ∧ (⋃ i', ψ i' '' ball 0 1) = univ := 79 | let h := nice_atlas E' I' (fun _ : Unit ↦ @isOpen_univ M' _) (by simp [eq_univ_iff_forall]) 80 | ⟨h.choose, h.choose_spec.choose, h.choose_spec.choose_spec.2⟩ 81 | 82 | /-- A collection of charts on a manifold `M'` which are smooth open embeddings with domain the whole 83 | model space, and which cover the manifold when restricted in each case to the unit ball. -/ 84 | def targetCharts (i' : IndexType (nice_atlas_target E' I' M').choose) : 85 | OpenSmoothEmbedding 𝓘(ℝ, E') E' I' M' := 86 | (nice_atlas_target E' I' M').choose_spec.choose i' 87 | 88 | theorem targetCharts_cover : (⋃ i', targetCharts E' I' M' i' '' ball (0 : E') 1) = univ := 89 | (nice_atlas_target E' I' M').choose_spec.choose_spec.2 90 | 91 | variable (E) {M'} 92 | variable {f : M → M'} (hf : Continuous f) 93 | 94 | include hf in 95 | theorem nice_atlas_domain : 96 | ∃ n, 97 | ∃ φ : IndexType n → OpenSmoothEmbedding 𝓘(ℝ, E) E I M, 98 | (∀ i, ∃ i', range (φ i) ⊆ f ⁻¹' (targetCharts E' I' M' i' '' ball (0 : E') 1)) ∧ 99 | (LocallyFinite fun i ↦ range (φ i)) ∧ (⋃ i, φ i '' ball 0 1) = univ := 100 | nice_atlas E I 101 | (fun i' ↦ ((targetCharts E' I' M' i').isOpenMap (ball 0 1) isOpen_ball).preimage hf) 102 | (by rw [← preimage_iUnion, targetCharts_cover, preimage_univ]) 103 | 104 | /-- Lemma `lem:ex_localisation` 105 | Any continuous map between manifolds has some localisation data. -/ 106 | def stdLocalisationData : LocalisationData I I' f where 107 | cont := hf 108 | N := (nice_atlas_domain E I E' I' hf).choose 109 | ι' := IndexType (nice_atlas_target E' I' M').choose 110 | φ := (nice_atlas_domain E I E' I' hf).choose_spec.choose 111 | ψ := targetCharts E' I' M' 112 | j i := ((nice_atlas_domain E I E' I' hf).choose_spec.choose_spec.1 i).choose 113 | h₁ := (nice_atlas_domain E I E' I' hf).choose_spec.choose_spec.2.2 114 | h₂ := targetCharts_cover E' I' M' 115 | h₃ i := by 116 | rw [range_comp] 117 | rintro - ⟨y, hy, rfl⟩ 118 | exact ((nice_atlas_domain E I E' I' hf).choose_spec.choose_spec.1 i).choose_spec hy 119 | h₄ := (nice_atlas_target E' I' M').choose_spec.choose_spec.1 120 | lf_φ := (nice_atlas_domain E I E' I' hf).choose_spec.choose_spec.2.1 121 | 122 | variable {E E' I I'} 123 | 124 | section 125 | 126 | omit [FiniteDimensional ℝ E] [SigmaCompactSpace M] [LocallyCompactSpace M] [T2Space M] 127 | [I.Boundaryless] [Nonempty M] [IsManifold I ∞ M] [I'.Boundaryless] 128 | [SigmaCompactSpace M'] [LocallyCompactSpace M'] [Nonempty M'] [IsManifold I' ∞ M'] 129 | 130 | /-- Lemma `lem:localisation_stability`. -/ 131 | theorem localisation_stability {f : M → M'} (ld : LocalisationData I I' f) : 132 | ∃ (ε : M → ℝ) (_hε : ∀ m, 0 < ε m) (_hε' : Continuous ε), 133 | ∀ (g : M → M') (_hg : ∀ m, dist (g m) (f m) < ε m) (i), 134 | range (g ∘ ld.φ i) ⊆ range (ld.ψj i) := by 135 | let K : ld.ι' → Set M' := fun i ↦ ld.ψ i '' closedBall 0 1 136 | let U : ld.ι' → Set M' := fun i ↦ range <| ld.ψ i 137 | have hK : ∀ i, IsClosed (K i) := fun i ↦ 138 | IsCompact.isClosed (IsCompact.image (isCompact_closedBall 0 1) (ld.ψ i).continuous) 139 | have hK' : LocallyFinite K := ld.h₄.subset fun i ↦ image_subset_range (ld.ψ i) (closedBall 0 1) 140 | have hU : ∀ i, IsOpen (U i) := fun i ↦ (ld.ψ i).isOpen_range 141 | have hKU : ∀ i, K i ⊆ U i := fun i ↦ image_subset_range _ _ 142 | obtain ⟨δ, hδ₀, hδ₁⟩ := exists_continuous_real_forall_closedBall_subset hK hU hKU hK' 143 | have := ld.cont 144 | refine ⟨δ ∘ f, fun m ↦ hδ₀ (f m), by fun_prop, fun g hg i ↦ ?_⟩ 145 | rintro - ⟨e, rfl⟩ 146 | have hi : f (ld.φ i e) ∈ K (ld.j i) := 147 | image_subset _ ball_subset_closedBall (ld.h₃ i (mem_range_self e)) 148 | exact hδ₁ (ld.j i) (f <| ld.φ i e) hi (le_of_lt (hg _)) 149 | 150 | namespace LocalisationData 151 | 152 | protected def ε (ld : LocalisationData I I' f) : M → ℝ := 153 | (localisation_stability ld).choose 154 | 155 | theorem ε_pos (ld : LocalisationData I I' f) : ∀ m, 0 < ld.ε m := 156 | (localisation_stability ld).choose_spec.choose 157 | 158 | theorem ε_cont (ld : LocalisationData I I' f) : Continuous ld.ε := 159 | (localisation_stability ld).choose_spec.choose_spec.choose 160 | 161 | theorem ε_spec (ld : LocalisationData I I' f) : 162 | ∀ (g : M → M') (_hg : ∀ m, dist (g m) (f m) < ld.ε m) (i : ld.ι), 163 | range (g ∘ ld.φ i) ⊆ range (ld.ψj i) := 164 | (localisation_stability ld).choose_spec.choose_spec.choose_spec 165 | 166 | end LocalisationData 167 | end 168 | 169 | variable (I I') 170 | 171 | open LocalisationData in 172 | theorem _root_.exists_stability_dist {f : M → M'} (hf : Continuous f) : 173 | ∃ ε : M → ℝ, (∀ m, 0 < ε m) ∧ Continuous ε ∧ 174 | ∀ x : M, 175 | ∃ φ : OpenSmoothEmbedding 𝓘(ℝ, E) E I M, 176 | ∃ ψ : OpenSmoothEmbedding 𝓘(ℝ, E') E' I' M', 177 | x ∈ range φ ∧ 178 | ∀ (g : M → M'), (∀ m, dist (g m) (f m) < ε m) → range (g ∘ φ) ⊆ range ψ := by 179 | let L := stdLocalisationData E I E' I' hf 180 | use L.ε, L.ε_pos, L.ε_cont 181 | intro x 182 | rcases mem_iUnion.mp <| eq_univ_iff_forall.mp L.h₁ x with ⟨i, hi⟩ 183 | use L.φ i, L.ψj i, mem_range_of_mem_image (φ L i) _ hi, ?_ 184 | have := L.ε_spec 185 | tauto 186 | 187 | end 188 | -------------------------------------------------------------------------------- /SphereEversion/Global/LocalizedConstruction.lean: -------------------------------------------------------------------------------- 1 | import SphereEversion.Global.Localisation 2 | import SphereEversion.Local.HPrinciple 3 | 4 | noncomputable section 5 | 6 | open Set Filter ModelWithCorners Metric 7 | 8 | open scoped Topology Manifold ContDiff 9 | 10 | variable {EM : Type*} [NormedAddCommGroup EM] [NormedSpace ℝ EM] [FiniteDimensional ℝ EM] 11 | {HM : Type*} [TopologicalSpace HM] {IM : ModelWithCorners ℝ EM HM} 12 | {M : Type*} [TopologicalSpace M] [ChartedSpace HM M] [IsManifold IM ∞ M] 13 | [T2Space M] 14 | {EX : Type*} [NormedAddCommGroup EX] [NormedSpace ℝ EX] [FiniteDimensional ℝ EX] 15 | {HX : Type*} [TopologicalSpace HX] {IX : ModelWithCorners ℝ EX HX} 16 | {X : Type*} [MetricSpace X] [ChartedSpace HX X] [IsManifold IX ∞ X] 17 | 18 | theorem OpenSmoothEmbedding.improve_formalSol (φ : OpenSmoothEmbedding 𝓘(ℝ, EM) EM IM M) 19 | (ψ : OpenSmoothEmbedding 𝓘(ℝ, EX) EX IX X) {R : RelMfld IM M IX X} (hRample : R.Ample) 20 | (hRopen : IsOpen R) {C : Set M} (hC : IsClosed C) {δ : M → ℝ} (hδ_pos : ∀ x, 0 < δ x) 21 | (hδ_cont : Continuous δ) {F : FormalSol R} (hFφψ : F.bs '' range φ ⊆ range ψ) 22 | (hFC : ∀ᶠ x near C, F.IsHolonomicAt x) {K₀ K₁ : Set EM} (hK₀ : IsCompact K₀) 23 | (hK₁ : IsCompact K₁) (hK₀K₁ : K₀ ⊆ interior K₁) : 24 | ∃ F' : HtpyFormalSol R, 25 | (∀ᶠ t near Iic (0 : ℝ), F' t = F) ∧ 26 | (∀ᶠ t near Ici (1 : ℝ), F' t = F' 1) ∧ 27 | (∀ᶠ x near C, ∀ t, F' t x = F x) ∧ 28 | (∀ t, ∀ x ∉ φ '' K₁, F' t x = F x) ∧ 29 | (∀ t x, dist ((F' t).bs x) (F.bs x) < δ x) ∧ 30 | ∀ᶠ x near C ∪ φ '' K₀, (F' 1).IsHolonomicAt x := by 31 | let Rloc : RelLoc EM EX := (R.localize φ ψ).relLoc 32 | have hRloc_op : IsOpen Rloc := 33 | isOpen_of_isOpen _ (hRopen.preimage <| OneJetBundle.continuous_transfer _ _) 34 | have hRloc_ample : Rloc.IsAmple := ample_of_ample _ (hRample.localize _ _) 35 | -- TODO: try to be consistent about how to state the hFφψ condition 36 | replace hFφψ : range (F.bs ∘ φ) ⊆ range ψ := by 37 | rw [range_comp] 38 | exact hFφψ 39 | let p : ChartPair IM M IX X := 40 | { φ 41 | ψ 42 | K₁ 43 | hK₁ } 44 | rcases p.dist_update' hδ_pos hδ_cont hFφψ with ⟨τ, τ_pos, hτ⟩ 45 | let 𝓕 := F.localize p hFφψ 46 | let L : Landscape EM := 47 | { C := φ ⁻¹' C 48 | K₀ 49 | K₁ 50 | hC := hC.preimage φ.continuous 51 | hK₀ 52 | hK₁ 53 | h₀₁ := hK₀K₁ } 54 | have h𝓕C : ∀ᶠ x : EM near L.C, 𝓕.IsHolonomicAt x := by 55 | rw [eventually_nhdsSet_iff_forall] at hFC ⊢ 56 | intro e he 57 | rw [φ.isInducing.nhds_eq_comap, eventually_comap] 58 | apply (hFC _ he).mono 59 | rintro x hx e rfl 60 | exact F.isHolonomicLocalize p hFφψ e hx 61 | rcases 𝓕.improve_htpy' hRloc_op hRloc_ample L τ_pos h𝓕C with 62 | ⟨𝓕', h𝓕't0, h𝓕't1, h𝓕'relC, h𝓕'relK₁, h𝓕'dist, h𝓕'hol⟩ 63 | have hcompat : p.compat' F 𝓕' := ⟨hFφψ, h𝓕'relK₁⟩ 64 | let F' : HtpyFormalSol R := p.mkHtpy F 𝓕' 65 | have hF'relK₁ : ∀ t, ∀ x ∉ φ '' K₁, F' t x = F x := by apply p.mkHtpy_eq_of_not_mem 66 | have hF't0 : ∀ᶠ t : ℝ near Iic 0, F' t = F := by 67 | apply h𝓕't0.mono 68 | rintro t ht 69 | exact p.mkHtpy_eq_of_forall hcompat ht 70 | have hF't1 : ∀ᶠ t : ℝ near Ici 1, F' t = F' 1 := h𝓕't1.mono fun t ↦ p.mkHtpy_congr _ 71 | refine ⟨F', hF't0, hF't1, ?_, hF'relK₁, ?_, ?_⟩ 72 | · apply φ.forall_near hK₁ h𝓕'relC (Eventually.of_forall fun x hx t ↦ hF'relK₁ t x hx) 73 | · intro e he t 74 | rw [p.mkHtpy_eq_of_eq _ _ hcompat] 75 | exact he t 76 | · intro t x 77 | rcases Classical.em (x ∈ φ '' K₁) with (⟨e, he, rfl⟩ | hx) 78 | · by_cases ht : t ∈ (Icc 0 1 : Set ℝ) 79 | · exact hτ hcompat e he t ht (h𝓕'dist e t) 80 | · rw [mem_Icc, not_and_or, not_le, not_le] at ht 81 | obtain (ht | ht) := ht 82 | · erw [hF't0.self_of_nhdsSet t ht.le, dist_self] 83 | apply hδ_pos 84 | · rw [hF't1.self_of_nhdsSet t ht.le] 85 | exact hτ hcompat e he 1 (right_mem_Icc.mpr zero_le_one) (h𝓕'dist e 1) 86 | · change dist (F' t x).1.2 (F.bs x) < δ x 87 | erw [p.mkHtpy_eq_of_not_mem _ _ hx, dist_self] 88 | apply hδ_pos 89 | · have h𝓕'holC : ∀ᶠ x : EM near L.C, (𝓕' 1).IsHolonomicAt x := by 90 | apply (h𝓕'relC.eventually_nhdsSet.and h𝓕C).mono 91 | rintro x ⟨hx, hx'⟩ 92 | exact JetSec.IsHolonomicAt.congr hx' (hx.mono fun x' hx' ↦ (hx' 1).symm) 93 | have : ∀ᶠ x near φ ⁻¹' C ∪ K₀, (𝓕' 1).IsHolonomicAt x := h𝓕'holC.union h𝓕'hol 94 | rw [← preimage_image_eq K₀ φ.injective, ← preimage_union] at this 95 | apply φ.forall_near hK₁ this 96 | · apply Filter.Eventually.union 97 | · apply hFC.mono 98 | intro x hx hx' 99 | apply hx.congr 100 | symm 101 | have : ∀ᶠ y in 𝓝 x, y ∈ (φ '' K₁)ᶜ := 102 | isOpen_iff_mem_nhds.mp (hK₁.image φ.continuous).isClosed.isOpen_compl x hx' 103 | apply this.mono 104 | exact hF'relK₁ _ 105 | · have : ∀ᶠ x near φ '' K₀, x ∈ p.φ '' K₁ := by 106 | suffices ∀ᶠ x near φ '' K₀, x ∈ interior (p.φ '' K₁) from this.mono interior_subset 107 | exact isOpen_interior.mem_nhdsSet.mpr 108 | ((image_subset φ hK₀K₁).trans (φ.isOpenMap.image_interior_subset K₁)) 109 | exact this.mono (fun a hx hx' ↦ (hx' hx).elim) 110 | · exact fun _ ↦ (p.mkHtpy_isHolonomicAt_iff hcompat).mpr 111 | -------------------------------------------------------------------------------- /SphereEversion/Global/TwistOneJetSec.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2022 Heather Macbeth. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Heather Macbeth 5 | 6 | ! This file was ported from Lean 3 source module global.twist_one_jet_sec 7 | -/ 8 | import SphereEversion.Global.OneJetSec 9 | 10 | noncomputable section 11 | 12 | open Set Equiv Bundle ContinuousLinearMap 13 | 14 | open scoped Manifold Bundle Topology ContDiff 15 | 16 | section ArbitraryField 17 | universe u v 18 | 19 | variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] {E : Type u} [NormedAddCommGroup E] 20 | [NormedSpace 𝕜 E] {H : Type*} [TopologicalSpace H] (I : ModelWithCorners 𝕜 E H) (M : Type*) 21 | [TopologicalSpace M] [ChartedSpace H M] [IsManifold I ∞ M] {F : Type*} 22 | [NormedAddCommGroup F] [NormedSpace 𝕜 F] {G : Type*} [TopologicalSpace G] 23 | {J : ModelWithCorners 𝕜 F G} {N : Type*} [TopologicalSpace N] [ChartedSpace G N] 24 | (V : Type*) [NormedAddCommGroup V] [NormedSpace 𝕜 V] 25 | (V' : Type*) [NormedAddCommGroup V'] [NormedSpace 𝕜 V'] 26 | 27 | section Smoothness 28 | 29 | 30 | notation "J¹[" 𝕜 ", " E ", " I ", " M ", " V "]" => TotalSpace (E →L[𝕜] V) 31 | (Bundle.ContinuousLinearMap (RingHom.id 𝕜) (TangentSpace I : M → _) (Bundle.Trivial M V)) 32 | variable {I M V} 33 | variable {f : N → J¹[𝕜, E, I, M, V]} 34 | 35 | -- todo: remove or use to prove `contMDiff_one_jet_eucl_bundle` 36 | theorem contMDiffAt_one_jet_eucl_bundle' {x₀ : N} : 37 | ContMDiffAt J (I.prod 𝓘(𝕜, E →L[𝕜] V)) ∞ f x₀ ↔ ContMDiffAt J I ∞ (fun x ↦ (f x).1) x₀ ∧ 38 | ContMDiffAt J 𝓘(𝕜, E →L[𝕜] V) ∞ (fun x ↦ show E →L[𝕜] V from 39 | (f x).2 ∘L (trivializationAt E (TangentSpace I : M → _) (f x₀).1).symmL 𝕜 (f x).1) x₀ := by 40 | simp_rw [contMDiffAt_hom_bundle, inCoordinates, Trivial.trivializationAt, 41 | Trivial.trivialization_continuousLinearMapAt] 42 | dsimp only [Bundle.Trivial] 43 | simp_rw [ContinuousLinearMap.id_comp] 44 | 45 | theorem contMDiffAt_one_jet_eucl_bundle {x₀ : N} : 46 | ContMDiffAt J (I.prod 𝓘(𝕜, E →L[𝕜] V)) ∞ f x₀ ↔ 47 | ContMDiffAt J I ∞ (fun x ↦ (f x).1) x₀ ∧ 48 | ContMDiffAt J 𝓘(𝕜, E →L[𝕜] V) ∞ (fun x ↦ show E →L[𝕜] V from 49 | (f x).2 ∘L (trivializationAt E (TangentSpace I) (f x₀).proj).symmL 𝕜 (f x).proj) x₀ := by 50 | rw [contMDiffAt_hom_bundle, and_congr_right_iff] 51 | intro hf 52 | refine Filter.EventuallyEq.contMDiffAt_iff ?_ 53 | have := 54 | hf.continuousAt.preimage_mem_nhds 55 | (((tangentBundleCore I M).isOpen_baseSet (achart H (f x₀).proj)).mem_nhds 56 | ((tangentBundleCore I M).mem_baseSet_at (f x₀).proj)) 57 | filter_upwards [this] with x _ 58 | simp_rw [inCoordinates, Trivial.trivializationAt, 59 | Trivial.trivialization_continuousLinearMapAt, ← ContinuousLinearMap.comp_assoc] 60 | dsimp only [Bundle.Trivial] 61 | simp_rw [ContinuousLinearMap.id_comp] 62 | 63 | theorem ContMDiffAt.one_jet_eucl_bundle_mk' {f : N → M} {ϕ : N → E →L[𝕜] V} {x₀ : N} 64 | (hf : ContMDiffAt J I ∞ f x₀) 65 | (hϕ : ContMDiffAt J 𝓘(𝕜, E →L[𝕜] V) ∞ (fun x ↦ show E →L[𝕜] V from 66 | ϕ x ∘L (trivializationAt E (TangentSpace I : M → _) (f x₀)).symmL 𝕜 (f x)) x₀) : 67 | ContMDiffAt J (I.prod 𝓘(𝕜, E →L[𝕜] V)) ∞ 68 | (fun x ↦ Bundle.TotalSpace.mk (f x) (ϕ x) : N → J¹[𝕜, E, I, M, V]) x₀ := 69 | contMDiffAt_one_jet_eucl_bundle'.mpr ⟨hf, hϕ⟩ 70 | 71 | theorem ContMDiffAt.one_jet_eucl_bundle_mk {f : N → M} {ϕ : N → E →L[𝕜] V} {x₀ : N} 72 | (hf : ContMDiffAt J I ∞ f x₀) 73 | (hϕ : ContMDiffAt J 𝓘(𝕜, E →L[𝕜] V) ∞ (fun x ↦ show E →L[𝕜] V from 74 | ϕ x ∘L (trivializationAt E (TangentSpace I) (f x₀)).symmL 𝕜 (f x)) x₀) : 75 | ContMDiffAt J (I.prod 𝓘(𝕜, E →L[𝕜] V)) ∞ 76 | (fun x ↦ Bundle.TotalSpace.mk (f x) (ϕ x) : N → J¹[𝕜, E, I, M, V]) x₀ := 77 | contMDiffAt_one_jet_eucl_bundle.mpr ⟨hf, hϕ⟩ 78 | 79 | end Smoothness 80 | 81 | section Sections 82 | 83 | /-- A section of a 1-jet bundle seen as a bundle over the source manifold. -/ 84 | @[ext] 85 | structure OneJetEuclSec where 86 | toFun : M → J¹[𝕜, E, I, M, V] 87 | is_sec' : ∀ p, (toFun p).1 = p 88 | smooth' : ContMDiff I (I.prod 𝓘(𝕜, E →L[𝕜] V)) ∞ toFun 89 | 90 | variable {I M V} 91 | 92 | instance : DFunLike (OneJetEuclSec I M V) M fun _ ↦ J¹[𝕜, E, I, M, V] where 93 | coe := OneJetEuclSec.toFun 94 | coe_injective' := by 95 | intro S T h 96 | ext x <;> rw [h] 97 | 98 | @[simp] 99 | theorem OneJetEuclSec.is_sec (s : OneJetEuclSec I M V) (p : M) : (s p).1 = p := 100 | s.is_sec' p 101 | 102 | @[simp] 103 | theorem OneJetEuclSec.smooth (s : OneJetEuclSec I M V) : ContMDiff I (I.prod 𝓘(𝕜, E →L[𝕜] V)) ∞ s := 104 | s.smooth' 105 | 106 | end Sections 107 | 108 | section proj 109 | 110 | instance piBugInstanceRestatement (x : M) : TopologicalSpace 111 | (Bundle.ContinuousLinearMap (RingHom.id 𝕜) (TangentSpace I) (Trivial M V) x) := by 112 | infer_instance 113 | 114 | instance piBugInstanceRestatement2 (x : M × V) : TopologicalSpace (OneJetSpace I 𝓘(𝕜, V) x) := by 115 | infer_instance 116 | 117 | /- Given a smooth manifold `M` and a normed space `V`, there is a canonical projection from the 118 | one-jet bundle of maps from `M` to `V` to the bundle of homomorphisms from `TM` to `V`. This is 119 | constructed using the fact that each tangent space to `V` is canonically isomorphic to `V`. -/ 120 | def proj (v : OneJetBundle I M 𝓘(𝕜, V) V) : J¹[𝕜, E, I, M, V] := 121 | ⟨v.1.1, v.2⟩ 122 | 123 | theorem smooth_proj : ContMDiff ((I.prod 𝓘(𝕜, V)).prod 𝓘(𝕜, E →L[𝕜] V)) (I.prod 𝓘(𝕜, E →L[𝕜] V)) 124 | ∞ (proj I M V) := by 125 | intro x₀ 126 | have : ContMDiffAt ((I.prod 𝓘(𝕜, V)).prod 𝓘(𝕜, E →L[𝕜] V)) _ ∞ id x₀ := contMDiffAt_id 127 | simp_rw +unfoldPartialApp [contMDiffAt_oneJetBundle, inTangentCoordinates, inCoordinates, 128 | TangentBundle.continuousLinearMapAt_model_space, ContinuousLinearMap.one_def, 129 | TangentSpace, ContinuousLinearMap.id_comp] at this 130 | exact this.1.one_jet_eucl_bundle_mk this.2.2 131 | 132 | variable {I M V} 133 | 134 | def drop (s : OneJetSec I M 𝓘(𝕜, V) V) : OneJetEuclSec I M V where 135 | toFun := (proj I M V).comp s 136 | is_sec' _ := rfl 137 | smooth' := (smooth_proj I M V).comp s.smooth 138 | 139 | end proj 140 | 141 | section incl 142 | 143 | /- Given a smooth manifold `M` and a normed space `V`, there is a canonical map from the 144 | the product with V of the bundle of homomorphisms from `TM` to `V` to the one-jet bundle of maps 145 | from `M` to `V`. In fact this map is a diffeomorphism. This is constructed using the fact that each 146 | tangent space to `V` is canonically isomorphic to `V`. -/ 147 | def incl (v : J¹[𝕜, E, I, M, V] × V) : OneJetBundle I M 𝓘(𝕜, V) V := 148 | ⟨(v.1.1, v.2), v.1.2⟩ 149 | 150 | theorem smooth_incl : ContMDiff ((I.prod 𝓘(𝕜, E →L[𝕜] V)).prod 𝓘(𝕜, V)) 151 | ((I.prod 𝓘(𝕜, V)).prod 𝓘(𝕜, E →L[𝕜] V)) ∞ (incl I M V) := by 152 | intro x₀ 153 | have : ContMDiffAt ((I.prod 𝓘(𝕜, E →L[𝕜] V)).prod 𝓘(𝕜, V)) _ ∞ Prod.fst x₀ := contMDiffAt_fst 154 | rw [contMDiffAt_one_jet_eucl_bundle] at this 155 | refine this.1.oneJetBundle_mk contMDiffAt_snd ?_ 156 | unfold inTangentCoordinates inCoordinates TangentSpace 157 | simp_rw [TangentBundle.continuousLinearMapAt_model_space, ContinuousLinearMap.one_def, 158 | ContinuousLinearMap.id_comp] 159 | exact this.2 160 | 161 | omit [IsManifold I ∞ M] in 162 | @[simp] 163 | theorem incl_fst_fst (v : J¹[𝕜, E, I, M, V] × V) : (incl I M V v).1.1 = v.1.1 := 164 | rfl 165 | 166 | omit [IsManifold I ∞ M] in 167 | @[simp] 168 | theorem incl_snd (v : J¹[𝕜, E, I, M, V] × V) : (incl I M V v).1.2 = v.2 := 169 | rfl 170 | 171 | end incl 172 | 173 | end ArbitraryField 174 | 175 | section familyTwist 176 | 177 | variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] {H : Type*} [TopologicalSpace H] 178 | (I : ModelWithCorners ℝ E H) (M : Type*) [TopologicalSpace M] [ChartedSpace H M] 179 | [IsManifold I ∞ M] (V : Type*) [NormedAddCommGroup V] [NormedSpace ℝ V] 180 | (V' : Type*) [NormedAddCommGroup V'] [NormedSpace ℝ V'] {F : Type*} [NormedAddCommGroup F] 181 | [NormedSpace ℝ F] {G : Type*} [TopologicalSpace G] (J : ModelWithCorners ℝ F G) (N : Type*) 182 | [TopologicalSpace N] [ChartedSpace G N] 183 | 184 | /-- A section of a 1-jet bundle seen as a bundle over the source manifold. -/ 185 | @[ext] 186 | structure FamilyOneJetEuclSec where 187 | toFun : N × M → J¹[ℝ, E, I, M, V] 188 | is_sec' : ∀ p, (toFun p).1 = p.2 189 | smooth' : ContMDiff (J.prod I) (I.prod 𝓘(ℝ, E →L[ℝ] V)) ∞ toFun 190 | 191 | instance : FunLike (FamilyOneJetEuclSec I M V J N) (N × M) J¹[ℝ, E, I, M, V] where 192 | coe := FamilyOneJetEuclSec.toFun 193 | coe_injective' := by 194 | intro S T h 195 | ext x <;> rw [h] 196 | 197 | variable {I M V J N} 198 | 199 | @[simp] 200 | theorem FamilyOneJetEuclSec.is_sec (s : FamilyOneJetEuclSec I M V J N) (p : N × M) : 201 | (s p).1 = p.2 := 202 | s.is_sec' p 203 | 204 | @[simp] 205 | theorem FamilyOneJetEuclSec.smooth (s : FamilyOneJetEuclSec I M V J N) : 206 | ContMDiff (J.prod I) (I.prod 𝓘(ℝ, E →L[ℝ] V)) ∞ s := 207 | s.smooth' 208 | 209 | variable {V'} 210 | 211 | def familyJoin {f : N × M → V} (hf : ContMDiff (J.prod I) 𝓘(ℝ, V) ∞ f) 212 | (s : FamilyOneJetEuclSec I M V J N) : FamilyOneJetSec I M 𝓘(ℝ, V) V J N 213 | where 214 | bs n m := (incl I M V (s (n, m), f (n, m))).1.2 215 | ϕ n m := (incl I M V (s (n, m), f (n, m))).2 216 | smooth' := by 217 | convert (smooth_incl I M V).comp (s.smooth.prodMk hf) 218 | ext : 1 <;> simp 219 | 220 | def familyTwist (s : OneJetEuclSec I M V) (i : N × M → V →L[ℝ] V') 221 | (i_smooth : ∀ x₀ : N × M, ContMDiffAt (J.prod I) 𝓘(ℝ, V →L[ℝ] V') ∞ i x₀) : 222 | FamilyOneJetEuclSec I M V' J N 223 | where 224 | toFun p := ⟨p.2, (i p).comp (s p.2).2⟩ 225 | is_sec' p := rfl 226 | smooth' := by 227 | refine fun x₀ ↦ contMDiffAt_snd.one_jet_eucl_bundle_mk' ?_ 228 | simp_rw [ContinuousLinearMap.comp_assoc] 229 | have : ContMDiffAt (J.prod I) _ ∞ (fun x : N × M ↦ _) x₀ := s.smooth.comp contMDiff_snd x₀ 230 | rw [contMDiffAt_one_jet_eucl_bundle'] at this 231 | refine (i_smooth x₀).clm_comp ?_ 232 | convert this.2 <;> simp [s.is_sec] 233 | 234 | end familyTwist 235 | -------------------------------------------------------------------------------- /SphereEversion/Indexing.lean: -------------------------------------------------------------------------------- 1 | import Mathlib.Order.Interval.Finset.Fin 2 | import Mathlib.Data.Fin.SuccPred 3 | import Mathlib.Data.Nat.SuccPred 4 | import Mathlib.SetTheory.Cardinal.Basic 5 | import SphereEversion.ToMathlib.Data.Nat.Basic 6 | /-! 7 | # Indexing types 8 | 9 | This file introduces `IndexType : ℕ → Type` such that `IndexType 0 = ℕ` and 10 | `IndexType (N+1) = Fin (N+1)`. Each `IndexType N` has a total order and and inductive principle 11 | together with supporting lemmas. 12 | -/ 13 | 14 | 15 | open Fin Set 16 | 17 | /-- Our model indexing type depending on `n : ℕ` is `ℕ` if `n = 0` and `Fin n` otherwise -/ 18 | def IndexType (n : ℕ) : Type := 19 | Nat.casesOn n ℕ fun k ↦ Fin <| k + 1 20 | 21 | def IndexType.fromNat : ∀ {N : ℕ}, ℕ → IndexType N 22 | | 0 => id 23 | | N + 1 => (Nat.cast : ℕ → Fin (N + 1)) 24 | 25 | def IndexType.toNat : ∀ {N}, IndexType N → ℕ 26 | | 0 => id 27 | | _ + 1 => Fin.val 28 | 29 | @[simp] 30 | theorem indexType_zero : IndexType 0 = ℕ := 31 | rfl 32 | 33 | @[simp] 34 | theorem indexType_succ (n : ℕ) : IndexType (n + 1) = Fin (n + 1) := 35 | rfl 36 | 37 | @[simp] 38 | theorem indexType_of_zero_lt {n : ℕ} (h : 0 < n) : IndexType n = Fin n := by 39 | rw [← Nat.succ_pred_eq_of_pos h, indexType_succ] 40 | 41 | set_option hygiene false in 42 | run_cmd 43 | for n in [`LinearOrder, `LocallyFiniteOrder, `OrderBot, `Zero, `SuccOrder].map Lean.mkIdent do 44 | Lean.Elab.Command.elabCommand (← `( 45 | instance : ∀ n, $n (IndexType n) 46 | | 0 => inferInstanceAs ($n ℕ) 47 | | (N + 1) => inferInstanceAs ($n (Fin (N + 1))) 48 | )) 49 | 50 | theorem Set.countable_iff_exists_nonempty_indexType_equiv {α : Type*} {s : Set α} 51 | (hne : s.Nonempty) : s.Countable ↔ ∃ n, Nonempty (IndexType n ≃ s) := by 52 | -- Huge golfing opportunity. 53 | obtain (h | h) := s.finite_or_infinite 54 | · refine ⟨fun _ ↦ ⟨h.toFinset.card, ?_⟩, fun _ ↦ h.countable⟩ 55 | have : 0 < h.toFinset.card := by 56 | rw [Finset.card_pos] 57 | exact (Set.Finite.toFinset_nonempty h).mpr hne 58 | simp only [this, indexType_of_zero_lt] 59 | have e₁ := Fintype.equivFin h.toFinset 60 | rw [Fintype.card_coe, h.coeSort_toFinset] at e₁ 61 | exact ⟨e₁.symm⟩ 62 | · refine ⟨fun hh ↦ ⟨0, ?_⟩, ?_⟩ 63 | · simp only [indexType_zero] 64 | obtain ⟨_i⟩ := Set.countable_infinite_iff_nonempty_denumerable.mp ⟨hh, h⟩ 65 | haveI := _i 66 | exact ⟨(Denumerable.eqv s).symm⟩ 67 | · rintro ⟨n, ⟨fn⟩⟩ 68 | have hn : n = 0 := by 69 | by_contra hn 70 | replace hn : 0 < n := Nat.pos_iff_ne_zero.mpr hn 71 | simp only [hn, indexType_of_zero_lt] at fn 72 | exact h (Finite.intro fn.symm) 73 | simp only [hn, indexType_zero] at fn 74 | exact Set.countable_iff_exists_injective.mpr ⟨fn.symm, fn.symm.injective⟩ 75 | 76 | variable {n : ℕ} 77 | 78 | theorem IndexType.not_lt_zero (j : IndexType n) : ¬j < 0 := 79 | Nat.casesOn n Nat.not_lt_zero (fun _ ↦ Fin.not_lt_zero) j 80 | 81 | theorem IndexType.zero_le (i : IndexType n) : 0 ≤ i := by cases n <;> dsimp at * <;> simp 82 | 83 | def IndexType.succ {N : ℕ} : IndexType N → IndexType N := Order.succ 84 | 85 | theorem IndexType.succ_castSuccEmb (i : Fin n) : 86 | @IndexType.succ (n + 1) i.castSucc = i.succ := 87 | (Fin.orderSucc_apply _).trans (lastCases_castSucc i) 88 | 89 | theorem IndexType.succ_eq (i : IndexType n) : i.succ = i ↔ IsMax i := 90 | Order.succ_eq_iff_isMax 91 | 92 | theorem IndexType.lt_succ (i : IndexType n) (h : ¬IsMax i) : i < i.succ := 93 | Order.lt_succ_of_not_isMax h 94 | 95 | theorem IndexType.le_max {i : IndexType n} (h : IsMax i) (j) : j ≤ i := 96 | h.isTop j 97 | 98 | nonrec theorem IndexType.le_of_lt_succ (i : IndexType n) {j : IndexType n} 99 | (h : i < j.succ) : i ≤ j := 100 | Order.le_of_lt_succ h 101 | 102 | theorem IndexType.exists_castSucc_eq (i : IndexType (n + 1)) (hi : ¬IsMax i) : 103 | ∃ i' : Fin n, i'.castSucc = i := by 104 | revert hi 105 | refine Fin.lastCases ?_ (fun i _ ↦ ⟨_, rfl⟩) i 106 | intro hi; apply hi.elim; intro i _; exact le_last i 107 | 108 | theorem IndexType.toNat_succ (i : IndexType n) (hi : ¬IsMax i) : 109 | i.succ.toNat = i.toNat + 1 := by 110 | cases n; · rfl 111 | rcases i.exists_castSucc_eq hi with ⟨i, rfl⟩ 112 | rw [IndexType.succ_castSuccEmb] 113 | exact val_succ i 114 | 115 | -- @[simp] can prove this 116 | theorem IndexType.not_isMax (n : IndexType 0) : ¬IsMax n := 117 | not_isMax_of_lt <| Nat.lt_succ_self n 118 | 119 | @[elab_as_elim] 120 | theorem IndexType.induction_from {P : IndexType n → Prop} {i₀ : IndexType n} (h₀ : P i₀) 121 | (ih : ∀ i ≥ i₀, ¬IsMax i → P i → P i.succ) : ∀ i ≥ i₀, P i := by 122 | cases n 123 | · intro i h 124 | induction' h with i hi₀i hi ih 125 | · exact h₀ 126 | exact ih i hi₀i (IndexType.not_isMax i) hi 127 | intro i 128 | refine Fin.induction ?_ ?_ i 129 | · intro hi; convert h₀; exact (hi.le.antisymm <| Fin.zero_le _).symm 130 | · intro i hi hi₀i 131 | rcases hi₀i.le.eq_or_lt with (rfl | hi₀i) 132 | · exact h₀ 133 | rw [← IndexType.succ_castSuccEmb] 134 | refine ih _ ?_ ?_ ?_ 135 | · rwa [ge_iff_le, le_castSucc_iff] 136 | · exact not_isMax_of_lt (castSucc_lt_succ i) 137 | · apply hi; rwa [ge_iff_le, le_castSucc_iff] 138 | 139 | @[elab_as_elim] 140 | theorem IndexType.induction {P : IndexType n → Prop} (h₀ : P 0) 141 | (ih : ∀ i, ¬IsMax i → P i → P i.succ) : ∀ i, P i := fun i ↦ 142 | IndexType.induction_from h₀ (fun i _ ↦ ih i) i i.zero_le 143 | 144 | -- We make `P` and `Q` explicit to help the elaborator when applying the lemma 145 | -- (elab_as_eliminator isn't enough). 146 | theorem IndexType.exists_by_induction {α : Type*} (P : IndexType n → α → Prop) 147 | (Q : IndexType n → α → α → Prop) (h₀ : ∃ a, P 0 a) 148 | (ih : ∀ n a, P n a → ¬IsMax n → ∃ a', P n.succ a' ∧ Q n a a') : 149 | ∃ f : IndexType n → α, ∀ n, P n (f n) ∧ (¬IsMax n → Q n (f n) (f n.succ)) := by 150 | revert P Q h₀ ih 151 | cases' n with N 152 | · intro P Q h₀ ih 153 | rcases exists_by_induction' P Q h₀ (by simpa using ih) with ⟨f, hf⟩ 154 | exact ⟨f, fun n ↦ ⟨(hf n).1, fun _ ↦ (hf n).2⟩⟩ 155 | · -- dsimp only [IndexType, IndexType.succ] 156 | intro P Q h₀ ih 157 | choose f₀ hf₀ using h₀ 158 | choose! F hF hF' using ih 159 | let G := fun i : Fin N ↦ F i.castSucc 160 | let f : Fin (N + 1) → α := fun i ↦ Fin.induction f₀ G i 161 | have key : ∀ i, P i (f i) := by 162 | refine fun i ↦ Fin.induction hf₀ ?_ i 163 | intro i hi 164 | simp_rw [f, induction_succ, ← IndexType.succ_castSuccEmb] 165 | exact hF _ _ hi (not_isMax_of_lt (castSucc_lt_succ i)) 166 | refine ⟨f, fun i ↦ ⟨key i, fun hi ↦ ?_⟩⟩ 167 | convert hF' _ _ (key i) hi 168 | rcases i.exists_castSucc_eq hi with ⟨i, rfl⟩ 169 | simp_rw [IndexType.succ_castSuccEmb, f, induction_succ] 170 | rfl 171 | -------------------------------------------------------------------------------- /SphereEversion/Local/AmpleRelation.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021 Patrick Massot. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Patrick Massot 5 | -/ 6 | import Mathlib.Analysis.Convex.AmpleSet 7 | import SphereEversion.Local.DualPair 8 | import SphereEversion.Local.Relation 9 | 10 | /-! # Slices of first order relations 11 | 12 | Recal that a first order partial differential relation for maps between real normed vector spaces 13 | `E` and `F` is a set `R` in `OneJet E F := E × F × (E →L[ℝ] F)`. In this file we study slices 14 | of such relations. The word slice is meant to convey the idea of intersecting with an affine 15 | subspace. Here we fix `(x, y, φ) : OneJet E F` and some hyperplane `H` in `E`. The points 16 | `x` and `y` are fixed and we will take a slice in `E →L[ℝ] F` by intersecting `R` with the affine 17 | subspace of linear maps that coincide with `φ` on `H`. 18 | 19 | It will be convenient for convex integration purposes to identify this slice with `F`. There is 20 | no natural identification but we can build one by fixing more data that a hyperplane in `E`. 21 | Namely we fix `p : DualPair E` (where `ker p.π` is the relevant hyperplane) and reformulate 22 | "linear map that coincides with `φ` on `H`" as `p.update φ w` for some `w : F`. 23 | 24 | This `slice` definition allows to define `RelLoc.isAmple`, the ampleness condition for first 25 | order relations: a relation is ample if all its slices are ample sets. 26 | 27 | At the end of the file we consider 1-jet sections and slices corresponding to points in their image. 28 | -/ 29 | 30 | 31 | variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] 32 | {F : Type*} [NormedAddCommGroup F] [NormedSpace ℝ F] 33 | {X : Type*} [NormedAddCommGroup X] [InnerProductSpace ℝ X] 34 | 35 | variable {R : RelLoc E F} 36 | 37 | open Set 38 | 39 | /-! ## Slices and ampleness -/ 40 | 41 | 42 | namespace RelLoc 43 | 44 | /-- The slice of a local relation `R : RelLoc E F` for a dual pair `p` at a jet `θ` is 45 | the set of `w` in `F` such that updating `θ` using `p` and `w` leads to a jet in `R`. -/ 46 | def slice (R : RelLoc E F) (p : DualPair E) (θ : E × F × (E →L[ℝ] F)) : Set F := 47 | {w | (θ.1, θ.2.1, p.update θ.2.2 w) ∈ R} 48 | 49 | theorem mem_slice (R : RelLoc E F) {p : DualPair E} {θ : E × F × (E →L[ℝ] F)} {w : F} : 50 | w ∈ R.slice p θ ↔ (θ.1, θ.2.1, p.update θ.2.2 w) ∈ R := 51 | Iff.rfl 52 | 53 | /-- A relation is ample if all its slices are ample. -/ 54 | def IsAmple (R : RelLoc E F) : Prop := 55 | ∀ (p : DualPair E) (θ : E × F × (E →L[ℝ] F)), AmpleSet (R.slice p θ) 56 | 57 | theorem IsAmple.mem_hull (h : IsAmple R) {θ : E × F × (E →L[ℝ] F)} (hθ : θ ∈ R) (v : F) (p) : 58 | v ∈ hull (connectedComponentIn (R.slice p θ) (θ.2.2 p.v)) := by 59 | rw [h p θ (θ.2.2 p.v)] 60 | · exact mem_univ _ 61 | rw [mem_slice, p.update_self] 62 | exact hθ 63 | 64 | theorem slice_update {θ : E × F × (E →L[ℝ] F)} {p : DualPair E} (x : F) : 65 | R.slice p (θ.1, θ.2.1, p.update θ.2.2 x) = R.slice p θ := by 66 | ext1 w 67 | dsimp [slice] 68 | rw [p.update_update] 69 | 70 | /-- In order to check ampleness, it suffices to consider slices through elements of the relation. -/ 71 | theorem isAmple_iff : 72 | R.IsAmple ↔ ∀ ⦃θ : OneJet E F⦄ (p : DualPair E), θ ∈ R → AmpleSet (R.slice p θ) := by 73 | refine ⟨fun h θ p _ ↦ h p θ, fun h p θ w hw ↦ ?_⟩ 74 | dsimp [slice] at hw 75 | have := h p hw 76 | rw [slice_update] at this 77 | exact this w hw 78 | 79 | open scoped Pointwise 80 | 81 | theorem slice_of_ker_eq_ker {θ : OneJet E F} {p p' : DualPair E} (hpp' : p.π = p'.π) : 82 | R.slice p θ = θ.2.2 (p.v - p'.v) +ᵥ R.slice p' θ := by 83 | rcases θ with ⟨x, y, φ⟩ 84 | have key : ∀ w, p'.update φ w = p.update φ (w + φ (p.v - p'.v)) := fun w ↦ by 85 | simp only [DualPair.update, hpp', map_sub, add_right_inj] 86 | congr 2 87 | abel 88 | ext w 89 | simp only [slice, mem_setOf_eq, map_sub, vadd_eq_add, mem_vadd_set_iff_neg_vadd_mem, key] 90 | have : -(φ p.v - φ p'.v) + w + (φ p.v - φ p'.v) = w := by abel 91 | rw [this] 92 | 93 | theorem ample_slice_of_ample_slice {θ : OneJet E F} {p p' : DualPair E} (hpp' : p.π = p'.π) 94 | (h : AmpleSet (R.slice p θ)) : AmpleSet (R.slice p' θ) := by 95 | rw [slice_of_ker_eq_ker hpp'.symm] 96 | exact h.vadd --h 97 | 98 | theorem ample_slice_of_forall (R : RelLoc E F) {x y φ} (p : DualPair E) 99 | (h : ∀ w, (x, y, p.update φ w) ∈ R) : AmpleSet (R.slice p (x, y, φ)) := by 100 | rw [show R.slice p (x, y, φ) = univ from eq_univ_of_forall h] 101 | exact ampleSet_univ 102 | 103 | end RelLoc 104 | 105 | open RelLoc 106 | 107 | /-! ## Slices for 1-jet sections and formal solutions. -/ 108 | 109 | namespace JetSec 110 | 111 | /-- The slice associated to a jet section and a dual pair at some point. -/ 112 | def sliceAt (𝓕 : JetSec E F) (R : RelLoc E F) (p : DualPair E) (x : E) : Set F := 113 | R.slice p (x, 𝓕.f x, 𝓕.φ x) 114 | 115 | /-- A 1-jet section `𝓕` is short for a dual pair `p` at a point `x` if the derivative of 116 | the function `𝓕.f` at `x` is in the convex hull of the relevant connected component of the 117 | corresponding slice. -/ 118 | def IsShortAt (𝓕 : JetSec E F) (R : RelLoc E F) (p : DualPair E) (x : E) : Prop := 119 | D 𝓕.f x p.v ∈ hull (connectedComponentIn (𝓕.sliceAt R p x) <| 𝓕.φ x p.v) 120 | 121 | end JetSec 122 | 123 | namespace RelLoc.FormalSol 124 | 125 | /-- The slice associated to a formal solution and a dual pair at some point. -/ 126 | def sliceAt (𝓕 : FormalSol R) (p : DualPair E) (x : E) : Set F := 127 | R.slice p (x, 𝓕.f x, 𝓕.φ x) 128 | 129 | /-- A formal solution `𝓕` is short for a dual pair `p` at a point `x` if the derivative of 130 | the function `𝓕.f` at `x` is in the convex hull of the relevant connected component of the 131 | corresponding slice. -/ 132 | def IsShortAt (𝓕 : FormalSol R) (p : DualPair E) (x : E) : Prop := 133 | D 𝓕.f x p.v ∈ hull (connectedComponentIn (𝓕.sliceAt p x) <| 𝓕.φ x p.v) 134 | 135 | end RelLoc.FormalSol 136 | 137 | theorem RelLoc.IsAmple.isShortAt (hR : IsAmple R) (𝓕 : FormalSol R) (p : DualPair E) (x : E) : 138 | 𝓕.IsShortAt p x := 139 | hR.mem_hull (𝓕.is_sol x) _ p 140 | -------------------------------------------------------------------------------- /SphereEversion/Local/DualPair.lean: -------------------------------------------------------------------------------- 1 | import Mathlib.Analysis.Calculus.ContDiff.FiniteDimension 2 | import Mathlib.Analysis.Normed.Module.Completion 3 | import Mathlib.Data.Complex.Norm 4 | import Mathlib.LinearAlgebra.Dual.Lemmas 5 | import SphereEversion.Notations 6 | import SphereEversion.ToMathlib.Analysis.NormedSpace.OperatorNorm.Prod 7 | import SphereEversion.ToMathlib.LinearAlgebra.Basic 8 | 9 | /-! # Dual pairs 10 | 11 | Convex integration improves maps by successively modify their derivative in a direction of a 12 | vector field (almost) without changing their derivative along some complementary hyperplace. 13 | 14 | In this file we introduce the linear algebra part of the story of such modifications. 15 | The main gadget here are dual pairs on a real topological vector space `E`. 16 | Each `p : DualPair E` is made of a (continuous) linear form `p.π : E →L[ℝ] ℝ` and a vector 17 | `p.v : E` such that `p.π p.v = 1`. 18 | 19 | Those pairs are used to build the updating operation turning `φ : E →L[ℝ] F` into 20 | `p.update φ w` which equals `φ` on `ker p.π` and sends `v` to the given `w : F`. 21 | 22 | After formalizing the above definitions, we first record a number of trivial algebraic lemmas. 23 | Then we prove a lemma which is geometrically obvious but not so easy to formalize: 24 | `injective_update_iff` states, starting with an injective `φ`, that the updated 25 | map `p.update φ w` is injective if and only if `w` isn't in the image of `ker p.π` under `φ`. 26 | This is crucial in order to apply convex integration to immersions. 27 | 28 | Then we prove continuity and smoothness lemmas for this operation. 29 | -/ 30 | 31 | 32 | noncomputable section 33 | 34 | open Function ContinuousLinearMap 35 | 36 | open LinearMap (ker) 37 | 38 | open scoped ContDiff 39 | 40 | /-! ## General theory -/ 41 | 42 | section NoNorm 43 | 44 | variable (E : Type*) {E' F G : Type*} 45 | 46 | variable [AddCommGroup E] [Module ℝ E] [TopologicalSpace E] 47 | 48 | variable [AddCommGroup E'] [Module ℝ E'] [TopologicalSpace E'] 49 | 50 | variable [NormedAddCommGroup F] [NormedSpace ℝ F] [NormedAddCommGroup G] [NormedSpace ℝ G] 51 | 52 | /-- A continuous linear form `π` and a vector `v` that pair to one. In particular `ker π` is a 53 | hyperplane and `v` spans a complement of this hyperplane. -/ 54 | structure DualPair where 55 | π : E →L[ℝ] ℝ 56 | v : E 57 | pairing : π v = 1 58 | 59 | namespace DualPair 60 | 61 | variable {E} 62 | 63 | attribute [local simp] ContinuousLinearMap.toSpanSingleton_apply 64 | 65 | theorem pi_ne_zero (p : DualPair E) : p.π ≠ 0 := fun H ↦ by 66 | simpa [H] using p.pairing 67 | 68 | theorem ker_pi_ne_top (p : DualPair E) : ker p.π ≠ ⊤ := fun H ↦ by 69 | have : p.π.toLinearMap p.v = 1 := p.pairing 70 | simpa [LinearMap.ker_eq_top.mp H] 71 | 72 | theorem v_ne_zero (p : DualPair E) : p.v ≠ 0 := fun H ↦ by simpa [H] using p.pairing 73 | 74 | /-- Given a dual pair `p`, `p.span_v` is the line spanned by `p.v`. -/ 75 | def spanV (p : DualPair E) : Submodule ℝ E := 76 | Submodule.span ℝ {p.v} 77 | 78 | theorem mem_spanV (p : DualPair E) {u : E} : u ∈ p.spanV ↔ ∃ t : ℝ, u = t • p.v := by 79 | simp [DualPair.spanV, Submodule.mem_span_singleton, eq_comm] 80 | 81 | /-- Update a continuous linear map `φ : E →L[ℝ] F` using a dual pair `p` on `E` and a 82 | vector `w : F`. The new map coincides with `φ` on `ker p.π` and sends `p.v` to `w`. -/ 83 | def update (p : DualPair E) (φ : E →L[ℝ] F) (w : F) : E →L[ℝ] F := 84 | φ + (w - φ p.v) ⬝ p.π 85 | 86 | @[simp] 87 | theorem update_ker_pi (p : DualPair E) (φ : E →L[ℝ] F) (w : F) {u} (hu : u ∈ ker p.π) : 88 | p.update φ w u = φ u := by 89 | rw [LinearMap.mem_ker] at hu 90 | simp [update, hu] 91 | 92 | @[simp] 93 | theorem update_v (p : DualPair E) (φ : E →L[ℝ] F) (w : F) : p.update φ w p.v = w := by 94 | simp [update, p.pairing] 95 | 96 | @[simp] 97 | theorem update_self (p : DualPair E) (φ : E →L[ℝ] F) : p.update φ (φ p.v) = φ := by 98 | simp only [update, add_zero, ContinuousLinearMap.toSpanSingleton_zero, 99 | ContinuousLinearMap.zero_comp, sub_self] 100 | 101 | @[simp] 102 | theorem update_update (p : DualPair E) (φ : E →L[ℝ] F) (w w' : F) : 103 | p.update (p.update φ w') w = p.update φ w := by 104 | simp_rw [update, add_apply, coe_comp', (· ∘ ·), toSpanSingleton_apply, p.pairing, one_smul, 105 | add_sub_cancel, add_assoc, ← ContinuousLinearMap.add_comp, ← toSpanSingleton_add, 106 | sub_add_eq_add_sub, add_sub_cancel] 107 | 108 | theorem inf_eq_bot (p : DualPair E) : ker p.π ⊓ p.spanV = ⊥ := bot_unique <| fun x hx ↦ by 109 | have : p.π x = 0 ∧ ∃ a : ℝ, a • p.v = x := by 110 | simpa [DualPair.spanV, Submodule.mem_span_singleton] using hx 111 | rcases this with ⟨H, t, rfl⟩ 112 | rw [p.π.map_smul, p.pairing, Algebra.id.smul_eq_mul, mul_one] at H 113 | simp [H] 114 | 115 | theorem sup_eq_top (p : DualPair E) : ker p.π ⊔ p.spanV = ⊤ := by 116 | rw [Submodule.sup_eq_top_iff] 117 | intro x 118 | refine ⟨x - p.π x • p.v, ?_, p.π x • p.v, ?_, ?_⟩ <;> 119 | simp [DualPair.spanV, Submodule.mem_span_singleton, p.pairing] 120 | 121 | theorem decomp (p : DualPair E) (e : E) : ∃ u ∈ ker p.π, ∃ t : ℝ, e = u + t • p.v := by 122 | have : e ∈ ker p.π ⊔ p.spanV := by 123 | rw [p.sup_eq_top] 124 | exact Submodule.mem_top 125 | simp_rw [Submodule.mem_sup, DualPair.mem_spanV] at this 126 | rcases this with ⟨u, hu, -, ⟨t, rfl⟩, rfl⟩ 127 | exact ⟨u, hu, t, rfl⟩ 128 | 129 | -- useful with `DualPair.decomp` 130 | theorem update_apply (p : DualPair E) (φ : E →L[ℝ] F) {w : F} {t : ℝ} {u} (hu : u ∈ ker p.π) : 131 | p.update φ w (u + t • p.v) = φ u + t • w := by 132 | rw [map_add, map_smul, p.update_v, p.update_ker_pi _ _ hu] 133 | 134 | /-- Map a dual pair under a linear equivalence. -/ 135 | @[simps] 136 | def map (p : DualPair E) (L : E ≃L[ℝ] E') : DualPair E' := 137 | ⟨p.π ∘L ↑L.symm, L p.v, (congr_arg p.π <| L.symm_apply_apply p.v).trans p.pairing⟩ 138 | 139 | theorem update_comp_left (p : DualPair E) (ψ' : F →L[ℝ] G) (φ : E →L[ℝ] F) (w : F) : 140 | p.update (ψ' ∘L φ) (ψ' w) = ψ' ∘L p.update φ w := by 141 | ext1 u 142 | simp only [update, add_apply, ContinuousLinearMap.comp_apply, toSpanSingleton_apply, map_add, 143 | map_smul, map_sub] 144 | 145 | theorem map_update_comp_right (p : DualPair E) (ψ' : E →L[ℝ] F) (φ : E ≃L[ℝ] E') (w : F) : 146 | (p.map φ).update (ψ' ∘L ↑φ.symm) w = p.update ψ' w ∘L ↑φ.symm := by 147 | ext1 u 148 | simp [update] 149 | 150 | /-! ## Injectivity criterion -/ 151 | 152 | 153 | @[simp] 154 | theorem injective_update_iff (p : DualPair E) {φ : E →L[ℝ] F} (hφ : Injective φ) {w : F} : 155 | Injective (p.update φ w) ↔ w ∉ (ker p.π).map φ := by 156 | constructor 157 | · rintro h ⟨u, hu, rfl⟩ 158 | have : p.update φ (φ u) p.v = φ u := p.update_v φ (φ u) 159 | conv_rhs at this => rw [← p.update_ker_pi φ (φ u) hu] 160 | rw [← h this] at hu 161 | simp only [SetLike.mem_coe, LinearMap.mem_ker] at hu 162 | rw [p.pairing] at hu 163 | linarith 164 | · intro hw 165 | refine (injective_iff_map_eq_zero (p.update φ w)).mpr fun x hx ↦ ?_ 166 | rcases p.decomp x with ⟨u, hu, t, rfl⟩ 167 | rw [map_add, map_smul, update_v, p.update_ker_pi _ _ hu] at hx 168 | obtain rfl : t = 0 := by 169 | by_contra ht 170 | apply hw 171 | refine ⟨-t⁻¹ • u, (ker p.π).smul_mem _ hu, ?_⟩ 172 | rw [map_smul] 173 | have : -t⁻¹ • (φ u + t • w) + w = -t⁻¹ • (0 : F) + w := congr_arg (-t⁻¹ • · + w) hx 174 | rwa [smul_add, neg_smul, neg_smul, inv_smul_smul₀ ht, smul_zero, zero_add, 175 | neg_add_cancel_right, ← neg_smul] at this 176 | rw [zero_smul, add_zero] at hx ⊢ 177 | exact (injective_iff_map_eq_zero φ).mp hφ u hx 178 | 179 | end DualPair 180 | 181 | end NoNorm 182 | 183 | /-! ## Smoothness and continuity -/ 184 | 185 | 186 | namespace DualPair 187 | 188 | variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] {F : Type*} [NormedAddCommGroup F] 189 | [NormedSpace ℝ F] 190 | 191 | /- In the next two lemmas, finite dimensionality of `E` is clearly uneeded, but allows 192 | to use `contDiff_clm_apply_iff` and `continuous_clm_apply`. -/ 193 | theorem smooth_update [FiniteDimensional ℝ E] (p : DualPair E) {G : Type*} [NormedAddCommGroup G] 194 | [NormedSpace ℝ G] {φ : G → E →L[ℝ] F} (hφ : 𝒞 ∞ φ) {w : G → F} (hw : 𝒞 ∞ w) : 195 | 𝒞 ∞ fun g ↦ p.update (φ g) (w g) := by 196 | apply hφ.add 197 | rw [contDiff_clm_apply_iff] 198 | intro y 199 | exact (hw.sub (contDiff_clm_apply_iff.mp hφ p.v)).const_smul _ 200 | 201 | theorem continuous_update [FiniteDimensional ℝ E] (p : DualPair E) {X : Type*} [TopologicalSpace X] 202 | {φ : X → E →L[ℝ] F} (hφ : Continuous φ) {w : X → F} (hw : Continuous w) : 203 | Continuous fun g ↦ p.update (φ g) (w g) := by 204 | apply hφ.add 205 | rw [continuous_clm_apply] 206 | intro y 207 | exact (hw.sub (continuous_clm_apply.mp hφ p.v)).const_smul _ 208 | 209 | end DualPair 210 | 211 | variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] {F : Type*} [NormedAddCommGroup F] 212 | [NormedSpace ℝ F] 213 | 214 | /-- Given a finite basis `e : basis ι ℝ E`, and `i : ι`, 215 | `e.DualPair i` is given by the `i`th basis element and its dual. -/ 216 | def Basis.dualPair [FiniteDimensional ℝ E] {ι : Type*} [Fintype ι] [DecidableEq ι] 217 | (e : Basis ι ℝ E) (i : ι) : DualPair E where 218 | π := LinearMap.toContinuousLinearMap (e.dualBasis i) 219 | v := e i 220 | pairing := by simp 221 | -------------------------------------------------------------------------------- /SphereEversion/Local/Relation.lean: -------------------------------------------------------------------------------- 1 | import Mathlib.Topology.MetricSpace.HausdorffDistance 2 | import SphereEversion.Local.OneJet 3 | 4 | /-! 5 | # Local partial differential relations and their formal solutions 6 | 7 | This file defines `RelLoc E F`, the type of first order partial differential relations 8 | for maps between two real normed spaces `E` and `F`. 9 | 10 | To any `R : RelLoc E F` we associate the type `sol R` of maps `f : E → F` of 11 | solutions of `R`, and its formal counterpart `FormalSol R`. 12 | (FIXME(grunweg): `sol` is never mentioned; is this docstring outdated?) 13 | 14 | The h-principle question is whether we can deform any formal solution into a solution. 15 | The type of deformations is `HtpyJetSet E F` (homotopies of 1-jet sections). 16 | -/ 17 | 18 | 19 | noncomputable section 20 | 21 | open Set Function Real Filter 22 | 23 | open scoped unitInterval Topology 24 | 25 | variable (E : Type*) [NormedAddCommGroup E] [NormedSpace ℝ E] 26 | (F : Type*) [NormedAddCommGroup F] [NormedSpace ℝ F] 27 | (P : Type*) [NormedAddCommGroup P] [NormedSpace ℝ P] 28 | 29 | /-- A first order relation for maps between real vector spaces. -/ 30 | abbrev RelLoc := 31 | Set (OneJet E F) 32 | 33 | instance : Membership (E × F × (E →L[ℝ] F)) (RelLoc E F) := by delta RelLoc; infer_instance 34 | 35 | variable {E F} 36 | 37 | /-- A predicate stating that a 1-jet section is a formal solution to a first order relation for 38 | maps between vector spaces. -/ 39 | def JetSec.IsFormalSol (𝓕 : JetSec E F) (R : RelLoc E F) : Prop := 40 | ∀ x, (x, 𝓕.f x, 𝓕.φ x) ∈ R 41 | 42 | namespace RelLoc 43 | 44 | /-- A formal solution to a local relation `R`. -/ 45 | @[ext] 46 | structure FormalSol (R : RelLoc E F) extends JetSec E F where 47 | is_sol : ∀ x, (x, f x, φ x) ∈ R 48 | 49 | attribute [coe] FormalSol.toJetSec 50 | 51 | instance (R : RelLoc E F) : CoeOut (FormalSol R) (JetSec E F) := 52 | ⟨FormalSol.toJetSec⟩ 53 | 54 | -- Note: syntactic tautology 55 | @[simp] 56 | theorem FormalSol.toJetSec_eq_coe {R : RelLoc E F} (𝓕 : FormalSol R) : 57 | 𝓕.toJetSec = (𝓕 : JetSec E F) := 58 | rfl 59 | 60 | @[simp] 61 | theorem FormalSol.coeIsFormalSol {R : RelLoc E F} (𝓕 : FormalSol R) : 62 | (𝓕 : JetSec E F).IsFormalSol R := 63 | 𝓕.is_sol 64 | 65 | /-- Bundling a formal solution from a 1-jet section that is a formal solution. -/ 66 | def _root_.JetSec.IsFormalSol.formalSol {𝓕 : JetSec E F} {R : RelLoc E F} (h : 𝓕.IsFormalSol R) : 67 | FormalSol R := 68 | { 𝓕 with is_sol := h } 69 | 70 | instance (R : RelLoc E F) : FunLike (FormalSol R) E (F × (E →L[ℝ] F)) := 71 | ⟨fun 𝓕 x ↦ (𝓕.f x, 𝓕.φ x), 72 | by 73 | intros 𝓕 𝓕' h 74 | ext x : 2 <;> replace h := Prod.mk_inj.mp <| congrFun h x 75 | exacts [h.1, h.2]⟩ 76 | 77 | @[simp] 78 | theorem FormalSol.coe_apply {R : RelLoc E F} (𝓕 : FormalSol R) (x : E) : (𝓕 : JetSec E F) x = 𝓕 x := 79 | rfl 80 | 81 | variable {R : RelLoc E F} 82 | 83 | theorem FormalSol.eq_iff {𝓕 𝓕' : FormalSol R} {x : E} : 84 | 𝓕 x = 𝓕' x ↔ 𝓕.f x = 𝓕'.f x ∧ 𝓕.φ x = 𝓕'.φ x := 85 | JetSec.eq_iff 86 | 87 | /-- A formal solution (f, φ) is holonomic at `x` if the differential of `f` at `x` is `φ x`. -/ 88 | def FormalSol.IsHolonomicAt (𝓕 : FormalSol R) (x : E) : Prop := 89 | D 𝓕.f x = 𝓕.φ x 90 | 91 | -- TODO: this should come from a lemma about `JetSec` 92 | theorem FormalSol.isHolonomicAt_congr (𝓕 𝓕' : FormalSol R) {s : Set E} 93 | (h : ∀ᶠ x near s, 𝓕 x = 𝓕' x) : ∀ᶠ x near s, 𝓕.IsHolonomicAt x ↔ 𝓕'.IsHolonomicAt x := by 94 | apply h.eventually_nhdsSet.mono 95 | intro x hx 96 | have hf : 𝓕.f =ᶠ[𝓝 x] 𝓕'.f := by 97 | apply hx.mono 98 | simp_rw [RelLoc.FormalSol.eq_iff] 99 | tauto 100 | unfold RelLoc.FormalSol.IsHolonomicAt 101 | rw [hf.fderiv_eq, (RelLoc.FormalSol.eq_iff.mp hx.self_of_nhds).2] 102 | 103 | /-- A family of formal solutions is a 1-parameter family of formal solutions. -/ 104 | @[ext] 105 | structure FamilyFormalSol (R : RelLoc E F) extends FamilyJetSec E F P where 106 | is_sol : ∀ t x, (x, f t x, φ t x) ∈ R 107 | 108 | /-- A homotopy of formal solutions is a 1-parameter family of formal solutions. -/ 109 | @[reducible] 110 | def HtpyFormalSol (R : RelLoc E F) := 111 | R.FamilyFormalSol ℝ 112 | 113 | def HtpyFormalSol.toHtpyJetSec {R : RelLoc E F} (𝓕 : R.HtpyFormalSol) : HtpyJetSec E F := 114 | 𝓕.toFamilyJetSec 115 | 116 | open RelLoc 117 | 118 | instance (R : RelLoc E F) : FunLike (FamilyFormalSol P R) P (JetSec E F) := 119 | ⟨fun S ↦ S.toFamilyJetSec, by 120 | intros S S' h 121 | ext p x : 3 <;> replace h := congrFun h p 122 | exacts [congrFun (JetSec.ext_iff.1 h).1 x, congrFun (JetSec.ext_iff.1 h).2 x]⟩ 123 | 124 | end RelLoc 125 | -------------------------------------------------------------------------------- /SphereEversion/Main.lean: -------------------------------------------------------------------------------- 1 | import SphereEversion.Global.Immersion 2 | 3 | open Metric FiniteDimensional Set ModelWithCorners 4 | 5 | open scoped Manifold Topology ContDiff 6 | 7 | /-! # The sphere eversion project 8 | 9 | The goal of this project was to formalize Gromov's flexibility theorem for open and ample 10 | partial differential relation. A famous corollary of this theorem is Smale's sphere eversion 11 | theorem: you can turn a sphere inside-out. Let us state this corollary first. 12 | -/ 13 | 14 | section Smale 15 | 16 | -- As usual, we denote by `ℝ³` the Euclidean 3-space and by `𝕊²` its unit sphere. 17 | notation "ℝ³" => EuclideanSpace ℝ (Fin 3) 18 | 19 | notation "𝕊²" => sphere (0 : ℝ³) 1 20 | 21 | -- In the following statements, notation involving `𝓘` and `𝓡` denote smooth structures in the 22 | -- sense of differential geometry. 23 | theorem Smale : 24 | -- There exists a family of functions `f t` indexed by `ℝ` going from `𝕊²` to `ℝ³` such that 25 | ∃ f : ℝ → 𝕊² → ℝ³, 26 | -- it is smooth in both variables (for the obvious smooth structures on `ℝ × 𝕊²` and `ℝ³`) and 27 | ContMDiff (𝓘(ℝ, ℝ).prod (𝓡 2)) 𝓘(ℝ, ℝ³) ∞ ↿f ∧ 28 | (-- `f 0` is the inclusion map, sending `x` to `x` and 29 | f 0 = fun x : 𝕊² ↦ (x : ℝ³)) ∧ 30 | (-- `f 1` is the antipodal map, sending `x` to `-x` and 31 | f 1 = fun x : 𝕊² ↦ -(x : ℝ³)) ∧ 32 | -- every `f t` is an immersion. 33 | ∀ t, Immersion (𝓡 2) 𝓘(ℝ, ℝ³) (f t) ∞ := 34 | sphere_eversion ℝ³ 35 | 36 | end Smale 37 | 38 | /- 39 | The above result is an easy corollary of a much more general theorem by Gromov. 40 | The next three paragraphs lines simply introduce three real dimensional manifolds 41 | `M`, `M'` and `P` and assume they are separated and σ-compact. They are rather verbose because 42 | mathlib has a very general theory of manifolds (including manifolds with boundary and corners). 43 | We will consider families of maps from `M` to `M'` parametrized by `P`. 44 | Actually `M'` is assumed to be equipped with a preferred metric space structure in order to make it 45 | easier to state the `𝓒⁰` approximation property. 46 | -/ 47 | section Gromov 48 | 49 | variable (n n' d : ℕ) 50 | {M : Type*} [TopologicalSpace M] [ChartedSpace (ℝ^n) M] 51 | [IsManifold (𝓡 n) ∞ M] [T2Space M] [SigmaCompactSpace M] 52 | {M' : Type*} [MetricSpace M'] [ChartedSpace (ℝ^n') M'] [IsManifold (𝓡 n') ∞ M'] 53 | [SigmaCompactSpace M'] 54 | {P : Type*} [TopologicalSpace P] [ChartedSpace (ℝ^d) P] 55 | [IsManifold (𝓡 d) ∞ P] [T2Space P] [SigmaCompactSpace P] 56 | 57 | /-- Gromov's flexibility theorem for open and ample first order partial differential relations 58 | for maps between manifolds. -/ 59 | theorem Gromov 60 | -- Let `R` be an open and ample first order partial differential relation for maps 61 | -- from `M` to `M'`. 62 | {R : RelMfld 𝓘(ℝ, ℝ^n) M 𝓘(ℝ, ℝ^n') M'} 63 | (hRample : R.Ample) (hRopen : IsOpen R) 64 | -- Let `C` be a closed subset in `P × M` 65 | {C : Set (P × M)} 66 | (hC : IsClosed C) 67 | -- Let `ε` be a continuous positive function on `M` 68 | {ε : M → ℝ} 69 | (ε_pos : ∀ x, 0 < ε x) (ε_cont : Continuous ε) 70 | -- Let `𝓕₀` be a family of formal solutions to `R` parametrized by `P` 71 | (𝓕₀ : FamilyFormalSol 𝓘(ℝ, ℝ^d) P R) 72 | -- Assume it is holonomic near `C`. 73 | (hhol : 74 | ∀ᶠ p : P × M in 𝓝ˢ C, 75 | (𝓕₀ p.1).IsHolonomicAt p.2) :-- then there is a homotopy of such families 76 | ∃ 𝓕 : FamilyFormalSol (𝓘(ℝ, ℝ).prod 𝓘(ℝ, ℝ^d)) (ℝ × P) R, 77 | (-- that agrees with `𝓕₀` at time `t = 0` for every parameter and every point in the source 78 | ∀ (p : P) (x : M), 𝓕 (0, p) x = 𝓕₀ p x) ∧ 79 | (-- is holonomic everywhere for `t = 1`, 80 | ∀ p : P, (𝓕 (1, p)).toOneJetSec.IsHolonomic) ∧ 81 | (-- agrees with `𝓕₀` near `C`, 82 | ∀ᶠ p : P × M in 𝓝ˢ C, 83 | ∀ t : ℝ, 𝓕 (t, p.1) p.2 = 𝓕₀ p.1 p.2) ∧ 84 | -- and whose underlying maps are `ε`--close to `𝓕₀`. 85 | ∀ (t : ℝ) (p : P) (x : M), dist ((𝓕 (t, p)).bs x) ((𝓕₀ p).bs x) ≤ ε x := by 86 | apply RelMfld.Ample.satisfiesHPrincipleWith <;> assumption 87 | 88 | end Gromov 89 | -------------------------------------------------------------------------------- /SphereEversion/Notations.lean: -------------------------------------------------------------------------------- 1 | import Mathlib.Analysis.Calculus.ContDiff.Basic 2 | 3 | open scoped Topology ContDiff 4 | 5 | notation "𝒞" => ContDiff ℝ 6 | 7 | notation "hull" => convexHull ℝ 8 | 9 | notation "D" => fderiv ℝ 10 | 11 | notation "smooth_on" => ContDiffOn ℝ ∞ 12 | 13 | -- `∀ᶠ x near s, p x` means property `p` holds at every point in a neighborhood of the set `s`. 14 | notation3 (prettyPrint := false) 15 | "∀ᶠ " (...)" near "s", "r:(scoped p => Filter.Eventually p <| 𝓝ˢ s) => r 16 | 17 | notation:70 u " ⬝ " φ:65 => ContinuousLinearMap.comp (ContinuousLinearMap.toSpanSingleton ℝ u) φ 18 | -------------------------------------------------------------------------------- /SphereEversion/ToMathlib/Algebra/Ring/Periodic.lean: -------------------------------------------------------------------------------- 1 | import Mathlib.Analysis.Normed.Order.Lattice 2 | import Mathlib.Algebra.Ring.Periodic 3 | import SphereEversion.ToMathlib.Topology.Separation.Hausdorff 4 | 5 | -- TODO: the file this references doesn't exist in mathlib any more; rename this one appropriately! 6 | 7 | /-! 8 | 9 | # Boundedness property of periodic function 10 | 11 | The main purpose of that file it to prove 12 | ``` 13 | lemma Continuous.bounded_of_onePeriodic_of_isCompact {f : X → ℝ → E} (cont : Continuous ↿f) 14 | (hper : ∀ x, OnePeriodic (f x)) {K : Set X} (hK : IsCompact K) (hfK : ∀ x ∉ K, f x = 0) : 15 | ∃ C, ∀ x t, ‖f x t‖ ≤ C 16 | ``` 17 | 18 | This is done by introducing the quotient 𝕊₁ = ℝ/ℤ as a compact topological space. 19 | Patrick is not sure this is the optimal version. 20 | 21 | In the first part, generalize many lemmas to any period and add to `Algebra.Ring.Periodic.lean`? 22 | -/ 23 | 24 | 25 | noncomputable section 26 | 27 | open Set Function Filter TopologicalSpace Int 28 | 29 | open scoped Topology 30 | 31 | section OnePeriodic 32 | 33 | variable {α : Type*} 34 | 35 | /-- The integers as an additive subgroup of the reals. -/ 36 | def ℤSubℝ : AddSubgroup ℝ := AddMonoidHom.range (Int.castAddHom ℝ) 37 | 38 | /-- The equivalence relation on `ℝ` corresponding to its partition as cosets of `ℤ`. -/ 39 | def transOne : Setoid ℝ := QuotientAddGroup.leftRel ℤSubℝ 40 | 41 | /-- The proposition that a function on `ℝ` is periodic with period `1`. -/ 42 | def OnePeriodic (f : ℝ → α) : Prop := Periodic f 1 43 | 44 | theorem OnePeriodic.add_nat {f : ℝ → α} (h : OnePeriodic f) (k : ℕ) (x : ℝ) : f (x + k) = f x := by 45 | simpa using h.nat_mul k x 46 | 47 | theorem OnePeriodic.add_int {f : ℝ → α} (h : OnePeriodic f) (k : ℤ) (x : ℝ) : f (x + k) = f x := by 48 | simpa using h.int_mul k x 49 | 50 | /-- The circle `𝕊₁ := ℝ/ℤ`. 51 | 52 | TODO [Yury]: use `AddCircle`. -/ 53 | def 𝕊₁ := 54 | Quotient transOne 55 | deriving TopologicalSpace, Inhabited 56 | 57 | theorem transOne_rel_iff {a b : ℝ} : transOne.r a b ↔ ∃ k : ℤ, b = a + k := by 58 | refine QuotientAddGroup.leftRel_apply.trans (exists_congr fun k ↦ ?_) 59 | rw [coe_castAddHom, eq_neg_add_iff_add_eq, eq_comm] 60 | 61 | section 62 | 63 | attribute [local instance] transOne 64 | 65 | /-- The quotient map from the reals to the circle `ℝ ⧸ ℤ`. -/ 66 | def proj𝕊₁ : ℝ → 𝕊₁ := 67 | Quotient.mk' 68 | 69 | @[simp] 70 | theorem proj𝕊₁_add_int (t : ℝ) (k : ℤ) : proj𝕊₁ (t + k) = proj𝕊₁ t := 71 | (Quotient.sound (transOne_rel_iff.mpr ⟨k, rfl⟩)).symm 72 | 73 | /-- The unique representative in the half-open interval `[0, 1)` for each coset of `ℤ` in `ℝ`, 74 | regarded as a map from the circle `𝕊₁ → ℝ`. -/ 75 | def 𝕊₁.repr (x : 𝕊₁) : ℝ := 76 | let t := Quotient.out x 77 | fract t 78 | 79 | theorem 𝕊₁.repr_mem (x : 𝕊₁) : x.repr ∈ (Ico 0 1 : Set ℝ) := 80 | ⟨fract_nonneg _, fract_lt_one _⟩ 81 | 82 | theorem 𝕊₁.proj_repr (x : 𝕊₁) : proj𝕊₁ x.repr = x := by 83 | symm 84 | conv_lhs => rw [← Quotient.out_eq x] 85 | rw [← fract_add_floor (Quotient.out x)] 86 | apply proj𝕊₁_add_int 87 | 88 | theorem image_proj𝕊₁_Ico : proj𝕊₁ '' Ico 0 1 = univ := by 89 | rw [eq_univ_iff_forall] 90 | exact fun x ↦ ⟨x.repr, x.repr_mem, x.proj_repr⟩ 91 | 92 | theorem image_proj𝕊₁_Icc : proj𝕊₁ '' Icc 0 1 = univ := 93 | eq_univ_of_subset (image_subset proj𝕊₁ Ico_subset_Icc_self) image_proj𝕊₁_Ico 94 | 95 | @[continuity, fun_prop] 96 | theorem continuous_proj𝕊₁ : Continuous proj𝕊₁ := continuous_quotient_mk' 97 | 98 | theorem isOpenMap_proj𝕊₁ : IsOpenMap proj𝕊₁ := QuotientAddGroup.isOpenMap_coe 99 | 100 | theorem quotientMap_id_proj𝕊₁ {X : Type*} [TopologicalSpace X] : 101 | Topology.IsQuotientMap fun p : X × ℝ ↦ (p.1, proj𝕊₁ p.2) := 102 | (IsOpenMap.id.prodMap isOpenMap_proj𝕊₁).isQuotientMap (by fun_prop) 103 | (surjective_id.prodMap Quotient.exists_rep) 104 | 105 | /-- A one-periodic function on `ℝ` descends to a function on the circle `ℝ ⧸ ℤ`. -/ 106 | def OnePeriodic.lift {f : ℝ → α} (h : OnePeriodic f) : 𝕊₁ → α := 107 | Quotient.lift f (by intro a b hab; rcases transOne_rel_iff.mp hab with ⟨k, rfl⟩; rw [h.add_int]) 108 | 109 | end 110 | 111 | @[inherit_doc] local notation "π" => proj𝕊₁ 112 | 113 | instance : CompactSpace 𝕊₁ := 114 | ⟨by rw [← image_proj𝕊₁_Icc]; exact isCompact_Icc.image continuous_proj𝕊₁⟩ 115 | 116 | theorem isClosed_int : IsClosed (range ((↑) : ℤ → ℝ)) := 117 | Int.isClosedEmbedding_coe_real.isClosed_range 118 | 119 | instance : T2Space 𝕊₁ := by 120 | have πcont : Continuous π := continuous_quotient_mk' 121 | rw [t2Space_iff_of_continuous_surjective_open πcont Quotient.mk''_surjective isOpenMap_proj𝕊₁] 122 | have : {q : ℝ × ℝ | π q.fst = π q.snd} = {q : ℝ × ℝ | ∃ k : ℤ, q.2 = q.1 + k} := by 123 | ext ⟨a, b⟩ 124 | exact Quotient.eq''.trans transOne_rel_iff 125 | have : {q : ℝ × ℝ | π q.fst = π q.snd} = 126 | (fun q : ℝ × ℝ ↦ q.2 - q.1) ⁻¹' (range <| ((↑) : ℤ → ℝ)) := by 127 | rw [this] 128 | ext ⟨a, b⟩ 129 | refine exists_congr fun k ↦ ?_ 130 | conv_rhs => rw [eq_comm, sub_eq_iff_eq_add'] 131 | rw [this] 132 | exact IsClosed.preimage (continuous_snd.sub continuous_fst) isClosed_int 133 | 134 | variable {X E : Type*} [TopologicalSpace X] [NormedAddCommGroup E] 135 | 136 | theorem Continuous.bounded_on_compact_of_onePeriodic {f : X → ℝ → E} (cont : Continuous ↿f) 137 | (hper : ∀ x, OnePeriodic (f x)) {K : Set X} (hK : IsCompact K) : 138 | ∃ C, ∀ x ∈ K, ∀ t, ‖f x t‖ ≤ C := by 139 | let F : X × 𝕊₁ → E := fun p : X × 𝕊₁ ↦ (hper p.1).lift p.2 140 | have Fcont : Continuous F := by 141 | have qm : Topology.IsQuotientMap fun p : X × ℝ ↦ (p.1, π p.2) := quotientMap_id_proj𝕊₁ 142 | -- avoid weird elaboration issue 143 | have : ↿f = F ∘ fun p : X × ℝ ↦ (p.1, π p.2) := by ext p; rfl 144 | rwa [this, ← qm.continuous_iff] at cont 145 | obtain ⟨C, hC⟩ := 146 | (hK.prod isCompact_univ).bddAbove_image (continuous_norm.comp Fcont).continuousOn 147 | exact ⟨C, fun x x_in t ↦ hC ⟨(x, π t), ⟨x_in, mem_univ _⟩, rfl⟩⟩ 148 | 149 | theorem Continuous.bounded_of_onePeriodic_of_compact {f : X → ℝ → E} (cont : Continuous ↿f) 150 | (hper : ∀ x, OnePeriodic (f x)) {K : Set X} (hK : IsCompact K) 151 | (hfK : ∀ x ∉ K, f x = 0) : ∃ C, ∀ x t, ‖f x t‖ ≤ C := by 152 | obtain ⟨C, hC⟩ := cont.bounded_on_compact_of_onePeriodic hper hK 153 | refine ⟨max C 0, fun x t ↦ ?_⟩ 154 | by_cases hx : x ∈ K 155 | · exact le_max_of_le_left (hC x hx t) 156 | · simp [hfK, hx] 157 | 158 | end OnePeriodic 159 | -------------------------------------------------------------------------------- /SphereEversion/ToMathlib/Analysis/Calculus.lean: -------------------------------------------------------------------------------- 1 | import Mathlib.Analysis.Normed.Module.Completion 2 | import Mathlib.Analysis.SpecialFunctions.SmoothTransition 3 | import SphereEversion.ToMathlib.Topology.Misc 4 | 5 | noncomputable section 6 | 7 | open Set Function Filter 8 | 9 | open scoped Topology ContDiff 10 | 11 | namespace Real 12 | 13 | theorem smoothTransition_projI {x : ℝ} : smoothTransition (projI x) = smoothTransition x := 14 | smoothTransition.projIcc 15 | 16 | end Real 17 | 18 | section Calculus 19 | 20 | open ContinuousLinearMap 21 | 22 | variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] {E : Type*} [NormedAddCommGroup E] 23 | [NormedSpace 𝕜 E] {E₁ : Type*} [NormedAddCommGroup E₁] [NormedSpace 𝕜 E₁] {E₂ : Type*} 24 | [NormedAddCommGroup E₂] [NormedSpace 𝕜 E₂] {E' : Type*} [NormedAddCommGroup E'] 25 | [NormedSpace 𝕜 E'] {F : Type*} [NormedAddCommGroup F] [NormedSpace 𝕜 F] {G : Type*} 26 | [NormedAddCommGroup G] [NormedSpace 𝕜 G] {n : WithTop ℕ∞} 27 | 28 | theorem fderiv_prod_left {x₀ : E} {y₀ : F} : 29 | fderiv 𝕜 (fun x ↦ (x, y₀)) x₀ = ContinuousLinearMap.inl 𝕜 E F := 30 | ((hasFDerivAt_id _).prodMk (hasFDerivAt_const _ _)).fderiv 31 | 32 | theorem fderiv_prod_right {x₀ : E} {y₀ : F} : 33 | fderiv 𝕜 (fun y ↦ (x₀, y)) y₀ = ContinuousLinearMap.inr 𝕜 E F := 34 | ((hasFDerivAt_const _ _).prodMk (hasFDerivAt_id _)).fderiv 35 | 36 | theorem HasFDerivAt.partial_fst {φ : E → F → G} {φ' : E × F →L[𝕜] G} {e₀ : E} {f₀ : F} 37 | (h : HasFDerivAt (uncurry φ) φ' (e₀, f₀)) : 38 | HasFDerivAt (fun e ↦ φ e f₀) (φ'.comp (inl 𝕜 E F)) e₀ := by 39 | exact h.comp e₀ <| hasFDerivAt_prodMk_left (𝕜 := 𝕜) e₀ f₀ 40 | 41 | theorem HasFDerivAt.partial_snd {φ : E → F → G} {φ' : E × F →L[𝕜] G} {e₀ : E} {f₀ : F} 42 | (h : HasFDerivAt (uncurry φ) φ' (e₀, f₀)) : 43 | HasFDerivAt (fun f ↦ φ e₀ f) (φ'.comp (inr 𝕜 E F)) f₀ := 44 | h.comp f₀ <| hasFDerivAt_prodMk_right e₀ f₀ 45 | 46 | theorem fderiv_prod_eq_add {f : E × F → G} {p : E × F} (hf : DifferentiableAt 𝕜 f p) : 47 | fderiv 𝕜 f p = 48 | fderiv 𝕜 (fun z : E × F ↦ f (z.1, p.2)) p + fderiv 𝕜 (fun z : E × F ↦ f (p.1, z.2)) p := by 49 | have H₁ : fderiv 𝕜 (fun z : E × F ↦ f (z.1, p.2)) p = 50 | (fderiv 𝕜 f p).comp (.comp (.inl 𝕜 E F) (.fst 𝕜 E F)) := 51 | (hf.hasFDerivAt.comp p 52 | ((hasFDerivAt_fst (𝕜 := 𝕜) (E := E) (F := F)).prodMk (hasFDerivAt_const p.2 _))).fderiv 53 | have H₂ : fderiv 𝕜 (fun z : E × F ↦ f (p.1, z.2)) p = 54 | (fderiv 𝕜 f p).comp (.comp (.inr 𝕜 E F) (.snd 𝕜 E F)) := 55 | (hf.hasFDerivAt.comp _ ((hasFDerivAt_const p.1 _).prodMk 56 | (hasFDerivAt_snd (𝕜 := 𝕜) (E := E) (F := F)))).fderiv 57 | rw [H₁, H₂, ← comp_add, comp_fst_add_comp_snd, coprod_inl_inr, ContinuousLinearMap.comp_id] 58 | 59 | variable (𝕜) 60 | 61 | /-- The first partial derivative of a binary function. -/ 62 | def partialFDerivFst {F : Type*} (φ : E → F → G) : E → F → E →L[𝕜] G := fun (e₀ : E) (f₀ : F) ↦ 63 | fderiv 𝕜 (fun e ↦ φ e f₀) e₀ 64 | 65 | /-- The second partial derivative of a binary function. -/ 66 | def partialFDerivSnd {E : Type*} (φ : E → F → G) : E → F → F →L[𝕜] G := fun (e₀ : E) (f₀ : F) ↦ 67 | fderiv 𝕜 (fun f ↦ φ e₀ f) f₀ 68 | 69 | @[inherit_doc] local notation "∂₁" => partialFDerivFst 70 | 71 | @[inherit_doc] local notation "∂₂" => partialFDerivSnd 72 | 73 | variable {𝕜} 74 | 75 | theorem fderiv_partial_fst {φ : E → F → G} {φ' : E × F →L[𝕜] G} {e₀ : E} {f₀ : F} 76 | (h : HasFDerivAt (uncurry φ) φ' (e₀, f₀)) : ∂₁ 𝕜 φ e₀ f₀ = φ'.comp (inl 𝕜 E F) := 77 | h.partial_fst.fderiv 78 | 79 | theorem fderiv_partial_snd {φ : E → F → G} {φ' : E × F →L[𝕜] G} {e₀ : E} {f₀ : F} 80 | (h : HasFDerivAt (uncurry φ) φ' (e₀, f₀)) : ∂₂ 𝕜 φ e₀ f₀ = φ'.comp (inr 𝕜 E F) := 81 | h.partial_snd.fderiv 82 | 83 | theorem DifferentiableAt.hasFDerivAt_partial_fst {φ : E → F → G} {e₀ : E} {f₀ : F} 84 | (h : DifferentiableAt 𝕜 (uncurry φ) (e₀, f₀)) : 85 | HasFDerivAt (fun e ↦ φ e f₀) (partialFDerivFst 𝕜 φ e₀ f₀) e₀ := by 86 | apply (h.comp e₀ <| differentiableAt_id.prodMk <| differentiableAt_const f₀).hasFDerivAt (𝕜 := 𝕜) 87 | 88 | theorem DifferentiableAt.hasFDerivAt_partial_snd {φ : E → F → G} {e₀ : E} {f₀ : F} 89 | (h : DifferentiableAt 𝕜 (uncurry φ) (e₀, f₀)) : 90 | HasFDerivAt (fun f ↦ φ e₀ f) (partialFDerivSnd 𝕜 φ e₀ f₀) f₀ := by 91 | rw [fderiv_partial_snd h.hasFDerivAt] 92 | exact h.hasFDerivAt.partial_snd 93 | 94 | theorem ContDiff.partial_fst {φ : E → F → G} {n : ℕ∞} (h : ContDiff 𝕜 n <| uncurry φ) (f₀ : F) : 95 | ContDiff 𝕜 n fun e ↦ φ e f₀ := 96 | h.comp <| contDiff_prodMk_left f₀ 97 | 98 | theorem ContDiff.partial_snd {φ : E → F → G} {n : ℕ∞} (h : ContDiff 𝕜 n <| uncurry φ) (e₀ : E) : 99 | ContDiff 𝕜 n fun f ↦ φ e₀ f := 100 | h.comp <| contDiff_prodMk_right e₀ 101 | 102 | /-- Precomposition by a continuous linear map as a continuous linear map between spaces of 103 | continuous linear maps. -/ 104 | def ContinuousLinearMap.compRightL (φ : E →L[𝕜] F) : (F →L[𝕜] G) →L[𝕜] E →L[𝕜] G := 105 | precomp G φ 106 | 107 | /-- Postcomposition by a continuous linear map as a continuous linear map between spaces of 108 | continuous linear maps. -/ 109 | def ContinuousLinearMap.compLeftL (φ : F →L[𝕜] G) : (E →L[𝕜] F) →L[𝕜] E →L[𝕜] G := 110 | postcomp E φ 111 | 112 | nonrec theorem Differentiable.fderiv_partial_fst {φ : E → F → G} 113 | (hF : Differentiable 𝕜 (uncurry φ)) : 114 | ↿(∂₁ 𝕜 φ) = precomp G (inl 𝕜 E F) ∘ (fderiv 𝕜 <| uncurry φ) := by 115 | ext1 ⟨y, t⟩; exact fderiv_partial_fst (hF ⟨y, t⟩).hasFDerivAt 116 | 117 | nonrec theorem Differentiable.fderiv_partial_snd {φ : E → F → G} 118 | (hF : Differentiable 𝕜 (uncurry φ)) : 119 | ↿(∂₂ 𝕜 φ) = precomp G (inr 𝕜 E F) ∘ (fderiv 𝕜 <| uncurry φ) := by 120 | ext1 ⟨y, t⟩; exact fderiv_partial_snd (hF ⟨y, t⟩).hasFDerivAt 121 | 122 | /-- The first partial derivative of `φ : 𝕜 → F → G` seen as a function from `𝕜 → F → G` -/ 123 | def partialDerivFst (φ : 𝕜 → F → G) : 𝕜 → F → G := fun k f ↦ ∂₁ 𝕜 φ k f 1 124 | 125 | /-- The second partial derivative of `φ : E → 𝕜 → G` seen as a function from `E → 𝕜 → G` -/ 126 | def partialDerivSnd (φ : E → 𝕜 → G) : E → 𝕜 → G := fun e k ↦ ∂₂ 𝕜 φ e k 1 127 | 128 | omit [NormedAddCommGroup F] [NormedSpace 𝕜 F] in 129 | theorem partialFDerivFst_eq_smulRight (φ : 𝕜 → F → G) (k : 𝕜) (f : F) : 130 | ∂₁ 𝕜 φ k f = smulRight (1 : 𝕜 →L[𝕜] 𝕜) (partialDerivFst φ k f) := 131 | deriv_fderiv.symm 132 | 133 | omit [NormedAddCommGroup F] [NormedSpace 𝕜 F] in 134 | @[simp] 135 | theorem partialFDerivFst_one (φ : 𝕜 → F → G) (k : 𝕜) (f : F) : 136 | ∂₁ 𝕜 φ k f 1 = partialDerivFst φ k f := by 137 | simp [partialFDerivFst_eq_smulRight] 138 | 139 | omit [NormedAddCommGroup E] [NormedSpace 𝕜 E] in 140 | theorem partialFDerivSnd_eq_smulRight (φ : E → 𝕜 → G) (e : E) (k : 𝕜) : 141 | ∂₂ 𝕜 φ e k = smulRight (1 : 𝕜 →L[𝕜] 𝕜) (partialDerivSnd φ e k) := 142 | deriv_fderiv.symm 143 | 144 | omit [NormedAddCommGroup E] [NormedSpace 𝕜 E] in 145 | theorem partialFDerivSnd_one (φ : E → 𝕜 → G) (e : E) (k : 𝕜) : 146 | ∂₂ 𝕜 φ e k 1 = partialDerivSnd φ e k := by 147 | simp [partialFDerivSnd_eq_smulRight] 148 | 149 | @[to_additive] 150 | nonrec theorem WithTop.le_mul_self {α : Type*} [CommMonoid α] [PartialOrder α] [IsOrderedMonoid α] 151 | [CanonicallyOrderedMul α] (n m : α) : 152 | (n : WithTop α) ≤ (m * n : α) := 153 | WithTop.coe_le_coe.mpr le_mul_self 154 | 155 | @[to_additive] 156 | nonrec theorem WithTop.le_self_mul {α : Type*} [CommMonoid α] [PartialOrder α] [IsOrderedMonoid α] 157 | [CanonicallyOrderedMul α] (n m : α) : 158 | (n : WithTop α) ≤ (n * m : α) := 159 | WithTop.coe_le_coe.mpr le_self_mul 160 | 161 | theorem ContDiff.contDiff_partial_fst {φ : E → F → G} {n : ℕ} 162 | (hF : ContDiff 𝕜 (n + 1) (uncurry φ)) : ContDiff 𝕜 n ↿(∂₁ 𝕜 φ) := 163 | ContDiff.fderiv (hF.comp <| contDiff_snd.prodMk contDiff_fst.snd) contDiff_fst le_rfl 164 | 165 | theorem ContDiff.contDiff_partial_fst_apply {φ : E → F → G} {n : ℕ} 166 | (hF : ContDiff 𝕜 (n + 1) (uncurry φ)) {x : E} : ContDiff 𝕜 n ↿fun x' y ↦ ∂₁ 𝕜 φ x' y x := 167 | (ContinuousLinearMap.apply 𝕜 G x).contDiff.comp hF.contDiff_partial_fst 168 | 169 | theorem ContDiff.continuous_partial_fst {φ : E → F → G} {n : ℕ} 170 | (h : ContDiff 𝕜 ((n + 1 : ℕ) : ℕ∞) <| uncurry φ) : Continuous ↿(∂₁ 𝕜 φ) := 171 | h.contDiff_partial_fst.continuous 172 | 173 | -- XXX: fix this! 174 | theorem ContDiff.contDiff_top_partial_fst {φ : E → F → G} (hF : ContDiff 𝕜 ∞ (uncurry φ)) : 175 | ContDiff 𝕜 ∞ ↿(∂₁ 𝕜 φ) := 176 | contDiff_infty.mpr fun n ↦ (contDiff_infty.mp hF (n + 1)).contDiff_partial_fst 177 | 178 | theorem ContDiff.contDiff_partial_snd {φ : E → F → G} {n : ℕ} 179 | (hF : ContDiff 𝕜 (n + 1) (uncurry φ)) : ContDiff 𝕜 n ↿(∂₂ 𝕜 φ) := 180 | ContDiff.fderiv (hF.comp <| contDiff_fst.fst.prodMk contDiff_snd) contDiff_snd le_rfl 181 | 182 | theorem ContDiff.contDiff_partial_snd_apply {φ : E → F → G} {n : ℕ} 183 | (hF : ContDiff 𝕜 (n + 1) (uncurry φ)) {y : F} : ContDiff 𝕜 n ↿fun x y' ↦ ∂₂ 𝕜 φ x y' y := 184 | (ContinuousLinearMap.apply 𝕜 G y).contDiff.comp hF.contDiff_partial_snd 185 | 186 | theorem ContDiff.continuous_partial_snd {φ : E → F → G} {n : ℕ} 187 | (h : ContDiff 𝕜 ((n + 1 : ℕ) : ℕ∞) <| uncurry φ) : Continuous ↿(∂₂ 𝕜 φ) := 188 | h.contDiff_partial_snd.continuous 189 | 190 | -- FIXME: upgrade again to include analyticity 191 | theorem ContDiff.contDiff_top_partial_snd {φ : E → F → G} (hF : ContDiff 𝕜 ∞ (uncurry φ)) : 192 | ContDiff 𝕜 ∞ ↿(∂₂ 𝕜 φ) := 193 | contDiff_infty.mpr fun n ↦ (contDiff_infty.mp (hF.of_le (by simp)) (n + 1)).contDiff_partial_snd 194 | 195 | end Calculus 196 | 197 | section RealCalculus -- PRed in #12673 198 | 199 | open ContinuousLinearMap 200 | 201 | variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] {F : Type*} [NormedAddCommGroup F] 202 | [NormedSpace ℝ F] 203 | 204 | theorem ContDiff.lipschitzOnWith {s : Set E} {f : E → F} {n} (hf : ContDiff ℝ n f) (hn : 1 ≤ n) 205 | (hs : Convex ℝ s) (hs' : IsCompact s) : ∃ K, LipschitzOnWith K f s := by 206 | rcases (bddAbove_iff_exists_ge (0 : ℝ)).mp (hs'.image (hf.continuous_fderiv hn).norm).bddAbove 207 | with ⟨M, M_nonneg, hM⟩ 208 | simp_rw [forall_mem_image] at hM 209 | use ⟨M, M_nonneg⟩ 210 | exact Convex.lipschitzOnWith_of_nnnorm_fderiv_le (fun x _ ↦ hf.differentiable hn x) hM hs 211 | 212 | end RealCalculus 213 | 214 | open Filter 215 | 216 | theorem const_mul_one_div_lt {ε : ℝ} (ε_pos : 0 < ε) (C : ℝ) : ∀ᶠ N : ℝ in atTop, C * ‖1 / N‖ < ε := 217 | have h : Tendsto (fun N : ℝ ↦ C * ‖1 / N‖) atTop (𝓝 (C * ‖(0 : ℝ)‖)) := 218 | tendsto_const_nhds.mul (tendsto_const_nhds.div_atTop tendsto_id).norm 219 | Filter.Tendsto.eventually_lt h tendsto_const_nhds <| by simpa 220 | -------------------------------------------------------------------------------- /SphereEversion/ToMathlib/Analysis/Calculus/AddTorsor/AffineMap.lean: -------------------------------------------------------------------------------- 1 | import Mathlib.Analysis.Calculus.AddTorsor.AffineMap 2 | 3 | /-! 4 | 5 | # Lemmas about interaction of `ball` and `homothety` 6 | 7 | TODO Generalise these lemmas appropriately. 8 | 9 | -/ 10 | 11 | 12 | open Set Function Metric AffineMap 13 | 14 | variable {F : Type*} [NormedAddCommGroup F] 15 | 16 | -- Unused 17 | -- @[simp] 18 | -- theorem range_affine_equiv_ball {p c : F} {s r : ℝ} (hr : 0 < r) : 19 | -- (range fun x : ball p s ↦ c +ᵥ homothety p r (x : F)) = ball (c + p) (r * s) := by 20 | -- ext 21 | -- simp only [homothety_apply, dist_eq_norm, vsub_eq_sub, vadd_eq_add, mem_range, SetCoe.exists, 22 | -- mem_ball, Subtype.coe_mk, exists_prop] 23 | -- refine ⟨?_, fun h ↦ ⟨p + r⁻¹ • (x - (c + p)), ?_, ?_⟩⟩ 24 | -- · rintro ⟨y, h, rfl⟩ 25 | -- simpa [norm_smul, abs_eq_self.mpr hr.le] using (mul_lt_mul_left hr).mpr h 26 | -- · simpa [norm_smul, abs_eq_self.mpr hr.le] using (inv_mul_lt_iff hr).mpr h 27 | -- · simp [← smul_assoc, hr.ne.symm.is_unit.mul_inv_cancel]; abel 28 | 29 | @[simp] 30 | theorem norm_coe_ball_lt (r : ℝ) (x : ball (0 : F) r) : ‖(x : F)‖ < r := by 31 | obtain ⟨x, hx⟩ := x 32 | simpa using hx 33 | 34 | variable [NormedSpace ℝ F] 35 | 36 | theorem mapsTo_homothety_ball (c : F) {r : ℝ} (hr : 0 < r) : 37 | MapsTo (fun y ↦ homothety c r⁻¹ y -ᵥ c) (ball c r) (ball 0 1) := fun y hy ↦ by 38 | replace hy : r⁻¹ * ‖y - c‖ < 1 := by 39 | rw [← mul_lt_mul_left hr, ← mul_assoc, mul_inv_cancel₀ hr.ne.symm, mul_one, one_mul] 40 | simpa [dist_eq_norm] using hy 41 | simp only [homothety_apply, vsub_eq_sub, vadd_eq_add, add_sub_cancel_right, mem_ball_zero_iff, 42 | norm_smul, Real.norm_eq_abs, abs_eq_self.2 (inv_pos.mpr hr).le, hy] 43 | 44 | theorem contDiff_homothety {n : ℕ∞} (c : F) (r : ℝ) : ContDiff ℝ n (homothety c r) := 45 | (⟨homothety c r, homothety_continuous c r⟩ : F →ᴬ[ℝ] F).contDiff 46 | -------------------------------------------------------------------------------- /SphereEversion/ToMathlib/Analysis/Convex/Basic.lean: -------------------------------------------------------------------------------- 1 | import Mathlib.Analysis.Convex.Combination 2 | import Mathlib.Algebra.Module.BigOperators 3 | import Mathlib.Algebra.Order.Hom.Ring 4 | 5 | open Function Set 6 | 7 | -- TODO: move this lemma and the following one 8 | theorem map_finsum {β α γ : Type*} [AddCommMonoid β] [AddCommMonoid γ] 9 | {G : Type*} [FunLike G β γ] [AddMonoidHomClass G β γ] (g : G) 10 | {f : α → β} (hf : (Function.support f).Finite) : 11 | g (∑ᶠ i, f i) = ∑ᶠ i, g (f i) := 12 | (g : β →+ γ).map_finsum hf 13 | 14 | @[to_additive] 15 | theorem finprod_eq_prod_of_mulSupport_subset_of_finite {α M} [CommMonoid M] (f : α → M) {s : Set α} 16 | (h : mulSupport f ⊆ s) (hs : s.Finite) : ∏ᶠ i, f i = ∏ i ∈ hs.toFinset, f i := by 17 | apply finprod_eq_prod_of_mulSupport_subset f; rwa [Set.Finite.coe_toFinset] 18 | 19 | section 20 | 21 | variable {𝕜 𝕜' E E₂ E' : Type*} [Semiring 𝕜] [PartialOrder 𝕜] [IsOrderedRing 𝕜] 22 | [AddCommMonoid E] [Module 𝕜 E] [AddCommMonoid E₂] [Module 𝕜 E₂] [AddCommMonoid E'] 23 | [Semiring 𝕜'] [PartialOrder 𝕜'] [IsOrderedRing 𝕜'] [Module 𝕜' E'] (σ : 𝕜 →+*o 𝕜') 24 | 25 | def reallyConvexHull (𝕜 : Type*) {E : Type*} 26 | [Semiring 𝕜] [PartialOrder 𝕜] [IsOrderedRing 𝕜] [AddCommMonoid E] [SMul 𝕜 E] 27 | (s : Set E) : Set E := 28 | {e | ∃ w : E → 𝕜, 0 ≤ w ∧ support w ⊆ s ∧ ∑ᶠ x, w x = 1 ∧ e = ∑ᶠ x, w x • x} 29 | 30 | -- https://xkcd.com/927/ 31 | theorem finsum.exists_ne_zero_of_sum_ne_zero {β α : Type*} {s : Finset α} {f : α → β} 32 | [AddCommMonoid β] : ∑ᶠ x ∈ s, f x ≠ 0 → ∃ a ∈ s, f a ≠ 0 := by 33 | rw [finsum_mem_finset_eq_sum] 34 | exact Finset.exists_ne_zero_of_sum_ne_zero 35 | 36 | -- rename: `mul_support_finite_of_finprod_ne_one`? 37 | @[to_additive] 38 | theorem finite_of_finprod_ne_one {M : Type*} {ι : Sort _} [CommMonoid M] {f : ι → M} 39 | (h : ∏ᶠ i, f i ≠ 1) : (mulSupport f).Finite := by 40 | classical 41 | rw [finprod_def] at h 42 | contrapose h 43 | rw [Classical.not_not, dif_neg h] 44 | 45 | theorem support_finite_of_finsum_eq_of_neZero {M : Type*} {ι : Sort _} [AddCommMonoid M] 46 | {f : ι → M} {x : M} [NeZero x] (h : ∑ᶠ i, f i = x) : (support f).Finite := by 47 | apply finite_of_finsum_ne_zero 48 | rw [h] 49 | apply NeZero.ne 50 | 51 | @[to_additive] 52 | theorem Subsingleton.mulSupport_eq {α β} [Subsingleton β] [One β] (f : α → β) : 53 | mulSupport f = ∅ := by 54 | rw [mulSupport_eq_empty_iff]; apply Subsingleton.elim 55 | 56 | theorem support_finite_of_finsum_eq_one {M : Type*} {ι : Sort _} [NonAssocSemiring M] {f : ι → M} 57 | (h : ∑ᶠ i, f i = 1) : (support f).Finite := by 58 | cases subsingleton_or_nontrivial M 59 | · simp_rw [Subsingleton.support_eq, finite_empty] 60 | exact support_finite_of_finsum_eq_of_neZero h 61 | 62 | theorem finsum_sum_filter {α β M : Type*} [AddCommMonoid M] (f : β → α) (s : Finset β) 63 | [DecidableEq α] (g : β → M) : 64 | ∑ᶠ x : α, ∑ y ∈ Finset.filter (fun j : β ↦ f j = x) s, g y = ∑ k ∈ s, g k := by 65 | rw [finsum_eq_finset_sum_of_support_subset] 66 | · rw [Finset.sum_image'] 67 | exact fun _ _ ↦ rfl 68 | · intro x hx 69 | rw [mem_support] at hx 70 | obtain ⟨a, h, -⟩ := Finset.exists_ne_zero_of_sum_ne_zero hx 71 | simp at h ⊢ 72 | exact ⟨a, h⟩ 73 | 74 | theorem sum_mem_reallyConvexHull {s : Set E} {ι : Type*} {t : Finset ι} {w : ι → 𝕜} {z : ι → E} 75 | (h₀ : ∀ i ∈ t, 0 ≤ w i) (h₁ : ∑ i ∈ t, w i = 1) (hz : ∀ i ∈ t, z i ∈ s) : 76 | ∑ i ∈ t, w i • z i ∈ reallyConvexHull 𝕜 s := by 77 | classical 78 | refine ⟨fun e ↦ ∑ᶠ i ∈ t.filter fun j ↦ z j = e, w i, ?_, ?_, ?_, ?_⟩ 79 | · rw [Pi.le_def] 80 | exact fun e ↦ finsum_nonneg fun i ↦ finsum_nonneg fun hi ↦ h₀ _ (Finset.mem_of_mem_filter i hi) 81 | · intro e he 82 | rw [mem_support] at he 83 | obtain ⟨a, h, ha⟩ := finsum.exists_ne_zero_of_sum_ne_zero he 84 | rw [Finset.mem_filter] at h 85 | rcases h with ⟨h, rfl⟩ 86 | exact hz a h 87 | · simp_rw [← h₁, finsum_mem_finset_eq_sum, finsum_sum_filter z ..] 88 | · simp_rw [finsum_mem_finset_eq_sum, Finset.sum_smul] 89 | rw [← finsum_sum_filter z] 90 | congr with x 91 | rw [Finset.sum_congr rfl] 92 | intro y hy 93 | rw [Finset.mem_filter] at hy 94 | rw [hy.2] 95 | 96 | theorem reallyConvexHull_mono : Monotone (reallyConvexHull 𝕜 : Set E → Set E) := by 97 | rintro s t h _ ⟨w, w_pos, supp_w, sum_w, rfl⟩ 98 | exact ⟨w, w_pos, supp_w.trans h, sum_w, rfl⟩ 99 | 100 | /-- Generalization of `Convex` to semirings. We only add the `s = ∅` clause if `𝕜` is trivial. -/ 101 | def ReallyConvex (𝕜 : Type*) {E : Type*} [Semiring 𝕜] [PartialOrder 𝕜] [IsOrderedRing 𝕜] 102 | [AddCommMonoid E] [Module 𝕜 E] 103 | (s : Set E) : Prop := 104 | s = ∅ ∨ ∀ w : E → 𝕜, 0 ≤ w → support w ⊆ s → ∑ᶠ x, w x = 1 → ∑ᶠ x, w x • x ∈ s 105 | 106 | variable {s : Set E} 107 | 108 | @[simp] 109 | theorem reallyConvex_empty : ReallyConvex 𝕜 (∅ : Set E) := Or.inl rfl 110 | 111 | @[simp] 112 | theorem reallyConvex_univ : ReallyConvex 𝕜 (univ : Set E) := Or.inr fun _ _ _ _ ↦ mem_univ _ 113 | 114 | -- for every lemma that requires `Nontrivial` should we also add a lemma that has the condition 115 | -- `s.Nonempty` (or even `Nontrivial 𝕜 ∨ s.Nonempty`)? 116 | theorem Nontrivial.reallyConvex_iff [Nontrivial 𝕜] : 117 | ReallyConvex 𝕜 s ↔ ∀ w : E → 𝕜, 0 ≤ w → support w ⊆ s → ∑ᶠ x, w x = 1 → ∑ᶠ x, w x • x ∈ s := by 118 | rw [ReallyConvex, or_iff_right_iff_imp] 119 | rintro rfl w hw h2w h3w 120 | obtain rfl : w = 0 := by ext; simp [imp_false] at h2w; simp [h2w] 121 | simp at h3w 122 | 123 | theorem Subsingleton.reallyConvex [Subsingleton 𝕜] : ReallyConvex 𝕜 s := by 124 | rcases eq_empty_or_nonempty s with (rfl | ⟨z, hz⟩) 125 | · exact reallyConvex_empty 126 | · refine Or.inr fun w _ _ _ ↦ ?_ 127 | convert hz 128 | haveI := Module.subsingleton 𝕜 E 129 | exact Subsingleton.elim .. 130 | 131 | theorem reallyConvex_iff_hull [Nontrivial 𝕜] : ReallyConvex 𝕜 s ↔ reallyConvexHull 𝕜 s ⊆ s := by 132 | rw [Nontrivial.reallyConvex_iff] 133 | constructor 134 | · rintro h _ ⟨w, w_pos, supp_w, sum_w, rfl⟩ 135 | exact h w w_pos supp_w sum_w 136 | · rintro h w w_pos supp_w sum_w 137 | exact h ⟨w, w_pos, supp_w, sum_w, rfl⟩ 138 | 139 | -- turn this into an iff 140 | theorem ReallyConvex.sum_mem [Nontrivial 𝕜] (hs : ReallyConvex 𝕜 s) {ι : Type*} {t : Finset ι} 141 | {w : ι → 𝕜} {z : ι → E} (h₀ : ∀ i ∈ t, 0 ≤ w i) (h₁ : ∑ i ∈ t, w i = 1) 142 | (hz : ∀ i ∈ t, z i ∈ s) : ∑ i ∈ t, w i • z i ∈ s := 143 | reallyConvex_iff_hull.mp hs (sum_mem_reallyConvexHull h₀ h₁ hz) 144 | 145 | theorem ReallyConvex.finsum_mem [Nontrivial 𝕜] (hs : ReallyConvex 𝕜 s) {ι : Type*} {w : ι → 𝕜} 146 | {z : ι → E} (h₀ : ∀ i, 0 ≤ w i) (h₁ : ∑ᶠ i, w i = 1) (hz : ∀ i ∈ support w, z i ∈ s) : 147 | ∑ᶠ i, w i • z i ∈ s := by 148 | rw [finsum_eq_sum_of_support_subset_of_finite _ _ (support_finite_of_finsum_eq_one h₁)] 149 | swap; · exact support_smul_subset_left w z 150 | apply hs.sum_mem fun i _ ↦ h₀ i 151 | · rw [← finsum_eq_sum, h₁] 152 | · simp_rw [Set.Finite.mem_toFinset]; exact hz 153 | 154 | theorem ReallyConvex.add_mem (hs : ReallyConvex 𝕜 s) {w₁ w₂ : 𝕜} {z₁ z₂ : E} 155 | (hw₁ : 0 ≤ w₁) (hw₂ : 0 ≤ w₂) (hw : w₁ + w₂ = 1) (hz₁ : z₁ ∈ s) (hz₂ : z₂ ∈ s) : 156 | w₁ • z₁ + w₂ • z₂ ∈ s := by 157 | cases subsingleton_or_nontrivial 𝕜 158 | · have := Module.subsingleton 𝕜 E 159 | rwa [Subsingleton.mem_iff_nonempty] at hz₁ ⊢ 160 | suffices ∑ b : Bool, cond b w₁ w₂ • cond b z₁ z₂ ∈ s by simpa using this 161 | apply hs.sum_mem <;> simp [*] 162 | 163 | theorem ReallyConvex.inter {t : Set E} (hs : ReallyConvex 𝕜 s) (ht : ReallyConvex 𝕜 t) : 164 | ReallyConvex 𝕜 (s ∩ t) := by 165 | rcases hs with (rfl | hs); · simp 166 | rcases ht with (rfl | ht); · simp 167 | refine Or.inr fun w w_pos supp_w sum_w ↦ ?_ 168 | cases Set.subset_inter_iff.mp supp_w 169 | constructor 170 | · apply hs <;> assumption 171 | · apply ht <;> assumption 172 | 173 | theorem ReallyConvex.preimageₛₗ (f : E →ₛₗ[σ.toRingHom] E') {s : Set E'} (hs : ReallyConvex 𝕜' s) : 174 | ReallyConvex 𝕜 (f ⁻¹' s) := by 175 | -- this proof would be easier by casing on `s = ∅`, and 176 | cases subsingleton_or_nontrivial 𝕜' 177 | · haveI : Subsingleton E' := Module.subsingleton 𝕜' E' 178 | refine Subsingleton.set_cases ?_ ?_ s 179 | · simp_rw [preimage_empty, reallyConvex_empty] 180 | · simp_rw [preimage_univ, reallyConvex_univ] 181 | · refine Or.inr fun w hw h2w h3w ↦ ?_ 182 | have h4w : (support w).Finite := support_finite_of_finsum_eq_one h3w 183 | have : (support fun x ↦ w x • x).Finite := h4w.subset (support_smul_subset_left w id) 184 | simp_rw [mem_preimage, map_finsum f this, map_smulₛₗ f] 185 | apply hs.finsum_mem 186 | · intro i; rw [← map_zero σ]; apply σ.monotone'; apply hw 187 | · rw [← map_finsum _ h4w, h3w, map_one] 188 | · intro i hi; apply h2w; rw [mem_support] at hi ⊢; contrapose! hi; rw [hi, map_zero] 189 | 190 | theorem ReallyConvex.preimage (f : E →ₗ[𝕜] E₂) {s : Set E₂} (hs : ReallyConvex 𝕜 s) : 191 | ReallyConvex 𝕜 (f ⁻¹' s) := 192 | ReallyConvex.preimageₛₗ (OrderRingHom.id 𝕜) f hs 193 | 194 | /- The next lemma would also be nice to have. 195 | lemma reallyConvex_reallyConvexHull (s : Set E) : reallyConvex 𝕜 (reallyConvexHull 𝕜 s) := sorry 196 | -/ 197 | end 198 | 199 | section 200 | 201 | variable (𝕜 : Type*) {E : Type*} [Field 𝕜] [LinearOrder 𝕜] [IsStrictOrderedRing 𝕜] [AddCommGroup E] 202 | [Module 𝕜 E] 203 | 204 | theorem reallyConvex_iff_convex {s : Set E} : ReallyConvex 𝕜 s ↔ Convex 𝕜 s := by 205 | refine ⟨fun h ↦ ?_, fun h ↦ ?_⟩ 206 | · intro x hx y hy v w hv hw hvw; apply ReallyConvex.add_mem <;> assumption 207 | · exact Or.inr fun w hw h2w h3w ↦ h.finsum_mem hw h3w fun i hi ↦ h2w <| mem_support.mpr hi 208 | 209 | end 210 | -------------------------------------------------------------------------------- /SphereEversion/ToMathlib/Analysis/CutOff.lean: -------------------------------------------------------------------------------- 1 | import Mathlib.Geometry.Manifold.PartitionOfUnity 2 | 3 | open Set Filter 4 | 5 | open scoped Manifold Topology ContDiff 6 | 7 | -- The bundled versions of this lemma has been merged to mathlib. 8 | -- TODO: add the unbundled version, or (better) re-write sphere-eversion accordingly. 9 | theorem exists_contDiff_one_nhds_of_interior {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] 10 | [FiniteDimensional ℝ E] {s t : Set E} (hs : IsClosed s) (hd : s ⊆ interior t) : 11 | ∃ f : E → ℝ, ContDiff ℝ ∞ f ∧ (∀ᶠ x in 𝓝ˢ s, f x = 1) ∧ (∀ x ∉ t, f x = 0) ∧ 12 | ∀ x, f x ∈ Icc (0 : ℝ) 1 := 13 | let ⟨f, hfs, hft, hf01⟩ := exists_smooth_one_nhds_of_subset_interior 𝓘(ℝ, E) hs hd 14 | ⟨f, f.contMDiff.contDiff, hfs, hft, hf01⟩ 15 | -------------------------------------------------------------------------------- /SphereEversion/ToMathlib/Analysis/InnerProductSpace/CrossProduct.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2022 Heather Macbeth. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Heather Macbeth 5 | 6 | ! This file was ported from Lean 3 source module to_mathlib.analysis.inner_product_space.cross_product 7 | -/ 8 | import Mathlib.Analysis.InnerProductSpace.Dual 9 | import Mathlib.Analysis.InnerProductSpace.Orientation 10 | import Mathlib.LinearAlgebra.Alternating.Curry 11 | 12 | /-! # The cross-product on an oriented real inner product space of dimension three -/ 13 | 14 | noncomputable section 15 | 16 | open scoped RealInnerProductSpace 17 | 18 | open Module 19 | 20 | set_option synthInstance.checkSynthOrder false 21 | attribute [local instance] FiniteDimensional.of_fact_finrank_eq_succ 22 | set_option synthInstance.checkSynthOrder true 23 | 24 | variable (E : Type*) [NormedAddCommGroup E] [InnerProductSpace ℝ E] 25 | 26 | /-- The identification of a finite-dimensional inner product space with its algebraic dual. -/ 27 | private def to_dual [FiniteDimensional ℝ E] : E ≃ₗ[ℝ] E →ₗ[ℝ] ℝ := 28 | (InnerProductSpace.toDual ℝ E).toLinearEquiv ≪≫ₗ LinearMap.toContinuousLinearMap.symm 29 | 30 | namespace Orientation 31 | 32 | variable {E} 33 | variable [Fact (finrank ℝ E = 3)] (ω : Orientation ℝ E (Fin 3)) 34 | 35 | /-- Linear map from `E` to `E →ₗ[ℝ] E` constructed from a 3-form `Ω` on `E` and an identification of 36 | `E` with its dual. Effectively, the Hodge star operation. (Under appropriate hypotheses it turns 37 | out that the image of this map is in `𝔰𝔬(E)`, the skew-symmetric operators, which can be identified 38 | with `Λ²E`.) -/ 39 | def crossProduct : E →ₗ[ℝ] E →ₗ[ℝ] E := by 40 | let z : AlternatingMap ℝ E ℝ (Fin 0) ≃ₗ[ℝ] ℝ := 41 | AlternatingMap.constLinearEquivOfIsEmpty.symm 42 | let y : AlternatingMap ℝ E ℝ (Fin 1) →ₗ[ℝ] E →ₗ[ℝ] ℝ := 43 | LinearMap.llcomp ℝ E (AlternatingMap ℝ E ℝ (Fin 0)) ℝ z ∘ₗ AlternatingMap.curryLeftLinearMap 44 | let y' : AlternatingMap ℝ E ℝ (Fin 1) →ₗ[ℝ] E := 45 | (LinearMap.llcomp ℝ (AlternatingMap ℝ E ℝ (Fin 1)) (E →ₗ[ℝ] ℝ) E (to_dual E).symm) y 46 | let u : AlternatingMap ℝ E ℝ (Fin 2) →ₗ[ℝ] E →ₗ[ℝ] E := 47 | LinearMap.llcomp ℝ E (AlternatingMap ℝ E ℝ (Fin 1)) _ y' ∘ₗ AlternatingMap.curryLeftLinearMap 48 | exact u ∘ₗ AlternatingMap.curryLeftLinearMap (n := 2) ω.volumeForm 49 | 50 | local infixl:100 "×₃" => ω.crossProduct 51 | 52 | theorem crossProduct_apply_self (v : E) : v×₃v = 0 := by simp [crossProduct] 53 | 54 | theorem inner_crossProduct_apply (u v w : E) : ⟪u×₃v, w⟫ = ω.volumeForm ![u, v, w] := by 55 | simp only [crossProduct, to_dual, LinearEquiv.trans_symm, LinearEquiv.symm_symm, 56 | LinearIsometryEquiv.toLinearEquiv_symm, AlternatingMap.curryLeftLinearMap_apply, 57 | LinearMap.coe_comp, Function.comp_apply, LinearMap.llcomp_apply, LinearEquiv.coe_coe, 58 | LinearEquiv.trans_apply, LinearIsometryEquiv.coe_toLinearEquiv, LinearMap.coe_comp] 59 | rw [InnerProductSpace.toDual_symm_apply] 60 | simp 61 | 62 | theorem inner_crossProduct_apply_self (u : E) (v : (ℝ ∙ u)ᗮ) : ⟪u×₃v, u⟫ = 0 := by 63 | rw [ω.inner_crossProduct_apply u v u] 64 | exact ω.volumeForm.map_eq_zero_of_eq ![u, v, u] (by simp) (by norm_num; decide : (0 : Fin 3) ≠ 2) 65 | 66 | theorem inner_crossProduct_apply_apply_self (u : E) (v : (ℝ ∙ u)ᗮ) : ⟪u×₃v, v⟫ = 0 := by 67 | rw [ω.inner_crossProduct_apply u v v] 68 | exact ω.volumeForm.map_eq_zero_of_eq ![u, v, v] (by simp) (by norm_num; decide : (1 : Fin 3) ≠ 2) 69 | 70 | /-- The map `crossProduct`, upgraded from linear to continuous-linear; useful for calculus. -/ 71 | def crossProduct' : E →L[ℝ] E →L[ℝ] E := 72 | LinearMap.toContinuousLinearMap 73 | (↑(LinearMap.toContinuousLinearMap : (E →ₗ[ℝ] E) ≃ₗ[ℝ] E →L[ℝ] E) ∘ₗ ω.crossProduct) 74 | 75 | @[simp] 76 | theorem crossProduct'_apply (v : E) : 77 | ω.crossProduct' v = LinearMap.toContinuousLinearMap (ω.crossProduct v) := 78 | rfl 79 | 80 | theorem norm_crossProduct (u : E) (v : (ℝ ∙ u)ᗮ) : ‖u×₃v‖ = ‖u‖ * ‖v‖ := by 81 | classical 82 | refine le_antisymm ?_ ?_ 83 | · obtain (h | h) := eq_or_lt_of_le (norm_nonneg (u×₃v)) 84 | · rw [← h] 85 | positivity 86 | · refine le_of_mul_le_mul_right ?_ h 87 | rw [← real_inner_self_eq_norm_mul_norm] 88 | simpa [inner_crossProduct_apply, Fin.mk_zero, Fin.prod_univ_succ, Matrix.cons_val_zero, 89 | mul_assoc] using ω.volumeForm_apply_le ![u, v, u×₃v] 90 | let K : Submodule ℝ E := Submodule.span ℝ ({u, ↑v} : Set E) 91 | have : Nontrivial Kᗮ := by 92 | apply Module.nontrivial_of_finrank_pos (R := ℝ) 93 | have : finrank ℝ K ≤ Finset.card {u, (v : E)} := by 94 | simpa [Set.toFinset_singleton] using finrank_span_le_card ({u, ↑v} : Set E) 95 | have : Finset.card {u, (v : E)} ≤ Finset.card {(v : E)} + 1 := Finset.card_insert_le u {↑v} 96 | have : Finset.card {(v : E)} = 1 := Finset.card_singleton (v : E) 97 | have : finrank ℝ K + finrank ℝ Kᗮ = finrank ℝ E := K.finrank_add_finrank_orthogonal 98 | have : finrank ℝ E = 3 := Fact.out 99 | linarith 100 | obtain ⟨w, hw⟩ : ∃ w : Kᗮ, w ≠ 0 := exists_ne 0 101 | have H : Pairwise fun i j ↦ ⟪![u, v, w] i, ![u, v, w] j⟫ = 0 := by 102 | intro i j hij 103 | have h1 : ⟪u, v⟫ = 0 := v.2 _ (Submodule.mem_span_singleton_self _) 104 | have h2 : ⟪(v : E), w⟫ = 0 := w.2 _ (Submodule.subset_span (by simp)) 105 | have h3 : ⟪u, w⟫ = 0 := w.2 _ (Submodule.subset_span (by simp)) 106 | fin_cases i <;> fin_cases j <;> norm_num at hij <;> simp [h1, h2, h3] <;> 107 | rw [real_inner_comm] <;> 108 | assumption 109 | refine le_of_mul_le_mul_right ?_ (by rwa [norm_pos_iff] : 0 < ‖w‖) 110 | -- Cauchy-Schwarz inequality for `u ×₃ v` and `w` 111 | simpa [inner_crossProduct_apply, ω.abs_volumeForm_apply_of_pairwise_orthogonal H, 112 | Fin.prod_univ_succ, Fintype.univ_ofSubsingleton, Matrix.cons_val_fin_one, Matrix.cons_val_succ, 113 | mul_assoc] using abs_real_inner_le_norm (u×₃v) w 114 | 115 | theorem isometry_on_crossProduct (u : Metric.sphere (0 : E) 1) (v : (ℝ ∙ (u : E))ᗮ) : 116 | ‖u×₃v‖ = ‖v‖ := by simp [norm_crossProduct] 117 | 118 | end Orientation 119 | -------------------------------------------------------------------------------- /SphereEversion/ToMathlib/Analysis/InnerProductSpace/Dual.lean: -------------------------------------------------------------------------------- 1 | import Mathlib.Analysis.InnerProductSpace.Dual 2 | import SphereEversion.ToMathlib.Analysis.InnerProductSpace.Projection 3 | 4 | open scoped RealInnerProductSpace 5 | 6 | open Submodule InnerProductSpace 7 | 8 | open LinearMap (ker) 9 | 10 | variable {E : Type*} [NormedAddCommGroup E] [InnerProductSpace ℝ E] [CompleteSpace E] 11 | 12 | @[inherit_doc] local notation "Δ" => spanLine 13 | 14 | @[inherit_doc] local notation "{." x "}ᗮ" => spanOrthogonal x 15 | 16 | @[inherit_doc] local notation "pr[" x "]ᗮ" => projSpanOrthogonal x 17 | 18 | theorem orthogonal_span_toDual_symm (π : E →L[ℝ] ℝ) : 19 | {.(InnerProductSpace.toDual ℝ E).symm π}ᗮ = ker π := by 20 | ext x 21 | suffices (∀ a : ℝ, ⟪a • (toDual ℝ E).symm π, x⟫ = 0) ↔ π x = 0 by 22 | simp only [orthogonal, mem_mk, Set.mem_setOf_eq, LinearMap.mem_ker, ← toDual_symm_apply] 23 | change (∀ (u : E), u ∈ span ℝ {(LinearIsometryEquiv.symm (toDual ℝ E)) π} → inner _ u x = 0) ↔ _ 24 | simpa 25 | refine ⟨fun h ↦ ?_, fun h _ ↦ ?_⟩ 26 | · simpa using h 1 27 | · simp [inner_smul_left, h] 28 | -------------------------------------------------------------------------------- /SphereEversion/ToMathlib/Analysis/InnerProductSpace/Rotation.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2022 Heather Macbeth. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Heather Macbeth 5 | 6 | ! This file was ported from Lean 3 source module to_mathlib.analysis.inner_product_space.rotation 7 | -/ 8 | import Mathlib.Analysis.SpecialFunctions.Trigonometric.Deriv 9 | import SphereEversion.ToMathlib.Analysis.ContDiff 10 | import SphereEversion.ToMathlib.LinearAlgebra.Basic 11 | import SphereEversion.ToMathlib.Analysis.InnerProductSpace.CrossProduct 12 | 13 | /-! # Rotation about an axis, considered as a function in that axis -/ 14 | 15 | noncomputable section 16 | 17 | open scoped RealInnerProductSpace 18 | 19 | open Module Submodule 20 | 21 | set_option synthInstance.checkSynthOrder false 22 | attribute [local instance] FiniteDimensional.of_fact_finrank_eq_succ 23 | set_option synthInstance.checkSynthOrder true 24 | 25 | variable {E : Type*} [NormedAddCommGroup E] [InnerProductSpace ℝ E] [Fact (finrank ℝ E = 3)] 26 | (ω : Orientation ℝ E (Fin 3)) 27 | 28 | local infixl:100 "×₃" => ω.crossProduct 29 | 30 | namespace Orientation 31 | 32 | /-- A family of endomorphisms of `E`, parametrized by `ℝ × E`. The idea is that for nonzero `v : E` 33 | and `t : ℝ` this endomorphism will be the rotation by the angle `t` about the axis spanned by `v`. 34 | -/ 35 | def rot (p : ℝ × E) : E →L[ℝ] E := 36 | (ℝ ∙ p.2).subtypeL ∘L (orthogonalProjection (ℝ ∙ p.2) : E →L[ℝ] ℝ ∙ p.2) + 37 | Real.cos (p.1 * Real.pi) • 38 | (ℝ ∙ p.2)ᗮ.subtypeL ∘L (orthogonalProjection (ℝ ∙ p.2)ᗮ : E →L[ℝ] (ℝ ∙ p.2)ᗮ) + 39 | Real.sin (p.1 * Real.pi) • LinearMap.toContinuousLinearMap (ω.crossProduct p.2) 40 | 41 | /-- Alternative form of the construction `rot`, convenient for the smoothness calculation. -/ 42 | def rotAux (p : ℝ × E) : E →L[ℝ] E := 43 | Real.cos (p.1 * Real.pi) • ContinuousLinearMap.id ℝ E + 44 | ((1 - Real.cos (p.1 * Real.pi)) • 45 | (ℝ ∙ p.2).subtypeL ∘L (orthogonalProjection (ℝ ∙ p.2) : E →L[ℝ] ℝ ∙ p.2) + 46 | Real.sin (p.1 * Real.pi) • ω.crossProduct' p.2) 47 | 48 | theorem rot_eq_aux : ω.rot = ω.rotAux := by 49 | ext1 p 50 | dsimp [rot, rotAux] 51 | rw [id_eq_sum_orthogonalProjection_self_orthogonalComplement (ℝ ∙ p.2)] 52 | simp only [smul_add, sub_smul, one_smul] 53 | abel 54 | 55 | /-- The map `rot` is smooth on `ℝ × (E \ {0})`. -/ 56 | theorem contDiff_rot {p : ℝ × E} (hp : p.2 ≠ 0) : ContDiffAt ℝ ⊤ ω.rot p := by 57 | simp only [rot_eq_aux] 58 | refine (contDiffAt_fst.mul_const.cos.smul contDiffAt_const).add ?_ 59 | refine ((contDiffAt_const.sub contDiffAt_fst.mul_const.cos).smul ?_).add 60 | (contDiffAt_fst.mul_const.sin.smul ?_) 61 | · exact (contDiffAt_orthogonalProjection_singleton hp).comp _ contDiffAt_snd 62 | · exact ω.crossProduct'.contDiff.contDiffAt.comp _ contDiffAt_snd 63 | 64 | /-- The map `rot` sends `{0} × E` to the identity. -/ 65 | theorem rot_zero (v : E) : ω.rot (0, v) = ContinuousLinearMap.id ℝ E := by 66 | ext w 67 | simp [rot] 68 | 69 | /-- The map `rot` sends `(1, v)` to a transformation which on `(ℝ ∙ v)ᗮ` acts as the negation. -/ 70 | theorem rot_one (v : E) {w : E} (hw : w ∈ (ℝ ∙ v)ᗮ) : ω.rot (1, v) w = -w := by 71 | suffices (orthogonalProjection (Submodule.span ℝ {v}) w : E) + 72 | -(orthogonalProjection (Submodule.span ℝ {v})ᗮ w) = -w by simpa [rot] 73 | simp [orthogonalProjection_eq_self_iff.mpr hw, 74 | orthogonalProjection_mem_subspace_orthogonalComplement_eq_zero hw] 75 | 76 | /-- The map `rot` sends `(v, t)` to a transformation fixing `v`. -/ 77 | @[simp] 78 | theorem rot_self (p : ℝ × E) : ω.rot p (no_index p.2) = p.2 := by 79 | have H : orthogonalProjection (ℝ ∙ p.2) p.2 = p.2 := 80 | orthogonalProjection_eq_self_iff.mpr (Submodule.mem_span_singleton_self p.2) 81 | simp [rot, crossProduct_apply_self, orthogonalProjection_orthogonalComplement_singleton_eq_zero,H] 82 | 83 | /-- The map `rot` sends `(t, v)` to a transformation preserving `span v`. -/ 84 | theorem rot_eq_of_mem_span (p : ℝ × E) {x : E} (hx : x ∈ ℝ ∙ p.2) : ω.rot p x = x := by 85 | obtain ⟨a, rfl⟩ := Submodule.mem_span_singleton.mp hx; simp_rw [map_smul, rot_self] 86 | 87 | /-- The map `rot` sends `(v, t)` to a transformation preserving the subspace `(ℝ ∙ v)ᗮ`. -/ 88 | theorem inner_rot_apply_self (p : ℝ × E) (w : E) (hw : w ∈ (ℝ ∙ p.2)ᗮ) : 89 | ⟪ω.rot p w, no_index p.2⟫ = 0 := by 90 | have H₁ : orthogonalProjection (ℝ ∙ p.2) w = 0 := 91 | orthogonalProjection_mem_subspace_orthogonalComplement_eq_zero hw 92 | have H₂ : (orthogonalProjection (ℝ ∙ p.2)ᗮ w : E) = w := 93 | congr_arg (_ : (ℝ ∙ p.2)ᗮ → E) (orthogonalProjection_mem_subspace_eq_self ⟨w, hw⟩) 94 | have H₃ : ⟪w, p.2⟫ = 0 := by 95 | simpa only [real_inner_comm] using hw p.2 (Submodule.mem_span_singleton_self _) 96 | have H₄ : ⟪p.2×₃w, p.2⟫ = 0 := ω.inner_crossProduct_apply_self p.2 ⟨w, hw⟩ 97 | simp [rot, H₁, H₂, H₃, H₄, inner_smul_left, inner_add_left] 98 | 99 | theorem isometry_on_rot (t : ℝ) (v : Metric.sphere (0 : E) 1) (w : (ℝ ∙ (v : E))ᗮ) : 100 | ‖ω.rot (t, v) w‖ = ‖(w : E)‖ := by 101 | have h1 : ⟪v×₃w, v×₃w⟫ = ⟪w, w⟫ := by 102 | simp only [inner_self_eq_norm_sq_to_K, ω.isometry_on_crossProduct v w] 103 | have h2 : ⟪v×₃w, w⟫ = 0 := ω.inner_crossProduct_apply_apply_self v w 104 | have h3 : ⟪(w : E), v×₃w⟫ = 0 := by rwa [real_inner_comm] 105 | have : ‖Real.cos (t * Real.pi) • (w : E) + Real.sin (t * Real.pi) • v×₃w‖ = ‖(w : E)‖ := by 106 | simp only [@norm_eq_sqrt_re_inner ℝ] 107 | congr 2 108 | simp only [inner_add_left, inner_add_right, inner_smul_left, inner_smul_right, h1, h2, h3, 109 | RCLike.conj_to_real, Submodule.coe_inner] 110 | linear_combination ⟪(w : E), w⟫ * Real.cos_sq_add_sin_sq (t * Real.pi) 111 | dsimp [rot] 112 | simp [orthogonalProjection_mem_subspace_orthogonalComplement_eq_zero w.prop, this] 113 | 114 | theorem isometry_rot (t : ℝ) (v : Metric.sphere (0 : E) 1) : Isometry (ω.rot (t, v)) := by 115 | rw [AddMonoidHomClass.isometry_iff_norm] 116 | intro w 117 | obtain ⟨a, ha, w, hw, rfl⟩ := (ℝ ∙ (v : E)).exists_add_mem_mem_orthogonal w 118 | rw [Submodule.mem_span_singleton] at ha 119 | obtain ⟨s, rfl⟩ := ha 120 | rw [← sq_eq_sq₀ (norm_nonneg _) (norm_nonneg _), sq, sq, map_add, 121 | @norm_add_sq_eq_norm_sq_add_norm_sq_of_inner_eq_zero ℝ, 122 | @norm_add_sq_eq_norm_sq_add_norm_sq_of_inner_eq_zero ℝ] 123 | · have hvw : ‖ω.rot (t, v) w‖ = ‖w‖ := ω.isometry_on_rot t v ⟨w, hw⟩ 124 | simp [hvw] 125 | · simp [inner_smul_left, hw v (Submodule.mem_span_singleton_self _)] 126 | rw [real_inner_comm] 127 | simp [inner_smul_right, ω.inner_rot_apply_self (t, v) w hw] 128 | 129 | open Real Submodule 130 | 131 | open scoped Real 132 | 133 | set_option linter.style.multiGoal false in 134 | theorem injOn_rot_of_ne (t : ℝ) {x : E} (hx : x ≠ 0) : Set.InjOn (ω.rot (t, x)) (ℝ ∙ x)ᗮ := by 135 | change Set.InjOn (ω.rot (t, x)).toLinearMap (ℝ ∙ x)ᗮ 136 | simp_rw [← LinearMap.ker_inf_eq_bot, Submodule.eq_bot_iff, Submodule.mem_inf] 137 | rintro y ⟨hy, hy'⟩ 138 | rw [LinearMap.mem_ker] at hy 139 | change 140 | ↑((orthogonalProjection (span ℝ {x})) y) + 141 | cos (t * Real.pi) • ↑((orthogonalProjection (span ℝ {x})ᗮ) y) + 142 | Real.sin (t * Real.pi) • x×₃y = 143 | 0 at hy 144 | rw [orthogonalProjection_mem_subspace_orthogonalComplement_eq_zero hy', 145 | orthogonalProjection_eq_self_iff.mpr hy', coe_zero, zero_add] at hy 146 | apply_fun fun x ↦ ‖x‖ ^ 2 at hy 147 | rw [pow_two, @norm_add_sq_eq_norm_sq_add_norm_sq_of_inner_eq_zero ℝ] at hy 148 | simp_rw [← pow_two, norm_smul, mul_pow] at hy 149 | change _ + _ * ‖x×₃(⟨y, hy'⟩ : (span ℝ {x})ᗮ)‖ ^ 2 = ‖(0 : E)‖ ^ 2 at hy 150 | rw [norm_crossProduct] at hy 151 | simp +decide only [norm_eq_abs, Even.pow_abs, coe_mk, norm_zero, zero_pow, Ne, Nat.one_ne_zero, 152 | not_false_iff] at hy 153 | change _ + _ * (_ * ‖y‖) ^ 2 = 0 at hy 154 | rw [mul_pow, ← mul_assoc, ← add_mul, mul_eq_zero, or_iff_not_imp_left] at hy 155 | have : (0 : ℝ) < cos (t * π) ^ 2 + sin (t * π) ^ 2 * ‖x‖ ^ 2 := by 156 | have : (0 : ℝ) < ‖x‖ ^ 2 := pow_pos (norm_pos_iff.mpr hx) 2 157 | rw [cos_sq'] 158 | rw [show (1 : ℝ) - sin (t * π) ^ 2 + sin (t * π) ^ 2 * ‖x‖ ^ 2 = 159 | (1 : ℝ) + sin (t * π) ^ 2 * (‖x‖ ^ 2 - 1) by ring] 160 | have I₂ : sin (t * π) ^ 2 ≤ 1 := sin_sq_le_one (t * π) 161 | have I₃ : (0 : ℝ) ≤ sin (t * π) ^ 2 := sq_nonneg _ 162 | rcases I₃.eq_or_lt with (H | H) 163 | · rw [← H] 164 | norm_num 165 | · nlinarith 166 | · replace hy := hy this.ne' 167 | exact norm_eq_zero.mp (pow_eq_zero hy) 168 | · rw [inner_smul_left, inner_smul_right] 169 | have := inner_crossProduct_apply_apply_self ω x ⟨y, hy'⟩ 170 | change ⟪x×₃y, y⟫ = 0 at this 171 | rw [real_inner_comm, this] 172 | simp 173 | 174 | theorem injOn_rot (t : ℝ) (x : Metric.sphere (0 : E) 1) : 175 | Set.InjOn (ω.rot (t, x)) (ℝ ∙ (x : E))ᗮ := by 176 | intro v hv w hw h 177 | simpa [sub_eq_zero, h] using (ω.isometry_on_rot t x (⟨v, hv⟩ - ⟨w, hw⟩)).symm 178 | 179 | end Orientation 180 | -------------------------------------------------------------------------------- /SphereEversion/ToMathlib/Analysis/Normed/Module/FiniteDimension.lean: -------------------------------------------------------------------------------- 1 | import Mathlib.Analysis.Normed.Module.FiniteDimension 2 | 3 | -- being affinely independent is an open condition 4 | section 5 | variable (𝕜 E : Type*) {ι : Type*} [NontriviallyNormedField 𝕜] 6 | [NormedAddCommGroup E] [NormedSpace 𝕜 E] [CompleteSpace 𝕜] [Finite ι] 7 | 8 | theorem isOpen_affineIndependent : IsOpen {p : ι → E | AffineIndependent 𝕜 p} := by 9 | classical 10 | rcases isEmpty_or_nonempty ι with h | ⟨⟨i₀⟩⟩ 11 | · exact isOpen_discrete _ 12 | · simp_rw [affineIndependent_iff_linearIndependent_vsub 𝕜 _ i₀] 13 | let ι' := { x // x ≠ i₀ } 14 | cases nonempty_fintype ι 15 | haveI : Fintype ι' := Subtype.fintype _ 16 | convert_to 17 | IsOpen ((fun (p : ι → E) (i : ι') ↦ p i -ᵥ p i₀) ⁻¹' {p : ι' → E | LinearIndependent 𝕜 p}) 18 | refine isOpen_setOf_linearIndependent.preimage ?_ 19 | exact continuous_pi fun i' ↦ 20 | (continuous_apply (π := fun _ : ι ↦ E) i'.1).vsub <| continuous_apply i₀ 21 | 22 | end 23 | -------------------------------------------------------------------------------- /SphereEversion/ToMathlib/Analysis/NormedSpace/Misc.lean: -------------------------------------------------------------------------------- 1 | import Mathlib.Analysis.InnerProductSpace.EuclideanDist 2 | 3 | variable {F : Type*} [NormedAddCommGroup F] [NormedSpace ℝ F] 4 | 5 | open Set Metric 6 | open scoped ContDiff 7 | 8 | noncomputable section 9 | 10 | variable [FiniteDimensional ℝ F] (c : F) (r : ℝ) 11 | 12 | theorem PartialHomeomorph.exists_contDiff_source_univ_target_subset_ball : 13 | ∃ f : PartialHomeomorph F F, ContDiff ℝ ∞ f ∧ ContDiffOn ℝ ∞ f.symm f.target ∧ 14 | f.source = univ ∧ (0 < r → f.target ⊆ ball c r) ∧ f 0 = c := by 15 | by_cases hr : 0 < r 16 | · let e := toEuclidean (E := F) 17 | let e' := e.toHomeomorph 18 | rcases Euclidean.nhds_basis_ball.mem_iff.1 (ball_mem_nhds c hr) with ⟨ε, ε₀, hε⟩ 19 | set f := (e'.transPartialHomeomorph (.univBall (e c) ε)).transHomeomorph e'.symm 20 | have hf : f.target = Euclidean.ball c ε := by 21 | rw [transHomeomorph_target, Homeomorph.transPartialHomeomorph_target, univBall_target _ ε₀] 22 | rfl 23 | refine ⟨f, ?_, ?_, ?_, fun _ ↦ by rwa [hf], by simp [e, e', f]⟩ 24 | · exact e.symm.contDiff.comp <| contDiff_univBall.comp e.contDiff 25 | · exact e.symm.contDiff.comp_contDiffOn <| contDiffOn_univBall_symm.comp 26 | e.contDiff.contDiffOn hf.subset 27 | · rw [transHomeomorph_source, Homeomorph.transPartialHomeomorph_source, univBall_source]; rfl 28 | · use (IsometryEquiv.vaddConst c).toHomeomorph.toPartialHomeomorph, 29 | contDiff_id.add contDiff_const, contDiffOn_id.sub contDiffOn_const 30 | simp [hr] 31 | 32 | /-- A variant of `InnerProductSpace.diffeomorphToNhd` which provides a function satisfying the 33 | weaker condition `range_diffeomorph_to_nhd_subset_ball` but which applies to any `NormedSpace`. 34 | 35 | In fact one could demand this stronger property but it would be more work and we don't need it. -/ 36 | def diffeomorphToNhd (c : F) (r : ℝ) : PartialHomeomorph F F := 37 | (PartialHomeomorph.exists_contDiff_source_univ_target_subset_ball c r).choose 38 | 39 | @[simp] 40 | theorem diffeomorphToNhd_source : (diffeomorphToNhd c r).source = univ := 41 | (PartialHomeomorph.exists_contDiff_source_univ_target_subset_ball c r).choose_spec.2.2.1 42 | 43 | @[simp] 44 | theorem diffeomorphToNhd_apply_zero : diffeomorphToNhd c r 0 = c := 45 | (PartialHomeomorph.exists_contDiff_source_univ_target_subset_ball c r).choose_spec.2.2.2.2 46 | 47 | variable {r} in 48 | theorem target_diffeomorphToNhd_subset_ball (hr : 0 < r) : 49 | (diffeomorphToNhd c r).target ⊆ ball c r := 50 | (PartialHomeomorph.exists_contDiff_source_univ_target_subset_ball c r).choose_spec.2.2.2.1 hr 51 | 52 | variable {r} in 53 | @[simp] 54 | theorem range_diffeomorphToNhd_subset_ball (hr : 0 < r) : 55 | range (diffeomorphToNhd c r) ⊆ ball c r := by 56 | erw [← image_univ, ← diffeomorphToNhd_source c r, PartialEquiv.image_source_eq_target] 57 | exact target_diffeomorphToNhd_subset_ball c hr 58 | 59 | @[simp] 60 | theorem contDiff_diffeomorphToNhd {n : ℕ∞} : 61 | ContDiff ℝ n <| diffeomorphToNhd c r := 62 | (PartialHomeomorph.exists_contDiff_source_univ_target_subset_ball c r).choose_spec.1.of_le 63 | (mod_cast le_top) 64 | 65 | @[simp] 66 | theorem contDiffOn_diffeomorphToNhd_inv {n : ℕ∞} : 67 | ContDiffOn ℝ n (diffeomorphToNhd c r).symm (diffeomorphToNhd c r).target := 68 | (PartialHomeomorph.exists_contDiff_source_univ_target_subset_ball c r).choose_spec.2.1.of_le 69 | (mod_cast le_top) 70 | -------------------------------------------------------------------------------- /SphereEversion/ToMathlib/Analysis/NormedSpace/OperatorNorm/Prod.lean: -------------------------------------------------------------------------------- 1 | import Mathlib.Analysis.InnerProductSpace.Basic 2 | import Mathlib.Analysis.Normed.Operator.BoundedLinearMaps 3 | import Mathlib.Analysis.NormedSpace.OperatorNorm.Prod 4 | 5 | noncomputable section 6 | 7 | @[inherit_doc] local notation:70 u " ⬝ " φ:65 => 8 | ContinuousLinearMap.comp (ContinuousLinearMap.toSpanSingleton ℝ u) φ 9 | 10 | variable {𝕜 E F G Fₗ Gₗ X : Type*} [NontriviallyNormedField 𝕜] [NormedAddCommGroup E] 11 | [NormedAddCommGroup Fₗ] [NormedAddCommGroup Gₗ] [NormedAddCommGroup F] [NormedAddCommGroup G] 12 | [NormedSpace 𝕜 E] [NormedSpace 𝕜 Fₗ] [NormedSpace 𝕜 Gₗ] [NormedSpace 𝕜 F] [NormedSpace 𝕜 G] 13 | [TopologicalSpace X] 14 | 15 | theorem ContinuousLinearMap.le_opNorm_of_le' {𝕜 : Type*} {𝕜₂ : Type*} {E : Type*} {F : Type*} 16 | [NormedAddCommGroup E] [SeminormedAddCommGroup F] [NontriviallyNormedField 𝕜] 17 | [NontriviallyNormedField 𝕜₂] [NormedSpace 𝕜 E] [NormedSpace 𝕜₂ F] {σ₁₂ : 𝕜 →+* 𝕜₂} 18 | [RingHomIsometric σ₁₂] (f : E →SL[σ₁₂] F) {x : E} (hx : x ≠ 0) {C : ℝ} (h : C * ‖x‖ ≤ ‖f x‖) : 19 | C ≤ ‖f‖ := by 20 | apply le_of_mul_le_mul_right (h.trans (f.le_opNorm x)) 21 | rwa [norm_pos_iff] 22 | 23 | @[simp] 24 | theorem ContinuousLinearMap.toSpanSingleton_zero (𝕜 : Type*) {E : Type*} 25 | [SeminormedAddCommGroup E] [NontriviallyNormedField 𝕜] [NormedSpace 𝕜 E] : 26 | ContinuousLinearMap.toSpanSingleton 𝕜 (0 : E) = 0 := by 27 | ext 28 | simp only [ContinuousLinearMap.toSpanSingleton_apply, ContinuousLinearMap.zero_apply, smul_zero] 29 | 30 | @[simp] 31 | theorem ContinuousLinearMap.comp_toSpanSingleton_apply {E : Type*} [NormedAddCommGroup E] 32 | [NormedSpace ℝ E] {F : Type*} [NormedAddCommGroup F] [NormedSpace ℝ F] (φ : E →L[ℝ] ℝ) (v : E) 33 | (u : F) : (u ⬝ φ) v = φ v • u := 34 | rfl 35 | 36 | universe u₁ u₂ u₃ u₄ 37 | 38 | /-- The natural linear map `(M →ₗ[R] M₃) × (M₂ →ₗ[R] M₃) →ₗ[R] M × M₂ →ₗ[R] M₃` for `R`-modules `M`, 39 | `M₂`, `M₃` over a commutative ring `R`. 40 | 41 | If `f : M →ₗ[R] M₃` and `g : M₂ →ₗ[R] M₃` then `linear_map.coprodₗ (f, g)` is the map 42 | `(m, n) ↦ f m + g n`. -/ 43 | def LinearMap.coprodₗ (R : Type u₁) (M : Type u₂) (M₂ : Type u₃) (M₃ : Type u₄) [CommRing R] 44 | [AddCommMonoid M] [AddCommMonoid M₂] [AddCommMonoid M₃] [Module R M] [Module R M₂] 45 | [Module R M₃] : (M →ₗ[R] M₃) × (M₂ →ₗ[R] M₃) →ₗ[R] M × M₂ →ₗ[R] M₃ where 46 | toFun p := p.1.coprod p.2 47 | map_add' p q := by 48 | apply LinearMap.coe_injective 49 | ext x 50 | simp only [Prod.fst_add, LinearMap.coprod_apply, LinearMap.add_apply, Prod.snd_add] 51 | module 52 | map_smul' r p := by 53 | apply LinearMap.coe_injective 54 | ext x 55 | simp only [Prod.smul_fst, Prod.smul_snd, LinearMap.coprod_apply, LinearMap.smul_apply, 56 | RingHom.id_apply, smul_add] 57 | 58 | theorem isBoundedLinearMap_coprod (𝕜 : Type*) [NontriviallyNormedField 𝕜] (E : Type*) 59 | [NormedAddCommGroup E] [NormedSpace 𝕜 E] (F : Type*) [NormedAddCommGroup F] [NormedSpace 𝕜 F] 60 | (G : Type*) [NormedAddCommGroup G] [NormedSpace 𝕜 G] : 61 | IsBoundedLinearMap 𝕜 fun p : (E →L[𝕜] G) × (F →L[𝕜] G) ↦ p.1.coprod p.2 := 62 | { map_add := by 63 | intros 64 | apply ContinuousLinearMap.coeFn_injective 65 | ext u 66 | simp only [Prod.fst_add, Prod.snd_add, ContinuousLinearMap.coprod_apply, 67 | ContinuousLinearMap.add_apply] 68 | ac_rfl 69 | map_smul := by 70 | intro r p 71 | apply ContinuousLinearMap.coeFn_injective 72 | ext x 73 | simp only [Prod.smul_fst, Prod.smul_snd, ContinuousLinearMap.coprod_apply, 74 | ContinuousLinearMap.coe_smul', Pi.smul_apply, smul_add] 75 | bound := by 76 | refine ⟨2, zero_lt_two, fun ⟨φ, ψ⟩ ↦ ?_⟩ 77 | apply ContinuousLinearMap.opNorm_le_bound _ (by positivity) 78 | rintro ⟨e, f⟩ 79 | calc 80 | ‖φ e + ψ f‖ ≤ ‖φ e‖ + ‖ψ f‖ := norm_add_le _ _ 81 | _ ≤ ‖φ‖ * ‖e‖ + ‖ψ‖ * ‖f‖ := (add_le_add (φ.le_opNorm e) (ψ.le_opNorm f)) 82 | _ ≤ 2 * max ‖φ‖ ‖ψ‖ * max ‖e‖ ‖f‖ := by 83 | rw [two_mul, add_mul] 84 | gcongr <;> first | apply le_max_left | apply le_max_right } 85 | 86 | /-- The natural continuous linear map `((E →L[𝕜] G) × (F →L[𝕜] G)) →L[𝕜] (E × F →L[𝕜] G)` for 87 | normed spaces `E`, `F`, `G` over a normed field `𝕜`. 88 | 89 | If `g₁ : E →L[𝕜] G` and `g₂ : F →L[𝕜] G` then `continuous_linear_map.coprodL (g₁, g₂)` is the map 90 | `(e, f) ↦ g₁ e + g₂ f`. -/ 91 | def ContinuousLinearMap.coprodL : (E →L[𝕜] G) × (F →L[𝕜] G) →L[𝕜] E × F →L[𝕜] G := 92 | (isBoundedLinearMap_coprod 𝕜 E F G).toContinuousLinearMap 93 | 94 | @[continuity, fun_prop] 95 | theorem Continuous.coprodL {f : X → E →L[𝕜] G} {g : X → F →L[𝕜] G} (hf : Continuous f) 96 | (hg : Continuous g) : Continuous fun x ↦ (f x).coprod (g x) := 97 | ContinuousLinearMap.coprodL.continuous.comp₂ hf hg 98 | 99 | theorem Continuous.prodL' {𝕜 : Type*} {E : Type*} {Fₗ : Type*} {Gₗ : Type*} 100 | [SeminormedAddCommGroup E] [SeminormedAddCommGroup Fₗ] [SeminormedAddCommGroup Gₗ] 101 | [NontriviallyNormedField 𝕜] [NormedSpace 𝕜 E] [NormedSpace 𝕜 Fₗ] [NormedSpace 𝕜 Gₗ] 102 | {X : Type*} [TopologicalSpace X] 103 | {f : X → E →L[𝕜] Fₗ} {g : X → E →L[𝕜] Gₗ} (hf : Continuous f) (hg : Continuous g) : 104 | Continuous fun x ↦ (f x).prod (g x) := 105 | (ContinuousLinearMap.prodₗᵢ 𝕜).continuous.comp₂ hf hg 106 | 107 | @[continuity, fun_prop] 108 | theorem Continuous.prodL {𝕜 : Type*} {E : Type*} {Fₗ : Type*} {Gₗ : Type*} 109 | [SeminormedAddCommGroup E] [SeminormedAddCommGroup Fₗ] [SeminormedAddCommGroup Gₗ] 110 | [NontriviallyNormedField 𝕜] [NormedSpace 𝕜 E] [NormedSpace 𝕜 Fₗ] [NormedSpace 𝕜 Gₗ] {X : Type*} 111 | [TopologicalSpace X] {f : X → E →L[𝕜] Fₗ} {g : X → E →L[𝕜] Gₗ} (hf : Continuous f) 112 | (hg : Continuous g) : Continuous fun x ↦ (f x).prod (g x) := 113 | hf.prodL' hg 114 | 115 | @[continuity, fun_prop] 116 | theorem ContinuousAt.compL {f : X → Fₗ →L[𝕜] Gₗ} {g : X → E →L[𝕜] Fₗ} {x₀ : X} 117 | (hf : ContinuousAt f x₀) (hg : ContinuousAt g x₀) : 118 | ContinuousAt (fun x ↦ (f x).comp (g x)) x₀ := 119 | ((ContinuousLinearMap.compL 𝕜 E Fₗ Gₗ).continuous₂.tendsto (f x₀, g x₀)).comp 120 | (hf.prodMk_nhds hg) 121 | 122 | @[continuity, fun_prop] 123 | theorem Continuous.compL {f : X → Fₗ →L[𝕜] Gₗ} {g : X → E →L[𝕜] Fₗ} (hf : Continuous f) 124 | (hg : Continuous g) : Continuous fun x ↦ (f x).comp (g x) := 125 | (ContinuousLinearMap.compL 𝕜 E Fₗ Gₗ).continuous₂.comp₂ hf hg 126 | -------------------------------------------------------------------------------- /SphereEversion/ToMathlib/Data/Nat/Basic.lean: -------------------------------------------------------------------------------- 1 | import Mathlib.Logic.Function.Basic 2 | import Mathlib.Tactic.Choose 3 | 4 | -- The next lemma won't be used, it's a warming up exercise for the one below. 5 | -- It could go to mathlib. 6 | theorem exists_by_induction {α : Type*} {P : ℕ → α → Prop} (h₀ : ∃ a, P 0 a) 7 | (ih : ∀ n a, P n a → ∃ a', P (n + 1) a') : ∃ f : ℕ → α, ∀ n, P n (f n) := by 8 | choose f₀ hf₀ using h₀ 9 | choose! F hF using ih 10 | exact ⟨fun n ↦ Nat.recOn n f₀ F, fun n ↦ Nat.rec hf₀ (fun n ih ↦ hF n _ ih) n⟩ 11 | 12 | -- We make `P` and `Q` explicit to help the elaborator when applying the lemma 13 | -- (elab_as_eliminator isn't enough). 14 | theorem exists_by_induction' {α : Type*} (P : ℕ → α → Prop) (Q : ℕ → α → α → Prop) 15 | (h₀ : ∃ a, P 0 a) (ih : ∀ n a, P n a → ∃ a', P (n + 1) a' ∧ Q n a a') : 16 | ∃ f : ℕ → α, ∀ n, P n (f n) ∧ Q n (f n) (f <| n + 1) := by 17 | choose f₀ hf₀ using h₀ 18 | choose! F hF hF' using ih 19 | have key : ∀ n, P n (Nat.recOn n f₀ F) := fun n ↦ Nat.rec hf₀ (fun n ih ↦ hF n _ ih) n 20 | exact ⟨fun n ↦ Nat.recOn n f₀ F, fun n ↦ ⟨key n, hF' n _ (key n)⟩⟩ 21 | -------------------------------------------------------------------------------- /SphereEversion/ToMathlib/Geometry/Manifold/Immersion.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2024 Michael Rothgang. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Michael Rothgang 5 | -/ 6 | import Mathlib.Geometry.Manifold.ContMDiff.Defs 7 | import Mathlib.Geometry.Manifold.MFDeriv.Defs 8 | 9 | /-! ## Smooth immersions 10 | 11 | In this file, we define immersions and prove some of their basic properties. 12 | 13 | ## Main definitions 14 | * `Immersion I I' f n`: a `C^n` map `f : M → M'` is an immersion iff 15 | its `mfderiv` is injective at each point 16 | * `InjImmersion`: an immersion which is also injective 17 | 18 | ## Main results 19 | 20 | 21 | ## Implementation notes 22 | The initial design of immersions only assumed injectivity of the differential. 23 | This is not sufficient to ensure immersions are `C^n`: 24 | While mathlib's `mfderiv` has junk value zero when `f` is not differentiable and the zero map is 25 | only injective if `M` has dimension zero (in which case `f` is always `C^n`), this argument only 26 | shows immersions are `MDifferentiable`, not `C^n`. 27 | 28 | **NOTE.** This is **not** the correct definition in the infinite-dimensional context, 29 | but in finite dimensions, the general definition is equivalent to the one in this file. 30 | 31 | ## Tags 32 | manifold, immersion 33 | 34 | -/ 35 | noncomputable section 36 | 37 | open Set Function 38 | 39 | open scoped Manifold ContDiff 40 | 41 | -- Let `M` be a manifold with corners over the pair `(E, H)`. 42 | -- Let `M'` be a manifold with corners over the pair `(E', H')`. 43 | variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] 44 | {E : Type*} [instE: NormedAddCommGroup E] [instE': NormedSpace 𝕜 E] 45 | {H : Type*} [TopologicalSpace H] (I : ModelWithCorners 𝕜 E H) 46 | {M : Type*} [TopologicalSpace M] [ChartedSpace H M] 47 | {E' : Type*} [NormedAddCommGroup E'] [NormedSpace 𝕜 E'] 48 | {H' : Type*} [TopologicalSpace H'] (I' : ModelWithCorners 𝕜 E' H') 49 | {M' : Type*} [TopologicalSpace M'] [ChartedSpace H' M'] 50 | 51 | section Definition 52 | 53 | /-- A `C^n` immersion `f : M → M` is a `C^n` map whose differential is injective at every point. -/ 54 | structure Immersion (f : M → M') (n : WithTop ℕ∞) : Prop where 55 | contMDiff : ContMDiff I I' n f 56 | diff_injective : ∀ p, Injective (mfderiv I I' f p) 57 | 58 | /-- An injective `C^n` immersion -/ 59 | structure InjImmersion (f : M → M') (n : WithTop ℕ∞) : Prop extends Immersion I I' f n where 60 | injective : Injective f 61 | 62 | end Definition 63 | -------------------------------------------------------------------------------- /SphereEversion/ToMathlib/Geometry/Manifold/IsManifold/ExtChartAt.lean: -------------------------------------------------------------------------------- 1 | import Mathlib.Geometry.Manifold.IsManifold.ExtChartAt 2 | 3 | open scoped Topology 4 | 5 | open Metric hiding mem_nhds_iff ball 6 | 7 | open Set 8 | 9 | section 10 | 11 | variable {𝕜 E M H : Type*} [NontriviallyNormedField 𝕜] 12 | [TopologicalSpace H] [TopologicalSpace M] [ChartedSpace H M] 13 | [NormedAddCommGroup E] [NormedSpace 𝕜 E] (I : ModelWithCorners 𝕜 E H) 14 | 15 | namespace ChartedSpace 16 | 17 | /-- If `M` is a `ChartedSpace` we can use the preferred chart at any point to transfer a 18 | ball in coordinate space into a set in `M`. These can be a useful neighbourhood basis. -/ 19 | def ball (x : M) (r : ℝ) := 20 | (extChartAt I x).symm '' Metric.ball (extChartAt I x x) r 21 | 22 | theorem nhds_hasBasis_balls_of_open_cov [I.Boundaryless] (x : M) {ι : Type*} {s : ι → Set M} 23 | (s_op : ∀ j, IsOpen <| s j) (cov : (⋃ j, s j) = univ) : 24 | (𝓝 x).HasBasis (fun r ↦ 0 < r ∧ Metric.ball (extChartAt I x x) r ⊆ (extChartAt I x).target ∧ 25 | ∃ j, ChartedSpace.ball I x r ⊆ s j) 26 | (ChartedSpace.ball I x) := by 27 | -- TODO golf etc 28 | obtain ⟨j, hj⟩ : ∃ j, x ∈ s j := by simpa only [mem_iUnion, ← cov] using mem_univ x 29 | replace hj : s j ∈ 𝓝 x := mem_nhds_iff.mpr ⟨s j, Subset.rfl, s_op j, hj⟩ 30 | have hx : (extChartAt I x).source ∈ 𝓝 x := extChartAt_source_mem_nhds x 31 | refine Filter.hasBasis_iff.mpr fun n ↦ ⟨fun hn ↦ ?_, ?_⟩ 32 | · let m := s j ∩ n ∩ (extChartAt I x).source 33 | have hm : m ∈ 𝓝 x := Filter.inter_mem (Filter.inter_mem hj hn) hx 34 | replace hm : extChartAt I x '' m ∈ 𝓝 (extChartAt I x x) := 35 | extChartAt_image_nhds_mem_nhds_of_boundaryless hm 36 | obtain ⟨r, hr₀, hr₁⟩ := 37 | (Filter.hasBasis_iff.mp (@nhds_basis_ball E _ (extChartAt I x x)) _).mp hm 38 | refine ⟨r, ⟨hr₀, hr₁.trans ?_, ⟨j, ?_⟩⟩, ?_⟩ 39 | · exact ((extChartAt I x).mapsTo.mono inter_subset_right Subset.rfl).image_subset 40 | · suffices m ⊆ s j by 41 | refine Subset.trans ?_ this 42 | convert monotone_image (f := (extChartAt I x).symm) hr₁ 43 | exact (PartialEquiv.symm_image_image_of_subset_source _ inter_subset_right).symm 44 | exact inter_subset_left.trans inter_subset_left 45 | · suffices m ⊆ n by 46 | refine Subset.trans ?_ this 47 | convert monotone_image (f := (extChartAt I x).symm) hr₁ 48 | exact (PartialEquiv.symm_image_image_of_subset_source _ inter_subset_right).symm 49 | exact inter_subset_left.trans inter_subset_right 50 | · rintro ⟨r, ⟨hr₀, hr₁, -⟩, hr₂⟩ 51 | replace hr₀ : Metric.ball (extChartAt I x x) r ∈ 𝓝 (extChartAt I x x) := ball_mem_nhds _ hr₀ 52 | rw [← map_extChartAt_nhds_of_boundaryless, Filter.mem_map] at hr₀ 53 | replace hr₀ := Filter.inter_mem hx hr₀ 54 | rw [← (extChartAt I x).symm_image_eq_source_inter_preimage hr₁] at hr₀ 55 | filter_upwards [hr₀] using hr₂ 56 | 57 | end ChartedSpace 58 | 59 | end 60 | -------------------------------------------------------------------------------- /SphereEversion/ToMathlib/Geometry/Manifold/Metrizable.lean: -------------------------------------------------------------------------------- 1 | import Mathlib.Geometry.Manifold.Metrizable 2 | 3 | noncomputable section 4 | 5 | section 6 | 7 | variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] [FiniteDimensional ℝ E] {H : Type*} 8 | [TopologicalSpace H] (I : ModelWithCorners ℝ E H) (M : Type*) [TopologicalSpace M] 9 | [ChartedSpace H M] 10 | 11 | /-- A metric defining the topology on a σ-compact T₂ real manifold. -/ 12 | def manifoldMetric [T2Space M] [SigmaCompactSpace M] : MetricSpace M := 13 | TopologicalSpace.metrizableSpaceMetric (h := Manifold.metrizableSpace I M) 14 | 15 | end 16 | -------------------------------------------------------------------------------- /SphereEversion/ToMathlib/Geometry/Manifold/MiscManifold.lean: -------------------------------------------------------------------------------- 1 | import Mathlib.Geometry.Manifold.ContMDiff.Constructions 2 | 3 | open Set Function Filter 4 | 5 | open scoped Manifold Topology ContDiff 6 | 7 | noncomputable section 8 | 9 | section IsManifold 10 | 11 | open IsManifold 12 | 13 | variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] 14 | {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] 15 | {E' : Type*} [NormedAddCommGroup E'] [NormedSpace 𝕜 E'] 16 | {F : Type*} [NormedAddCommGroup F] [NormedSpace 𝕜 F] 17 | {F' : Type*} [NormedAddCommGroup F'] [NormedSpace 𝕜 F'] 18 | {H : Type*} [TopologicalSpace H] {I : ModelWithCorners 𝕜 E H} 19 | {H' : Type*} [TopologicalSpace H'] {I' : ModelWithCorners 𝕜 E' H'} 20 | {G : Type*} [TopologicalSpace G] {J : ModelWithCorners 𝕜 F G} 21 | {G' : Type*} [TopologicalSpace G'] {J' : ModelWithCorners 𝕜 F' G'} 22 | {M : Type*} [TopologicalSpace M] [ChartedSpace H M] 23 | {M' : Type*} [TopologicalSpace M'] [ChartedSpace H' M'] 24 | {N : Type*} [TopologicalSpace N] [ChartedSpace G N] 25 | {N' : Type*} [TopologicalSpace N'] [ChartedSpace G' N'] 26 | {F'' : Type*} [NormedAddCommGroup F''] [NormedSpace 𝕜 F''] 27 | {E'' : Type*} [NormedAddCommGroup E''] [NormedSpace 𝕜 E''] 28 | {H'' : Type*} [TopologicalSpace H''] {I'' : ModelWithCorners 𝕜 E'' H''} 29 | {M'' : Type*} [TopologicalSpace M''] [ChartedSpace H'' M''] 30 | {e : PartialHomeomorph M H} {f : M → M'} {m n : WithTop ℕ∞} {s : Set M} {x x' : M} 31 | 32 | theorem contMDiff_prod {f : M → M' × N'} : 33 | ContMDiff I (I'.prod J') n f ↔ 34 | (ContMDiff I I' n fun x ↦ (f x).1) ∧ ContMDiff I J' n fun x ↦ (f x).2 := 35 | ⟨fun h ↦ ⟨h.fst, h.snd⟩, fun h ↦ h.1.prodMk h.2⟩ 36 | 37 | theorem contMDiffAt_prod {f : M → M' × N'} {x : M} : 38 | ContMDiffAt I (I'.prod J') n f x ↔ 39 | ContMDiffAt I I' n (fun x ↦ (f x).1) x ∧ ContMDiffAt I J' n (fun x ↦ (f x).2) x := 40 | ⟨fun h ↦ ⟨h.fst, h.snd⟩, fun h ↦ h.1.prodMk h.2⟩ 41 | 42 | theorem smooth_prod {f : M → M' × N'} : 43 | ContMDiff I (I'.prod J') ∞ f ↔ 44 | (ContMDiff I I' ∞ fun x ↦ (f x).1) ∧ ContMDiff I J' ∞ fun x ↦ (f x).2 := 45 | contMDiff_prod 46 | 47 | theorem smoothAt_prod {f : M → M' × N'} {x : M} : 48 | ContMDiffAt I (I'.prod J') ∞ f x ↔ 49 | ContMDiffAt I I' ∞ (fun x ↦ (f x).1) x ∧ ContMDiffAt I J' ∞ (fun x ↦ (f x).2) x := 50 | contMDiffAt_prod 51 | 52 | end IsManifold 53 | -------------------------------------------------------------------------------- /SphereEversion/ToMathlib/Geometry/Manifold/VectorBundle/Misc.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2022 Floris van Doorn. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Floris van Doorn 5 | 6 | ! This file was ported from Lean 3 source module to_mathlib.geometry.manifold.vector_bundle.misc 7 | -/ 8 | import Mathlib.Geometry.Manifold.VectorBundle.Basic 9 | import Mathlib.Topology.VectorBundle.Hom 10 | 11 | /-! 12 | # Various operations on and properties of smooth vector bundles 13 | -/ 14 | 15 | noncomputable section 16 | 17 | open Bundle Set 18 | 19 | namespace FiberBundle 20 | 21 | variable {𝕜 B B' F M : Type*} {E : B → Type*} 22 | 23 | variable [TopologicalSpace F] [TopologicalSpace (TotalSpace F E)] [∀ x, TopologicalSpace (E x)] 24 | {HB : Type*} [TopologicalSpace HB] [TopologicalSpace B] [ChartedSpace HB B] [FiberBundle F E] 25 | 26 | theorem chartedSpace_chartAt_fst' (x y : TotalSpace F E) : 27 | (chartAt (ModelProd HB F) x y).1 = chartAt HB x.proj (trivializationAt F E x.proj y).1 := by 28 | rw [chartedSpace_chartAt]; rfl 29 | 30 | theorem chartedSpace_chartAt_fst {x y : TotalSpace F E} 31 | (hy : y.proj ∈ (trivializationAt F E x.proj).baseSet) : 32 | (chartAt (ModelProd HB F) x y).1 = chartAt HB x.proj y.proj := by 33 | rw [chartedSpace_chartAt_fst', (trivializationAt F E x.proj).coe_fst' hy] 34 | 35 | theorem chartedSpace_chartAt_snd (x y : TotalSpace F E) : 36 | (chartAt (ModelProd HB F) x y).2 = (trivializationAt F E x.proj y).2 := by 37 | rw [chartedSpace_chartAt]; rfl 38 | 39 | end FiberBundle 40 | 41 | section VectorBundle 42 | 43 | variable {𝕜 B F F₁ F₂ : Type*} {E : B → Type*} {E₁ : B → Type*} {E₂ : B → Type*} 44 | [NontriviallyNormedField 𝕜] [∀ x, AddCommMonoid (E x)] [∀ x, Module 𝕜 (E x)] 45 | [NormedAddCommGroup F] [NormedSpace 𝕜 F] [TopologicalSpace (TotalSpace F E)] 46 | [∀ x, TopologicalSpace (E x)] [∀ x, AddCommMonoid (E₁ x)] [∀ x, Module 𝕜 (E₁ x)] 47 | [NormedAddCommGroup F₁] [NormedSpace 𝕜 F₁] [TopologicalSpace (TotalSpace F₁ E₁)] 48 | [∀ x, TopologicalSpace (E₁ x)] [∀ x, AddCommMonoid (E₂ x)] [∀ x, Module 𝕜 (E₂ x)] 49 | [NormedAddCommGroup F₂] [NormedSpace 𝕜 F₂] [TopologicalSpace (TotalSpace F₂ E₂)] 50 | [∀ x, TopologicalSpace (E₂ x)] [TopologicalSpace B] {n : ℕ∞} [FiberBundle F₁ E₁] 51 | [VectorBundle 𝕜 F₁ E₁] [FiberBundle F₂ E₂] [VectorBundle 𝕜 F₂ E₂] 52 | {e₁ e₁' : Trivialization F₁ (π F₁ E₁)} {e₂ e₂' : Trivialization F₂ (π F₂ E₂)} 53 | 54 | end VectorBundle 55 | 56 | namespace VectorBundleCore 57 | 58 | variable {R 𝕜 B F ι : Type*} [NontriviallyNormedField R] [NormedAddCommGroup F] [NormedSpace R F] 59 | [TopologicalSpace B] (Z : VectorBundleCore R B F ι) 60 | 61 | /-- `Z.coord_change j i` is a partial inverse of `Z.coord_change i j`. -/ 62 | theorem coordChange_comp_eq_self {i j : ι} {x : B} (hx : x ∈ Z.baseSet i ∩ Z.baseSet j) (v : F) : 63 | Z.coordChange j i x (Z.coordChange i j x v) = v := by 64 | rw [Z.coordChange_comp i j i x ⟨hx, hx.1⟩, Z.coordChange_self i x hx.1] 65 | 66 | end VectorBundleCore 67 | 68 | namespace Bundle.Trivial 69 | 70 | variable {𝕜 B F : Type*} 71 | 72 | variable [NontriviallyNormedField 𝕜] [NormedAddCommGroup F] [NormedSpace 𝕜 F] [TopologicalSpace B] 73 | 74 | @[simp, mfld_simps] 75 | protected theorem trivializationAt (x : B) : 76 | trivializationAt F (Trivial B F) x = Trivial.trivialization B F := 77 | rfl 78 | 79 | @[simp, mfld_simps] 80 | theorem trivialization_continuousLinearMapAt (x : B) : 81 | (Trivial.trivialization B F).continuousLinearMapAt 𝕜 x = ContinuousLinearMap.id 𝕜 F := by 82 | ext v 83 | simp_rw [Trivialization.continuousLinearMapAt_apply, Trivialization.coe_linearMapAt] 84 | rw [if_pos] 85 | exacts [rfl, mem_univ _] 86 | 87 | end Bundle.Trivial 88 | 89 | section Hom 90 | 91 | variable {𝕜₁ : Type*} [NontriviallyNormedField 𝕜₁] {𝕜₂ : Type*} [NontriviallyNormedField 𝕜₂] 92 | (σ : 𝕜₁ →+* 𝕜₂) 93 | 94 | variable {B : Type*} [TopologicalSpace B] 95 | 96 | variable (F₁ : Type*) [NormedAddCommGroup F₁] [NormedSpace 𝕜₁ F₁] (E₁ : B → Type*) 97 | [∀ x, AddCommGroup (E₁ x)] [∀ x, Module 𝕜₁ (E₁ x)] [TopologicalSpace (TotalSpace F₁ E₁)] 98 | 99 | variable (F₂ : Type*) [NormedAddCommGroup F₂] [NormedSpace 𝕜₂ F₂] (E₂ : B → Type*) 100 | [∀ x, AddCommGroup (E₂ x)] [∀ x, Module 𝕜₂ (E₂ x)] [TopologicalSpace (TotalSpace F₂ E₂)] 101 | 102 | variable [RingHomIsometric σ] 103 | 104 | variable [∀ x : B, TopologicalSpace (E₁ x)] [FiberBundle F₁ E₁] [VectorBundle 𝕜₁ F₁ E₁] 105 | 106 | variable [∀ x : B, TopologicalSpace (E₂ x)] [FiberBundle F₂ E₂] [VectorBundle 𝕜₂ F₂ E₂] 107 | 108 | variable [∀ x, IsTopologicalAddGroup (E₂ x)] [∀ x, ContinuousSMul 𝕜₂ (E₂ x)] 109 | 110 | @[simp, mfld_simps] 111 | theorem continuousLinearMap_trivializationAt (x : B) : 112 | trivializationAt (F₁ →SL[σ] F₂) (Bundle.ContinuousLinearMap σ E₁ E₂) x = 113 | (trivializationAt F₁ E₁ x).continuousLinearMap σ (trivializationAt F₂ E₂ x) := 114 | rfl 115 | 116 | end Hom 117 | 118 | section Pullback 119 | 120 | /-- We need some instances like this to work with negation on pullbacks -/ 121 | instance {B B'} {E : B → Type*} {f : B' → B} {x : B'} [∀ x', AddCommGroup (E x')] : 122 | AddCommGroup ((f *ᵖ E) x) := by delta Bundle.Pullback; infer_instance 123 | 124 | instance {B B'} {E : B → Type*} {f : B' → B} {x : B'} [∀ x', Zero (E x')] : Zero ((f *ᵖ E) x) := by 125 | delta Bundle.Pullback; infer_instance 126 | 127 | variable {B F B' K : Type*} {E : B → Type*} {f : K} [TopologicalSpace B'] 128 | [TopologicalSpace (TotalSpace F E)] [TopologicalSpace F] [TopologicalSpace B] [∀ b, Zero (E b)] 129 | [FunLike K B' B] [ContinuousMapClass K B' B] 130 | 131 | namespace Trivialization 132 | 133 | -- attribute [simps base_set] trivialization.pullback 134 | theorem pullback_symm (e : Trivialization F (π F E)) (x : B') : 135 | (e.pullback f).symm x = e.symm (f x) := by 136 | ext y 137 | simp_rw [Trivialization.symm, Pretrivialization.symm] 138 | congr; ext (hx : f x ∈ e.toPretrivialization.baseSet) 139 | change cast _ (e.symm (f x) y) = cast _ (e.toPartialHomeomorph.symm (f x, y)).2 140 | simp_rw [Trivialization.symm, Pretrivialization.symm, dif_pos hx, cast_cast] 141 | rfl 142 | 143 | end Trivialization 144 | 145 | variable [∀ x, TopologicalSpace (E x)] [FiberBundle F E] 146 | 147 | theorem pullback_trivializationAt {x : B'} : 148 | trivializationAt F (f *ᵖ E) x = (trivializationAt F E (f x)).pullback f := 149 | rfl 150 | 151 | end Pullback 152 | 153 | section PullbackVb 154 | 155 | variable {R 𝕜 B F B' : Type*} {E : B → Type*} 156 | 157 | variable [TopologicalSpace B'] [TopologicalSpace (TotalSpace F E)] [NontriviallyNormedField 𝕜] 158 | [NormedAddCommGroup F] [NormedSpace 𝕜 F] [TopologicalSpace B] [∀ x, AddCommMonoid (E x)] 159 | [∀ x, Module 𝕜 (E x)] [∀ x, TopologicalSpace (E x)] [FiberBundle F E] {K : Type*} 160 | [FunLike K B' B] [ContinuousMapClass K B' B] (f : K) 161 | 162 | namespace Trivialization 163 | 164 | theorem pullback_symmL (e : Trivialization F (π F E)) [e.IsLinear 𝕜] (x : B') : 165 | (e.pullback f).symmL 𝕜 x = e.symmL 𝕜 (f x) := by 166 | ext y 167 | simp only [symmL_apply, pullback_symm] 168 | rfl 169 | 170 | end Trivialization 171 | 172 | end PullbackVb 173 | -------------------------------------------------------------------------------- /SphereEversion/ToMathlib/LinearAlgebra/Basic.lean: -------------------------------------------------------------------------------- 1 | import Mathlib.Algebra.Module.Submodule.Ker 2 | import Mathlib.LinearAlgebra.Span.Defs 3 | 4 | /-! Note: some results should go to `LinearAlgebra.Span`. -/ 5 | 6 | 7 | open Submodule Function 8 | 9 | section 10 | 11 | variable {R : Type*} {R₂ : Type*} {M : Type*} {M₂ : Type*} [Semiring R] [Semiring R₂] 12 | [AddCommMonoid M] [AddCommMonoid M₂] {σ₁₂ : R →+* R₂} [Module R M] [Module R₂ M₂] 13 | 14 | theorem Submodule.sup_eq_span_union (s t : Submodule R M) : s ⊔ t = span R (s ∪ t) := by 15 | rw [span_union, span_eq s, span_eq t] 16 | 17 | end 18 | 19 | section 20 | 21 | variable {R : Type*} {R₂ : Type*} {M : Type*} {M₂ : Type*} [Ring R] [Ring R₂] [AddCommGroup M] 22 | [AddCommGroup M₂] {σ₁₂ : R →+* R₂} [Module R M] [Module R₂ M₂] {p q : Submodule R M} 23 | [RingHomSurjective σ₁₂] 24 | 25 | theorem LinearMap.injective_iff_of_direct_sum (f : M →ₛₗ[σ₁₂] M₂) (hpq : p ⊓ q = ⊥) 26 | (hpq' : p ⊔ q = ⊤) : 27 | Injective f ↔ Disjoint p (LinearMap.ker f) ∧ Disjoint q (LinearMap.ker f) ∧ 28 | Disjoint (map f p) (map f q) := by 29 | constructor 30 | · intro h 31 | simp [disjoint_iff_inf_le, LinearMap.ker_eq_bot.mpr h, ← Submodule.map_inf _ h, hpq] 32 | · rintro ⟨hp, hq, h⟩ 33 | rw [LinearMap.disjoint_ker] at * 34 | rw [← LinearMap.ker_eq_bot, ← @inf_top_eq _ _ _ (LinearMap.ker f), ← hpq'] 35 | rw [← le_bot_iff] 36 | rintro x ⟨hx, hx' : x ∈ p ⊔ q⟩ 37 | rcases mem_sup.mp hx' with ⟨u, hu, v, hv, rfl⟩ 38 | rw [mem_bot] 39 | erw [← sub_neg_eq_add, LinearMap.sub_mem_ker_iff] at hx 40 | have hu' : f u ∈ map f p := mem_map_of_mem hu 41 | have hv' : f (-v) ∈ map f q := mem_map_of_mem (q.neg_mem hv) 42 | rw [← hx] at hv' 43 | have H : f u ∈ map f p ⊓ map f q := mem_inf.mpr ⟨hu', hv'⟩ 44 | rw [disjoint_iff_inf_le] at h 45 | rw [hp u hu (h H), zero_add] 46 | rw [hp u hu (h H), f.map_zero, f.map_neg, eq_comm, neg_eq_zero] at hx 47 | exact hq v hv hx 48 | 49 | end 50 | 51 | theorem LinearMap.ker_inf_eq_bot {R : Type*} {R₂ : Type*} {M : Type*} {M₂ : Type*} [Ring R] 52 | [Ring R₂] [AddCommGroup M] [AddCommGroup M₂] [Module R M] [Module R₂ M₂] {τ₁₂ : R →+* R₂} 53 | {f : M →ₛₗ[τ₁₂] M₂} {S : Submodule R M} : LinearMap.ker f ⊓ S = ⊥ ↔ Set.InjOn f S := by 54 | rw [Set.injOn_iff_injective, inf_comm, ← disjoint_iff, LinearMap.disjoint_ker'] 55 | constructor 56 | · intro h x y hxy 57 | exact Subtype.coe_injective (h x x.prop y y.prop hxy) 58 | · intro h x hx y hy hxy 59 | have : (S : Set M).restrict f ⟨x, hx⟩ = (S : Set M).restrict f ⟨y, hy⟩ := hxy 60 | cases h this 61 | rfl 62 | -------------------------------------------------------------------------------- /SphereEversion/ToMathlib/LinearAlgebra/FiniteDimensional.lean: -------------------------------------------------------------------------------- 1 | import Mathlib.LinearAlgebra.FiniteDimensional.Lemmas 2 | 3 | open Module Submodule 4 | 5 | variable {𝕜 : Type*} [Field 𝕜] {E : Type*} [AddCommGroup E] [Module 𝕜 E] {E' : Type*} 6 | [AddCommGroup E'] [Module 𝕜 E'] 7 | 8 | theorem one_lt_rank_of_rank_lt_rank [FiniteDimensional 𝕜 E] [FiniteDimensional 𝕜 E'] {π : E →ₗ[𝕜] 𝕜} 9 | (hπ : LinearMap.ker π ≠ ⊤) (h : finrank 𝕜 E < finrank 𝕜 E') (φ : E →ₗ[𝕜] E') : 10 | 1 < Module.rank 𝕜 (E' ⧸ Submodule.map φ (LinearMap.ker π)) := by 11 | suffices 2 ≤ finrank 𝕜 (E' ⧸ (LinearMap.ker π).map φ) by 12 | rw [← finrank_eq_rank] 13 | exact_mod_cast this 14 | apply le_of_add_le_add_right 15 | rw [Submodule.finrank_quotient_add_finrank ((LinearMap.ker π).map φ)] 16 | have := calc finrank 𝕜 ((LinearMap.ker π).map φ) 17 | _ ≤ finrank 𝕜 (LinearMap.ker π) := finrank_map_le φ (LinearMap.ker π) 18 | _ < finrank 𝕜 E := Submodule.finrank_lt hπ 19 | linarith 20 | -------------------------------------------------------------------------------- /SphereEversion/ToMathlib/MeasureTheory/BorelSpace.lean: -------------------------------------------------------------------------------- 1 | import Mathlib.MeasureTheory.Constructions.BorelSpace.Basic 2 | 3 | variable (X : Type*) [TopologicalSpace X] 4 | 5 | scoped[Borelize] attribute [instance] borel 6 | 7 | theorem borelSpace_borel : @BorelSpace X _ (borel X) := 8 | letI := borel X; ⟨rfl⟩ 9 | 10 | scoped[Borelize] attribute [instance] borelSpace_borel 11 | -------------------------------------------------------------------------------- /SphereEversion/ToMathlib/Order/Filter/Basic.lean: -------------------------------------------------------------------------------- 1 | import Mathlib.Order.Filter.Basic 2 | 3 | theorem Filter.EventuallyEq.eventuallyEq_ite {X Y : Type*} {l : Filter X} {f g : X → Y} 4 | {P : X → Prop} [DecidablePred P] (h : f =ᶠ[l] g) : 5 | (fun x ↦ if P x then f x else g x) =ᶠ[l] f := 6 | h.mono fun x hx ↦ by simp [hx] 7 | -------------------------------------------------------------------------------- /SphereEversion/ToMathlib/Partition.lean: -------------------------------------------------------------------------------- 1 | import Mathlib.Geometry.Manifold.PartitionOfUnity 2 | import SphereEversion.ToMathlib.Analysis.Convex.Basic 3 | import SphereEversion.ToMathlib.Geometry.Manifold.Algebra.SmoothGerm 4 | 5 | noncomputable section 6 | 7 | open scoped Topology 8 | open Set Function Filter 9 | 10 | variable {ι : Type*} 11 | 12 | variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] 13 | {H : Type*} [TopologicalSpace H] {I : ModelWithCorners ℝ E H} 14 | {M : Type*} [TopologicalSpace M] [ChartedSpace H M] 15 | 16 | section 17 | 18 | variable {F : Type*} [AddCommGroup F] [Module ℝ F] 19 | 20 | namespace SmoothPartitionOfUnity 21 | 22 | variable {s : Set M} (ρ : SmoothPartitionOfUnity ι I M s) {x : M} 23 | 24 | theorem sum_germ (hx : x ∈ interior s := by simp) : 25 | ∑ i ∈ ρ.fintsupport x, (ρ i : smoothGerm I x) = 1 := by 26 | have : ∀ᶠ y in 𝓝 x, y ∈ interior s := isOpen_interior.eventually_mem hx 27 | have : ∀ᶠ y in 𝓝 x, (⇑(∑ i ∈ ρ.fintsupport x, ρ i)) y = 1 := by 28 | filter_upwards [ρ.eventually_finsupport_subset x, this] with y hy hy' 29 | rw [SmoothMap.coe_sum, Finset.sum_apply] 30 | apply ρ.toPartitionOfUnity.sum_finsupport' (interior_subset hy') hy 31 | rw [← smoothGerm.coe_sum] 32 | exact smoothGerm.coe_eq_coe _ _ 1 this 33 | 34 | def combine (ρ : SmoothPartitionOfUnity ι I M s) (φ : ι → M → F) (x : M) : F := 35 | ∑ᶠ i, ρ i x • φ i x 36 | 37 | theorem germ_combine_mem (φ : ι → M → F) (hx : x ∈ interior s := by simp) : 38 | (ρ.combine φ : Germ (𝓝 x) F) ∈ 39 | reallyConvexHull (smoothGerm I x) ((fun i ↦ (φ i : Germ (𝓝 x) F)) '' ρ.fintsupport x) := by 40 | change x ∈ interior s at hx 41 | have : (ρ.combine φ : Germ (𝓝 x) F) = 42 | ∑ i ∈ ρ.fintsupport x, (ρ i : smoothGerm I x) • (φ i : Germ (𝓝 x) F) := by 43 | suffices (ρ.combine φ : Germ (𝓝 x) F) = 44 | ↑(∑ i ∈ ρ.fintsupport x, ((ρ i : M → ℝ) • φ i : M → F)) by 45 | rw [this, Germ.coe_sum]; rfl 46 | rw [Germ.coe_eq] 47 | filter_upwards [ρ.eventually_finsupport_subset x] with x' hx' 48 | simp_rw [SmoothPartitionOfUnity.combine, Finset.sum_apply, Pi.smul_apply'] 49 | rw [finsum_eq_sum_of_support_subset] 50 | refine Subset.trans ?_ (Finset.coe_subset.mpr hx') 51 | rw [SmoothPartitionOfUnity.finsupport, PartitionOfUnity.finsupport, Finite.coe_toFinset] 52 | apply support_smul_subset_left 53 | rw [this] 54 | refine sum_mem_reallyConvexHull ?_ (ρ.sum_germ hx) (fun i hi ↦ mem_image_of_mem _ hi) 55 | · intro i _ 56 | filter_upwards using fun a ↦ ρ.nonneg i a 57 | 58 | end SmoothPartitionOfUnity 59 | 60 | end 61 | -------------------------------------------------------------------------------- /SphereEversion/ToMathlib/SmoothBarycentric.lean: -------------------------------------------------------------------------------- 1 | import Mathlib.Analysis.Calculus.AddTorsor.Coord 2 | import Mathlib.Analysis.Matrix 3 | import Mathlib.LinearAlgebra.AffineSpace.Matrix 4 | 5 | noncomputable section 6 | 7 | open Set Function 8 | 9 | open scoped Affine Matrix ContDiff 10 | 11 | section BarycentricDet 12 | 13 | variable (ι R k P : Type*) {M : Type*} [Ring R] [AddCommGroup M] [Module R M] [AffineSpace M P] 14 | 15 | /-- The set of affine bases for an affine space. -/ 16 | def affineBases : Set (ι → P) := 17 | {v | AffineIndependent R v ∧ affineSpan R (range v) = ⊤} 18 | 19 | theorem affineBases_findim [Fintype ι] [Field k] [Module k M] [FiniteDimensional k M] 20 | (h : Fintype.card ι = Module.finrank k M + 1) : 21 | affineBases ι k P = {v | AffineIndependent k v} := by 22 | ext v 23 | simp only [affineBases, mem_setOf_eq, and_iff_left_iff_imp] 24 | exact fun h_ind ↦ h_ind.affineSpan_eq_top_iff_card_eq_finrank_add_one.mpr h 25 | 26 | theorem mem_affineBases_iff [Fintype ι] [DecidableEq ι] [Nontrivial R] (b : AffineBasis ι R P) 27 | (v : ι → P) : v ∈ affineBases ι R P ↔ IsUnit (b.toMatrix v) := 28 | (b.isUnit_toMatrix_iff v).symm 29 | 30 | /-- If `P` is an affine space over the ring `R`, `v : ι → P` is an affine basis (for some indexing 31 | set `ι`) and `p : P` is a point, then `eval_barycentric_coords ι R P p v` are the barycentric 32 | coordinates of `p` with respect to the affine basis `v`. 33 | 34 | Actually for convenience `eval_barycentric_coords` is defined even when `v` is not an affine basis. 35 | In this case its value should be regarded as "junk". -/ 36 | def evalBarycentricCoords [DecidablePred (· ∈ affineBases ι R P)] (p : P) (v : ι → P) : ι → R := 37 | if h : v ∈ affineBases ι R P then ((AffineBasis.mk v h.1 h.2).coords p : ι → R) else 0 38 | 39 | @[simp] 40 | theorem evalBarycentricCoords_apply_of_mem_bases [DecidablePred (· ∈ affineBases ι R P)] (p : P) 41 | {v : ι → P} (h : v ∈ affineBases ι R P) : 42 | evalBarycentricCoords ι R P p v = (AffineBasis.mk v h.1 h.2).coords p := 43 | dif_pos h 44 | 45 | @[simp] 46 | theorem evalBarycentricCoords_apply_of_not_mem_bases [DecidablePred (· ∈ affineBases ι R P)] (p : P) 47 | {v : ι → P} (h : v ∉ affineBases ι R P) : evalBarycentricCoords ι R P p v = 0 := 48 | dif_neg h 49 | 50 | variable {ι R P} 51 | 52 | theorem evalBarycentricCoords_eq_det [Fintype ι] [DecidableEq ι] (S : Type*) [Field S] [Module S M] 53 | [∀ v, Decidable (v ∈ affineBases ι S P)] (b : AffineBasis ι S P) (p : P) (v : ι → P) : 54 | evalBarycentricCoords ι S P p v = 55 | (b.toMatrix v).det⁻¹ • (b.toMatrix v)ᵀ.cramer (b.coords p) := by 56 | ext i 57 | by_cases h : v ∈ affineBases ι S P 58 | · simp only [evalBarycentricCoords, h, dif_pos, Algebra.id.smul_eq_mul, Pi.smul_apply, 59 | AffineBasis.coords_apply] 60 | erw [← b.det_smul_coords_eq_cramer_coords ⟨v, h.1, h.2⟩ p] 61 | simp only [Pi.smul_apply, AffineBasis.coords_apply, Algebra.id.smul_eq_mul] 62 | have hu := b.isUnit_toMatrix ⟨v, h.1, h.2⟩ 63 | rw [Matrix.isUnit_iff_isUnit_det] at hu 64 | erw [← mul_assoc, ← Ring.inverse_eq_inv, Ring.inverse_mul_cancel _ hu, one_mul] 65 | · simp only [evalBarycentricCoords, h, Algebra.id.smul_eq_mul, Pi.zero_apply, inv_eq_zero, 66 | dif_neg, not_false_iff, zero_eq_mul, Pi.smul_apply] 67 | left 68 | rwa [mem_affineBases_iff ι S P b v, Matrix.isUnit_iff_isUnit_det, isUnit_iff_ne_zero, 69 | Classical.not_not] at h 70 | 71 | end BarycentricDet 72 | 73 | namespace Matrix 74 | 75 | variable (ι k : Type*) [Fintype ι] [DecidableEq ι] [NontriviallyNormedField k] 76 | 77 | attribute [instance] Matrix.normedAddCommGroup Matrix.normedSpace 78 | 79 | theorem smooth_det (m : WithTop ℕ∞) : ContDiff k m (det : Matrix ι ι k → k) := by 80 | suffices ∀ n : ℕ, ContDiff k m (det : Matrix (Fin n) (Fin n) k → k) by 81 | have h : (det : Matrix ι ι k → k) = det ∘ reindex (Fintype.equivFin ι) (Fintype.equivFin ι) := 82 | by ext; simp 83 | rw [h] 84 | apply (this (Fintype.card ι)).comp 85 | exact contDiff_pi.mpr fun i ↦ contDiff_pi.mpr fun j ↦ contDiff_apply_apply _ _ _ _ 86 | intro n 87 | induction' n with n ih 88 | · rw [coe_det_isEmpty] 89 | exact contDiff_const 90 | change ContDiff k m fun A : Matrix (Fin n.succ) (Fin n.succ) k ↦ A.det 91 | simp_rw [det_succ_column_zero] 92 | apply ContDiff.sum fun l _ ↦ ?_ 93 | apply ContDiff.mul 94 | · refine contDiff_const.mul ?_ 95 | apply contDiff_apply_apply 96 | · apply ih.comp 97 | refine contDiff_pi.mpr fun i ↦ contDiff_pi.mpr fun j ↦ ?_ 98 | simp only [submatrix_apply] 99 | apply contDiff_apply_apply 100 | 101 | end Matrix 102 | 103 | section smooth_barycentric 104 | 105 | variable (ι 𝕜 F : Type*) 106 | 107 | variable [Fintype ι] [NontriviallyNormedField 𝕜] [CompleteSpace 𝕜] 108 | 109 | variable [NormedAddCommGroup F] [NormedSpace 𝕜 F] 110 | 111 | -- An alternative approach would be to prove the affine version of `contDiffAt_map_inverse` 112 | -- and prove that barycentric coordinates give a continuous affine equivalence to 113 | -- `{ f : ι →₀ 𝕜 | f.sum = 1 }`. 114 | -- This should obviate the need for the finite-dimensionality assumption. 115 | theorem smooth_barycentric [DecidablePred (· ∈ affineBases ι 𝕜 F)] [FiniteDimensional 𝕜 F] 116 | (h : Fintype.card ι = Module.finrank 𝕜 F + 1) {n : WithTop ℕ∞} : 117 | ContDiffOn 𝕜 n (uncurry (evalBarycentricCoords ι 𝕜 F)) (@univ F ×ˢ affineBases ι 𝕜 F) := by 118 | classical 119 | obtain ⟨b⟩ : Nonempty (AffineBasis ι 𝕜 F) := AffineBasis.exists_affineBasis_of_finiteDimensional h 120 | simp_rw [uncurry_def, contDiffOn_pi, evalBarycentricCoords_eq_det 𝕜 b] 121 | intro i 122 | change ContDiffOn 𝕜 n 123 | (fun x : F × (ι → F) ↦ 124 | (Matrix.det (AffineBasis.toMatrix b x.snd))⁻¹ • 125 | (Matrix.cramer (AffineBasis.toMatrix b x.snd)ᵀ : (ι → 𝕜) → ι → 𝕜) 126 | ((AffineBasis.coords b : F → ι → 𝕜) x.1) i) 127 | (univ ×ˢ affineBases ι 𝕜 F) 128 | simp only [Pi.smul_apply, Matrix.cramer_transpose_apply] 129 | have hcont : ContDiff 𝕜 n fun x : ι → F ↦ b.toMatrix x := 130 | contDiff_pi.mpr fun j ↦ contDiff_pi.mpr fun j' ↦ 131 | ((smooth_barycentric_coord b j').of_le le_top).comp (contDiff_apply 𝕜 F j) 132 | have h_snd : ContDiff 𝕜 n fun x : F × (ι → F) ↦ b.toMatrix x.snd := hcont.comp contDiff_snd 133 | apply ContDiffOn.mul 134 | · apply ((Matrix.smooth_det ι 𝕜 n).comp h_snd).contDiffOn.inv 135 | rintro ⟨p, v⟩ hpv 136 | have hv : IsUnit (b.toMatrix v) := by simpa [mem_affineBases_iff ι 𝕜 F b v] using hpv 137 | rw [← isUnit_iff_ne_zero, comp_apply, ← Matrix.isUnit_iff_isUnit_det] 138 | exact hv 139 | · refine ((Matrix.smooth_det ι 𝕜 n).comp ?_).contDiffOn 140 | refine contDiff_pi.mpr fun j ↦ contDiff_pi.mpr fun j' ↦ ?_ 141 | simp only [Matrix.updateRow_apply] 142 | simp only [AffineBasis.toMatrix_apply, AffineBasis.coords_apply] 143 | by_cases hij : j = i 144 | · simp only [hij, if_true, eq_self_iff_true] 145 | exact (smooth_barycentric_coord b j').fst'.of_le le_top 146 | · simp only [hij, if_false] 147 | exact ((smooth_barycentric_coord b j').of_le le_top).comp (contDiff_pi.mp contDiff_snd j) 148 | 149 | end smooth_barycentric 150 | -------------------------------------------------------------------------------- /SphereEversion/ToMathlib/Topology/Algebra/Module.lean: -------------------------------------------------------------------------------- 1 | import Mathlib.Topology.Algebra.Module.Equiv 2 | 3 | namespace ContinuousLinearMap 4 | 5 | variable {R₁ M₁ M₂ M₃ : Type*} [Semiring R₁] 6 | 7 | variable [TopologicalSpace M₁] [AddCommMonoid M₁] 8 | [TopologicalSpace M₂] [AddCommMonoid M₂] 9 | [TopologicalSpace M₃] [AddCommMonoid M₃] 10 | [Module R₁ M₁] [Module R₁ M₂] [Module R₁ M₃] 11 | 12 | theorem fst_prod_zero_add_zero_prod_snd [ContinuousAdd M₁] [ContinuousAdd M₂] : 13 | (ContinuousLinearMap.fst R₁ M₁ M₂).prod 0 + 14 | ContinuousLinearMap.prod 0 (ContinuousLinearMap.snd R₁ M₁ M₂) = 15 | ContinuousLinearMap.id R₁ (M₁ × M₂) := by 16 | rw [ContinuousLinearMap.ext_iff] 17 | intro x 18 | simp_rw [ContinuousLinearMap.add_apply, ContinuousLinearMap.id_apply, 19 | ContinuousLinearMap.prod_apply, ContinuousLinearMap.coe_fst', ContinuousLinearMap.coe_snd', 20 | ContinuousLinearMap.zero_apply, Prod.mk_add_mk, add_zero, zero_add] 21 | 22 | end ContinuousLinearMap 23 | 24 | variable {R₁ R₂ R₃ : Type*} [Semiring R₁] [Semiring R₂] [Semiring R₃] 25 | {σ₁₂ : R₁ →+* R₂} {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] 26 | {σ₂₃ : R₂ →+* R₃} {σ₃₂ : R₃ →+* R₂} [RingHomInvPair σ₂₃ σ₃₂] [RingHomInvPair σ₃₂ σ₂₃] 27 | {σ₁₃ : R₁ →+* R₃} {σ₃₁ : R₃ →+* R₁} [RingHomInvPair σ₁₃ σ₃₁] [RingHomInvPair σ₃₁ σ₁₃] 28 | [RingHomCompTriple σ₁₂ σ₂₃ σ₁₃] [RingHomCompTriple σ₃₂ σ₂₁ σ₃₁] 29 | {M₁ : Type*} [TopologicalSpace M₁] [AddCommMonoid M₁] 30 | {M'₁ : Type*} [TopologicalSpace M'₁] [AddCommMonoid M'₁] 31 | {M₂ : Type*} [TopologicalSpace M₂] [AddCommMonoid M₂] 32 | {M₃ : Type*} [TopologicalSpace M₃] [AddCommMonoid M₃] 33 | [Module R₁ M₁] [Module R₁ M'₁] [Module R₂ M₂] [Module R₃ M₃] 34 | 35 | section 36 | 37 | theorem Function.Surjective.clm_comp_injective {g : M₁ →SL[σ₁₂] M₂} (hg : Function.Surjective g) : 38 | Function.Injective fun f : M₂ →SL[σ₂₃] M₃ ↦ f.comp g := fun f f' hff' ↦ by 39 | rw [ContinuousLinearMap.ext_iff] at hff' ⊢ 40 | intro x 41 | obtain ⟨y, rfl⟩ := hg x 42 | exact hff' y 43 | 44 | end 45 | 46 | namespace ContinuousLinearEquiv 47 | 48 | theorem cancel_right {f f' : M₂ →SL[σ₂₃] M₃} {e : M₁ ≃SL[σ₁₂] M₂} : 49 | f.comp (e : M₁ →SL[σ₁₂] M₂) = f'.comp (e : M₁ →SL[σ₁₂] M₂) ↔ f = f' := by 50 | constructor 51 | · simp_rw [ContinuousLinearMap.ext_iff, ContinuousLinearMap.comp_apply, coe_coe] 52 | intro h v; rw [← e.apply_symm_apply v, h] 53 | · rintro rfl; rfl 54 | 55 | theorem cancel_left {e : M₂ ≃SL[σ₂₃] M₃} {f f' : M₁ →SL[σ₁₂] M₂} : 56 | (e : M₂ →SL[σ₂₃] M₃).comp f = (e : M₂ →SL[σ₂₃] M₃).comp f' ↔ f = f' := by 57 | constructor 58 | · simp_rw [ContinuousLinearMap.ext_iff, ContinuousLinearMap.comp_apply, coe_coe] 59 | intro h v; rw [← e.symm_apply_apply (f v), h, e.symm_apply_apply] 60 | · rintro rfl; rfl 61 | 62 | end ContinuousLinearEquiv 63 | -------------------------------------------------------------------------------- /SphereEversion/ToMathlib/Topology/Paracompact.lean: -------------------------------------------------------------------------------- 1 | import Mathlib.Topology.Separation.Hausdorff 2 | import Mathlib.Data.Real.Basic 3 | import Mathlib.Order.Interval.Finset.Nat 4 | 5 | open scoped Topology 6 | 7 | open Set Function 8 | 9 | /-- We could generalise and replace `ι × ℝ` with a dependent family of types but it doesn't seem 10 | worth it. Proof partly based on `refinement_of_locallyCompact_sigmaCompact_of_nhds_basis_set`. -/ 11 | theorem exists_countable_locallyFinite_cover {ι X : Type*} [TopologicalSpace X] [T2Space X] 12 | [LocallyCompactSpace X] [SigmaCompactSpace X] {c : ι → X} {W : ι → ℝ → Set X} 13 | {B : ι → ℝ → Set X} {p : ι → ℝ → Prop} (hc : Surjective c) (hW₀ : ∀ i r, p i r → c i ∈ W i r) 14 | (hW₁ : ∀ i r, p i r → IsOpen (W i r)) (hB : ∀ i, (𝓝 (c i)).HasBasis (p i) (B i)) : 15 | ∃ s : Set (ι × ℝ), s.Countable ∧ (∀ z ∈ s, (↿p) z) ∧ (⋃ z ∈ s, (↿W) z) = univ ∧ 16 | LocallyFinite (↿B ∘ ((↑) : s → ι × ℝ)) := by 17 | let K' := CompactExhaustion.choice X 18 | let K := K'.shiftr.shiftr 19 | let C : ℕ → Set X := fun n ↦ K (n + 2) \ interior (K (n + 1)) 20 | let U : ℕ → Set X := fun n ↦ interior (K (n + 3)) \ K n 21 | have hCU : ∀ n, C n ⊆ U n := fun n x hx ↦ 22 | ⟨K.subset_interior_succ _ hx.1, mt (fun hx₃ ↦ K.subset_interior_succ _ hx₃) hx.2⟩ 23 | have hC : ∀ n, IsCompact (C n) := fun n ↦ (K.isCompact _).diff isOpen_interior 24 | have hC' : (⋃ n, C n) = univ := by 25 | refine Set.univ_subset_iff.mp fun x _ ↦ mem_iUnion.mpr ⟨K'.find x, ?_⟩ 26 | simpa only [K'.find_shiftr] using 27 | diff_subset_diff_right interior_subset (K'.shiftr.mem_diff_shiftr_find x) 28 | have hU : ∀ n, IsOpen (U n) := fun n ↦ 29 | isOpen_interior.sdiff <| IsCompact.isClosed <| K.isCompact _ 30 | have hU' : ∀ n, {m | (U m ∩ U n).Nonempty}.Finite := fun n ↦ by 31 | suffices {m | (U m ∩ U n).Nonempty} ⊆ Icc (n - 2) (n + 2) by exact (finite_Icc _ _).subset this 32 | rintro m ⟨x, ⟨⟨hx₁, hx₂⟩, ⟨hx₃, hx₄⟩⟩⟩ 33 | simp only [mem_Icc, tsub_le_iff_right] 34 | suffices ∀ {a b : ℕ}, x ∉ K a → x ∈ interior (K b.succ) → a ≤ b by 35 | exact ⟨this hx₄ hx₁, this hx₂ hx₃⟩ 36 | intro a b ha hb 37 | by_contra hab 38 | replace hab : b + 1 ≤ a := by simpa using hab 39 | exact Set.Nonempty.ne_empty (⟨x, interior_subset hb, ha⟩ : (K b.succ \ K a).Nonempty) 40 | (Set.diff_eq_empty.mpr (K.subset hab)) 41 | have hU'' : ∀ n x, x ∈ C n → U n ∈ 𝓝 x := fun n x hx ↦ 42 | mem_nhds_iff.mpr ⟨U n, Subset.rfl, hU n, hCU n hx⟩ 43 | have : ∀ (n) (x : C n), ∃ i r, ↑x ∈ W i r ∧ B i r ⊆ U n ∧ p i r := fun n ⟨x, hx⟩ ↦ by 44 | obtain ⟨i, rfl⟩ := hc x 45 | obtain ⟨r, hr₁, hr₂⟩ := (hB i).mem_iff.mp (hU'' n _ hx) 46 | exact ⟨i, r, hW₀ i r hr₁, hr₂, hr₁⟩ 47 | choose i r h₁ h₂ h₃ using fun n ↦ this n 48 | let V : ∀ n, C n → Set X := fun n x ↦ W (i n x) (r n x) 49 | have hV₁ : ∀ n x, IsOpen (V n x) := fun n x ↦ hW₁ _ _ (h₃ n x) 50 | have hV₂ : ∀ n, C n ⊆ ⋃ x : C n, V n x := fun n x hx ↦ mem_iUnion.mpr ⟨⟨x, hx⟩, h₁ _ _⟩ 51 | choose f hf using fun n ↦ (hC n).elim_finite_subcover (V n) (hV₁ n) (hV₂ n) 52 | classical 53 | let s : Set (ι × ℝ) := ⋃ n, (f n).image (Pi.prod (i n) (r n)) 54 | refine ⟨s, countable_iUnion fun n ↦ Finset.countable_toSet _, fun z hz ↦ ?_, 55 | Set.univ_subset_iff.mp fun x _ ↦ ?_, fun x ↦ ?_⟩ 56 | · simp only [Pi.prod, mem_iUnion, Finset.coe_image, mem_image, Finset.mem_coe, 57 | SetCoe.exists, s] at hz 58 | obtain ⟨n, x, hx, -, rfl⟩ := hz 59 | apply h₃ 60 | · obtain ⟨n, hn⟩ := iUnion_eq_univ_iff.mp hC' x 61 | specialize hf n hn 62 | simp only [iUnion_coe_set, mem_iUnion, exists_prop] at hf 63 | obtain ⟨y, hy₁, hy₂, hy₃⟩ := hf 64 | simp only [Pi.prod, mem_iUnion, Finset.mem_coe, Finset.mem_image, exists_prop, SetCoe.exists, 65 | iUnion_exists, exists_and_right, Prod.exists, Prod.mk_inj, s] 66 | exact ⟨i n ⟨y, hy₁⟩, r n ⟨y, hy₁⟩, ⟨n, y, hy₁, hy₂, rfl, rfl⟩, hy₃⟩ 67 | · obtain ⟨n, hn⟩ := iUnion_eq_univ_iff.mp hC' x 68 | refine ⟨U n, hU'' n x hn, ?_⟩ 69 | let P : ι × ℝ → Prop := fun z ↦ ((↿B) (z : ι × ℝ) ∩ U n).Nonempty 70 | erw [(Equiv.Set.sep s P).symm.set_finite_iff] 71 | simp only [Set.iUnion_inter, ← inter_setOf_eq_sep, s] 72 | refine (hU' n).iUnion (fun m _ ↦ Set.toFinite _) fun m hm ↦ ?_ 73 | rw [Set.eq_empty_iff_forall_notMem] 74 | intro z 75 | simp only [Pi.prod, Finset.coe_image, mem_inter_iff, mem_image, Finset.mem_coe, SetCoe.exists, 76 | mem_setOf_eq, not_and, exists₂_imp, and_imp] 77 | rintro x hx₁ - rfl 78 | rw [Set.not_nonempty_iff_eq_empty] 79 | have := Set.inter_subset_inter_left (U n) (h₂ m ⟨x, hx₁⟩) 80 | rwa [Set.not_nonempty_iff_eq_empty.mp hm, Set.subset_empty_iff] at this 81 | -------------------------------------------------------------------------------- /SphereEversion/ToMathlib/Topology/Path.lean: -------------------------------------------------------------------------------- 1 | import Mathlib.Analysis.Normed.Field.Basic 2 | import Mathlib.Analysis.Normed.Order.Lattice 3 | import Mathlib.Topology.Connected.PathConnected 4 | 5 | open Set Function 6 | 7 | open scoped Topology unitInterval 8 | 9 | noncomputable section 10 | 11 | variable {X X' Y Z : Type*} [TopologicalSpace X] 12 | [TopologicalSpace X'] [TopologicalSpace Y] [TopologicalSpace Z] 13 | 14 | namespace Path 15 | 16 | variable {x : X} {γ γ' : Path x x} 17 | 18 | /-- A loop evaluated at `t / t` is equal to its endpoint. Note that `t / t = 0` for `t = 0`. -/ 19 | @[simp] 20 | theorem extend_div_self (γ : Path x x) (t : ℝ) : γ.extend (t / t) = x := by 21 | by_cases h : t = 0 <;> simp [h] 22 | 23 | /-- Concatenation of two loops which moves through the first loop on `[0, t₀]` and 24 | through the second one on `[t₀, 1]`. All endpoints are assumed to be the same so that this 25 | function is also well-defined for `t₀ ∈ {0, 1}`. 26 | `strans` stands either for a *s*kewed transitivity, or a transitivity with different *s*peeds. -/ 27 | def strans (γ γ' : Path x x) (t₀ : I) : Path x x where 28 | toFun t := if t ≤ t₀ then γ.extend (t / t₀) else γ'.extend ((t - t₀) / (1 - t₀)) 29 | continuous_toFun := by 30 | refine Continuous.if_le ?_ ?_ continuous_id continuous_const ?_ 31 | · continuity 32 | · continuity 33 | · simp only [extend_div_self, Icc.mk_zero, zero_le_one, id, zero_div, forall_eq, 34 | extend_extends, Path.source, left_mem_Icc, sub_self] 35 | source' := by simp 36 | target' := by 37 | simp +contextual only [unitInterval.le_one'.le_iff_eq.trans eq_comm, 38 | extend_div_self, Icc.coe_one, imp_true_iff, eq_self_iff_true, comp_apply, ite_eq_right_iff] 39 | 40 | /-- Reformulate `strans` without using `extend`. This is useful to not have to prove that the 41 | arguments to `γ` lie in `I` after this. -/ 42 | theorem strans_def {x : X} {t₀ t : I} (γ γ' : Path x x) : 43 | γ.strans γ' t₀ t = 44 | if h : t ≤ t₀ then γ ⟨t / t₀, unitInterval.div_mem t.2.1 t₀.2.1 h⟩ 45 | else γ' ⟨(t - t₀) / (1 - t₀), 46 | unitInterval.div_mem (sub_nonneg.mpr <| le_of_not_le h) (sub_nonneg.mpr t₀.2.2) 47 | (sub_le_sub_right t.2.2 t₀)⟩ := by 48 | split_ifs with h <;> simp [strans, h, ← extend_extends] 49 | 50 | @[simp] 51 | theorem strans_of_ge {t t₀ : I} (h : t₀ ≤ t) : 52 | γ.strans γ' t₀ t = γ'.extend ((t - t₀) / (1 - t₀)) := by 53 | simp only [Path.coe_mk_mk, Path.strans, ite_eq_right_iff] 54 | intro h2; obtain rfl := le_antisymm h h2; simp 55 | 56 | theorem unitInterval.zero_le (x : I) : 0 ≤ x := 57 | x.prop.1 58 | 59 | @[simp] 60 | theorem strans_zero (γ γ' : Path x x) : γ.strans γ' 0 = γ' := by 61 | ext t 62 | simp +contextual only [strans_of_ge (unitInterval.zero_le t), Icc.coe_zero, div_one, 63 | extend_extends', unitInterval.nonneg'.le_iff_eq, sub_zero, div_zero, extend_zero, 64 | ite_eq_right_iff, show (t : ℝ) = 0 ↔ t = 0 from (@Subtype.ext_iff _ _ t 0).symm, 65 | Path.source, eq_self_iff_true, imp_true_iff] 66 | 67 | @[simp] 68 | theorem strans_one {x : X} (γ γ' : Path x x) : γ.strans γ' 1 = γ := by 69 | ext t 70 | simp only [strans, unitInterval.le_one', Path.coe_mk_mk, if_pos, div_one, extend_extends', 71 | Icc.coe_one] 72 | 73 | theorem strans_self {x : X} (γ γ' : Path x x) (t₀ : I) : γ.strans γ' t₀ t₀ = x := by 74 | simp only [strans, Path.coe_mk_mk, extend_div_self, if_pos, le_rfl] 75 | 76 | @[simp] 77 | theorem refl_strans_refl {x : X} {t₀ : I} : (refl x).strans (refl x) t₀ = refl x := by 78 | ext s 79 | simp [strans] 80 | 81 | theorem subset_range_strans_left {x : X} {γ γ' : Path x x} {t₀ : I} (h : t₀ ≠ 0) : 82 | range γ ⊆ range (γ.strans γ' t₀) := by 83 | rintro _ ⟨t, rfl⟩ 84 | use t * t₀ 85 | field_simp [strans, unitInterval.mul_le_right, unitInterval.coe_ne_zero.mpr h] 86 | 87 | theorem subset_range_strans_right {x : X} {γ γ' : Path x x} {t₀ : I} (h : t₀ ≠ 1) : 88 | range γ' ⊆ range (γ.strans γ' t₀) := by 89 | rintro _ ⟨t, rfl⟩ 90 | have := mul_nonneg t.2.1 (sub_nonneg.mpr t₀.2.2) 91 | let t' : I := ⟨t₀ + t * (1 - t₀), add_nonneg t₀.2.1 this, by 92 | rw [add_comm, ← le_sub_iff_add_le] 93 | exact (mul_le_mul_of_nonneg_right t.2.2 <| sub_nonneg.mpr t₀.2.2).trans_eq (one_mul _)⟩ 94 | have h2 : t₀ ≤ t' := le_add_of_nonneg_right this 95 | have h3 := sub_ne_zero.mpr (unitInterval.coe_ne_one.mpr h).symm 96 | use t' 97 | simp [h2, unitInterval.coe_ne_one.mpr h, h3, t'] 98 | 99 | theorem range_strans_subset {x : X} {γ γ' : Path x x} {t₀ : I} : 100 | range (γ.strans γ' t₀) ⊆ range γ ∪ range γ' := by 101 | rintro _ ⟨t, rfl⟩ 102 | by_cases h : t ≤ t₀ 103 | · rw [strans_def, dif_pos h]; exact Or.inl (mem_range_self _) 104 | · rw [strans_def, dif_neg h]; exact Or.inr (mem_range_self _) 105 | 106 | theorem Continuous.path_strans {X Y : Type*} [UniformSpace X] 107 | [LocallyCompactSpace X] [UniformSpace Y] {f : X → Y} {t : X → I} {s : X → I} 108 | {γ γ' : ∀ x, Path (f x) (f x)} (hγ : Continuous ↿γ) (hγ' : Continuous ↿γ') 109 | (hγ0 : ∀ ⦃x s⦄, t x = 0 → γ x s = f x) (hγ'1 : ∀ ⦃x s⦄, t x = 1 → γ' x s = f x) 110 | (ht : Continuous t) (hs : Continuous s) : 111 | Continuous fun x ↦ strans (γ x) (γ' x) (t x) (s x) := by 112 | have hγ0 : ∀ {x₀}, t x₀ = 0 → 113 | TendstoUniformly (fun x ↦ γ x) (fun _ ↦ f x₀) (𝓝 x₀) := fun h₀ ↦ by 114 | convert Continuous.tendstoUniformly (fun x ↦ γ x) hγ _; rw [hγ0 h₀] 115 | have hγ'1 : ∀ {x₀}, t x₀ = 1 → 116 | TendstoUniformly (fun x ↦ γ' x) (fun _ ↦ f x₀) (𝓝 x₀) := fun h₀ ↦ by 117 | convert Continuous.tendstoUniformly (fun x ↦ γ' x) hγ' _; rw [hγ'1 h₀] 118 | refine Continuous.if_le ?_ ?_ hs ht ?_ 119 | · rw [continuous_iff_continuousAt] 120 | intro x 121 | refine (continuous_subtype_val.comp hs).continuousAt.comp_div_cases (fun x s ↦ (γ x).extend s) 122 | (continuous_subtype_val.comp ht).continuousAt ?_ ?_ 123 | · intro _ 124 | refine ContinuousAt.pathExtend _ ?_ continuousAt_snd 125 | exact hγ.continuousAt.comp (continuousAt_fst.fst.prodMk continuousAt_snd) 126 | · intro h 127 | have ht : t x = 0 := Subtype.ext h 128 | apply Filter.Tendsto.pathExtend 129 | dsimp only; rw [(projIcc_surjective _).filter_map_top, extend_zero] 130 | exact tendsto_prod_top_iff.mpr (hγ0 ht) 131 | · rw [continuous_iff_continuousAt] 132 | intro x 133 | refine (continuous_subtype_val.comp hs).sub (continuous_subtype_val.comp ht) 134 | |>.continuousAt.comp_div_cases (fun x s ↦ (γ' x).extend s) 135 | (continuous_const.sub <| continuous_subtype_val.comp ht).continuousAt ?_ ?_ 136 | · intro _ 137 | refine ContinuousAt.pathExtend _ ?_ continuousAt_snd 138 | exact hγ'.continuousAt.comp (continuousAt_fst.fst.prodMk continuousAt_snd) 139 | · intro h 140 | have ht : t x = 1 := Subtype.ext (sub_eq_zero.mp h).symm 141 | apply Filter.Tendsto.pathExtend 142 | dsimp only; rw [(projIcc_surjective _).filter_map_top, extend_zero] 143 | exact tendsto_prod_top_iff.mpr (hγ'1 ht) 144 | · rintro x h; rw [h, sub_self, zero_div, extend_div_self, extend_zero] 145 | 146 | end Path 147 | -------------------------------------------------------------------------------- /SphereEversion/ToMathlib/Topology/Separation/Hausdorff.lean: -------------------------------------------------------------------------------- 1 | import Mathlib.Topology.Separation.Hausdorff 2 | open Set Function 3 | 4 | variable {X Y : Type*} [TopologicalSpace X] [TopologicalSpace Y] 5 | 6 | /- 7 | TODO? State a specialized version for quotient maps? Note the open map assumption is still 8 | a separate assumption there, because there is no `QuotientMap.prod_map`. 9 | -/ 10 | theorem t2Space_iff_of_continuous_surjective_open {α β : Type*} [TopologicalSpace α] 11 | [TopologicalSpace β] {π : α → β} (hcont : Continuous π) (hsurj : Surjective π) 12 | (hop : IsOpenMap π) : T2Space β ↔ IsClosed {q : α × α | π q.1 = π q.2} := by 13 | rw [t2_iff_isClosed_diagonal] 14 | constructor <;> intro H 15 | · exact H.preimage (hcont.prodMap hcont) 16 | · simp_rw [← isOpen_compl_iff] at H ⊢ 17 | convert hop.prodMap hop _ H 18 | exact ((hsurj.prodMap hsurj).image_preimage _).symm 19 | -------------------------------------------------------------------------------- /SphereEversion/ToMathlib/Unused/Fin.lean: -------------------------------------------------------------------------------- 1 | import Mathlib.Data.Fin.Basic 2 | 3 | -- not directly used 4 | theorem Fin.coe_succ_le_iff_le {n : ℕ} {j k : Fin n} : (j : Fin <| n + 1) ≤ k ↔ j ≤ k := by 5 | simp only [Fin.coe_eq_castSucc]; rfl 6 | 7 | -- not directly used 8 | theorem Fin.coe_lt_succ {n : ℕ} (k : Fin n) : (k : Fin <| n + 1) < k.succ := by 9 | rw [Fin.coe_eq_castSucc] 10 | exact Nat.lt_succ_self _ 11 | -------------------------------------------------------------------------------- /SphereEversion/ToMathlib/Unused/IntervalIntegral.lean: -------------------------------------------------------------------------------- 1 | import Mathlib.MeasureTheory.Integral.IntervalIntegral 2 | import Mathlib.MeasureTheory.Integral.Periodic 3 | 4 | noncomputable section 5 | 6 | open MeasureTheory Set 7 | 8 | section 9 | 10 | variable {E : Type*} [NormedAddCommGroup E] 11 | 12 | -- lemma interval_integrable_of_integral_ne_zero 13 | -- [complete_space E] [normed_space ℝ E] {a b : ℝ} 14 | -- {f : ℝ → E} {μ : measure ℝ} (h : ∫ x in a..b, f x ∂μ ≠ 0) : 15 | -- interval_integrable f μ a b := 16 | -- begin 17 | -- contrapose! h, 18 | -- exact interval_integral.integral_undef h, 19 | -- end 20 | -- lemma interval_integrable_norm_iff {f : ℝ → E} {μ : measure ℝ} {a b : ℝ} 21 | -- (hf : ae_strongly_measurable f (μ.restrict (Ι a b))) : 22 | -- interval_integrable (λ t, ‖f t‖) μ a b ↔ interval_integrable f μ a b := 23 | -- begin 24 | -- simp_rw [interval_integrable_iff, integrable_on], 25 | -- exact integrable_norm_iff hf 26 | -- end 27 | -- not ported 28 | theorem intervalIntegrable_of_nonneg_of_le {f g : ℝ → ℝ} {μ : Measure ℝ} {a b : ℝ} 29 | (hf : AEStronglyMeasurable f <| μ.restrict (Ι a b)) 30 | (h : ∀ᵐ t ∂μ.restrict <| Ι a b, 0 ≤ f t ∧ f t ≤ g t) (hg : IntervalIntegrable g μ a b) : 31 | IntervalIntegrable f μ a b := by 32 | rw [intervalIntegrable_iff] at * 33 | apply Integrable.mono' hg hf (h.mono _) 34 | rintro t ⟨H, H'⟩ 35 | change abs (f t) ≤ _ 36 | rwa [abs_of_nonneg H] 37 | 38 | variable [CompleteSpace E] [NormedSpace ℝ E] {a b : ℝ} {f : ℝ → E} {bound : ℝ → ℝ} {μ : Measure ℝ} 39 | 40 | -- next 4 lemmas not ported (also unused) 41 | namespace intervalIntegral 42 | 43 | theorem integral_mono_of_le {f g : ℝ → ℝ} {a b : ℝ} {μ : Measure ℝ} (hab : a ≤ b) 44 | (hf : IntervalIntegrable f μ a b) (hg : IntervalIntegrable g μ a b) 45 | (hfg : f ≤ᵐ[μ.restrict (Ι a b)] g) : ∫ u in a..b, f u ∂μ ≤ ∫ u in a..b, g u ∂μ := by 46 | rw [uIoc_of_le hab] at hfg 47 | let H := hfg.filter_mono (ae_mono le_rfl) 48 | simpa only [integral_of_le hab] using setIntegral_mono_ae_restrict hf.1 hg.1 H 49 | 50 | theorem integral_mono_of_le_of_nonneg {f g : ℝ → ℝ} {a b : ℝ} {μ : Measure ℝ} (hab : a ≤ b) 51 | (hf : AEStronglyMeasurable f <| μ.restrict (Ι a b)) 52 | (hfnonneg : ∀ᵐ t ∂μ.restrict <| Ι a b, 0 ≤ f t) (hg : IntervalIntegrable g μ a b) 53 | (hfg : f ≤ᵐ[μ.restrict (Ι a b)] g) : ∫ u in a..b, f u ∂μ ≤ ∫ u in a..b, g u ∂μ := by 54 | apply integral_mono_of_le hab _ hg hfg 55 | have : ∀ᵐ t ∂μ.restrict <| Ι a b, 0 ≤ f t ∧ f t ≤ g t := hfnonneg.and hfg 56 | apply intervalIntegrable_of_nonneg_of_le hf this hg 57 | 58 | theorem integral_antimono_of_le {f g : ℝ → ℝ} {a b : ℝ} {μ : Measure ℝ} (hab : b ≤ a) 59 | (hf : IntervalIntegrable f μ a b) (hg : IntervalIntegrable g μ a b) 60 | (hfg : f ≤ᵐ[μ.restrict (Ι a b)] g) : ∫ u in a..b, g u ∂μ ≤ ∫ u in a..b, f u ∂μ := by 61 | cases' hab.eq_or_lt with hab hab 62 | · simp [hab] 63 | · rw [uIoc_of_ge hab.le] at hfg 64 | rw [integral_symm b a, integral_symm b a] 65 | apply neg_le_neg 66 | apply integral_mono_of_le hab.le hf.symm hg.symm 67 | have : Ioc b a = uIoc b a := by rw [uIoc_of_le hab.le] 68 | rw [← this] 69 | exact hfg 70 | 71 | theorem integral_antimono_of_le_of_nonneg {f g : ℝ → ℝ} {a b : ℝ} {μ : Measure ℝ} (hab : b ≤ a) 72 | (hf : AEStronglyMeasurable f <| μ.restrict (Ι a b)) 73 | (hfnonneg : ∀ᵐ t ∂μ.restrict <| Ι a b, 0 ≤ f t) (hg : IntervalIntegrable g μ a b) 74 | (hfg : f ≤ᵐ[μ.restrict (Ι a b)] g) : ∫ u in a..b, g u ∂μ ≤ ∫ u in a..b, f u ∂μ := by 75 | apply integral_antimono_of_le hab _ hg hfg 76 | have : ∀ᵐ t ∂μ.restrict <| Ι a b, 0 ≤ f t ∧ f t ≤ g t := hfnonneg.and hfg 77 | apply intervalIntegrable_of_nonneg_of_le hf this hg 78 | 79 | end intervalIntegral 80 | 81 | section 82 | 83 | open intervalIntegral 84 | 85 | theorem norm_intervalIntegral_eq (f : ℝ → E) (a b : ℝ) (μ : Measure ℝ) : 86 | ‖∫ x in a..b, f x ∂μ‖ = ‖∫ x in Ι a b, f x ∂μ‖ := by 87 | simp_rw [intervalIntegral_eq_integral_uIoc, norm_smul] 88 | split_ifs <;> simp only [norm_neg, norm_one, one_mul] 89 | 90 | end 91 | 92 | theorem abs_intervalIntegral_eq (f : ℝ → ℝ) (a b : ℝ) (μ : Measure ℝ) : 93 | |∫ x in a..b, f x ∂μ| = |∫ x in Ι a b, f x ∂μ| := 94 | norm_intervalIntegral_eq f a b μ 95 | 96 | -- not ported 97 | theorem intervalIntegrable_of_norm_sub_le {β : Type*} [NormedAddCommGroup β] {f₀ f₁ : ℝ → β} 98 | {g : ℝ → ℝ} {a b : ℝ} (hf₁_m : AEStronglyMeasurable f₁ (μ.restrict <| Ι a b)) 99 | (hf₀_i : IntervalIntegrable f₀ μ a b) (hg_i : IntervalIntegrable g μ a b) 100 | (h : ∀ᵐ a ∂μ.restrict <| Ι a b, ‖f₀ a - f₁ a‖ ≤ g a) : IntervalIntegrable f₁ μ a b := 101 | haveI : ∀ᵐ a ∂μ.restrict <| Ι a b, ‖f₁ a‖ ≤ ‖f₀ a‖ + g a := by 102 | apply h.mono 103 | intro a ha 104 | calc 105 | ‖f₁ a‖ ≤ ‖f₀ a‖ + ‖f₀ a - f₁ a‖ := norm_le_insert _ _ 106 | _ ≤ ‖f₀ a‖ + g a := add_le_add_left ha _ 107 | (hf₀_i.norm.add hg_i).mono_fun' hf₁_m this 108 | 109 | end 110 | 111 | -- section 112 | -- variables {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] [complete_space E] 113 | -- open interval_integral 114 | -- lemma integral_comp_add_right' {f : ℝ → E} (a b : ℝ) : 115 | -- ∫ t in a..(b + a), f t = ∫ t in 0..b, f (t + a) := 116 | -- by simp [← integral_comp_add_right] 117 | -- lemma integral_comp_add_left' {f : ℝ → E} (a b : ℝ) : 118 | -- ∫ t in a..(a + b), f t = ∫ t in 0..b, f (t + a) := 119 | -- by simp [← integral_comp_add_left, add_comm] 120 | -- end 121 | -------------------------------------------------------------------------------- /SphereEversion/ToMathlib/Unused/LinearAlgebra/Multilinear.lean: -------------------------------------------------------------------------------- 1 | import Mathlib.Analysis.Calculus.BumpFunction.InnerProduct 2 | 3 | /- 4 | Multilinear map stuff that was meant as preliminaries for smooth functions gluing. 5 | 6 | We no longer intend to use this file in the sphere eversion project. 7 | -/ 8 | namespace Function 9 | 10 | variable {ι : Sort*} [DecidableEq ι] {α β : ι → Type*} 11 | 12 | /-- Special case of `function.apply_update`. Useful for `rw`/`simp`. -/ 13 | theorem update_fst (g : ∀ i, α i × β i) (i : ι) (v : α i × β i) (j : ι) : 14 | (update g i v j).fst = update (fun k ↦ (g k).fst) i v.fst j := 15 | apply_update (fun _ ↦ Prod.fst) g i v j 16 | 17 | /-- Special case of `function.apply_update`. Useful for `rw`/`simp`. -/ 18 | theorem update_snd (g : ∀ i, α i × β i) (i : ι) (v : α i × β i) (j : ι) : 19 | (update g i v j).snd = update (fun k ↦ (g k).snd) i v.snd j := 20 | apply_update (fun _ ↦ Prod.snd) g i v j 21 | 22 | end Function 23 | 24 | open Function 25 | 26 | namespace MultilinearMap 27 | 28 | variable {R ι ι' M₃ M₄ : Type*} {M₁ M₂ : ι → Type*} {N : ι' → Type*} 29 | 30 | variable [Semiring R] 31 | 32 | variable [∀ i, AddCommMonoid (M₁ i)] [∀ i, Module R (M₁ i)] 33 | 34 | variable [∀ i, AddCommMonoid (M₂ i)] [∀ i, Module R (M₂ i)] 35 | 36 | variable [∀ i, AddCommMonoid (N i)] [∀ i, Module R (N i)] 37 | 38 | variable [AddCommMonoid M₃] [Module R M₃] 39 | 40 | variable [AddCommMonoid M₄] [Module R M₄] 41 | 42 | /-- The coproduct of two multilinear maps. -/ 43 | @[simps!] 44 | def coprod (L₁ : MultilinearMap R M₁ M₃) (L₂ : MultilinearMap R M₂ M₃) : 45 | MultilinearMap R (fun i ↦ M₁ i × M₂ i) M₃ := 46 | (L₁.compLinearMap fun _ ↦ .fst ..) + L₂.compLinearMap fun _ ↦ .snd .. 47 | 48 | end MultilinearMap 49 | 50 | namespace ContinuousMultilinearMap 51 | 52 | variable {R ι ι' : Type*} {M₁ M₂ : ι → Type*} {M₃ M₄ : Type*} {N : ι' → Type*} 53 | 54 | variable [Semiring R] 55 | 56 | variable [∀ i, AddCommMonoid (M₁ i)] [∀ i, AddCommMonoid (M₂ i)] [AddCommMonoid M₃] 57 | 58 | variable [∀ i, Module R (M₁ i)] [∀ i, Module R (M₂ i)] [Module R M₃] 59 | 60 | variable [∀ i, TopologicalSpace (M₁ i)] [∀ i, TopologicalSpace (M₂ i)] 61 | 62 | variable [TopologicalSpace M₃] 63 | 64 | variable [AddCommMonoid M₄] [Module R M₄] [TopologicalSpace M₄] 65 | 66 | variable [∀ i, AddCommMonoid (N i)] [∀ i, Module R (N i)] [∀ i, TopologicalSpace (N i)] 67 | 68 | variable [ContinuousAdd M₃] 69 | 70 | @[simps!] 71 | def coprod (L₁ : ContinuousMultilinearMap R M₁ M₃) (L₂ : ContinuousMultilinearMap R M₂ M₃) : 72 | ContinuousMultilinearMap R (fun i ↦ M₁ i × M₂ i) M₃ := 73 | (L₁.compContinuousLinearMap fun _ ↦ .fst ..) + L₂.compContinuousLinearMap fun _ ↦ .snd .. 74 | 75 | @[simp] 76 | def zero_coprod_zero : 77 | (0 : ContinuousMultilinearMap R M₁ M₃).coprod (0 : ContinuousMultilinearMap R M₂ M₃) = 0 := by 78 | ext; simp 79 | 80 | end ContinuousMultilinearMap 81 | -------------------------------------------------------------------------------- /blueprint/.gitignore: -------------------------------------------------------------------------------- 1 | *.pdf 2 | *.paux 3 | *.aux 4 | print 5 | web 6 | __pycache__ 7 | *.fdb_latexmk 8 | 9 | -------------------------------------------------------------------------------- /blueprint/mathjax.sed: -------------------------------------------------------------------------------- 1 | s/⁰/^0/g 2 | s/⁻¹/^{-1}/g 3 | s/¹/^1/g 4 | s/²/^2/g 5 | s/³/^3/g 6 | s/⁴/^4/g 7 | s/⁵/^5/g 8 | s/⁶/^6/g 9 | s/⁷/^7/g 10 | s/⁸/^8/g 11 | s/⁹/^9/g 12 | s/₀/_0/g 13 | s/₁/_1/g 14 | s/₂/_2/g 15 | s/₃/_3/g 16 | s/₄/_4/g 17 | s/₅/_5/g 18 | s/₆/_6/g 19 | s/₇/_7/g 20 | s/₈/_8/g 21 | s/₉/_9/g 22 | -------------------------------------------------------------------------------- /blueprint/src/appendix.tex: -------------------------------------------------------------------------------- 1 | \chapter{Local sphere eversion} 2 | \label{chap:local_eversion} 3 | 4 | The local theory of \Cref{chap:local} is already enough to deduce 5 | Smale's sphere eversion theorem, although it is less natural than going through 6 | the general results of \Cref{chap:global}. The goal of this appendix is to 7 | explain how to do so. In this section $E$ denote a finite dimensional real 8 | vector space equipped with an inner product. Later we will assume it is 9 | 3-dimensional. We denote by $𝕊$ the unit sphere in $E$. 10 | 11 | Although we want to study immersions of $𝕊$ into $E$, we want to work only with 12 | functions defined on the whole $E$. So we introduce a slightly artificial relation. 13 | We denote by $B$ the open ball with radius $9/10$ around the origin in $E$ and set: 14 | 15 | \[ 16 | \Rel := \{(x, y, φ) ∈ J^1(E, E) \;|\; x ∉ B ⇒ \rst{φ}{x^⊥} \text{ is injective}\}. 17 | \] 18 | 19 | Of course solutions of this relation restrict to immersions of $𝕊$. 20 | 21 | \begin{lemma} 22 | \label{lem:loc_immersion_rel_open} 23 | \lean{loc_immersion_rel_open} 24 | \leanok 25 | The relation $\Rel$ above is open. 26 | \end{lemma} 27 | 28 | \begin{proof} 29 | \leanok 30 | The main task is to fix $x_0 \notin B$ 31 | and $\varphi_0 \in L(E, E)$ which is injective on $x_0^\perp$ and prove that, 32 | for every $x$ close to $x_0$ and $\varphi$ close to $\varphi_0$, $\varphi$ is 33 | injective on $x^\perp$. This is a typical situation where geometric intuition 34 | makes it feel like there is nothing to prove. 35 | 36 | One difficulty is that the subspace $x^\perp$ moves with $x$. We reduce to a fixed 37 | subspace by considering the restriction to $x_0^\perp$ of the orthogonal 38 | projection onto $x^\perp$. One can check this is an isomorphism as long as $x$ 39 | is not perpendicular to $x_0$. 40 | More precisely, we consider $f \co J^1(E, E) \to ℝ \times L(x_0^\perp, E)$ which sends 41 | $(x, y, \varphi)$ to $(\langle x_0, x\rangle, \varphi \circ \pr_{x^\perp} \circ j_0)$ 42 | where $j_0$ is the inclusion of $x_0^\perp$ into $E$. The set $U$ of injective 43 | linear maps is open in $L(x_0^\perp, E)$ and the map $f$ is continuous 44 | hence the preimage of $\{0\}^c \times U$ is open. This is good enough for us because 45 | injectivity of $\varphi \circ \pr_{x^\perp} \circ j_0$ implies injectivity of 46 | $\varphi$ on the image of $\pr_{x^\perp} \circ j_0$ which is $x^\perp$ whenever 47 | $\langle x_0, x\rangle \neq 0$. 48 | \end{proof} 49 | 50 | \begin{lemma} 51 | \label{lem:loc_immersion_rel_ample} 52 | \lean{loc_immersion_rel_ample} 53 | \uses{lem:ample_codim_two} 54 | \leanok 55 | The relation $\Rel$ above is ample. 56 | \end{lemma} 57 | 58 | \begin{proof} 59 | \leanok 60 | The core fact here is that if one fixes vector spaces $F$ and $F'$, a dual 61 | pair $(\pi, v)$ on $F$ and an injective linear map $\varphi \co F \to F'$ 62 | then the updated map $\Upd{p}{\varphi}{w}$ is injective if and only if $w$ 63 | is not in $\varphi(\ker\pi)$. First we assume $\Upd{p}{\varphi}{\varphi(u)}$ 64 | is injective for some $u$ in $\varphi(\ker\pi)$ and derive a contradiction. 65 | We have $\Upd{p}{\varphi}{\varphi(u)}v = \varphi(u)$ by the general 66 | definition of updating and also $\Upd{p}{\varphi}{\varphi(u)}u = \varphi(u)$ 67 | since $u$ is in $\ker \pi$. Hence injectivity of $\varphi$ ensure $u = v$, 68 | which is absurd since $\pi(u) = 0$ and $\pi(v) = 1$. Conversely assume $w$ 69 | is not in $\varphi(\ker\pi)$ and let us prove $\Upd{p}{\varphi}{w}$ is 70 | injective. Assume $x$ is in the kernel of $\Upd{p}{\varphi}{w}$. Decompose 71 | $x = u + tv$ with $u \in \ker\pi$ and $t$ a real number. We have 72 | $\Upd{p}{\varphi}{w}(x) = \varphi(u) + tw$. Hence our assumption on $x$ 73 | implies $t$ vanishes otherwise we would have $w = -t⁻¹\varphi(u)$ 74 | contradicting that $w$ isn't in $\varphi(\ker\pi)$. This vanishing and the 75 | assumption on $x$ then imply $\varphi(u) = 0$. Since $\varphi$ is injective 76 | we conclude that $u = 0$ and finally $x = 0$. 77 | 78 | We now turn to $\Rel$. It suffices to prove that for every $\sigma = (x, y, 79 | \varphi) \in \Rel$ and every dual pair $p = (\pi, v)$ on $E$, the slice 80 | $\Rel(\sigma, p)$ is ample. If $x$ is in $B$ then $\Rel(\sigma, p)$ is the 81 | whole $E$ which is obviously ample. So we assume $x$ is not in $B$. Since 82 | $\sigma$ is in $\Rel$, $\varphi$ is injective on $x^\perp$. The slice is the 83 | set of $w$ such that $\Upd{p}{\varphi}{w}$ is injective on $x^\perp$. Assume 84 | first $\ker\pi = x^\perp$. Then $\Upd{p}{\varphi}{w}$ coincides with 85 | $\varphi$ on $x^\perp$ hence the slice is the whole $E$. Assume now that 86 | $\ker\pi \neq x^\perp$. The slice is not very easy to picture in this case. 87 | But one should remember that, up to affine isomorphism, the slice depends 88 | only on $\ker \pi$. More specifically, if we keep $\pi$ but change $v$ then 89 | the slice is simply translated in $E$. Here we replace $v$ by the projection 90 | on $x^\perp$ of the vector dual to $\pi$ rescaled to keep the property 91 | $\pi(v) = 1$. What has been gained is that we now have $v \in x^\perp$ and 92 | $x^\perp = (x^\perp \cap \ker \pi) \oplus ℝ v$. Since $\varphi$ is 93 | injective on $x^\perp$, $\varphi(x^\perp \cap \ker \pi)$ is a hyperplane 94 | in $x^\perp$ and $\Upd{p}{\varphi}{w}$ is injective on $x^\perp$ if and only 95 | if $w$ is in the complement of $\varphi(x^\perp \cap \ker \pi)$ according to 96 | the core fact above. Since it is an hyperplane in $x^\perp$, it has 97 | codimension at least $2$ in $E$ hence its complement is ample. 98 | \end{proof} 99 | 100 | \begin{theorem}[Smale 1958] 101 | \lean{sphere_eversion_of_loc} 102 | \label{sphere_eversion_of_loc} 103 | \leanok 104 | There is a homotopy of immersion of $𝕊^2$ into $ℝ^3$ from the inclusion map to 105 | the antipodal map $a \co q ↦ -q$. 106 | \end{theorem} 107 | 108 | \begin{proof} 109 | \leanok 110 | \uses{lem:loc_immersion_rel_open, lem:loc_immersion_rel_ample, lem:h_principle_open_ample_loc, lem:param_trick} 111 | We denote by $ι$ the inclusion of $𝕊^2$ into $ℝ^3$. 112 | We set $j_t = (1-t)ι + ta$. 113 | This is a homotopy from $ι$ to $a$ (but not an immersion for $t=1/2$). 114 | Using the canonical trivialization of the tangent 115 | bundle of $ℝ^3$, we can set, for $(q, v) ∈ T𝕊^2$, 116 | $G_t(q, v) = \mathrm{Rot}_{Oq}^{πt}(v)$, the rotation around axis $Oq$ with 117 | angle $πt$. 118 | The family $σ : t ↦ (j_t, G_t)$ is a homotopy of formal immersions 119 | relating $j^1ι$ to $j^1a$. Those formal solutions are holonomic when $t$ is 120 | zero or one, so we can reparametrize the family to make such it is holonomic 121 | when $t$ is close to zero or one. Then we can extend it to a homotopy of 122 | formal solutions of $\Rel$ using a suitable cut-off ensuring smoothness 123 | near the origin. The relation $\Rel$ is ample according to 124 | \Cref{lem:loc_immersion_rel_ample} and then \Cref{lem:param_trick} ensures 125 | its 1-parameter version $\Rel^ℝ$ is also ample. The relation $\Rel$ is open according to 126 | \Cref{lem:loc_immersion_rel_open} hence $\Rel^ℝ$ is also ample. 127 | So we can use \Cref{lem:h_principle_open_ample_loc} to deform our family of 128 | formal solutions into a holonomic one. 129 | \end{proof} 130 | 131 | % vim: set expandtab sw=2 tabstop=2: 132 | -------------------------------------------------------------------------------- /blueprint/src/content.tex: -------------------------------------------------------------------------------- 1 | \input{introduction} 2 | \input{loops} 3 | \input{local_convex_integration} 4 | \input{global_convex_integration} 5 | \appendix 6 | \input{appendix} 7 | \input{local_to_global} 8 | -------------------------------------------------------------------------------- /blueprint/src/fix_mathjax.sh: -------------------------------------------------------------------------------- 1 | #/bin/bash 2 | 3 | for i in `ls *.tex` 4 | do sed -f ../mathjax.sed $i | sponge $i 5 | done 6 | -------------------------------------------------------------------------------- /blueprint/src/latexmkrc: -------------------------------------------------------------------------------- 1 | $pdf_mode = 1; 2 | $pdflatex = 'xelatex -synctex=1 -output-directory=../print/'; 3 | @default_files = ('print.tex'); 4 | -------------------------------------------------------------------------------- /blueprint/src/macros_common.tex: -------------------------------------------------------------------------------- 1 | \newcommand{\co}{\!:} 2 | \newcommand{\define}[1]{\index{#1}\emph{#1}} % définition du mot #1 3 | \newcommand{\subdefine}[2]{\index{#1!#2}\emph{#2}} % définition du mot #2 comme sous-entrée de #1 4 | \newcommand{\definex}[2]{\index{#2}\emph{#1}} % affiche #1 index #2 5 | \newcommand{\subdefinex}[3]{\index{#2!#3}\emph{#1}} % affiche #1 index #3 comme sous-entrée de #2 6 | 7 | \newcommand{\image}[1]{{\centering\includegraphics{images/#1}\par}} 8 | 9 | \newcommand{\A}{\mathcal{A}} % un atlas 10 | \newcommand{\E}{\mathcal{E}} % un espace affine 11 | \newcommand{\Cinf}{C^\infty} 12 | \newcommand{\Rel}{\mathcal{R}} % une relation différentielle 13 | 14 | \newcommand{\rst}[2]{#1|_{#2}} 15 | \newcommand{\bigO}[1]{O\left(#1\right)} 16 | 17 | \DeclareMathOperator{\Diff}{Diff} 18 | \DeclareMathOperator{\Loop}{\mathcal{L}} 19 | \DeclareMathOperator{\Bij}{Bij} 20 | \DeclareMathOperator{\Aut}{Aut} 21 | \DeclareMathOperator{\Id}{Id} 22 | \DeclareMathOperator{\rk}{rk} 23 | \DeclareMathOperator{\GL}{GL} 24 | \DeclareMathOperator{\SO}{SO} 25 | \DeclareMathOperator{\SU}{SU} 26 | \DeclareMathOperator{\Lin}{L} 27 | \DeclareMathOperator{\Hom}{Hom} 28 | \DeclareMathOperator{\Bsym}{B_\mathrm{s}} 29 | \DeclareMathOperator{\Hess}{Hess} 30 | \DeclareMathOperator{\im}{im} 31 | \DeclareMathOperator{\coker}{coker} 32 | \DeclareMathOperator{\supp}{supp} % support 33 | \DeclareMathOperator{\ind}{ind} % indice d'un point critique 34 | \DeclareMathOperator{\sgn}{ε} % signature d'une permutation 35 | \DeclareMathOperator{\Alt}{Alt} % antisymétrisation 36 | \DeclareMathOperator{\Div}{div} % divergence 37 | \DeclareMathOperator{\tr}{tr} % trace 38 | \DeclareMathOperator{\pr}{pr} % une projection 39 | \DeclareMathOperator{\Int}{Int} % intérieur 40 | \DeclareMathOperator{\Conv}{Conv} % enveloppe convexe 41 | \DeclareMathOperator{\IntConv}{IntConv} % intérieur enveloppe convexe 42 | \DeclareMathOperator{\Conn}{Conn} % connected component 43 | \DeclareMathOperator{\Span}{Span} % sev engendré 44 | \DeclareMathOperator{\Gr}{Gr} % Grassmanienne 45 | \DeclareMathOperator{\CP}{CP} % Corrugation process 46 | \DeclareMathOperator{\Sec}{Sec} % Sections 47 | \DeclareMathOperator{\Hol}{Hol} % Holonomic sections 48 | \DeclareMathOperator{\codim}{codim} % codimension 49 | \DeclareMathOperator{\set}{\mathcal{P}} % powerset 50 | 51 | % Cohomologie de de Rham, version avec indice dR ou non 52 | %\newcommand{\HdR}[1][*]{H^{#1}_{\!\scriptscriptstyle d\!R}} 53 | %\newcommand{\HdRc}[1][*]{H^{#1}_{\!\scriptscriptstyle d\!R, c}} 54 | \newcommand{\HdR}[1][*]{H^{#1}} 55 | \newcommand{\HdRc}[1][*]{H^{#1}_c} 56 | 57 | \newcommand{\transp}[1]{#1^T} 58 | 59 | \newcommand{\eE}{\mathcal{E}} 60 | \newcommand{\fF}{\mathcal{F}} 61 | \newcommand{\tT}{\mathcal{T}} 62 | 63 | \newcommand{\X}{\mathfrak{X}} % champs de vecteurs 64 | 65 | \newcommand{\C}{C} % cube 66 | \newcommand{\F}{\mathcal{F}} % non-holonomic section 67 | \newcommand{\HH}{\mathcal{H}} % homotopy of non-holonomic section 68 | \newcommand{\G}{\mathcal{G}} % non-holonomic section 69 | 70 | \newcommand{\pt}{\mathrm{pt}} % point 71 | 72 | \DeclareMathOperator{\Op}{Op} % non-specified open neighborhood 73 | \DeclareMathOperator{\dCzero}{d_{C^0}} % C^0 distance 74 | \DeclareMathOperator{\dCone}{d_{C^1}} % C^1 distance 75 | \DeclareMathOperator{\bs}{bs} % base map of a jet bundle section 76 | 77 | % Linear maps between vector spaces 78 | \renewcommand{\L}[2]{L(#1, #2)} 79 | % Update a linear map #2 using dual pair #1 and vector #3 80 | \newcommand{\Upd}[3]{\Upsilon_{#1}\left(#2, #3\right)} 81 | \newcommand{\IT}[1]{\mathcal{I}_{#1}} 82 | -------------------------------------------------------------------------------- /blueprint/src/macros_print.tex: -------------------------------------------------------------------------------- 1 | % Déclaration de théorème avec thmtools 2 | \declaretheorem[numberwithin=chapter]{theorem} 3 | \declaretheorem[sibling=theorem]{proposition} 4 | \declaretheorem[sibling=theorem]{corollary} 5 | \declaretheorem[sibling=theorem]{remark} 6 | \declaretheorem[sibling=theorem]{lemma} 7 | \declaretheorem[sibling=theorem]{definition} 8 | \declaretheorem[sibling=theorem]{example} 9 | 10 | \declaretheorem[numbered=no, name=Theorem]{theorem-intro} 11 | \declaretheorem[numbered=no, name=Proposition]{proposition-intro} 12 | \declaretheorem[numbered=no, name=Corollary]{corollary-intro} 13 | \declaretheorem[numbered=no, name=Lemma]{lemma-intro} 14 | \declaretheorem[numbered=no, name=Definition]{definition-intro} 15 | 16 | % Cache les commandes plastex 17 | \newcommand{\uses}[1]{} 18 | \newcommand{\proves}[1]{} 19 | \newcommand{\lean}[1]{} 20 | \newcommand{\leanok}{} 21 | 22 | %\newcommand{\D}[1]{\gls{D}[#1]} 23 | \newcommand{\D}[1]{\mathrm{Diff}(#1)} 24 | 25 | \newenvironment{diagram}[1][]{\[\begin{tikzcd}[#1]}{\end{tikzcd}\]} 26 | \newenvironment{dessin}[1][]{\[\begin{tikzpicture}[#1]}{\end{tikzpicture}\]} 27 | 28 | % Bugfix : 29 | % http://tex.stackexchange.com/questions/262617/issue-with-xrightarrow 30 | \makeatletter 31 | \patchcmd{\arrowfill@}{-7mu}{-14mu}{}{} 32 | \patchcmd{\arrowfill@}{-7mu}{-14mu}{}{} 33 | \patchcmd{\arrowfill@}{-2mu}{-4mu}{}{} 34 | \patchcmd{\arrowfill@}{-2mu}{-4mu}{}{} 35 | \makeatother 36 | 37 | \newcommand{\oiint}{\iint} 38 | 39 | \definecolor{fond}{RGB}{255,255,255} 40 | 41 | \newcommand{\N}[1]{\mathcal{N}_{\!\!#1}} % neighborhood 42 | -------------------------------------------------------------------------------- /blueprint/src/macros_web.tex: -------------------------------------------------------------------------------- 1 | \newtheorem{theorem}{Theorem}[chapter] 2 | \newtheorem{definition}[theorem]{Definition} 3 | \newtheorem{proposition}[theorem]{Proposition} 4 | \newtheorem{lemma}[theorem]{Lemma} 5 | \newtheorem{sublemma}[theorem]{Sub-lemma} 6 | \newtheorem{corollary}[theorem]{Corollary} 7 | \newtheorem{remark}[theorem]{Remark} 8 | \newtheorem{example}[theorem]{Example} 9 | 10 | \newtheorem*{theorem-intro}{Theorem}[chapter] 11 | \newtheorem*{definition-intro}[theorem]{Definition} 12 | \newtheorem*{proposition-intro}[theorem]{Proposition} 13 | \newtheorem*{lemma-intro}[theorem]{Lemma} 14 | \newtheorem*{corollary-intro}[theorem]{Corollary} 15 | 16 | \newcommand{\D}[1]{\mathrm{Diff}(#1)} 17 | 18 | \newenvironment{diagram}[1][]{\begin{tikzcd}[#1]}{\end{tikzcd}} 19 | \newenvironment{dessin}[1][]{\begin{tikzpicture}[#1]}{\end{tikzpicture}} 20 | \renewcommand\qedhere{} 21 | 22 | \newcommand{\intprod}{\,\lrcorner\,} 23 | 24 | \newcommand{\N}[1]{\mathcal{N}_{#1}} % neighborhood 25 | -------------------------------------------------------------------------------- /blueprint/src/plastex.cfg: -------------------------------------------------------------------------------- 1 | [general] 2 | renderer=HTML5 3 | copy-theme-extras=yes 4 | plugins=leanblueprint 5 | 6 | [document] 7 | toc-depth=2 8 | toc-non-files=True 9 | 10 | [files] 11 | directory=../web/ 12 | split-level=0 13 | 14 | [html5] 15 | localtoc-level=0 16 | extra-css=stylecours.css 17 | mathjax-dollars=True 18 | theme-css=blue 19 | -------------------------------------------------------------------------------- /blueprint/src/print.tex: -------------------------------------------------------------------------------- 1 | \documentclass[a4paper]{report} 2 | 3 | \usepackage[textwidth=14cm]{geometry} 4 | \usepackage{xfrac} 5 | \usepackage{polyglossia} 6 | \setdefaultlanguage{english} 7 | 8 | \usepackage{amsmath} 9 | \usepackage{enumitem} 10 | 11 | \usepackage{tikz-cd} 12 | 13 | \usepackage[warnings-off={mathtools-colon,mathtools-overbracket}]{unicode-math} 14 | \usepackage{fontspec} 15 | \setmathfont{Latin Modern Math} 16 | \setmathfont[range=\varnothing]{Asana-Math.otf} 17 | \setmathfont[range=\pitchfork]{Asana-Math.otf} 18 | \setmathfont[range=\intprod]{Asana-Math.otf} 19 | \setmathfont[range=\int]{Latin Modern Math} 20 | 21 | \usepackage[nameinlink, capitalize]{cleveref} 22 | 23 | \input{macros_common} 24 | 25 | \usepackage{amsthm} 26 | \usepackage{etexcmds} 27 | \usepackage{thmtools} 28 | 29 | \input{macros_print} 30 | 31 | \title{The sphere eversion project} 32 | \author{Patrick Massot \and Oliver Nash \and Floris van Doorn} 33 | 34 | \begin{document} 35 | \maketitle 36 | \input{content} 37 | \end{document} 38 | -------------------------------------------------------------------------------- /blueprint/src/stylecours.css: -------------------------------------------------------------------------------- 1 | div.theorem_thmcontent { 2 | border-left: .15rem solid black; 3 | } 4 | 5 | div.proposition_thmcontent { 6 | border-left: .15rem solid black; 7 | } 8 | 9 | div.lemma_thmcontent { 10 | border-left: .1rem solid black; 11 | } 12 | 13 | div.corollary_thmcontent { 14 | border-left: .1rem solid black; 15 | } 16 | 17 | div.proof_content { 18 | border-left: .08rem solid grey; 19 | } 20 | 21 | figure.subfloat span.subref { 22 | display: none; 23 | } 24 | 25 | nav.local_toc ul { 26 | font-size: 1.2rem; 27 | } 28 | 29 | @media (min-width:1024px) { 30 | nav.toc { 31 | width: 25vw; 32 | } 33 | } 34 | 35 | @media (min-width:1024px) { 36 | div.with-toc { 37 | margin-left:25vw; 38 | } 39 | } 40 | -------------------------------------------------------------------------------- /blueprint/src/web.tex: -------------------------------------------------------------------------------- 1 | \documentclass{report} 2 | 3 | \usepackage{amsmath, amsthm} 4 | \usepackage{graphicx} 5 | \DeclareGraphicsExtensions{.svg,.png,.jpg} 6 | \usepackage[capitalize]{cleveref} 7 | \usepackage[showmore, dep_graph]{blueprint} 8 | 9 | \usepackage{tikz-cd} 10 | \usepackage{tikz} 11 | 12 | \input{macros_common} 13 | \input{macros_web} 14 | 15 | \home{https://leanprover-community.github.io/sphere-eversion/} 16 | \github{https://github.com/leanprover-community/sphere-eversion/} 17 | \dochome{https://leanprover-community.github.io/sphere-eversion/docs} 18 | 19 | \title{The sphere eversion project} 20 | \author{Patrick Massot \and Oliver Nash \and Floris van Doorn} 21 | 22 | \begin{document} 23 | \maketitle 24 | \input{content} 25 | \end{document} 26 | -------------------------------------------------------------------------------- /blueprint/tasks.py: -------------------------------------------------------------------------------- 1 | import os 2 | from pathlib import Path 3 | import http.server 4 | import socketserver 5 | 6 | from invoke import run, task 7 | 8 | BP_DIR = Path(__file__).parent 9 | 10 | @task 11 | def print(ctx): 12 | cwd = os.getcwd() 13 | os.chdir(BP_DIR) 14 | run('mkdir -p print && cd src && xelatex -output-directory=../print print.tex') 15 | os.chdir(cwd) 16 | 17 | @task 18 | def bp(ctx): 19 | cwd = os.getcwd() 20 | os.chdir(BP_DIR) 21 | run('mkdir -p print && cd src && xelatex -output-directory=../print print.tex') 22 | run('cd src && xelatex -output-directory=../print print.tex') 23 | os.chdir(cwd) 24 | 25 | 26 | @task 27 | def web(ctx): 28 | cwd = os.getcwd() 29 | os.chdir(BP_DIR/'src') 30 | run('plastex -c plastex.cfg web.tex') 31 | os.chdir(cwd) 32 | 33 | @task 34 | def serve(ctx): 35 | cwd = os.getcwd() 36 | os.chdir(BP_DIR/'web') 37 | Handler = http.server.SimpleHTTPRequestHandler 38 | httpd = socketserver.TCPServer(("", 8000), Handler) 39 | httpd.serve_forever() 40 | os.chdir(cwd) 41 | 42 | -------------------------------------------------------------------------------- /docs_src/index.md: -------------------------------------------------------------------------------- 1 | --- 2 | title: The sphere eversion project 3 | --- 4 | 5 | This project is a formalization of the proof of existence of 6 | [sphere eversions](https://www.youtube.com/watch?v=wO61D9x6lNY) 7 | using the [Lean theorem prover](https://leanprover.github.io/), 8 | mainly developed at [Microsoft Research](https://www.microsoft.com/en-us/research/) 9 | by [Leonardo de Moura](https://leodemoura.github.io/). 10 | More precisely we formalized the full *h*-principle for open and 11 | ample first order differential relations, and deduce existence of sphere 12 | eversions as a corollary. 13 | 14 | The main motivations are: 15 | 16 | * Demonstrating the proof assistant can handle geometric topology, and 17 | not only algebra or abstract nonsense. Note that Fabian Immler and 18 | Yong Kiam Tan already pioneered this direction by formalizing 19 | Poincaré-Bendixon, but this project has much larger scale. 20 | * Exploring new infrastructure for collaborations on formalization 21 | projects, using the [interactive blueprint](blueprint/index.html). 22 | * Producing a bilingual informal/formal document by keeping the 23 | blueprint and the formalization in sync. 24 | 25 | The main statements that we formalized appear in [main.lean](https://github.com/leanprover-community/sphere-eversion/blob/master/src/main.lean). In particular the sphere eversion statement is: 26 | 27 | ```lean 28 | theorem Smale : 29 | -- There exists a family of functions `f t` indexed by `ℝ` going from `𝕊²` to `ℝ³` such that 30 | ∃ f : ℝ → 𝕊² → ℝ³, 31 | -- it is smooth in both variables (for the obvious smooth structures on `ℝ × 𝕊²` and `ℝ³`) and 32 | (cont_mdiff (𝓘(ℝ, ℝ).prod (𝓡 2)) 𝓘(ℝ, ℝ³) ∞ ↿f) ∧ 33 | -- `f 0` is the inclusion map, sending `x` to `x` and 34 | (f 0 = λ x, x) ∧ 35 | -- `f 1` is the antipodal map, sending `x` to `-x` and 36 | (f 1 = λ x, -x) ∧ 37 | -- every `f t` is an immersion. 38 | ∀ t, immersion (𝓡 2) 𝓘(ℝ, ℝ³) (f t) 39 | ``` 40 | -------------------------------------------------------------------------------- /docs_src/pygments.css: -------------------------------------------------------------------------------- 1 | .highlight .hll { background-color: #ffffcc } 2 | .highlight .c { color: #60a0b0; font-style: italic } /* Comment */ 3 | .highlight .err { border: 1px solid #FF0000 } /* Error */ 4 | .highlight .k { color: #007020; font-weight: bold } /* Keyword */ 5 | .highlight .o { color: #666666 } /* Operator */ 6 | .highlight .cm { color: #60a0b0; font-style: italic } /* Comment.Multiline */ 7 | .highlight .cp { color: #007020 } /* Comment.Preproc */ 8 | .highlight .c1 { color: #60a0b0; font-style: italic } /* Comment.Single */ 9 | .highlight .cs { color: #60a0b0; background-color: #fff0f0 } /* Comment.Special */ 10 | .highlight .gd { color: #A00000 } /* Generic.Deleted */ 11 | .highlight .ge { font-style: italic } /* Generic.Emph */ 12 | .highlight .gr { color: #FF0000 } /* Generic.Error */ 13 | .highlight .gh { color: #000080; font-weight: bold } /* Generic.Heading */ 14 | .highlight .gi { color: #00A000 } /* Generic.Inserted */ 15 | .highlight .go { color: #808080 } /* Generic.Output */ 16 | .highlight .gp { color: #c65d09; font-weight: bold } /* Generic.Prompt */ 17 | .highlight .gs { font-weight: bold } /* Generic.Strong */ 18 | .highlight .gu { color: #800080; font-weight: bold } /* Generic.Subheading */ 19 | .highlight .gt { color: #0040D0 } /* Generic.Traceback */ 20 | .highlight .kc { color: #007020; font-weight: bold } /* Keyword.Constant */ 21 | .highlight .kd { color: #007020; font-weight: bold } /* Keyword.Declaration */ 22 | .highlight .kn { color: #007020; font-weight: bold } /* Keyword.Namespace */ 23 | .highlight .kp { color: #007020 } /* Keyword.Pseudo */ 24 | .highlight .kr { color: #007020; font-weight: bold } /* Keyword.Reserved */ 25 | .highlight .kt { color: #902000 } /* Keyword.Type */ 26 | .highlight .m { color: #40a070 } /* Literal.Number */ 27 | .highlight .s { color: #4070a0 } /* Literal.String */ 28 | .highlight .na { color: #4070a0 } /* Name.Attribute */ 29 | .highlight .nb { color: #007020 } /* Name.Builtin */ 30 | .highlight .nc { color: #0e84b5; font-weight: bold } /* Name.Class */ 31 | .highlight .no { color: #60add5 } /* Name.Constant */ 32 | .highlight .nd { color: #555555; font-weight: bold } /* Name.Decorator */ 33 | .highlight .ni { color: #d55537; font-weight: bold } /* Name.Entity */ 34 | .highlight .ne { color: #007020 } /* Name.Exception */ 35 | .highlight .nf { color: #06287e } /* Name.Function */ 36 | .highlight .nl { color: #002070; font-weight: bold } /* Name.Label */ 37 | .highlight .nn { color: #0e84b5; font-weight: bold } /* Name.Namespace */ 38 | .highlight .nt { color: #062873; font-weight: bold } /* Name.Tag */ 39 | .highlight .nv { color: #bb60d5 } /* Name.Variable */ 40 | .highlight .ow { color: #007020; font-weight: bold } /* Operator.Word */ 41 | .highlight .w { color: #bbbbbb } /* Text.Whitespace */ 42 | .highlight .mf { color: #40a070 } /* Literal.Number.Float */ 43 | .highlight .mh { color: #40a070 } /* Literal.Number.Hex */ 44 | .highlight .mi { color: #40a070 } /* Literal.Number.Integer */ 45 | .highlight .mo { color: #40a070 } /* Literal.Number.Oct */ 46 | .highlight .sb { color: #4070a0 } /* Literal.String.Backtick */ 47 | .highlight .sc { color: #4070a0 } /* Literal.String.Char */ 48 | .highlight .sd { color: #4070a0; font-style: italic } /* Literal.String.Doc */ 49 | .highlight .s2 { color: #4070a0 } /* Literal.String.Double */ 50 | .highlight .se { color: #4070a0; font-weight: bold } /* Literal.String.Escape */ 51 | .highlight .sh { color: #4070a0 } /* Literal.String.Heredoc */ 52 | .highlight .si { color: #70a0d0; font-style: italic } /* Literal.String.Interpol */ 53 | .highlight .sx { color: #c65d09 } /* Literal.String.Other */ 54 | .highlight .sr { color: #235388 } /* Literal.String.Regex */ 55 | .highlight .s1 { color: #4070a0 } /* Literal.String.Single */ 56 | .highlight .ss { color: #517918 } /* Literal.String.Symbol */ 57 | .highlight .bp { color: #007020 } /* Name.Builtin.Pseudo */ 58 | .highlight .vc { color: #bb60d5 } /* Name.Variable.Class */ 59 | .highlight .vg { color: #bb60d5 } /* Name.Variable.Global */ 60 | .highlight .vi { color: #bb60d5 } /* Name.Variable.Instance */ 61 | .highlight .il { color: #40a070 } /* Literal.Number.Integer.Long */ 62 | -------------------------------------------------------------------------------- /docs_src/style.css: -------------------------------------------------------------------------------- 1 | /*! normalize.css v3.0.2 | MIT License | git.io/normalize */@import url("https://fonts.googleapis.com/css?family=Open+Sans:400,700");html{font-family:sans-serif;-ms-text-size-adjust:100%;-webkit-text-size-adjust:100%}body{margin:0}article,aside,details,figcaption,figure,footer,header,hgroup,main,menu,nav,section,summary{display:block}audio,canvas,progress,video{display:inline-block;vertical-align:baseline}audio:not([controls]){display:none;height:0}[hidden],template{display:none}a{background-color:transparent}a:active,a:hover{outline:0}abbr[title]{border-bottom:1px dotted}b,strong{font-weight:bold}dfn{font-style:italic}h1{font-size:2em;margin:0.67em 0}mark{background:#ff0;color:#000}small{font-size:80%}sub,sup{font-size:75%;line-height:0;position:relative;vertical-align:baseline}sup{top:-0.5em}sub{bottom:-0.25em}img{border:0}svg:not(:root){overflow:hidden}figure{margin:1em 40px}hr{box-sizing:content-box;height:0}pre{overflow:auto}code,kbd,pre,samp{font-family:monospace, monospace;font-size:1em}button,input,optgroup,select,textarea{color:inherit;font:inherit;margin:0}button{overflow:visible}button,select{text-transform:none}button,html input[type="button"],input[type="reset"],input[type="submit"]{-webkit-appearance:button;cursor:pointer}button[disabled],html input[disabled]{cursor:default}button::-moz-focus-inner,input::-moz-focus-inner{border:0;padding:0}input{line-height:normal}input[type="checkbox"],input[type="radio"]{box-sizing:border-box;padding:0}input[type="number"]::-webkit-inner-spin-button,input[type="number"]::-webkit-outer-spin-button{height:auto}input[type="search"]{-webkit-appearance:textfield;box-sizing:content-box}input[type="search"]::-webkit-search-cancel-button,input[type="search"]::-webkit-search-decoration{-webkit-appearance:none}fieldset{border:1px solid #c0c0c0;margin:0 2px;padding:0.35em 0.625em 0.75em}legend{border:0;padding:0}textarea{overflow:auto}optgroup{font-weight:bold}table{border-collapse:collapse;border-spacing:0}td,th{padding:0}.highlight table td{padding:5px}.highlight table pre{margin:0}.highlight .cm{color:#999988;font-style:italic}.highlight .cp{color:#999999;font-weight:bold}.highlight .c1{color:#999988;font-style:italic}.highlight .cs{color:#999999;font-weight:bold;font-style:italic}.highlight .c,.highlight .cd{color:#999988;font-style:italic}.highlight .err{color:#a61717;background-color:#e3d2d2}.highlight .gd{color:#000000;background-color:#ffdddd}.highlight .ge{color:#000000;font-style:italic}.highlight .gr{color:#aa0000}.highlight .gh{color:#999999}.highlight .gi{color:#000000;background-color:#ddffdd}.highlight .go{color:#888888}.highlight .gp{color:#555555}.highlight .gs{font-weight:bold}.highlight .gu{color:#aaaaaa}.highlight .gt{color:#aa0000}.highlight .kc{color:#000000;font-weight:bold}.highlight .kd{color:#000000;font-weight:bold}.highlight .kn{color:#000000;font-weight:bold}.highlight .kp{color:#000000;font-weight:bold}.highlight .kr{color:#000000;font-weight:bold}.highlight .kt{color:#445588;font-weight:bold}.highlight .k,.highlight .kv{color:#000000;font-weight:bold}.highlight .mf{color:#009999}.highlight .mh{color:#009999}.highlight .il{color:#009999}.highlight .mi{color:#009999}.highlight .mo{color:#009999}.highlight .m,.highlight .mb,.highlight .mx{color:#009999}.highlight .sb{color:#d14}.highlight .sc{color:#d14}.highlight .sd{color:#d14}.highlight .s2{color:#d14}.highlight .se{color:#d14}.highlight .sh{color:#d14}.highlight .si{color:#d14}.highlight .sx{color:#d14}.highlight .sr{color:#009926}.highlight .s1{color:#d14}.highlight .ss{color:#990073}.highlight .s{color:#d14}.highlight .na{color:#008080}.highlight .bp{color:#999999}.highlight .nb{color:#0086B3}.highlight .nc{color:#445588;font-weight:bold}.highlight .no{color:#008080}.highlight .nd{color:#3c5d5d;font-weight:bold}.highlight .ni{color:#800080}.highlight .ne{color:#990000;font-weight:bold}.highlight .nf{color:#990000;font-weight:bold}.highlight .nl{color:#990000;font-weight:bold}.highlight .nn{color:#555555}.highlight .nt{color:#000080}.highlight .vc{color:#008080}.highlight .vg{color:#008080}.highlight .vi{color:#008080}.highlight .nv{color:#008080}.highlight .ow{color:#000000;font-weight:bold}.highlight .o{color:#000000;font-weight:bold}.highlight .w{color:#bbbbbb}.highlight{background-color:#f8f8f8}*{box-sizing:border-box}body{padding:0;margin:0;font-family:"Open Sans", "Helvetica Neue", Helvetica, Arial, sans-serif;font-size:16px;line-height:1.5;color:#606c71}a{color:#1e6bb8;text-decoration:none}a:hover{text-decoration:underline}.btn{display:inline-block;margin-bottom:1rem;color:rgba(255,255,255,0.7);background-color:rgba(255,255,255,0.08);border-color:rgba(255,255,255,0.2);border-style:solid;border-width:1px;border-radius:0.3rem;transition:color 0.2s, background-color 0.2s, border-color 0.2s}.btn:hover{color:rgba(255,255,255,0.8);text-decoration:none;background-color:rgba(255,255,255,0.2);border-color:rgba(255,255,255,0.3)}.btn+.btn{margin-left:1rem}@media screen and (min-width: 64em){.btn{padding:0.75rem 1rem}}@media screen and (min-width: 42em) and (max-width: 64em){.btn{padding:0.6rem 0.9rem;font-size:0.9rem}}@media screen and (max-width: 42em){.btn{display:block;width:100%;padding:0.75rem;font-size:0.9rem}.btn+.btn{margin-top:1rem;margin-left:0}}.page-header{color:#fff;text-align:center;background-color:#159957;background-image:linear-gradient(120deg, #155799, #159957)}@media screen and (min-width: 64em){.page-header{padding:5rem 6rem}}@media screen and (min-width: 42em) and (max-width: 64em){.page-header{padding:3rem 4rem}}@media screen and (max-width: 42em){.page-header{padding:2rem 1rem}}.project-name{margin-top:0;margin-bottom:0.1rem}@media screen and (min-width: 64em){.project-name{font-size:3.25rem}}@media screen and (min-width: 42em) and (max-width: 64em){.project-name{font-size:2.25rem}}@media screen and (max-width: 42em){.project-name{font-size:1.75rem}}.project-tagline{margin-bottom:2rem;font-weight:normal;opacity:0.7}@media screen and (min-width: 64em){.project-tagline{font-size:1.25rem}}@media screen and (min-width: 42em) and (max-width: 64em){.project-tagline{font-size:1.15rem}}@media screen and (max-width: 42em){.project-tagline{font-size:1rem}}.main-content{word-wrap:break-word}.main-content :first-child{margin-top:0}@media screen and (min-width: 64em){.main-content{max-width:64rem;padding:2rem 6rem;margin:0 auto;font-size:1.1rem}}@media screen and (min-width: 42em) and (max-width: 64em){.main-content{padding:2rem 4rem;font-size:1.1rem}}@media screen and (max-width: 42em){.main-content{padding:2rem 1rem;font-size:1rem}}.main-content img{max-width:100%}.main-content h1,.main-content h2,.main-content h3,.main-content h4,.main-content h5,.main-content h6{margin-top:2rem;margin-bottom:1rem;font-weight:normal;color:#159957}.main-content p{margin-bottom:1em}.main-content code{padding:2px 4px;font-family:Consolas, "Liberation Mono", Menlo, Courier, monospace;font-size:0.9rem;color:#567482;background-color:#f3f6fa;border-radius:0.3rem}.main-content pre{padding:0.8rem;margin-top:0;margin-bottom:1rem;font:1rem Consolas, "Liberation Mono", Menlo, Courier, monospace;color:#567482;word-wrap:normal;background-color:#f3f6fa;border:solid 1px #dce6f0;border-radius:0.3rem}.main-content pre>code{padding:0;margin:0;font-size:0.9rem;color:#567482;word-break:normal;white-space:pre;background:transparent;border:0}.main-content .highlight{margin-bottom:1rem}.main-content .highlight pre{margin-bottom:0;word-break:normal}.main-content .highlight pre,.main-content pre{padding:0.8rem;overflow:auto;font-size:0.9rem;line-height:1.45;border-radius:0.3rem;-webkit-overflow-scrolling:touch}.main-content pre code,.main-content pre tt{display:inline;max-width:initial;padding:0;margin:0;overflow:initial;line-height:inherit;word-wrap:normal;background-color:transparent;border:0}.main-content pre code:before,.main-content pre code:after,.main-content pre tt:before,.main-content pre tt:after{content:normal}.main-content ul,.main-content ol{margin-top:0}.main-content blockquote{padding:0 1rem;margin-left:0;color:#819198;border-left:0.3rem solid #dce6f0}.main-content blockquote>:first-child{margin-top:0}.main-content blockquote>:last-child{margin-bottom:0}.main-content table{display:block;width:100%;overflow:auto;word-break:normal;word-break:keep-all;-webkit-overflow-scrolling:touch}.main-content table th{font-weight:bold}.main-content table th,.main-content table td{padding:0.5rem 1rem;border:1px solid #e9ebec}.main-content dl{padding:0}.main-content dl dt{padding:0;margin-top:1rem;font-size:1rem;font-weight:bold}.main-content dl dd{padding:0;margin-bottom:1rem}.main-content hr{height:2px;padding:0;margin:1rem 0;background-color:#eff0f1;border:0}.site-footer{padding-top:2rem;margin-top:2rem;border-top:solid 1px #eff0f1}@media screen and (min-width: 64em){.site-footer{font-size:1rem}}@media screen and (min-width: 42em) and (max-width: 64em){.site-footer{font-size:1rem}}@media screen and (max-width: 42em){.site-footer{font-size:0.9rem}}.site-footer-owner{display:block;font-weight:bold}.site-footer-credits{color:#819198}.page-header a {color: #fff; text-decoration: none} 2 | -------------------------------------------------------------------------------- /docs_src/template.html: -------------------------------------------------------------------------------- 1 | 2 | 3 |
4 | 5 | 6 |