├── .github └── workflows │ ├── main.yml │ └── push.yml ├── .gitignore ├── .vscode └── settings.json ├── FltRegular.lean ├── FltRegular ├── CaseI │ ├── AuxLemmas.lean │ └── Statement.lean ├── CaseII │ ├── AuxLemmas.lean │ ├── InductionStep.lean │ └── Statement.lean ├── FLT5.lean ├── FltRegular.lean ├── MayAssume │ └── Lemmas.lean └── NumberTheory │ ├── Cyclotomic │ ├── CaseI.lean │ ├── CyclRat.lean │ ├── CyclotomicUnits.lean │ ├── GaloisActionOnCyclo.lean │ ├── MoreLemmas.lean │ └── UnitLemmas.lean │ ├── CyclotomicRing.lean │ ├── GaloisPrime.lean │ ├── Hilbert90.lean │ ├── Hilbert92.lean │ ├── Hilbert94.lean │ ├── IdealNorm.lean │ ├── KummersLemma │ ├── Field.lean │ └── KummersLemma.lean │ ├── RegularPrimes.lean │ ├── SystemOfUnits.lean │ └── Unramified.lean ├── LICENSE ├── README.md ├── blueprint ├── print │ ├── print.aux │ ├── print.bbl │ ├── print.blg │ ├── print.log │ ├── print.out │ ├── print.pdf │ └── print.toc ├── requirements.txt ├── src │ ├── blueprint.sty │ ├── content.tex │ ├── demo.tex │ ├── plastex.cfg │ ├── preamble.tex │ ├── print.tex │ ├── references.bib │ ├── style.css │ └── web.tex └── tasks.py ├── flt_regular.json ├── lake-manifest.json ├── lakefile.toml ├── lean-toolchain ├── scripts └── noshake.json └── unseen.lean /.github/workflows/main.yml: -------------------------------------------------------------------------------- 1 | on: workflow_dispatch 2 | 3 | 4 | # Sets permissions of the GITHUB_TOKEN to allow deployment to GitHub Pages 5 | permissions: 6 | contents: write 7 | 8 | jobs: 9 | build_project: 10 | runs-on: ubuntu-latest 11 | name: Build project 12 | concurrency: 13 | group: ${{ github.workflow }}-${{ github.ref }} 14 | cancel-in-progress: true 15 | steps: 16 | - name: Checkout project 17 | uses: actions/checkout@v2 18 | with: 19 | fetch-depth: 0 20 | 21 | - name: Install elan 22 | run: curl https://raw.githubusercontent.com/leanprover/elan/master/elan-init.sh -sSf | sh -s -- -y --default-toolchain none 23 | 24 | - name: Get cache 25 | run: ~/.elan/bin/lake -Kenv=dev exe cache get 26 | 27 | - name: Build project 28 | run: ~/.elan/bin/lake -Kenv=dev build FltRegular 29 | 30 | - name: Build documentation 31 | run: ~/.elan/bin/lake -Kenv=dev build Std:docs Mathlib:docs FltRegular:docs 32 | 33 | - name: Install Python 34 | uses: actions/setup-python@v4 35 | with: 36 | python-version: '3.9' 37 | cache: 'pip' # caching pip dependencies 38 | 39 | - name: Install blueprint apt deps 40 | run: | 41 | sudo apt-get update 42 | sudo apt-get install -y graphviz libgraphviz-dev pdf2svg dvisvgm texlive-full 43 | 44 | - name: Install blueprint deps 45 | run: | 46 | cd blueprint && pip install -r requirements.txt 47 | 48 | - name: Build blueprint 49 | run: | 50 | cd blueprint && inv all 51 | 52 | - name: Copy documentation 53 | run: | 54 | mkdir -p _site 55 | mv build/doc _site/doc 56 | mv blueprint/web _site/blueprint 57 | 58 | - name: Deploy doc & blueprint 59 | uses: peaceiris/actions-gh-pages@v3 60 | with: 61 | github_token: ${{ secrets.github_token }} 62 | publish_dir: _site 63 | -------------------------------------------------------------------------------- /.github/workflows/push.yml: -------------------------------------------------------------------------------- 1 | on: 2 | push: 3 | branches: 4 | - master 5 | workflow_dispatch: 6 | 7 | # Sets permissions of the GITHUB_TOKEN to allow deployment to GitHub Pages 8 | permissions: 9 | contents: read 10 | pages: write 11 | id-token: write 12 | 13 | jobs: 14 | build_project: 15 | runs-on: ubuntu-latest 16 | name: Build project 17 | concurrency: 18 | group: ${{ github.workflow }}-${{ github.ref }} 19 | cancel-in-progress: true 20 | steps: 21 | - name: Free Disk Space (Ubuntu) 22 | uses: jlumbroso/free-disk-space@main 23 | with: 24 | # this might remove tools that are actually needed, 25 | # if set to "true" but frees about 6 GB 26 | tool-cache: false 27 | 28 | # all of these default to true, but feel free to set to 29 | # "false" if necessary for your workflow 30 | android: true 31 | dotnet: true 32 | haskell: true 33 | large-packages: true 34 | docker-images: true 35 | swap-storage: true 36 | 37 | - name: Checkout project 38 | uses: actions/checkout@v2 39 | with: 40 | fetch-depth: 0 41 | 42 | - name: Install elan 43 | run: curl https://raw.githubusercontent.com/leanprover/elan/master/elan-init.sh -sSf | sh -s -- -y --default-toolchain none 44 | 45 | - name: Get cache 46 | run: ~/.elan/bin/lake exe cache get || true 47 | 48 | - name: Build project 49 | run: ~/.elan/bin/lake -R -Kenv=dev build FltRegular 50 | 51 | - uses: actions/cache@v3 52 | name: Mathlib doc Cache 53 | with: 54 | path: build/doc/Mathlib 55 | key: DocGen4-${{ hashFiles('lake-manifest.json') }} 56 | restore-keys: | 57 | DocGen4- 58 | 59 | - name: Build documentation 60 | run: ~/.elan/bin/lake -R -Kenv=dev build FltRegular:docs 61 | 62 | - name: Install Python 63 | uses: actions/setup-python@v4 64 | with: 65 | python-version: '3.9' 66 | cache: 'pip' # caching pip dependencies 67 | 68 | - name: Install blueprint apt deps 69 | run: | 70 | sudo apt-get update 71 | sudo apt-get install -y graphviz libgraphviz-dev pdf2svg dvisvgm texlive-full 72 | 73 | - name: Install blueprint deps 74 | run: | 75 | cd blueprint && pip install -r requirements.txt 76 | 77 | - name: Build blueprint 78 | run: | 79 | cd blueprint && inv all 80 | 81 | - name: Copy documentation 82 | run: | 83 | mkdir -p _site 84 | mv .lake/build/doc _site/doc 85 | mv blueprint/web _site/blueprint 86 | 87 | - name: Upload doc & blueprint artifact 88 | uses: actions/upload-pages-artifact@v3 89 | with: 90 | path: '_site' 91 | 92 | - name: Deploy to GitHub Pages 93 | id: deployment 94 | uses: actions/deploy-pages@v4 95 | -------------------------------------------------------------------------------- /.gitignore: -------------------------------------------------------------------------------- 1 | # macOS leaves these files everywhere: 2 | .DS_Store 3 | # Created by `lake exe cache get` if no home directory is available 4 | /.cache 5 | # Prior to v4.3.0-rc2 lake stored files in these locations. 6 | # We'll leave them in the `.gitignore` for a while for users switching between toolchains. 7 | /build/ 8 | /lake-packages/ 9 | /lakefile.olean 10 | # After v4.3.0-rc2 lake stores its files here: 11 | /.lake/ 12 | -------------------------------------------------------------------------------- /.vscode/settings.json: -------------------------------------------------------------------------------- 1 | { 2 | "files.exclude": { 3 | "**/.git": true, 4 | "**/.DS_Store": true, 5 | "**/*.olean": true, 6 | "**/.DS_yml": true, 7 | "**/.gitpod.yml": true, 8 | "**/.vscode": true, 9 | ".docker": true, 10 | "build": true, 11 | "html": true, 12 | "lake-packages": true, 13 | ".gitignore": true, 14 | "lake-manifest.json": true, 15 | "lakefile.lean": true, 16 | "README.md": true, 17 | "Library.lean": true, 18 | "lean-toolchain": true, 19 | "img": true, 20 | ".lake": true, 21 | ".github": true, 22 | "LICENSE":true, 23 | ".devcontainer":true, 24 | "lean-tactics.tex":true, 25 | }, 26 | "editor.insertSpaces": true, 27 | "editor.tabSize": 2, 28 | "editor.rulers" : [100], 29 | "files.encoding": "utf8", 30 | "files.eol": "\n", 31 | "files.insertFinalNewline": true, 32 | "files.trimFinalNewlines": true, 33 | "files.trimTrailingWhitespace": true, 34 | "cSpell.enabledFileTypes": { 35 | "lean4": false 36 | }, 37 | } 38 | -------------------------------------------------------------------------------- /FltRegular.lean: -------------------------------------------------------------------------------- 1 | import FltRegular.CaseI.AuxLemmas 2 | import FltRegular.CaseI.Statement 3 | import FltRegular.CaseII.AuxLemmas 4 | import FltRegular.CaseII.InductionStep 5 | import FltRegular.CaseII.Statement 6 | import FltRegular.FLT5 7 | import FltRegular.FltRegular 8 | import FltRegular.MayAssume.Lemmas 9 | import FltRegular.NumberTheory.Cyclotomic.CaseI 10 | import FltRegular.NumberTheory.Cyclotomic.CyclRat 11 | import FltRegular.NumberTheory.Cyclotomic.CyclotomicUnits 12 | import FltRegular.NumberTheory.Cyclotomic.GaloisActionOnCyclo 13 | import FltRegular.NumberTheory.Cyclotomic.MoreLemmas 14 | import FltRegular.NumberTheory.Cyclotomic.UnitLemmas 15 | import FltRegular.NumberTheory.CyclotomicRing 16 | import FltRegular.NumberTheory.GaloisPrime 17 | import FltRegular.NumberTheory.Hilbert90 18 | import FltRegular.NumberTheory.Hilbert92 19 | import FltRegular.NumberTheory.Hilbert94 20 | import FltRegular.NumberTheory.IdealNorm 21 | import FltRegular.NumberTheory.KummersLemma.Field 22 | import FltRegular.NumberTheory.KummersLemma.KummersLemma 23 | import FltRegular.NumberTheory.RegularPrimes 24 | import FltRegular.NumberTheory.SystemOfUnits 25 | import FltRegular.NumberTheory.Unramified 26 | -------------------------------------------------------------------------------- /FltRegular/CaseI/AuxLemmas.lean: -------------------------------------------------------------------------------- 1 | import FltRegular.NumberTheory.Cyclotomic.UnitLemmas 2 | import FltRegular.NumberTheory.Cyclotomic.CyclRat 3 | 4 | open Finset Nat IsCyclotomicExtension Ideal Polynomial Int Basis 5 | 6 | open scoped BigOperators NumberField 7 | 8 | namespace FltRegular 9 | 10 | variable {p : ℕ} (hpri : p.Prime) 11 | 12 | local notation "K" => CyclotomicField p ℚ 13 | 14 | local notation "R" => 𝓞 K 15 | 16 | namespace CaseI 17 | 18 | theorem two_lt (hp5 : 5 ≤ p) : 2 < p := by linarith 19 | 20 | section Zerok₁ 21 | 22 | theorem aux_cong0k₁ {k : Fin p} (hcong : k ≡ -1 [ZMOD p]) : 23 | k = ⟨p.pred, pred_lt hpri.ne_zero⟩ := by 24 | refine Fin.ext ?_ 25 | rw [Fin.val_mk, ← ZMod.val_cast_of_lt (Fin.is_lt k)] 26 | suffices ((k : ℤ) : ZMod p).val = p.pred by simpa 27 | rw [← ZMod.intCast_eq_intCast_iff] at hcong 28 | rw [hcong, cast_neg, Int.cast_one, pred_eq_sub_one] 29 | haveI : NeZero p := ⟨hpri.ne_zero⟩ 30 | haveI : Fact p.Prime := ⟨hpri⟩ 31 | haveI : Fact (1 < p) := ⟨hpri.one_lt⟩ 32 | simp [ZMod.neg_val, ZMod.val_one] 33 | 34 | /-- Auxiliary function. -/ 35 | def f0k₁ (b : ℤ) (p : ℕ) : ℕ → ℤ := fun x => if x = 1 then b else if x = p.pred then -b else 0 36 | 37 | theorem auxf0k₁ (hp5 : 5 ≤ p) (b : ℤ) : ∃ i : Fin p, f0k₁ b p (i : ℕ) = 0 := by 38 | refine ⟨⟨2, two_lt hp5⟩, ?_⟩ 39 | have hpred : ((⟨2, two_lt hp5⟩ : Fin p) : ℕ) ≠ p.pred := by 40 | intro h 41 | simp only [Fin.ext_iff, Fin.val_mk] at h 42 | replace h := h.symm 43 | rw [Nat.pred_eq_sub_one] at h 44 | omega 45 | simp only [f0k₁, OfNat.ofNat_ne_one, ite_false, ite_eq_right_iff, neg_eq_zero] 46 | intro h2 47 | exfalso 48 | apply hpred 49 | simp [h2] 50 | 51 | include hpri in 52 | theorem aux0k₁ {a b c : ℤ} {ζ : R} (hp5 : 5 ≤ p) (hζ : IsPrimitiveRoot ζ p) 53 | (caseI : ¬↑p ∣ a * b * c) {k₁ k₂ : Fin p} (hcong : k₂ ≡ k₁ - 1 [ZMOD p]) 54 | (hdiv : ↑p ∣ ↑a + ↑b * ζ - ↑a * ζ ^ (k₁ : ℕ) - ↑b * ζ ^ (k₂ : ℕ)) : 0 ≠ (k₁ : ℕ) := by 55 | symm 56 | intro habs 57 | rw [show (k₁ : ℤ) = 0 by simpa using habs, zero_sub] at hcong 58 | rw [habs, _root_.pow_zero, mul_one, add_sub_cancel_left, aux_cong0k₁ hpri hcong] at hdiv 59 | nth_rw 1 [show ζ = ζ ^ ((⟨1, hpri.one_lt⟩ : Fin p) : ℕ) by simp] at hdiv 60 | have key : ↑(p : ℤ) ∣ ∑ j ∈ range p, f0k₁ b p j • ζ ^ j := by 61 | convert hdiv using 1 62 | have h : 1 ≠ p.pred := fun h => by linarith [pred_eq_succ_iff.1 h.symm] 63 | simp_rw [f0k₁, ite_smul, sum_ite, filter_filter, ← Ne.eq_def, ne_and_eq_iff_right h, 64 | Finset.range_filter_eq] 65 | simp [hpri.one_lt, Nat.sub_lt hpri.pos, sub_eq_add_neg] 66 | rw [sum_range] at key 67 | refine caseI (Dvd.dvd.mul_right (Dvd.dvd.mul_left ?_ _) _) 68 | have : NeZero p := ⟨hpri.ne_zero⟩ 69 | simpa [f0k₁] using dvd_coeff_cycl_integer hpri hζ (auxf0k₁ hp5 b) key ⟨1, hpri.one_lt⟩ 70 | 71 | end Zerok₁ 72 | 73 | section Zerok₂ 74 | 75 | /-- Auxiliary function -/ 76 | def f0k₂ (a b : ℤ) : ℕ → ℤ := fun x => if x = 0 then a - b else if x = 1 then b - a else 0 77 | 78 | theorem aux_cong0k₂ {k : Fin p} (hcong : k ≡ 1 [ZMOD p]) : k = ⟨1, hpri.one_lt⟩ := by 79 | refine Fin.ext ?_ 80 | rw [Fin.val_mk, ← ZMod.val_cast_of_lt (Fin.is_lt k)] 81 | suffices ((k : ℤ) : ZMod p).val = 1 by simpa 82 | rw [← ZMod.intCast_eq_intCast_iff] at hcong 83 | rw [hcong, Int.cast_one] 84 | haveI : Fact p.Prime := ⟨hpri⟩ 85 | simp [ZMod.val_one] 86 | 87 | theorem auxf0k₂ (hp5 : 5 ≤ p) (a b : ℤ) : ∃ i : Fin p, f0k₂ a b (i : ℕ) = 0 := 88 | ⟨⟨2, two_lt hp5⟩, rfl⟩ 89 | 90 | include hpri in 91 | theorem aux0k₂ {a b : ℤ} {ζ : R} (hp5 : 5 ≤ p) (hζ : IsPrimitiveRoot ζ p) (hab : ¬a ≡ b [ZMOD p]) 92 | {k₁ k₂ : Fin p} (hcong : k₂ ≡ k₁ - 1 [ZMOD p]) 93 | (hdiv : ↑p ∣ ↑a + ↑b * ζ - ↑a * ζ ^ (k₁ : ℕ) - ↑b * ζ ^ (k₂ : ℕ)) : (0 : ℕ) ≠ ↑k₂ := by 94 | symm 95 | intro habs 96 | replace hcong := hcong.symm 97 | rw [show (k₂ : ℤ) = 0 by simpa using habs, ← ZMod.intCast_eq_intCast_iff, Int.cast_sub, 98 | Int.cast_zero, sub_eq_zero, ZMod.intCast_eq_intCast_iff] at hcong 99 | rw [habs, _root_.pow_zero, mul_one, aux_cong0k₂ hpri hcong, Fin.val_mk, pow_one, add_sub_assoc, 100 | ← sub_mul, add_sub_right_comm, show ζ = ζ ^ ((⟨1, hpri.one_lt⟩ : Fin p) : ℕ) by simp] at hdiv 101 | have key : ↑(p : ℤ) ∣ ∑ j ∈ range p, f0k₂ a b j • ζ ^ j := by 102 | convert hdiv using 1 103 | simp_rw [f0k₂, ite_smul, sum_ite, filter_filter, ← Ne.eq_def, ne_and_eq_iff_right zero_ne_one, 104 | Finset.range_filter_eq] 105 | simp only [hpri.pos, hpri.one_lt, if_true, zsmul_eq_mul, Int.cast_sub, sum_singleton, 106 | _root_.pow_zero, mul_one, pow_one, Ne, zero_smul, sum_const_zero, add_zero, Fin.val_mk] 107 | rw [sum_range] at key 108 | refine hab ?_ 109 | symm 110 | rw [← ZMod.intCast_eq_intCast_iff, ZMod.intCast_eq_intCast_iff_dvd_sub] 111 | have : NeZero p := ⟨hpri.ne_zero⟩ 112 | simpa [f0k₂] using dvd_coeff_cycl_integer hpri hζ (auxf0k₂ hp5 a b) key ⟨0, hpri.pos⟩ 113 | 114 | end Zerok₂ 115 | 116 | section OnekOne 117 | 118 | theorem aux_cong1k₁ {k : Fin p} (hcong : k ≡ 0 [ZMOD p]) : k = ⟨0, hpri.pos⟩ := by 119 | refine Fin.ext ?_ 120 | rw [Fin.val_mk, ← ZMod.val_cast_of_lt (Fin.is_lt k)] 121 | suffices ((k : ℤ) : ZMod p).val = 0 by simpa 122 | rw [← ZMod.intCast_eq_intCast_iff] at hcong 123 | rw [hcong, Int.cast_zero] 124 | haveI : Fact p.Prime := ⟨hpri⟩ 125 | simp 126 | 127 | include hpri in 128 | theorem aux1k₁ {a b : ℤ} {ζ : R} (hp5 : 5 ≤ p) (hζ : IsPrimitiveRoot ζ p) (hab : ¬a ≡ b [ZMOD p]) 129 | {k₁ k₂ : Fin p} (hcong : k₂ ≡ k₁ - 1 [ZMOD p]) 130 | (hdiv : ↑p ∣ ↑a + ↑b * ζ - ↑a * ζ ^ (k₁ : ℕ) - ↑b * ζ ^ (k₂ : ℕ)) : (1 : ℕ) ≠ k₁ := by 131 | intro habs 132 | have h := aux0k₂ hpri hp5 hζ hab hcong hdiv 133 | rw [show (k₁ : ℤ) = 1 by simpa using habs.symm, sub_self] at hcong 134 | have := aux_cong1k₁ hpri hcong 135 | simp only [← Fin.val_eq_val, Fin.val_mk] at this 136 | exact h.symm this 137 | 138 | end OnekOne 139 | 140 | section OnekTwo 141 | 142 | /-- Auxiliary function -/ 143 | def f1k₂ (a : ℤ) : ℕ → ℤ := fun x => if x = 0 then a else if x = 2 then -a else 0 144 | 145 | theorem aux_cong1k₂ {k : Fin p} (hpri : p.Prime) (hp5 : 5 ≤ p) (hcong : k ≡ 1 + 1 [ZMOD p]) : 146 | k = ⟨2, two_lt hp5⟩ := by 147 | refine Fin.ext ?_ 148 | rw [Fin.val_mk, ← ZMod.val_cast_of_lt (Fin.is_lt k)] 149 | suffices ((k : ℤ) : ZMod p).val = 2 by simpa 150 | rw [← ZMod.intCast_eq_intCast_iff] at hcong 151 | rw [hcong] 152 | simp only [Int.cast_add, algebraMap.coe_one] 153 | haveI : Fact p.Prime := ⟨hpri⟩ 154 | have := congr_arg Nat.succ (Nat.succ_pred_eq_of_pos hpri.pred_pos) 155 | rw [succ_pred_prime hpri] at this 156 | rw [ZMod.val_add, Int.cast_one, ZMod.val_one, ← Nat.mod_add_mod, ← this, one_mod, this, 157 | Nat.mod_eq_of_lt] 158 | linarith 159 | 160 | include hpri in 161 | theorem auxf1k₂ (a : ℤ) : ∃ i : Fin p, f1k₂ a i = 0 := 162 | ⟨⟨1, hpri.one_lt⟩, rfl⟩ 163 | 164 | include hpri in 165 | theorem aux1k₂ {a b c : ℤ} {ζ : R} (hp5 : 5 ≤ p) (hζ : IsPrimitiveRoot ζ p) 166 | (caseI : ¬↑p ∣ a * b * c) {k₁ k₂ : Fin p} (hcong : k₂ ≡ k₁ - 1 [ZMOD p]) 167 | (hdiv : ↑p ∣ ↑a + ↑b * ζ - ↑a * ζ ^ (k₁ : ℕ) - ↑b * ζ ^ (k₂ : ℕ)) : (1 : ℕ) ≠ k₂ := by 168 | symm 169 | intro habs 170 | replace hcong := hcong.symm 171 | rw [show (k₂ : ℤ) = 1 by simpa using habs, ← ZMod.intCast_eq_intCast_iff, Int.cast_sub, 172 | sub_eq_iff_eq_add, ← Int.cast_add, ZMod.intCast_eq_intCast_iff] at hcong 173 | rw [habs, pow_one, aux_cong1k₂ hpri hp5 hcong] at hdiv 174 | ring_nf at hdiv 175 | have key : ↑(p : ℤ) ∣ ∑ j ∈ range p, f1k₂ a j • ζ ^ j := by 176 | suffices ∑ j ∈ range p, f1k₂ a j • ζ ^ j = ↑a - ↑a * ζ ^ 2 by 177 | rwa [this] 178 | simp_rw [f1k₂, ite_smul, sum_ite, filter_filter, ← Ne.eq_def, ne_and_eq_iff_right 179 | (show 0 ≠ 2 by norm_num), Finset.range_filter_eq] 180 | simp only [hpri.pos, ite_true, zsmul_eq_mul, sum_singleton, _root_.pow_zero, mul_one, 181 | two_lt hp5,neg_smul, sum_neg_distrib, ne_eq, mem_range, not_and, not_not, zero_smul, 182 | sum_const_zero, add_zero] 183 | ring 184 | rw [sum_range] at key 185 | refine caseI (Dvd.dvd.mul_right (Dvd.dvd.mul_right ?_ _) _) 186 | have : NeZero p := ⟨hpri.ne_zero⟩ 187 | simpa [f1k₂] using dvd_coeff_cycl_integer hpri hζ (auxf1k₂ hpri a) key ⟨0, hpri.pos⟩ 188 | 189 | end OnekTwo 190 | 191 | section KoneKtwo 192 | 193 | theorem auxk₁k₂ {k₁ k₂ : Fin p} (hpri : p.Prime) (hcong : k₂ ≡ k₁ - 1 [ZMOD p]) : 194 | (k₁ : ℕ) ≠ (k₂ : ℕ) := by 195 | haveI := (⟨hpri⟩ : Fact p.Prime) 196 | intro habs 197 | rw [habs, ← ZMod.intCast_eq_intCast_iff, Int.cast_sub, ← sub_eq_zero] at hcong 198 | simp at hcong 199 | 200 | end KoneKtwo 201 | 202 | end CaseI 203 | 204 | end FltRegular 205 | -------------------------------------------------------------------------------- /FltRegular/CaseI/Statement.lean: -------------------------------------------------------------------------------- 1 | import FltRegular.MayAssume.Lemmas 2 | import FltRegular.NumberTheory.Cyclotomic.CaseI 3 | import FltRegular.CaseI.AuxLemmas 4 | import FltRegular.NumberTheory.RegularPrimes 5 | import Mathlib.NumberTheory.FLT.Three 6 | 7 | open Finset Nat IsCyclotomicExtension Ideal Polynomial Int Basis FltRegular.CaseI 8 | 9 | open scoped NumberField 10 | 11 | namespace FltRegular 12 | 13 | variable {p : ℕ} 14 | 15 | namespace CaseI 16 | 17 | /-- Statement of case I with additional assumptions. -/ 18 | def SlightlyEasier : Prop := 19 | ∀ ⦃a b c : ℤ⦄ {p : ℕ} [Fact p.Prime], IsRegularPrime p → 5 ≤ p → 20 | ({a, b, c} : Finset ℤ).gcd id = 1 → ¬a ≡ b [ZMOD p] → ¬↑p ∣ a * b * c → a ^ p + b ^ p ≠ c ^ p 21 | 22 | /-- Statement of case I. -/ 23 | def Statement : Prop := 24 | ∀ ⦃a b c : ℤ⦄ {p : ℕ} [Fact p.Prime], IsRegularPrime p → ¬↑p ∣ a * b * c → a ^ p + b ^ p ≠ c ^ p 25 | 26 | theorem may_assume : SlightlyEasier → Statement := by 27 | intro Heasy 28 | intro a b c p hpri hreg hI H 29 | have hodd : p ≠ 2 := by 30 | intro h 31 | rw [h] at H hI 32 | refine hI <| Dvd.dvd.mul_left ?_ _ 33 | simp only [Nat.cast_ofNat] at hI ⊢ 34 | rw [← even_iff_two_dvd, Int.not_even_iff_odd] at hI 35 | rw [← even_iff_two_dvd, ← Int.even_pow' (show 2 ≠ 0 by norm_num), ← H] 36 | exact (Int.Odd.of_mul_left (Odd.of_mul_left hI)).pow.add_odd 37 | (Int.Odd.of_mul_right (Odd.of_mul_left hI)).pow 38 | have hprod : a * b * c ≠ 0 := by 39 | intro h 40 | simp [h] at hI 41 | have hp5 : 5 ≤ p := by 42 | by_contra! habs 43 | have : p ∈ Finset.Ioo 2 5 := 44 | (Finset.mem_Ioo).2 ⟨Nat.lt_of_le_of_ne hpri.out.two_le hodd.symm, by linarith⟩ 45 | fin_cases this 46 | · exact fermatLastTheoremFor_iff_int.1 fermatLastTheoremThree a b c 47 | (fun ha ↦ hprod <| by simp [ha]) (fun hb ↦ hprod <| by simp [hb]) 48 | (fun hc ↦ hprod <| by simp [hc]) H 49 | · rw [show 2 + 1 + 1 = 2 * 2 from rfl] at hpri 50 | exact Nat.not_prime_mul one_lt_two.ne' one_lt_two.ne' hpri.out 51 | let d := ({a, b, c} : Finset ℤ).gcd id 52 | have hdiv : ¬↑p ∣ a / d * (b / d) * (c / d) := by 53 | have hadiv : d ∣ a := gcd_dvd (by simp) 54 | have hbdiv : d ∣ b := gcd_dvd (by simp) 55 | have hcdiv : d ∣ c := gcd_dvd (by simp) 56 | intro hdiv 57 | replace hdiv := dvd_mul_of_dvd_right hdiv (d * d * d) 58 | rw [mul_assoc, ← mul_assoc d, ← mul_assoc d, Int.mul_ediv_cancel' hadiv, mul_assoc, mul_comm a, 59 | mul_assoc (b / d), ← mul_assoc _ (b / d), Int.mul_ediv_cancel' hbdiv, mul_comm, mul_assoc, 60 | mul_assoc, Int.ediv_mul_cancel hcdiv, mul_comm, mul_assoc, mul_comm c, ← mul_assoc] at hdiv 61 | exact hI hdiv 62 | rcases MayAssume.coprime H hprod with ⟨Hxyz, hunit, hprodxyx⟩ 63 | obtain ⟨X, Y, Z, H1, H2, H3, _, H5⟩ := a_not_cong_b hpri.out hp5 hprodxyx Hxyz hunit hdiv 64 | exact Heasy hreg hp5 H2 H3 (fun hfin => H5 hfin) H1 65 | 66 | end CaseI 67 | 68 | theorem ab_coprime {a b c : ℤ} (H : a ^ p + b ^ p = c ^ p) (hpzero : p ≠ 0) 69 | (hgcd : ({a, b, c} : Finset ℤ).gcd id = 1) : IsCoprime a b := by 70 | rw [isCoprime_iff_gcd_eq_one] 71 | by_contra! h 72 | obtain ⟨q, hqpri, hq⟩ := exists_prime_and_dvd h 73 | replace hqpri : Prime (q : ℤ) := prime_iff_natAbs_prime.2 (by simp [hqpri]) 74 | obtain ⟨n, hn⟩ := hq 75 | have haq : ↑q ∣ a := by 76 | obtain ⟨m, hm⟩ := @Int.gcd_dvd_left a b 77 | exact ⟨n * m, by rw [hm, hn]; simp [mul_assoc]⟩ 78 | have hbq : ↑q ∣ b := by 79 | obtain ⟨m, hm⟩ := @Int.gcd_dvd_right a b 80 | exact ⟨n * m, by rw [hm, hn]; simp [mul_assoc]⟩ 81 | have hcq : ↑q ∣ c := by 82 | suffices ↑q ∣ c ^ p by exact hqpri.dvd_of_dvd_pow this 83 | rw [← H] 84 | exact dvd_add (dvd_pow haq hpzero) (dvd_pow hbq hpzero) 85 | have Hq : ↑q ∣ ({a, b, c} : Finset ℤ).gcd id := by 86 | refine dvd_gcd fun x hx ↦ ?_ 87 | simp only [mem_insert, mem_singleton] at hx 88 | rcases hx with (H | H | H) <;> simpa [H] 89 | rw [hgcd] at Hq 90 | exact hqpri.not_unit (isUnit_of_dvd_one Hq) 91 | 92 | /-- Auxiliary function -/ 93 | def f (a b : ℤ) (k₁ k₂ : ℕ) : ℕ → ℤ := fun x => 94 | if x = 0 then a else if x = 1 then b else if x = k₁ then -a else if x = k₂ then -b else 0 95 | 96 | theorem auxf' (hp5 : 5 ≤ p) (a b : ℤ) (k₁ k₂ : Fin p) : 97 | ∃ i ∈ range p, f a b k₁ k₂ (i : ℕ) = 0 := by 98 | have h0 : 0 < p := by linarith 99 | have h1 : 1 < p := by linarith 100 | let s := ({0, 1, k₁.1, k₂.1} : Finset ℕ) 101 | have : s.card ≤ 4 := by 102 | repeat refine le_trans (card_insert_le _ _) (succ_le_succ ?_) 103 | exact rfl.ge 104 | replace this : s.card < 5 := lt_of_le_of_lt this (by norm_num) 105 | have hs : s ⊆ range p := insert_subset_iff.2 ⟨mem_range.2 h0, insert_subset_iff.2 106 | ⟨mem_range.2 h1, insert_subset_iff.2 ⟨mem_range.2 (Fin.is_lt _), 107 | singleton_subset_iff.2 (mem_range.2 (Fin.is_lt _))⟩⟩⟩ 108 | have hcard := card_sdiff hs 109 | replace hcard : (range p \ s).Nonempty := by 110 | rw [← Finset.card_pos, hcard, card_range] 111 | exact Nat.sub_pos_of_lt (lt_of_lt_of_le this hp5) 112 | obtain ⟨i, hi⟩ := hcard 113 | refine ⟨i, sdiff_subset hi, ?_⟩ 114 | have hi0 : i ≠ 0 := fun h => by simp [h, s] at hi 115 | have hi1 : i ≠ 1 := fun h => by simp [h, s] at hi 116 | have hik₁ : i ≠ k₁ := fun h => by simp [h, s] at hi 117 | have hik₂ : i ≠ k₂ := fun h => by simp [h, s] at hi 118 | simp [f, hi0, hi1, hik₁, hik₂] 119 | 120 | theorem auxf (hp5 : 5 ≤ p) (a b : ℤ) (k₁ k₂ : Fin p) : ∃ i : Fin p, f a b k₁ k₂ (i : ℕ) = 0 := 121 | by 122 | obtain ⟨i, hrange, hi⟩ := auxf' hp5 a b k₁ k₂ 123 | exact ⟨⟨i, mem_range.1 hrange⟩, hi⟩ 124 | 125 | variable [hpri : Fact p.Prime] 126 | 127 | local notation "K" => CyclotomicField p ℚ 128 | 129 | local notation "R" => 𝓞 K 130 | 131 | theorem exists_ideal {a b c : ℤ} (h5p : 5 ≤ p) (H : a ^ p + b ^ p = c ^ p) 132 | (hgcd : ({ a, b, c } : Finset ℤ).gcd id = 1) 133 | (caseI : ¬↑p ∣ a * b * c) {ζ : R} (hζ : ζ ∈ nthRootsFinset p 1) : 134 | ∃ I, span ({a + ζ * b} : Set R) = I ^ p := by 135 | classical 136 | have H₁ := congr_arg (algebraMap ℤ R) H 137 | simp only [eq_intCast, Int.cast_add, Int.cast_pow] at H₁ 138 | have hζ' := (zeta_spec p ℚ K).unit'_coe 139 | rw [hζ'.pow_add_pow_eq_prod_add_mul _ _ <| 140 | odd_iff.2 <| hpri.1.eq_two_or_odd.resolve_left fun h ↦ by simp [h] at h5p] at H₁ 141 | replace H₁ := congr_arg (fun x => span ({ x } : Set R)) H₁ 142 | simp only [← prod_span_singleton, ← span_singleton_pow] at H₁ 143 | refine exists_eq_pow_of_mul_eq_pow_of_coprime (fun η₁ hη₁ η₂ hη₂ hη => ?_) H₁ ζ hζ 144 | refine fltIdeals_coprime ?_ ?_ H (ab_coprime H hpri.out.ne_zero hgcd) hη₁ hη₂ hη caseI 145 | · exact hpri.out 146 | · exact h5p 147 | 148 | theorem is_principal_aux (K' : Type*) [Field K'] [CharZero K'] [IsCyclotomicExtension {p} ℚ K'] 149 | [Fintype (ClassGroup (𝓞 K'))] 150 | {a b : ℤ} {ζ : 𝓞 K'} (hreg : p.Coprime <| Fintype.card <| ClassGroup (𝓞 K')) 151 | (I : Ideal (𝓞 K')) (hI : span ({↑a + ζ * ↑b} : Set (𝓞 K')) = I ^ p) : 152 | ∃ (u : (𝓞 K')ˣ) (α : 𝓞 K'), ↑u * α ^ p = ↑a + ζ * ↑b := by 153 | letI : NumberField K' := IsCyclotomicExtension.numberField {p} ℚ K' 154 | obtain ⟨α, hα⟩ : I.IsPrincipal := by 155 | apply isPrincipal_of_isPrincipal_pow_of_coprime hreg 156 | constructor 157 | use ↑a + ζ * ↑b 158 | rw [submodule_span_eq, hI] 159 | replace hα := congr_arg (fun (J : Submodule _ _) => J ^ p) hα 160 | simp only [← hI, submodule_span_eq, span_singleton_pow, span_singleton_eq_span_singleton] at hα 161 | obtain ⟨u, hu⟩ := hα 162 | refine ⟨u⁻¹, α, ?_⟩ 163 | rw [← hu, mul_comm ((_ + ζ * _)), ← mul_assoc] 164 | simp only [Units.inv_mul, one_mul] 165 | 166 | theorem is_principal {a b c : ℤ} {ζ : R} (hreg : IsRegularPrime p) (hp5 : 5 ≤ p) 167 | (hgcd : ({ a, b, c } : Finset ℤ).gcd id = 1) (caseI : ¬↑p ∣ a * b * c) 168 | (H : a ^ p + b ^ p = c ^ p) (hζ : IsPrimitiveRoot ζ p) : 169 | ∃ (u : Rˣ) (α : R), ↑u * α ^ p = ↑a + ζ * ↑b := by 170 | haveI := CyclotomicField.isCyclotomicExtension p ℚ 171 | replace hζ := hζ.mem_nthRootsFinset hpri.out.pos 172 | obtain ⟨I, hI⟩ := exists_ideal hp5 H hgcd caseI hζ 173 | apply is_principal_aux 174 | · rwa [IsRegularPrime, IsRegularNumber] at hreg 175 | · exact hI 176 | 177 | theorem ex_fin_div {a b c : ℤ} {ζ : R} (hp5 : 5 ≤ p) (hreg : IsRegularPrime p) 178 | (hζ : IsPrimitiveRoot ζ p) (hgcd : ({a, b, c} : Finset ℤ).gcd id = 1) (caseI : ¬↑p ∣ a * b * c) 179 | (H : a ^ p + b ^ p = c ^ p) : 180 | ∃ k₁ k₂ : Fin p, 181 | k₂ ≡ k₁ - 1 [ZMOD p] ∧ ↑p ∣ ↑a + ↑b * ζ - ↑a * ζ ^ (k₁ : ℕ) - ↑b * ζ ^ (k₂ : ℕ) := by 182 | let ζ' := (ζ : K) 183 | have hζ' : IsPrimitiveRoot ζ' p := IsPrimitiveRoot.coe_submonoidClass_iff.2 hζ 184 | have h : ζ = (hζ'.unit' : R) := by rfl 185 | have hP : p ≠ 2 := by 186 | intro hP 187 | rw [hP] at hp5 188 | contradiction 189 | obtain ⟨u, α, hu⟩ := is_principal hreg hp5 hgcd caseI H hζ 190 | rw [h, mul_comm _ (↑b : R), ← pow_one hζ'.unit'] at hu 191 | obtain ⟨k, hk⟩ := FltRegular.CaseI.exists_int_sum_eq_zero hζ' hP hpri.out a b 1 hu.symm 192 | simp only [zpow_one, zpow_neg, PNat.mk_coe, mem_span_singleton, ← h] at hk 193 | have hpcoe : (p : ℤ) ≠ 0 := by simp [hpri.out.ne_zero] 194 | refine ⟨⟨(2 * k % p).natAbs, ?_⟩, ⟨((2 * k - 1) % p).natAbs, ?_⟩, ?_, ?_⟩ 195 | repeat' 196 | rw [← natAbs_natCast p] 197 | refine natAbs_lt_natAbs_of_nonneg_of_lt (emod_nonneg _ hpcoe) ?_ 198 | rw [natAbs_natCast] 199 | exact emod_lt_of_pos _ (by simp [hpri.out.pos]) 200 | · simp only [natAbs_of_nonneg (emod_nonneg _ hpcoe), ← ZMod.intCast_eq_intCast_iff, 201 | ZMod.intCast_mod, Int.cast_sub, Int.cast_mul, Int.cast_natCast, Int.cast_one] 202 | simp only [add_sub_assoc, sub_sub] at hk ⊢ 203 | convert hk using 3 204 | rw [mul_add, mul_comm (↑a : R), ← mul_assoc _ (↑b : R), mul_comm _ (↑b : R), mul_assoc (↑b : R)] 205 | congr 2 206 | · ext 207 | simp only [Fin.val_mk, map_pow, NumberField.Units.coe_zpow, ← h] 208 | refine eq_of_div_eq_one ?_ 209 | rw [← zpow_natCast, ← zpow_sub₀ (hζ'.ne_zero hpri.out.ne_zero), hζ'.zpow_eq_one_iff_dvd] 210 | simp only [natAbs_of_nonneg (emod_nonneg _ hpcoe), ← ZMod.intCast_zmod_eq_zero_iff_dvd, 211 | Int.cast_sub, ZMod.intCast_mod, Int.cast_mul, Int.cast_natCast, sub_self] 212 | · ext 213 | simp only [Fin.val_mk, map_pow, _root_.map_mul, NumberField.Units.coe_zpow, map_units_inv, ← h] 214 | refine eq_of_div_eq_one ?_ 215 | rw [← zpow_natCast, ← zpow_sub_one₀ (hζ'.ne_zero hpri.out.ne_zero), ← 216 | zpow_sub₀ (hζ'.ne_zero hpri.out.ne_zero), hζ'.zpow_eq_one_iff_dvd] 217 | simp only [natAbs_of_nonneg (emod_nonneg _ hpcoe), ← ZMod.intCast_zmod_eq_zero_iff_dvd, 218 | Int.cast_sub, ZMod.intCast_mod, Int.cast_mul, Int.cast_natCast, Int.cast_one, sub_self] 219 | 220 | /-- Case I with additional assumptions. -/ 221 | theorem caseI_easier {a b c : ℤ} (hreg : IsRegularPrime p) (hp5 : 5 ≤ p) 222 | (hgcd : ({a, b, c} : Finset ℤ).gcd id = 1) (hab : ¬a ≡ b [ZMOD p]) (caseI : ¬↑p ∣ a * b * c) : 223 | a ^ p + b ^ p ≠ c ^ p := by 224 | have hcycl : IsCyclotomicExtension {p} ℤ (𝓞 (CyclotomicField p ℚ)) := by 225 | apply @IsCyclotomicExtension.ring_of_integers' _ _ _ (by exact hpri) _ 226 | set ζ := zeta p ℤ R 227 | have hζ := zeta_spec p ℤ R 228 | intro H 229 | obtain ⟨k₁, k₂, hcong, hdiv⟩ := ex_fin_div hp5 hreg hζ hgcd caseI H 230 | have key : ↑(p : ℤ) ∣ ∑ j ∈ range p, f a b k₁ k₂ j • ζ ^ j := by 231 | convert hdiv using 1 232 | have h01 : 0 ≠ 1 := zero_ne_one 233 | have h0k₁ := aux0k₁ hpri.out hp5 hζ caseI hcong hdiv 234 | have h0k₂ := aux0k₂ hpri.out hp5 hζ hab hcong hdiv 235 | have h1k₁ := aux1k₁ hpri.out hp5 hζ hab hcong hdiv 236 | have h1k₂ := aux1k₂ hpri.out hp5 hζ caseI hcong hdiv 237 | have hk₁k₂ : (k₁ : ℕ) ≠ (k₂ : ℕ) := auxk₁k₂ hpri.out hcong 238 | simp_rw [f, ite_smul, sum_ite, filter_filter, ← Ne.eq_def, ne_and_eq_iff_right h01, and_assoc, 239 | ne_and_eq_iff_right h1k₁, ne_and_eq_iff_right h0k₁, ne_and_eq_iff_right hk₁k₂, 240 | ne_and_eq_iff_right h1k₂, ne_and_eq_iff_right h0k₂, Finset.range_filter_eq] 241 | simp only [hpri.out.pos, hpri.out.one_lt, if_true, zsmul_eq_mul, sum_singleton, _root_.pow_zero, 242 | mul_one, pow_one, Fin.is_lt, neg_smul, sum_neg_distrib, Ne, zero_smul, sum_const_zero, 243 | add_zero] 244 | ring 245 | rw [sum_range] at key 246 | refine caseI (Dvd.dvd.mul_right (Dvd.dvd.mul_right ?_ _) _) 247 | simpa [f] using dvd_coeff_cycl_integer (by exact hpri.out) hζ (auxf hp5 a b k₁ k₂) key 248 | ⟨0, hpri.out.pos⟩ 249 | 250 | /-- CaseI. -/ 251 | theorem caseI {a b c : ℤ} {p : ℕ} [Fact p.Prime] (hreg : IsRegularPrime p) 252 | (caseI : ¬↑p ∣ a * b * c) : a ^ p + b ^ p ≠ c ^ p := 253 | FltRegular.CaseI.may_assume 254 | (fun _ _ _ _ _ Hreg Hp5 Hunit Hxy HI H => caseI_easier Hreg Hp5 Hunit Hxy HI H) hreg 255 | caseI 256 | 257 | end FltRegular 258 | -------------------------------------------------------------------------------- /FltRegular/CaseII/AuxLemmas.lean: -------------------------------------------------------------------------------- 1 | import Mathlib.RingTheory.ClassGroup 2 | 3 | variable {K : Type*} {p : ℕ} [Field K] [CharZero K] 4 | 5 | variable {ζ : K} 6 | 7 | open scoped BigOperators nonZeroDivisors 8 | open Polynomial 9 | 10 | --TODO: fix the following proofs using new multiplicity API 11 | 12 | lemma WfDvdMonoid.multiplicity_finite_iff {M : Type*} [CancelCommMonoidWithZero M] [WfDvdMonoid M] 13 | {x y : M} : 14 | FiniteMultiplicity x y ↔ ¬IsUnit x ∧ y ≠ 0 := by 15 | constructor 16 | · rw [← not_imp_not, Ne, ← not_or, not_not] 17 | rintro (hx|hy) 18 | · exact fun ⟨n, hn⟩ ↦ hn (hx.pow _).dvd 19 | · simp [hy] 20 | · intro ⟨hx, hy⟩ 21 | exact FiniteMultiplicity.of_not_isUnit hx hy 22 | 23 | lemma dvd_iff_emultiplicity_le {M : Type*} 24 | [CancelCommMonoidWithZero M] [DecidableRel (fun a b : M ↦ a ∣ b)] [UniqueFactorizationMonoid M] 25 | {a b : M} (ha : a ≠ 0) : a ∣ b ↔ ∀ p : M, Prime p → emultiplicity p a ≤ emultiplicity p b := by 26 | constructor 27 | · intro hab p _ 28 | exact emultiplicity_le_emultiplicity_of_dvd_right hab 29 | · intro H 30 | by_cases hb : b = 0 31 | · exact hb ▸ dvd_zero a 32 | induction' a using WfDvdMonoid.induction_on_irreducible with u hu a q _ hq IH 33 | · exact (ha rfl).elim 34 | · exact hu.dvd 35 | · simp only [ne_eq, mul_eq_zero, not_or] at ha 36 | rw [UniqueFactorizationMonoid.irreducible_iff_prime] at hq 37 | obtain ⟨c, rfl⟩ : a ∣ b := by 38 | refine IH ha.2 (fun p hp ↦ (le_trans ?_ (H p hp))) 39 | rw [emultiplicity_mul hp] 40 | exact le_add_self 41 | rw [mul_comm] 42 | simp only [ne_eq, mul_eq_zero, not_or] at hb 43 | refine mul_dvd_mul_left _ ?_ 44 | rw [← pow_one q, pow_dvd_iff_le_emultiplicity] 45 | have := H q hq 46 | rw [emultiplicity_mul hq, emultiplicity_mul hq, 47 | FiniteMultiplicity.emultiplicity_eq_multiplicity (WfDvdMonoid.multiplicity_finite_iff.2 48 | ⟨hq.not_unit, hb.2⟩), FiniteMultiplicity.emultiplicity_eq_multiplicity 49 | (WfDvdMonoid.multiplicity_finite_iff.2 ⟨hq.not_unit, ha.2⟩), 50 | FiniteMultiplicity.emultiplicity_eq_multiplicity (WfDvdMonoid.multiplicity_finite_iff.2 51 | ⟨hq.not_unit, hq.ne_zero⟩), multiplicity_self, ← Nat.cast_add, ← Nat.cast_add, 52 | Nat.cast_le, add_comm, add_le_add_iff_left] at this 53 | exact le_emultiplicity_of_le_multiplicity this 54 | 55 | lemma pow_dvd_pow_iff_dvd {M : Type*} [CancelCommMonoidWithZero M] [UniqueFactorizationMonoid M] 56 | {a b : M} {x : ℕ} (h' : x ≠ 0) : a ^ x ∣ b ^ x ↔ a ∣ b := by 57 | classical 58 | by_cases ha : a = 0 59 | · simp [ha, h'] 60 | by_cases hb : b = 0 61 | · simp [hb, h'] 62 | have ha' : a ^ x ≠ 0 := (pow_ne_zero_iff h').mpr ha 63 | rw [dvd_iff_emultiplicity_le ha, dvd_iff_emultiplicity_le ha'] 64 | refine forall₂_congr (fun p hp ↦ ⟨fun h ↦ ?_, fun h ↦ ?_⟩) 65 | · rw [emultiplicity_pow hp, emultiplicity_pow hp, 66 | FiniteMultiplicity.emultiplicity_eq_multiplicity 67 | (WfDvdMonoid.multiplicity_finite_iff.2 ⟨hp.not_unit, ha⟩), 68 | FiniteMultiplicity.emultiplicity_eq_multiplicity 69 | (WfDvdMonoid.multiplicity_finite_iff.2 ⟨hp.not_unit, hb⟩), ← Nat.cast_mul, ← Nat.cast_mul, 70 | Nat.cast_le] at h 71 | rw [FiniteMultiplicity.emultiplicity_eq_multiplicity 72 | (WfDvdMonoid.multiplicity_finite_iff.2 ⟨hp.not_unit, ha⟩), 73 | FiniteMultiplicity.emultiplicity_eq_multiplicity 74 | (WfDvdMonoid.multiplicity_finite_iff.2 ⟨hp.not_unit, hb⟩), Nat.cast_le] 75 | exact le_of_nsmul_le_nsmul_right h' h 76 | · rw [emultiplicity_pow hp, emultiplicity_pow hp, 77 | FiniteMultiplicity.emultiplicity_eq_multiplicity 78 | (WfDvdMonoid.multiplicity_finite_iff.2 ⟨hp.not_unit, ha⟩), 79 | FiniteMultiplicity.emultiplicity_eq_multiplicity 80 | (WfDvdMonoid.multiplicity_finite_iff.2 ⟨hp.not_unit, hb⟩), ← Nat.cast_mul, ← Nat.cast_mul, 81 | Nat.cast_le] 82 | rw [FiniteMultiplicity.emultiplicity_eq_multiplicity 83 | (WfDvdMonoid.multiplicity_finite_iff.2 ⟨hp.not_unit, ha⟩), 84 | FiniteMultiplicity.emultiplicity_eq_multiplicity 85 | (WfDvdMonoid.multiplicity_finite_iff.2 ⟨hp.not_unit, hb⟩), 86 | Nat.cast_le] at h 87 | exact Nat.mul_le_mul_left x h 88 | 89 | theorem isPrincipal_of_isPrincipal_pow_of_Coprime' 90 | {A K: Type*} [CommRing A] [IsDedekindDomain A] [Fintype (ClassGroup A)] 91 | [Field K] [Algebra A K] [IsFractionRing A K] (p : ℕ) 92 | (H : p.Coprime <| Fintype.card <| ClassGroup A) (I : FractionalIdeal A⁰ K) 93 | (hI : (↑(I ^ p) : Submodule A K).IsPrincipal) : (I : Submodule A K).IsPrincipal := by 94 | by_cases Izero : I = 0 95 | · rw [Izero, FractionalIdeal.coe_zero] 96 | exact bot_isPrincipal 97 | rw [← Ne, ← isUnit_iff_ne_zero] at Izero 98 | show Submodule.IsPrincipal ((Izero.unit' : FractionalIdeal A⁰ K) : Submodule A K) 99 | rw [← ClassGroup.mk_eq_one_iff, ← orderOf_eq_one_iff, ← Nat.dvd_one, ← H, Nat.dvd_gcd_iff] 100 | refine ⟨?_, orderOf_dvd_card⟩ 101 | rw [orderOf_dvd_iff_pow_eq_one, ← map_pow, ClassGroup.mk_eq_one_iff] 102 | simp only [Units.val_pow_eq_pow_val, IsUnit.val_unit', hI] 103 | 104 | variable (hp : p ≠ 2) 105 | 106 | open FractionalIdeal in 107 | lemma exists_not_dvd_spanSingleton_eq {R : Type*} [CommRing R] [IsDomain R] [IsDedekindDomain R] 108 | {K : Type*} [Field K] [Algebra R K] [IsFractionRing R K] 109 | {x : R} (hx : Prime x) (I J : Ideal R) 110 | (hI : ¬ (Ideal.span <| singleton x) ∣ I) (hJ : ¬ (Ideal.span <| singleton x) ∣ J) 111 | (h : Submodule.IsPrincipal ((I / J : FractionalIdeal R⁰ K) : Submodule R K)) : 112 | ∃ a b : R, 113 | ¬(x ∣ a) ∧ ¬(x ∣ b) ∧ spanSingleton R⁰ (algebraMap R K a / algebraMap R K b) = I / J := by 114 | by_contra H1 115 | have hI' : (I : FractionalIdeal R⁰ K) ≠ 0 := 116 | by rw [← coeIdeal_bot, Ne, coeIdeal_inj]; rintro rfl; exact hI (dvd_zero _) 117 | have hJ' : (J : FractionalIdeal R⁰ K) ≠ 0 := 118 | by rw [← coeIdeal_bot, Ne, coeIdeal_inj]; rintro rfl; exact hJ (dvd_zero _) 119 | have : ∀ n : ℕ, (1 ≤ n) → ¬∃ a b : R, ¬(x ^ n ∣ a) ∧ ¬(x ^ n ∣ b) ∧ 120 | spanSingleton R⁰ (algebraMap R K a / algebraMap R K b) = I / J := by 121 | intro n hn 122 | induction' n, hn using Nat.le_induction with n' hn' IH 123 | · simp_rw [pow_one] 124 | exact H1 125 | · rintro ⟨a, b, ha, hb, e⟩ 126 | have e₀ := e 127 | rw [div_eq_mul_inv, ← spanSingleton_mul_spanSingleton, 128 | ← one_div_spanSingleton, ← mul_div_assoc, mul_one, div_eq_iff, 129 | ← mul_div_right_comm, eq_div_iff hJ', ← coeIdeal_span_singleton, ← coeIdeal_span_singleton, 130 | ← coeIdeal_mul, ← coeIdeal_mul, coeIdeal_inj] at e 131 | on_goal 2 => 132 | rw [Ne, spanSingleton_eq_zero_iff, ← (algebraMap R K).map_zero, 133 | (IsFractionRing.injective R K).eq_iff] 134 | rintro rfl 135 | apply hb (dvd_zero _) 136 | by_cases h : x ^ n' ∣ a 137 | · have ha' : x ∣ a := (dvd_pow_self _ (Nat.one_le_iff_ne_zero.mp hn')).trans h 138 | have hb' : x ∣ b := by 139 | have : gcd (Ideal.span <| singleton x) I = 1 := by 140 | rwa [Irreducible.gcd_eq_one_iff] 141 | · rwa [irreducible_iff_prime, Ideal.prime_iff_isPrime, Ideal.span_singleton_prime] 142 | · exact hx.ne_zero 143 | · rw [Ne, Ideal.span_singleton_eq_bot] 144 | exact hx.ne_zero 145 | rw [← Ideal.mem_span_singleton, ← Ideal.dvd_span_singleton] at ha' ⊢ 146 | replace h := ha'.trans (dvd_mul_right _ J) 147 | rwa [e, ← dvd_gcd_mul_iff_dvd_mul, this, one_mul] at h 148 | obtain ⟨a', rfl⟩ := ha' 149 | obtain ⟨b', rfl⟩ := hb' 150 | rw [pow_succ', mul_dvd_mul_iff_left hx.ne_zero] at ha hb 151 | rw [_root_.map_mul, _root_.map_mul, mul_div_mul_left] at e₀ 152 | · exact IH ⟨a', b', ha, hb, e₀⟩ 153 | · rw [Ne, ← (algebraMap R K).map_zero, (IsFractionRing.injective R K).eq_iff] 154 | exact hx.ne_zero 155 | · refine IH ⟨a, b, h, ?_, e₀⟩ 156 | intro hb 157 | apply h 158 | rw [← Ideal.mem_span_singleton, ← Ideal.dvd_span_singleton] at hb ⊢ 159 | replace hb := hb.trans (dvd_mul_left _ I) 160 | have : gcd (Ideal.span <| singleton <| x ^ n') J = 1 := by 161 | rwa [← Ideal.isCoprime_iff_gcd, ← Ideal.span_singleton_pow, 162 | IsCoprime.pow_left_iff, Ideal.isCoprime_iff_gcd, Irreducible.gcd_eq_one_iff] 163 | · rwa [irreducible_iff_prime, Ideal.prime_iff_isPrime, Ideal.span_singleton_prime] 164 | · exact hx.ne_zero 165 | · rw [Ne, Ideal.span_singleton_eq_bot] 166 | exact hx.ne_zero 167 | · rwa [Nat.pos_iff_ne_zero, ← Nat.one_le_iff_ne_zero] 168 | rwa [← e, mul_comm, ← dvd_gcd_mul_iff_dvd_mul, this, one_mul] at hb 169 | rw [isPrincipal_iff] at h 170 | obtain ⟨a, ha⟩ := h 171 | obtain ⟨s, t, rfl⟩ := IsLocalization.mk'_surjective R⁰ a 172 | by_cases h : s = 0 173 | · rw [div_eq_iff hJ', h, IsLocalization.mk'_zero, spanSingleton_zero, zero_mul] at ha 174 | exact hI' ha 175 | obtain ⟨n, hn⟩ := FiniteMultiplicity.of_not_isUnit hx.not_unit h 176 | obtain ⟨m, hm⟩ := FiniteMultiplicity.of_not_isUnit hx.not_unit (nonZeroDivisors.ne_zero t.prop) 177 | rw [IsFractionRing.mk'_eq_div] at ha 178 | refine this (n + m + 1) (Nat.le_add_left 1 (n + m)) ⟨s, t, ?_, ?_, ha.symm⟩ 179 | · intro hs 180 | refine hn (dvd_trans (pow_dvd_pow _ ?_) hs) 181 | linarith 182 | · intro ht 183 | refine hm (dvd_trans (pow_dvd_pow _ (Nat.le_add_left _ _)) ht) 184 | -------------------------------------------------------------------------------- /FltRegular/CaseII/Statement.lean: -------------------------------------------------------------------------------- 1 | import FltRegular.CaseII.InductionStep 2 | 3 | open scoped BigOperators nonZeroDivisors NumberField 4 | open Polynomial 5 | 6 | variable {K : Type*} {p : ℕ} [hpri : Fact p.Prime] [Field K] [NumberField K] 7 | [IsCyclotomicExtension {p} ℚ K] (hp : p ≠ 2) [Fintype (ClassGroup (𝓞 K))] 8 | (hreg : p.Coprime <| Fintype.card <| ClassGroup (𝓞 K)) 9 | 10 | variable {ζ : K} (hζ : IsPrimitiveRoot ζ p) 11 | 12 | namespace FltRegular 13 | 14 | include hp hreg in 15 | lemma not_exists_solution {m : ℕ} (hm : 1 ≤ m) : 16 | ¬∃ (x' y' z' : 𝓞 K) (ε₃ : (𝓞 K)ˣ), 17 | ¬((hζ.unit' : 𝓞 K) - 1 ∣ y') ∧ ¬((hζ.unit' : 𝓞 K) - 1 ∣ z') ∧ 18 | x' ^ p + y' ^ p = ε₃ * (((hζ.unit' : 𝓞 K) - 1) ^ m * z') ^ p := by 19 | induction' m, hm using Nat.le_induction with m' _ IH 20 | · rintro ⟨x, y, z, ε₃, hy, hz, e⟩ 21 | exact zero_lt_one.not_le (one_le_m hp hζ e hy hz) 22 | · rintro ⟨x, y, z, ε₃, hy, hz, e⟩ 23 | exact IH (exists_solution' hp hζ e hy hz hreg) 24 | 25 | include hp hreg in 26 | lemma not_exists_solution' : 27 | ¬∃ (x y z : 𝓞 K), ¬(hζ.unit' : 𝓞 K) - 1 ∣ y ∧ (hζ.unit' : 𝓞 K) - 1 ∣ z ∧ z ≠ 0 ∧ 28 | x ^ p + y ^ p = z ^ p := by 29 | letI : Fact (Nat.Prime p) := hpri 30 | letI : WfDvdMonoid (𝓞 K) := IsNoetherianRing.wfDvdMonoid 31 | rintro ⟨x, y, z, hy, hz, hz', e⟩ 32 | obtain ⟨m, z, hm, hz'', rfl⟩ : 33 | ∃ m z', 1 ≤ m ∧ ¬((hζ.unit' : 𝓞 K) - 1 ∣ z') ∧ z = ((hζ.unit' : 𝓞 K) - 1) ^ m * z' := by 34 | classical 35 | have H : FiniteMultiplicity ((hζ.unit' : 𝓞 K) - 1) z := FiniteMultiplicity.of_not_isUnit 36 | hζ.zeta_sub_one_prime'.not_unit hz' 37 | obtain ⟨z', h⟩ := pow_multiplicity_dvd ((hζ.unit' : 𝓞 K) - 1) z 38 | refine ⟨_, _, ?_, ?_, h⟩ 39 | · rwa [← Nat.cast_le (α := ENat), ← FiniteMultiplicity.emultiplicity_eq_multiplicity H, 40 | ← pow_dvd_iff_le_emultiplicity, pow_one] 41 | · intro h' 42 | have := mul_dvd_mul_left (((hζ.unit' : 𝓞 K) - 1) ^ (multiplicity ((hζ.unit' : 𝓞 K) - 1) z)) h' 43 | rw [← pow_succ, ← h] at this 44 | refine not_pow_dvd_of_emultiplicity_lt ?_ this 45 | rw [FiniteMultiplicity.emultiplicity_eq_multiplicity H, Nat.cast_lt] 46 | exact Nat.lt_succ_self _ 47 | refine not_exists_solution hp hreg hζ hm ⟨x, y, z, 1, hy, hz'', ?_⟩ 48 | rwa [Units.val_one, one_mul] 49 | 50 | lemma not_exists_Int_solution {p : ℕ} [hpri : Fact (Nat.Prime p)] (hreg : IsRegularPrime p) 51 | (hodd : p ≠ 2) : ¬∃ (x y z : ℤ), ¬↑p ∣ y ∧ ↑p ∣ z ∧ z ≠ 0 ∧ x ^ p + y ^ p = z ^ p := by 52 | haveI := CyclotomicField.isCyclotomicExtension p ℚ 53 | obtain ⟨ζ, hζ⟩ := IsCyclotomicExtension.exists_isPrimitiveRoot 54 | ℚ (B := (CyclotomicField p ℚ)) (Set.mem_singleton p) hpri.1.ne_zero 55 | have := fun n ↦ zeta_sub_one_dvd_Int_iff (K := CyclotomicField p ℚ) hζ (n := n) 56 | simp only [PNat.mk_coe] at this 57 | simp_rw [← this] 58 | rintro ⟨x, y, z, hy, hz, hz', e⟩ 59 | refine not_exists_solution' (K := CyclotomicField p ℚ) hodd ?_ 60 | hζ ⟨x, y, z, hy, hz, ?_, ?_⟩ 61 | · convert hreg 62 | · rwa [ne_eq, Int.cast_eq_zero] 63 | · dsimp 64 | simp_rw [← Int.cast_pow, ← Int.cast_add, e] 65 | 66 | lemma not_exists_Int_solution' {p : ℕ} [hpri : Fact (Nat.Prime p)] (hreg : IsRegularPrime p) 67 | (hodd : p ≠ 2) : 68 | ¬∃ (x y z : ℤ), ({x, y, z} : Finset ℤ).gcd id = 1 ∧ ↑p ∣ z ∧ z ≠ 0 ∧ x ^ p + y ^ p = z ^ p := by 69 | rintro ⟨x, y, z, hgcd, hz, hz', e⟩ 70 | refine not_exists_Int_solution hreg hodd ⟨x, y, z, ?_, hz, hz', e⟩ 71 | intro hy 72 | have := dvd_sub (dvd_pow hz hpri.out.ne_zero) (dvd_pow hy hpri.out.ne_zero) 73 | rw [← e, add_sub_cancel_right] at this 74 | replace this := (Nat.prime_iff_prime_int.mp hpri.out).dvd_of_dvd_pow this 75 | apply (Nat.prime_iff_prime_int.mp hpri.out).not_unit 76 | rw [isUnit_iff_dvd_one, ← hgcd] 77 | simp [dvd_gcd_iff, hz, hy, this] 78 | 79 | lemma Int.gcd_left_comm (a b c : ℤ) : Int.gcd a (Int.gcd b c) = Int.gcd b (Int.gcd a c) := by 80 | rw [← Int.gcd_assoc, ← Int.gcd_assoc, Int.gcd_comm a b] 81 | 82 | /-- CaseII. -/ 83 | theorem caseII {a b c : ℤ} {p : ℕ} [hpri : Fact p.Prime] (hreg : IsRegularPrime p) (hodd : p ≠ 2) 84 | (hprod : a * b * c ≠ 0) (hgcd : ({a, b, c} : Finset ℤ).gcd id = 1) 85 | (caseII : ↑p ∣ a * b * c) : a ^ p + b ^ p ≠ c ^ p := by 86 | intro e 87 | simp only [ne_eq, mul_eq_zero, not_or] at hprod 88 | obtain ⟨⟨a0, b0⟩, c0⟩ := hprod 89 | have hodd' := Nat.Prime.odd_of_ne_two hpri.out hodd 90 | obtain hab | hc := (Nat.prime_iff_prime_int.mp hpri.out).dvd_or_dvd caseII 91 | · obtain ha| hb := (Nat.prime_iff_prime_int.mp hpri.out).dvd_or_dvd hab 92 | · refine not_exists_Int_solution' hreg hodd ⟨b, -c, -a, ?_, ?_, ?_, ?_⟩ 93 | · simp only [← hgcd, Finset.mem_singleton, Finset.mem_insert, neg_inj, Finset.gcd_insert, 94 | id_eq, ← Int.coe_gcd, Int.neg_gcd, Nat.cast_inj, ← LawfulSingleton.insert_empty_eq, 95 | Finset.gcd_empty, Int.gcd_left_comm _ a] 96 | · rwa [dvd_neg] 97 | · rwa [ne_eq, neg_eq_zero] 98 | · simp [hodd'.neg_pow, ← e] 99 | · refine not_exists_Int_solution' hreg hodd ⟨-c, a, -b, ?_, ?_, ?_, ?_⟩ 100 | · simp only [← hgcd, Finset.mem_singleton, Finset.mem_insert, neg_inj, Finset.gcd_insert, 101 | id_eq, ← Int.coe_gcd, Int.neg_gcd, Nat.cast_inj, ← LawfulSingleton.insert_empty_eq, 102 | Finset.gcd_empty, Int.gcd_left_comm _ c] 103 | · rwa [dvd_neg] 104 | · rwa [ne_eq, neg_eq_zero] 105 | · simp [hodd'.neg_pow, ← e] 106 | · exact not_exists_Int_solution' hreg hodd ⟨a, b, c, hgcd, hc, c0, e⟩ 107 | 108 | end FltRegular 109 | -------------------------------------------------------------------------------- /FltRegular/FLT5.lean: -------------------------------------------------------------------------------- 1 | import FltRegular.FltRegular 2 | import Mathlib.NumberTheory.Cyclotomic.PID 3 | 4 | open Nat NumberField IsCyclotomicExtension 5 | 6 | theorem isRegularPrime_five : 7 | haveI : Fact (Nat.Prime 5) := ⟨Nat.prime_five⟩ 8 | IsRegularPrime 5 := by 9 | rw [IsRegularPrime, IsRegularNumber] 10 | convert coprime_one_right _ 11 | exact classNumber_eq_one_iff.2 (Rat.five_pid (CyclotomicField _ ℚ)) 12 | 13 | theorem fermatLastTheoremFive : FermatLastTheoremFor 5 := by 14 | have : Fact (Nat.Prime 5) := ⟨Nat.prime_five⟩ 15 | exact flt_regular isRegularPrime_five (by omega) 16 | -------------------------------------------------------------------------------- /FltRegular/FltRegular.lean: -------------------------------------------------------------------------------- 1 | import FltRegular.CaseI.Statement 2 | import FltRegular.CaseII.Statement 3 | import Mathlib.NumberTheory.FLT.Basic 4 | 5 | open FltRegular 6 | 7 | /-- Fermat's last theorem for regular primes. -/ 8 | theorem flt_regular {p : ℕ} [Fact p.Prime] (hreg : IsRegularPrime p) (hodd : p ≠ 2) : 9 | FermatLastTheoremFor p := by 10 | apply fermatLastTheoremFor_iff_int.mpr 11 | intro a b c ha hb hc e 12 | have hprod := mul_ne_zero (mul_ne_zero ha hb) hc 13 | obtain ⟨e', hgcd, hprod'⟩ := MayAssume.coprime e hprod 14 | let d := ({a, b, c} : Finset ℤ).gcd id 15 | by_cases case : ↑p ∣ (a / d) * (b / d) * (c / d) 16 | · exact caseII hreg hodd hprod' hgcd case e' 17 | · exact caseI hreg case e' 18 | -------------------------------------------------------------------------------- /FltRegular/MayAssume/Lemmas.lean: -------------------------------------------------------------------------------- 1 | import Mathlib.Algebra.GCDMonoid.Finset 2 | import Mathlib.FieldTheory.Finite.Basic 3 | import Mathlib.Analysis.Normed.Ring.Lemmas 4 | 5 | open Int Finset 6 | 7 | namespace FltRegular 8 | 9 | namespace MayAssume 10 | 11 | theorem coprime {a b c : ℤ} {n : ℕ} (H : a ^ n + b ^ n = c ^ n) (hprod : a * b * c ≠ 0) : 12 | let d := ({a, b, c} : Finset ℤ).gcd id 13 | (a / d) ^ n + (b / d) ^ n = (c / d) ^ n ∧ 14 | ({a / d, b / d, c / d} : Finset ℤ).gcd id = 1 ∧ a / d * (b / d) * (c / d) ≠ 0 := by 15 | have ha : a ≠ 0 := fun ha => by simp [ha] at hprod 16 | have hb : b ≠ 0 := fun hb => by simp [hb] at hprod 17 | have hc : c ≠ 0 := fun hc => by simp [hc] at hprod 18 | let s : Finset ℤ := {a, b, c} 19 | set d : ℤ := s.gcd id 20 | have hadiv : d ∣ a := gcd_dvd (by simp [s]) 21 | have hbdiv : d ∣ b := gcd_dvd (by simp [s]) 22 | have hcdiv : d ∣ c := gcd_dvd (by simp [s]) 23 | have hdzero : d ≠ 0 := fun hdzero => by 24 | simpa [ha] using Finset.gcd_eq_zero_iff.1 hdzero a (by simp [s]) 25 | have hdp : d ^ n ≠ 0 := fun hdn => hdzero (pow_eq_zero hdn) 26 | refine ⟨?_, ?_, fun habs => ?_⟩ 27 | · obtain ⟨na, hna⟩ := hadiv; obtain ⟨nb, hnb⟩ := hbdiv; obtain ⟨nc, hnc⟩ := hcdiv 28 | rwa [← mul_left_inj' hdp, add_mul, ← mul_pow, ← mul_pow, ← mul_pow, hna, hnb, hnc, 29 | Int.mul_ediv_cancel_left _ hdzero, Int.mul_ediv_cancel_left _ hdzero, 30 | Int.mul_ediv_cancel_left _ hdzero, mul_comm, ← hna, mul_comm, ← hnb, mul_comm, ← hnc] 31 | · simpa [gcd_eq_gcd_image, d] using 32 | Finset.gcd_div_id_eq_one (show a ∈ ({a, b, c} : Finset ℤ) by simp) ha 33 | · simp only [mul_eq_zero] at habs 34 | rcases habs with ((Ha | Hb) | Hc) 35 | · exact ha (Int.eq_zero_of_ediv_eq_zero (gcd_dvd (by simp [s])) Ha) 36 | · exact hb (Int.eq_zero_of_ediv_eq_zero (gcd_dvd (by simp [s])) Hb) 37 | · exact hc (Int.eq_zero_of_ediv_eq_zero (gcd_dvd (by simp [s])) Hc) 38 | 39 | end MayAssume 40 | 41 | theorem p_dvd_c_of_ab_of_anegc {p : ℕ} {a b c : ℤ} (hpri : p.Prime) (hp : p ≠ 3) 42 | (h : a ^ p + b ^ p = c ^ p) (hab : a ≡ b [ZMOD p]) (hbc : b ≡ -c [ZMOD p]) : ↑p ∣ c := by 43 | letI : Fact p.Prime := ⟨hpri⟩ 44 | replace h := congr_arg (fun n : ℤ => (n : ZMod p)) h 45 | simp only [Int.natCast_pow, Int.cast_add, Int.cast_pow, ZMod.pow_card] at h 46 | simp only [← ZMod.intCast_eq_intCast_iff, Int.cast_neg] at hbc hab 47 | rw [hab, hbc, ← sub_eq_zero, ← sub_eq_add_neg, ← Int.cast_neg, ← Int.cast_sub, 48 | ← Int.cast_sub] at h 49 | ring_nf at h 50 | simp only [Int.cast_neg, Int.cast_mul, Int.cast_one, Int.cast_ofNat, neg_eq_zero, 51 | mul_eq_zero] at h 52 | rw [← ZMod.intCast_zmod_eq_zero_iff_dvd] 53 | refine Or.resolve_right h fun h3 => ?_ 54 | rw [show (3 : ZMod p) = ((3 : ℕ) : ZMod p) by simp, ZMod.natCast_zmod_eq_zero_iff_dvd, 55 | Nat.dvd_prime Nat.prime_three] at h3 56 | rcases h3 with H₁ | H₂ 57 | · exact hpri.ne_one H₁ 58 | · exact hp H₂ 59 | 60 | theorem a_not_cong_b {p : ℕ} {a b c : ℤ} (hpri : p.Prime) (hp5 : 5 ≤ p) (hprod : a * b * c ≠ 0) 61 | (h : a ^ p + b ^ p = c ^ p) (hgcd : ({a, b, c} : Finset ℤ).gcd id = 1) 62 | (caseI : ¬↑p ∣ a * b * c) : 63 | ∃ x y z : ℤ, x ^ p + y ^ p = z ^ p ∧ 64 | ({x, y, z} : Finset ℤ).gcd id = 1 ∧ ¬x ≡ y [ZMOD p] ∧ x * y * z ≠ 0 ∧ ¬↑p ∣ x * y * z := by 65 | by_cases H : a ≡ b [ZMOD p] 66 | swap 67 | · exact ⟨a, b, c, ⟨h, hgcd, H, hprod, caseI⟩⟩ 68 | refine ⟨a, -c, -b, ⟨?_, ?_, fun habs => ?_, ?_, ?_⟩⟩ 69 | · have hpodd : p ≠ 2 := by linarith 70 | rw [neg_pow, (Or.resolve_left hpri.eq_two_or_odd' hpodd).neg_one_pow, neg_pow, 71 | (Or.resolve_left hpri.eq_two_or_odd' hpodd).neg_one_pow] 72 | simp only [← sub_eq_add_neg, sub_eq_iff_eq_add, neg_mul, one_mul] 73 | symm 74 | rw [neg_add_eq_iff_eq_add, add_comm] 75 | exact h.symm 76 | · convert hgcd using 1 77 | have : ({a, -c, -b} : Finset ℤ) = {a, -b, -c} := by 78 | refine Finset.ext fun x => ⟨fun hx => ?_, fun hx => ?_⟩ <;> 79 | · simp only [mem_insert, mem_singleton] at hx 80 | rcases hx with (H | H | H) <;> simp [H] 81 | rw [this] 82 | simp only [gcd_insert, id, gcd_singleton, normalize_apply, neg_mul] 83 | congr 1 84 | rw [← coe_gcd, ← coe_gcd, Int.gcd_eq_natAbs, Int.gcd_eq_natAbs] 85 | simp only [natAbs_neg, Nat.cast_inj] 86 | rcases Int.isUnit_iff.1 (normUnit (-c)).isUnit, Int.isUnit_iff.1 (normUnit c).isUnit with 87 | ⟨H₁ | H₂, H₃ | H₄⟩ 88 | · simp [H₁, H₃] 89 | · simp [H₁, H₄] 90 | · simp [H₂, H₃] 91 | · simp [H₂, H₄] 92 | · have hp3 : p ≠ 3 := by linarith 93 | rw [← ZMod.intCast_eq_intCast_iff] at habs H 94 | rw [H] at habs 95 | rw [ZMod.intCast_eq_intCast_iff] at habs H 96 | obtain ⟨n, hn⟩ := p_dvd_c_of_ab_of_anegc hpri hp3 h H habs 97 | refine caseI ⟨a * b * n, ?_⟩ 98 | rw [hn] 99 | ring 100 | · convert hprod using 1 101 | ring 102 | · ring_nf at caseI ⊢ 103 | exact caseI 104 | 105 | end FltRegular 106 | -------------------------------------------------------------------------------- /FltRegular/NumberTheory/Cyclotomic/CaseI.lean: -------------------------------------------------------------------------------- 1 | import FltRegular.NumberTheory.Cyclotomic.UnitLemmas 2 | import FltRegular.NumberTheory.Cyclotomic.CyclRat 3 | 4 | open scoped NumberField nonZeroDivisors 5 | 6 | variable {p : ℕ} [NeZero p] {K : Type*} [Field K] [CharZero K] [IsCyclotomicExtension {p} ℚ K] 7 | 8 | variable {ζ : K} (hζ : IsPrimitiveRoot ζ p) 9 | 10 | open FractionalIdeal 11 | 12 | variable (i : ℤ) 13 | 14 | namespace FltRegular.CaseI 15 | 16 | lemma coe_unitGalConj (x : (𝓞 K)ˣ) : ↑(unitGalConj K p x) = intGal (galConj K p) (x : 𝓞 K) := 17 | rfl 18 | 19 | theorem pow_sub_intGalConj_mem (hp : p.Prime) (α : 𝓞 K) : 20 | (α ^ p - intGal (galConj K p) (α ^ p)) ∈ Ideal.span ({(p : 𝓞 K)} : Set (𝓞 K)) := by 21 | have : Fact p.Prime := ⟨hp⟩ 22 | obtain ⟨a, ha⟩ := exists_int_sub_pow_prime_dvd p α 23 | rw [Ideal.mem_span_singleton] at ha ⊢ 24 | obtain ⟨γ, hγ⟩ := ha 25 | rw [sub_eq_iff_eq_add] at hγ 26 | rw [hγ, _root_.map_add, _root_.map_mul, map_natCast, map_intCast, add_sub_add_right_eq_sub, 27 | ← mul_sub] 28 | exact dvd_mul_right _ _ 29 | 30 | theorem exists_int_sum_eq_zero'_aux (x y i : ℤ) : 31 | intGal (galConj K p) (x + y * ↑(hζ.unit' ^ i) : 𝓞 K) = x + y * (hζ.unit' ^ (-i) : (𝓞 K)ˣ) := by 32 | ext1 33 | rw [intGal_apply_coe] 34 | simp only [_root_.map_add, map_intCast, _root_.map_mul, AlgHom.coe_coe, zpow_neg, map_units_inv, 35 | add_right_inj, mul_eq_mul_left_iff, Int.cast_eq_zero] 36 | simp_rw [NumberField.Units.coe_zpow] 37 | left 38 | simp only [map_zpow₀] 39 | rw [← inv_zpow] 40 | congr 41 | exact galConj_zeta_runity hζ 42 | 43 | theorem exists_int_sum_eq_zero' (hpodd : p ≠ 2) (hp : p.Prime) (x y i : ℤ) {u : (𝓞 K)ˣ} 44 | {α : 𝓞 K} (h : (x : 𝓞 K) + y * (hζ.unit' ^ i : (𝓞 K)ˣ) = u * α ^ p) : 45 | ∃ k : ℕ, (x : 𝓞 K) + y * (hζ.unit' ^ i : (𝓞 K)ˣ) - ((hζ.unit' ^ k) ^ 2 : (𝓞 K)ˣ) * 46 | (x + y * (hζ.unit' ^ (-i) : (𝓞 K)ˣ)) ∈ 47 | Ideal.span ({(p : 𝓞 K)} : Set (𝓞 K)) := by 48 | letI : NumberField K := IsCyclotomicExtension.numberField { p } ℚ _ 49 | have : Fact p.Prime := ⟨hp⟩ 50 | obtain ⟨k, H⟩ := unit_inv_conj_is_root_of_unity hζ hpodd hp u 51 | refine ⟨k, ?_⟩ 52 | rw [← exists_int_sum_eq_zero'_aux, h, ← H, Units.val_mul, mul_assoc, ← mul_sub, _root_.map_mul, 53 | ← coe_unitGalConj, ← mul_assoc, ← Units.val_mul, inv_mul_cancel, Units.val_one, one_mul] 54 | exact Ideal.mul_mem_left _ _ (pow_sub_intGalConj_mem hp α) 55 | 56 | theorem exists_int_sum_eq_zero (hpodd : p ≠ 2) (hp : p.Prime) (x y i : ℤ) {u : (𝓞 K)ˣ} 57 | {α : 𝓞 K} (h : (x : 𝓞 K) + y * (hζ.unit' ^ i : (𝓞 K)ˣ) = u * α ^ p) : 58 | ∃ k : ℤ, (x : 𝓞 K) + y * (hζ.unit' ^ i : (𝓞 K)ˣ) - (hζ.unit' ^ (2 * k) : (𝓞 K)ˣ) * 59 | (x + y * (hζ.unit' ^ (-i) : (𝓞 K)ˣ)) ∈ 60 | Ideal.span ({(p : 𝓞 K)} : Set (𝓞 K)) := by 61 | obtain ⟨k, hk⟩ := exists_int_sum_eq_zero' hζ hpodd hp x y i h 62 | use k 63 | convert hk 64 | rw [mul_comm, zpow_mul, zpow_ofNat] 65 | rfl 66 | 67 | end FltRegular.CaseI 68 | -------------------------------------------------------------------------------- /FltRegular/NumberTheory/Cyclotomic/CyclotomicUnits.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021 Alex J. Best. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Alex J. Best 5 | 6 | ! This file was ported from Lean 3 source module number_theory.cyclotomic.cyclotomic_units 7 | -/ 8 | import Mathlib.RingTheory.RootsOfUnity.PrimitiveRoots 9 | 10 | noncomputable section 11 | 12 | open scoped BigOperators nonZeroDivisors 13 | 14 | open Polynomial Finset Module Units Submodule 15 | 16 | universe u v w z 17 | 18 | variable (n : ℕ) (K : Type u) (L : Type v) (A : Type w) (B : Type z) 19 | 20 | variable [CommRing A] [CommRing B] [Algebra A B] 21 | 22 | variable [Field K] [Field L] [Algebra K L] 23 | 24 | variable [IsDomain A] [Algebra A K] [IsFractionRing A K] 25 | 26 | section CyclotomicUnit 27 | 28 | variable {n} 29 | 30 | namespace CyclotomicUnit 31 | 32 | -- I don't want to bother Leo, but I wonder if this can be automated in Lean4 33 | -- (if they were 0 < n → 1 < n, it would work already!) 34 | theorem Nat.one_lt_of_ne : ∀ n : ℕ, n ≠ 0 → n ≠ 1 → 1 < n 35 | | 0, h, _ => absurd rfl h 36 | | 1, _, h => absurd rfl h 37 | | n + 2, _, _ => n.one_lt_succ_succ 38 | 39 | -- this result holds for all primitive roots; dah. 40 | theorem associated_one_sub_pow_primitive_root_of_coprime {n j k : ℕ} {ζ : A} 41 | (hζ : IsPrimitiveRoot ζ n) (hk : k.Coprime n) (hj : j.Coprime n) : 42 | Associated (1 - ζ ^ j) (1 - ζ ^ k) := by 43 | suffices ∀ {j : ℕ}, j.Coprime n → Associated (1 - ζ) (1 - ζ ^ j) by 44 | exact (this hj).symm.trans (this hk) 45 | clear k j hk hj 46 | intro j hj 47 | refine associated_of_dvd_dvd ⟨∑ i ∈ range j, ζ ^ i, by rw [← geom_sum_mul_neg, mul_comm]⟩ ?_ 48 | -- is there an easier way to do this? 49 | rcases eq_or_ne n 0 with (rfl | hn') 50 | · simp [j.coprime_zero_right.mp hj] 51 | rcases eq_or_ne n 1 with (rfl | hn) 52 | · simp [IsPrimitiveRoot.one_right_iff.mp hζ] 53 | replace hn := Nat.one_lt_of_ne n hn' hn 54 | obtain ⟨m, hm⟩ := Nat.exists_mul_emod_eq_one_of_coprime hj hn 55 | use ∑ i ∈ range m, (ζ ^ j) ^ i 56 | have : ζ = (ζ ^ j) ^ m := by rw [← pow_mul, ←pow_mod_orderOf, ← hζ.eq_orderOf, hm, pow_one] 57 | nth_rw 1 [this] 58 | rw [← geom_sum_mul_neg, mul_comm] 59 | 60 | theorem IsPrimitiveRoot.sum_pow_unit {n k : ℕ} {ζ : A} (hn : 2 ≤ n) (hk : k.Coprime n) 61 | (hζ : IsPrimitiveRoot ζ n) : IsUnit (∑ i ∈ range k, ζ ^ i) := by 62 | have h1 : (1 : ℕ).Coprime n := Nat.coprime_one_left n 63 | have := associated_one_sub_pow_primitive_root_of_coprime _ hζ hk h1 64 | simp only [pow_one] at this 65 | rw [Associated] at this 66 | have h2 := mul_neg_geom_sum ζ k 67 | obtain ⟨u, hu⟩ := this 68 | have := u.isUnit 69 | convert this 70 | rw [← hu] at h2 71 | simp only [mul_eq_mul_left_iff] at h2 72 | rcases h2 with h2 | h2 73 | · exact h2 74 | · exfalso 75 | have hn1 : 1 < n := by linarith 76 | have hp := IsPrimitiveRoot.pow_ne_one_of_pos_of_lt hζ one_pos hn1 77 | rw [sub_eq_zero] at h2 78 | rw [← h2] at hp 79 | simp at hp 80 | 81 | theorem IsPrimitiveRoot.zeta_pow_sub_eq_unit_zeta_sub_one {p i j : ℕ} {ζ : A} (hn : 2 ≤ p) 82 | (hp : p.Prime) (hi : i < p) (hj : j < p) (hij : i ≠ j) (hζ : IsPrimitiveRoot ζ p) : 83 | ∃ u : Aˣ, ζ ^ i - ζ ^ j = u * (1 - ζ) := by 84 | by_cases hilj : j < i 85 | · have h1 : ζ ^ i - ζ ^ j = ζ ^ j * (ζ ^ (i - j) - 1) := by 86 | ring_nf 87 | rw [pow_mul_pow_sub _ hilj.le] 88 | rw [add_comm] 89 | ring 90 | rw [h1] 91 | have h2 := mul_neg_geom_sum ζ (i - j) 92 | have hic : (i - j).Coprime p := by 93 | rw [Nat.coprime_comm]; apply Nat.coprime_of_lt_prime _ _ hp 94 | apply Nat.sub_pos_of_lt hilj 95 | by_cases hj : 0 < j 96 | apply lt_trans _ hi 97 | apply Nat.sub_lt_of_pos_le hj hilj.le 98 | simp only [not_lt, _root_.le_zero_iff] at hj 99 | rw [hj] 100 | simp only [tsub_zero] 101 | exact hi 102 | have h3 : IsUnit (-ζ ^ j * ∑ k ∈ range (i - j), ζ ^ k) := by 103 | apply IsUnit.mul _ (IsPrimitiveRoot.sum_pow_unit _ hn hic hζ); apply IsUnit.neg 104 | apply IsUnit.pow; apply hζ.isUnit hp.pos 105 | obtain ⟨v, hv⟩ := h3 106 | use v 107 | rw [hv] 108 | rw [mul_comm] at h2 109 | rw [mul_assoc] 110 | rw [h2] 111 | ring 112 | · simp only [ne_eq, not_lt] at hij hilj 113 | have h1 : ζ ^ i - ζ ^ j = ζ ^ i * (1 - ζ ^ (j - i)) := by 114 | ring_nf 115 | rw [sub_right_inj, pow_mul_pow_sub _ hilj] 116 | rw [h1] 117 | have h2 := mul_neg_geom_sum ζ (j - i) 118 | have hjc : (j - i).Coprime p := by 119 | rw [Nat.coprime_comm] 120 | apply Nat.coprime_of_lt_prime _ _ hp 121 | have hilj' : i < j := by rw [lt_iff_le_and_ne]; simp [hij, hilj] 122 | apply Nat.sub_pos_of_lt hilj' 123 | by_cases hii : 0 < i 124 | apply lt_trans _ hj 125 | apply Nat.sub_lt_of_pos_le hii hilj 126 | simp only [not_lt, _root_.le_zero_iff] at hii 127 | rw [hii] 128 | simp only [tsub_zero] 129 | exact hj 130 | have h3 : IsUnit (ζ ^ i * ∑ k ∈ range (j - i), ζ ^ k) := by 131 | apply IsUnit.mul _ (IsPrimitiveRoot.sum_pow_unit _ hn hjc hζ); apply IsUnit.pow 132 | apply hζ.isUnit hp.pos 133 | obtain ⟨v, hv⟩ := h3 134 | use v 135 | rw [hv] 136 | rw [mul_comm] at h2 137 | rw [mul_assoc] 138 | rw [h2] 139 | 140 | end CyclotomicUnit 141 | 142 | lemma IsPrimitiveRoot.associated_sub_one {A : Type*} [CommRing A] [IsDomain A] 143 | {p : ℕ} (hp : p.Prime) {ζ : A} (hζ : IsPrimitiveRoot ζ p) {η₁ : A} 144 | (hη₁ : η₁ ∈ nthRootsFinset p 1) {η₂ : A} (hη₂ : η₂ ∈ nthRootsFinset p 1) (e : η₁ ≠ η₂) : 145 | Associated (ζ - 1) (η₁ - η₂) := by 146 | have : NeZero p := ⟨hp.ne_zero⟩ 147 | obtain ⟨i, ⟨hi, rfl⟩⟩ := 148 | hζ.eq_pow_of_pow_eq_one ((Polynomial.mem_nthRootsFinset hp.pos 1).1 hη₁) 149 | obtain ⟨j, ⟨hj, rfl⟩⟩ := 150 | hζ.eq_pow_of_pow_eq_one ((Polynomial.mem_nthRootsFinset hp.pos 1).1 hη₂) 151 | have : i ≠ j := ne_of_apply_ne _ e 152 | obtain ⟨u, h⟩ := CyclotomicUnit.IsPrimitiveRoot.zeta_pow_sub_eq_unit_zeta_sub_one A 153 | hp.two_le hp hi hj this hζ 154 | rw [h, associated_isUnit_mul_right_iff u.isUnit, ← associated_isUnit_mul_right_iff isUnit_one.neg, 155 | neg_one_mul, neg_sub] 156 | exact Associates.mk_eq_mk_iff_associated.mp (by simp) 157 | 158 | end CyclotomicUnit 159 | -------------------------------------------------------------------------------- /FltRegular/NumberTheory/Cyclotomic/GaloisActionOnCyclo.lean: -------------------------------------------------------------------------------- 1 | import Mathlib.NumberTheory.Cyclotomic.Gal 2 | import Mathlib.NumberTheory.NumberField.Units.Basic 3 | 4 | open FiniteDimensional 5 | 6 | open scoped NumberField 7 | 8 | 9 | variable (K : Type*) (p : ℕ) [NeZero p] [Field K] [CharZero K] [IsCyclotomicExtension {p} ℚ K] 10 | 11 | variable {ζ : K} (hζ : IsPrimitiveRoot ζ p) 12 | 13 | local notation "RR" => 𝓞 K 14 | 15 | open scoped NumberField Cyclotomic 16 | 17 | open IsCyclotomicExtension Polynomial 18 | 19 | noncomputable section 20 | 21 | /-- complex conjugation as a Galois automorphism -/ 22 | def galConj : K ≃ₐ[ℚ] K := 23 | (autEquivPow K (cyclotomic.irreducible_rat (NeZero.pos p))).symm (-1) 24 | 25 | variable {K p} 26 | 27 | theorem ZMod.val_neg_one' : ∀ {n : ℕ}, 0 < n → (-1 : ZMod n).val = n - 1 28 | | _ + 1, _ => ZMod.val_neg_one _ 29 | 30 | theorem galConj_zeta : galConj K p (zeta p ℚ K) = (zeta p ℚ K)⁻¹ := by 31 | let hζ := zeta_spec p ℚ K 32 | simp only [galConj, Units.coe_neg_one, autEquivPow_symm_apply, AlgEquiv.coe_algHom, 33 | PowerBasis.equivOfMinpoly_apply] 34 | convert (hζ.powerBasis ℚ).lift_gen (S' := K) _ _ 35 | rw [IsPrimitiveRoot.powerBasis_gen, ZMod.val_neg_one' (NeZero.pos p), 36 | pow_sub₀ _ (hζ.ne_zero (NeZero.ne p)) (NeZero.pos p), pow_one, hζ.pow_eq_one, one_mul] 37 | 38 | include hζ in 39 | @[simp] 40 | theorem galConj_zeta_runity : galConj K p ζ = ζ⁻¹ := by 41 | obtain ⟨t, _, rfl⟩ := (zeta_spec p ℚ K).eq_pow_of_pow_eq_one hζ.pow_eq_one 42 | rw [map_pow, galConj_zeta, inv_pow] 43 | 44 | include hζ in 45 | theorem galConj_zeta_runity_pow (n : ℕ) : galConj K p (ζ ^ n) = ζ⁻¹ ^ n := by 46 | rw [map_pow, galConj_zeta_runity hζ] 47 | 48 | open scoped ComplexConjugate 49 | 50 | include hζ in 51 | @[simp] 52 | theorem embedding_conj (x : K) (φ : K →+* ℂ) : conj (φ x) = φ (galConj K p x) := by 53 | change RingHom.comp conj φ x = (φ.comp (galConj K p)) x 54 | revert x 55 | suffices φ (galConj K p ζ) = conj (φ ζ) by 56 | rw [← funext_iff, DFunLike.coe_fn_eq, ← ((starRingEnd ℂ).comp φ).toRatAlgHom_toRingHom, 57 | ← (φ.comp ↑(galConj K p)).toRatAlgHom_toRingHom (R := K)] 58 | congr 1 59 | apply (hζ.powerBasis ℚ).algHom_ext 60 | exact this.symm 61 | rw [← Complex.inv_eq_conj, galConj_zeta_runity hζ, map_inv₀] 62 | exact Complex.norm_eq_one_of_pow_eq_one (by rw [← map_pow, hζ.pow_eq_one, map_one]) (NeZero.ne p) 63 | 64 | variable (p) 65 | 66 | theorem gal_map_mem_subtype (σ : K →ₐ[ℚ] K) (x : RR) : IsIntegral ℤ (σ x) := 67 | map_isIntegral_int _ x.2 68 | 69 | /-- Restriction of `σ : K →ₐ[ℚ] K` to the ring of integers. -/ 70 | def intGal (σ : K →ₐ[ℚ] K) : RR →ₐ[ℤ] RR := 71 | ((σ.restrictScalars ℤ).restrictDomain RR).codRestrict (integralClosure ℤ K) 72 | (gal_map_mem_subtype σ) 73 | 74 | @[simp] 75 | theorem intGal_apply_coe (σ : K →ₐ[ℚ] K) (x : RR) : (intGal σ x : K) = σ x := 76 | rfl 77 | 78 | /-- Restriction of `σ : K →ₐ[ℚ] K` to the units of the ring of integers. -/ 79 | def unitsGal (σ : K →ₐ[ℚ] K) : RRˣ →* RRˣ := 80 | Units.map <| intGal σ 81 | 82 | variable (K) 83 | 84 | /-- `unit_gal_conj` as a bundled hom. -/ 85 | def unitGalConj : RRˣ →* RRˣ := 86 | unitsGal (galConj K p) 87 | 88 | theorem unitGalConj_spec (u : RRˣ) : galConj K p u = unitGalConj K p u := rfl 89 | 90 | variable {K} 91 | 92 | theorem unit_lemma_val_one (u : RRˣ) (φ : K →+* ℂ) : 93 | ‖φ (u * (unitGalConj K p u)⁻¹)‖ = 1 := by 94 | rw [map_mul, norm_mul, ← zpow_neg_one, NumberField.Units.coe_zpow, 95 | zpow_neg_one, map_inv₀, ← unitGalConj_spec, 96 | ← embedding_conj <| zeta_spec p ℚ K] 97 | simp 98 | -------------------------------------------------------------------------------- /FltRegular/NumberTheory/Cyclotomic/MoreLemmas.lean: -------------------------------------------------------------------------------- 1 | import Mathlib.NumberTheory.Cyclotomic.Rat 2 | import FltRegular.NumberTheory.Cyclotomic.UnitLemmas 3 | import FltRegular.NumberTheory.Cyclotomic.CyclRat 4 | import Mathlib.RingTheory.Ideal.Norm.AbsNorm 5 | import Mathlib.RingTheory.NormTrace 6 | 7 | variable {K : Type*} {p : ℕ} [hpri : Fact p.Prime] [Field K] [CharZero K] 8 | [IsCyclotomicExtension {p} ℚ K] 9 | 10 | variable {ζ : K} (hζ : IsPrimitiveRoot ζ p) 11 | 12 | open scoped BigOperators nonZeroDivisors NumberField 13 | open Polynomial 14 | 15 | variable (hp : p ≠ 2) 16 | 17 | lemma IsPrimitiveRoot.prime_span_sub_one : 18 | Prime (Ideal.span <| singleton <| (hζ.unit' - 1 : 𝓞 K)) := by 19 | haveI : Fact (Nat.Prime p) := hpri 20 | letI := IsCyclotomicExtension.numberField {p} ℚ K 21 | rw [Ideal.prime_iff_isPrime, 22 | Ideal.span_singleton_prime (hζ.unit'_coe.sub_one_ne_zero hpri.out.one_lt)] 23 | · exact hζ.zeta_sub_one_prime' 24 | · rw [Ne, Ideal.span_singleton_eq_bot] 25 | exact hζ.unit'_coe.sub_one_ne_zero hpri.out.one_lt 26 | 27 | lemma associated_zeta_sub_one_pow_prime : Associated ((hζ.unit' - 1 : 𝓞 K) ^ (p - 1)) p := by 28 | letI := IsCyclotomicExtension.numberField {p} ℚ K 29 | haveI : Fact (Nat.Prime p) := hpri 30 | rw [← eval_one_cyclotomic_prime (R := 𝓞 K) (p := p), 31 | cyclotomic_eq_prod_X_sub_primitiveRoots hζ.unit'_coe, eval_prod] 32 | simp only [eval_sub, eval_X, eval_C] 33 | rw [← Nat.totient_prime this.out, ← hζ.unit'_coe.card_primitiveRoots, ← Finset.prod_const] 34 | apply Associated.prod 35 | intro η hη 36 | exact hζ.unit'_coe.associated_sub_one hpri.out 37 | (one_mem_nthRootsFinset this.out.pos) 38 | ((isPrimitiveRoot_of_mem_primitiveRoots hη).mem_nthRootsFinset hpri.out.pos) 39 | ((isPrimitiveRoot_of_mem_primitiveRoots hη).ne_one hpri.out.one_lt).symm 40 | 41 | lemma isCoprime_of_not_zeta_sub_one_dvd {x : 𝓞 K} (hx : ¬ (hζ.unit' : 𝓞 K) - 1 ∣ x) : 42 | IsCoprime ↑p x := by 43 | letI := IsCyclotomicExtension.numberField {p} ℚ K 44 | rwa [← Ideal.isCoprime_span_singleton_iff, 45 | ← Ideal.span_singleton_eq_span_singleton.mpr (associated_zeta_sub_one_pow_prime hζ), 46 | ← Ideal.span_singleton_pow, IsCoprime.pow_left_iff, Ideal.isCoprime_iff_gcd, 47 | hζ.prime_span_sub_one.irreducible.gcd_eq_one_iff, Ideal.dvd_span_singleton, 48 | Ideal.mem_span_singleton] 49 | · simpa only [ge_iff_le, tsub_pos_iff_lt] using hpri.out.one_lt 50 | 51 | lemma exists_zeta_sub_one_dvd_sub_Int (a : 𝓞 K) : ∃ b : ℤ, (hζ.unit' - 1 : 𝓞 K) ∣ a - b := by 52 | letI : Fact (Nat.Prime p) := hpri 53 | simp_rw [← Ideal.Quotient.eq_zero_iff_dvd, ← Ideal.Quotient.mk_eq_mk, Submodule.Quotient.mk_sub, 54 | sub_eq_zero, ← SModEq.def] 55 | exact hζ.subOneIntegralPowerBasis'_gen ▸ hζ.subOneIntegralPowerBasis'.exists_smodEq a 56 | 57 | include hp in 58 | lemma exists_dvd_pow_sub_Int_pow (a : 𝓞 K) : ∃ b : ℤ, ↑p ∣ a ^ p - (b : 𝓞 K) ^ p := by 59 | obtain ⟨ζ, hζ⟩ := IsCyclotomicExtension.exists_isPrimitiveRoot ℚ K (Set.mem_singleton p) 60 | hpri.1.ne_zero 61 | obtain ⟨b, k, e⟩ := exists_zeta_sub_one_dvd_sub_Int hζ a 62 | obtain ⟨r, hr⟩ := exists_add_pow_prime_eq hpri.out a (-b) 63 | obtain ⟨u, hu⟩ := (associated_zeta_sub_one_pow_prime hζ).symm 64 | rw [(Nat.Prime.odd_of_ne_two hpri.out hp).neg_pow, ← sub_eq_add_neg, e, 65 | mul_pow, ← sub_eq_add_neg] at hr 66 | use b, ↑u * ((hζ.unit' - 1 : 𝓞 K) * k ^ p) - r 67 | rw [← sub_eq_iff_eq_add.mpr hr, mul_sub, ← mul_assoc, ← mul_assoc, hu, ← pow_succ, 68 | Nat.sub_add_cancel (n := p) (m := 1) hpri.out.one_lt.le] 69 | 70 | section 71 | 72 | variable {α} [CommMonoidWithZero α] 73 | 74 | theorem prime_units_mul (a : αˣ) (b : α) : Prime (↑a * b) ↔ Prime b := by simp [Prime] 75 | 76 | end 77 | 78 | lemma zeta_sub_one_dvd_Int_iff {n : ℤ} : (hζ.unit' : 𝓞 K) - 1 ∣ n ↔ ↑p ∣ n := by 79 | letI := IsCyclotomicExtension.numberField {p} ℚ K 80 | by_cases hp : p = 2 81 | · subst hp 82 | have : IsCyclotomicExtension {2 ^ (0 + 1)} ℚ K := by rwa [zero_add, pow_one] 83 | have hζ' : IsPrimitiveRoot ζ (2 ^ (0 + 1)) := by simpa 84 | have := hζ'.norm_toInteger_pow_sub_one_of_two 85 | rw [pow_zero, pow_one, pow_one (-2)] at this 86 | replace this : (Algebra.norm ℤ) (hζ.toInteger - 1) = -2 := this 87 | simp only [PNat.val_ofNat, Nat.cast_ofNat] 88 | rw [← neg_dvd (a := (2 : ℤ)), ← this, Ideal.norm_dvd_iff] 89 | · rfl 90 | · rw [this] 91 | refine Prime.neg Int.prime_two 92 | rw [← hζ.norm_toInteger_sub_one_of_prime_ne_two' hp, Ideal.norm_dvd_iff] 93 | · rfl 94 | · rw [hζ.norm_toInteger_sub_one_of_prime_ne_two' hp, ← Nat.prime_iff_prime_int] 95 | exact hpri.1 96 | 97 | lemma IsPrimitiveRoot.sub_one_dvd_sub {A : Type*} [CommRing A] [IsDomain A] 98 | {p : ℕ} (hp : p.Prime) {ζ : A} (hζ : IsPrimitiveRoot ζ p) {η₁ : A} 99 | (hη₁ : η₁ ∈ nthRootsFinset p 1) {η₂ : A} (hη₂ : η₂ ∈ nthRootsFinset p 1) : 100 | ζ - 1 ∣ η₁ - η₂ := by 101 | by_cases h : η₁ = η₂ 102 | · rw [h, sub_self]; exact dvd_zero _ 103 | · exact (hζ.associated_sub_one hp hη₁ hη₂ h).dvd 104 | 105 | lemma quotient_zero_sub_one_comp_aut (σ : 𝓞 K →+* 𝓞 K) : 106 | (Ideal.Quotient.mk (Ideal.span {(hζ.unit' : 𝓞 K) - 1})).comp σ = Ideal.Quotient.mk _ := by 107 | have : Fact (Nat.Prime p) := hpri 108 | letI := IsCyclotomicExtension.numberField {p} ℚ K 109 | letI : AddGroup (𝓞 K ⧸ Ideal.span (singleton (hζ.unit' - 1: 𝓞 K))) := inferInstance 110 | apply RingHom.toIntAlgHom_injective 111 | apply hζ.integralPowerBasis'.algHom_ext 112 | have h : hζ.integralPowerBasis'.gen = hζ.unit' := by 113 | simp only [IsPrimitiveRoot.integralPowerBasis'_gen] 114 | rfl 115 | rw [h] 116 | simp only [RingHom.toIntAlgHom, RingHom.toMonoidHom_eq_coe, AlgHom.coe_mk, RingHom.coe_mk, 117 | MonoidHom.coe_coe, RingHom.coe_comp, RingHom.coe_coe, Function.comp_apply] 118 | rw [← sub_eq_zero, ← Ideal.Quotient.mk_eq_mk, ← Ideal.Quotient.mk_eq_mk, 119 | ← Submodule.Quotient.mk_sub, Ideal.Quotient.mk_eq_mk, Ideal.Quotient.eq_zero_iff_mem, 120 | Ideal.mem_span_singleton] 121 | apply hζ.unit'_coe.sub_one_dvd_sub hpri.out 122 | · rw [mem_nthRootsFinset (NeZero.pos p), ← map_pow, hζ.unit'_coe.pow_eq_one, map_one] 123 | · rw [mem_nthRootsFinset (NeZero.pos p), hζ.unit'_coe.pow_eq_one] 124 | 125 | set_option synthInstance.maxHeartbeats 80000 in 126 | -- Needed for `AddMonoidHomClass (𝓞 K →+* 𝓞 K ⧸ Ideal.span {↑hζ.unit' - 1}) ? ?` 127 | lemma zeta_sub_one_dvd_trace_sub_smul (x : 𝓞 K) : 128 | (hζ.unit' - 1 : 𝓞 K) ∣ Algebra.trace ℤ _ x - (p - 1) • x := by 129 | letI := IsCyclotomicExtension.numberField {p} ℚ K 130 | letI := IsCyclotomicExtension.isGalois p ℚ K 131 | have : (Algebra.trace ℤ _ x : 𝓞 K) = ∑ σ : K ≃ₐ[ℚ] K, (intGal σ).toRingHom x := by 132 | apply (show Function.Injective (algebraMap (𝓞 K) K) from Subtype.val_injective) 133 | rw [← eq_intCast (algebraMap ℤ (𝓞 K)), ← IsScalarTower.algebraMap_apply, 134 | IsScalarTower.algebraMap_apply ℤ ℚ K, eq_intCast, Algebra.coe_trace_int, 135 | trace_eq_sum_automorphisms, map_sum] 136 | rfl 137 | rw [← Ideal.mem_span_singleton, ← Ideal.Quotient.eq_zero_iff_mem, map_sub, this, 138 | map_sum] 139 | simp_rw [← RingHom.comp_apply, quotient_zero_sub_one_comp_aut] 140 | rw [Finset.sum_const, map_nsmul, sub_eq_zero, Finset.card_univ, IsGalois.card_aut_eq_finrank, 141 | IsCyclotomicExtension.finrank K (cyclotomic.irreducible_rat (NeZero.pos p)), 142 | Nat.totient_prime hpri.out] 143 | 144 | lemma zeta_sub_one_pow_dvd_norm_sub_pow (x : 𝓞 K) : 145 | (hζ.unit' - 1 : 𝓞 K) ^ p ∣ (Algebra.norm ℤ (1 + p • x) : 𝓞 K) - 1 + p • x := by 146 | letI := IsCyclotomicExtension.numberField {p} ℚ K 147 | obtain ⟨r, hr⟩ := Algebra.norm_one_add_smul (p : ℤ) x 148 | obtain ⟨s, hs⟩ := zeta_sub_one_dvd_trace_sub_smul hζ x 149 | obtain ⟨t, ht⟩ := (associated_zeta_sub_one_pow_prime hζ).dvd 150 | rw [sub_eq_iff_eq_add] at hs 151 | simp only [zsmul_eq_mul, Int.cast_natCast] at hr 152 | simp only [nsmul_eq_mul, hr, Int.cast_add, Int.cast_one, Int.cast_mul, hs, NeZero.pos p, 153 | Nat.cast_pred, Int.cast_natCast, Int.cast_pow] 154 | suffices (hζ.unit' - 1 : 𝓞 K) ^ p ∣ (hζ.unit' - 1) * p * s + (p : 𝓞 K) ^ 2 * (r + x) by 155 | convert this using 1; ring 156 | apply dvd_add 157 | · apply dvd_mul_of_dvd_left 158 | rw [ht, ← mul_assoc, ← pow_succ', tsub_add_cancel_of_le (Nat.Prime.one_lt hpri.out).le] 159 | exact dvd_mul_right _ _ 160 | · rw [ht, mul_pow, ← pow_mul, mul_assoc] 161 | apply dvd_mul_of_dvd_left 162 | apply pow_dvd_pow 163 | zify [(Nat.Prime.one_lt hpri.out).le] 164 | linarith only [Nat.Prime.two_le hpri.out] 165 | 166 | lemma norm_add_one_smul_of_isUnit {K} [Field K] [NumberField K] {p : ℕ} (hpri : p.Prime) 167 | (hp : p ≠ 2) (x : 𝓞 K) 168 | (hx : IsUnit (1 + p • x)) : Algebra.norm ℤ (1 + p • x) = 1 := by 169 | have H : Algebra.norm ℤ (1 + p • x) = 1 ∨ Algebra.norm ℤ (1 + p • x) = -1 := by 170 | apply Int.natAbs_eq_iff.mp 171 | apply (Int.cast_injective (α := ℚ)).comp Nat.cast_injective 172 | simp only [Int.cast_abs, Function.comp_apply, Nat.cast_one, Int.cast_one, ← Int.abs_eq_natAbs, 173 | Algebra.coe_norm_int, ← NumberField.isUnit_iff_norm.mp hx, RingOfIntegers.coe_norm] 174 | have : Algebra.norm ℤ (1 + p • x) ≠ -1 := by 175 | intro e; apply hp 176 | obtain ⟨r, hr⟩ := Algebra.norm_one_add_smul (p : ℤ) x 177 | have : (p : ℤ) * (- Algebra.trace ℤ _ x - r * p) = 2 := by 178 | rw [zsmul_eq_mul, Int.cast_natCast, ← nsmul_eq_mul, e, eq_comm, ← sub_eq_zero] at hr 179 | rw [eq_comm, ← sub_eq_zero, ← hr] 180 | ring 181 | exact (Nat.prime_two.eq_one_or_self_of_dvd _ 182 | (Int.natCast_dvd_natCast.mp ⟨_, this.symm⟩)).resolve_left hpri.ne_one 183 | exact H.resolve_right this 184 | -------------------------------------------------------------------------------- /FltRegular/NumberTheory/Cyclotomic/UnitLemmas.lean: -------------------------------------------------------------------------------- 1 | import FltRegular.NumberTheory.Cyclotomic.GaloisActionOnCyclo 2 | import Mathlib.NumberTheory.Cyclotomic.Rat 3 | 4 | variable {p : ℕ} [NeZero p] {K : Type*} [Field K] 5 | 6 | variable {ζ : K} (hζ : IsPrimitiveRoot ζ p) 7 | 8 | open scoped BigOperators nonZeroDivisors NumberField 9 | 10 | open IsCyclotomicExtension NumberField Polynomial 11 | 12 | local notation "R" => 𝓞 K 13 | 14 | --The whole file is now for a generic primitive root ζ, quite a lot of names should be changed. 15 | universe u 16 | 17 | noncomputable section 18 | 19 | /-- zeta now as a unit in the ring of integers. This way there are no coe issues. -/ 20 | def IsPrimitiveRoot.unit' {p : ℕ} [NeZero p] {K : Type*} [Field K] {ζ : K} 21 | (hζ : IsPrimitiveRoot ζ p) : (𝓞 K)ˣ where 22 | val := (⟨ζ, hζ.isIntegral (NeZero.pos p)⟩ : 𝓞 K) 23 | inv := (⟨ζ⁻¹, hζ.inv.isIntegral (NeZero.pos p)⟩ : 𝓞 K) 24 | val_inv := Subtype.ext <| mul_inv_cancel₀ <| hζ.ne_zero (NeZero.ne p) 25 | inv_val := Subtype.ext <| inv_mul_cancel₀ <| hζ.ne_zero (NeZero.ne p) 26 | 27 | set_option quotPrecheck false 28 | local notation "ζ1" => (hζ.unit' - 1 : 𝓞 K) 29 | 30 | set_option quotPrecheck false 31 | local notation "I" => (Ideal.span ({(hζ.unit' - 1 : 𝓞 K)} : Set (𝓞 K)) : Ideal (𝓞 K)) 32 | 33 | theorem IsPrimitiveRoot.unit'_pow : hζ.unit' ^ p = 1 := by 34 | ext; simpa using hζ.pow_eq_one 35 | 36 | theorem zeta_runity_pow_even (hpo : Odd p) (n : ℕ) : 37 | ∃ m : ℕ, hζ.unit' ^ n = hζ.unit' ^ (2 * m) := by 38 | rcases eq_or_ne n 0 with (rfl | _) 39 | · use 0 40 | obtain ⟨r, hr⟩ := hpo 41 | have he : 2 * (r + 1) * n = p * n + n := by rw [hr]; ring 42 | use (r + 1) * n 43 | rw [← mul_assoc, he, pow_add] 44 | convert (one_mul (M := (𝓞 K)ˣ) _).symm 45 | rw [pow_mul, hζ.unit'_pow, one_pow] 46 | 47 | variable [NumberField K] 48 | 49 | theorem IsPrimitiveRoot.unit'_coe : IsPrimitiveRoot hζ.unit'.1 p := by 50 | have z1 := hζ 51 | have : (algebraMap R K) (hζ.unit' : R) = ζ := rfl 52 | rw [← this] at z1 53 | exact z1.of_map_of_injective (IsFractionRing.injective _ _) 54 | 55 | theorem totient_le_one_dvd_two {a : ℕ} (han : 0 < a) (ha : a.totient ≤ 1) : a ∣ 2 := by 56 | rcases Nat.totient_eq_one_iff.1 (show a.totient = 1 by linarith [Nat.totient_pos.2 han]) with 57 | h | h <;> 58 | simp [h] 59 | 60 | theorem eq_one_mod_one_sub {A : Type*} [CommRing A] {t : A} : 61 | algebraMap A (A ⧸ Ideal.span ({t - 1} : Set A)) t = 1 := 62 | by 63 | rw [← map_one <| algebraMap A <| A ⧸ Ideal.span ({t - 1} : Set A), ← sub_eq_zero, ← map_sub, 64 | Ideal.Quotient.algebraMap_eq, Ideal.Quotient.eq_zero_iff_mem] 65 | apply Ideal.subset_span 66 | exact Set.mem_singleton _ 67 | 68 | theorem IsPrimitiveRoot.eq_one_mod_sub_of_pow {A : Type*} [CommRing A] [IsDomain A] {ζ : A} 69 | (hζ : IsPrimitiveRoot ζ p) {μ : A} (hμ : μ ^ p = 1) : 70 | (@DFunLike.coe _ A (fun _ => A ⧸ Ideal.span {ζ - 1}) _ 71 | (algebraMap A (A ⧸ Ideal.span {ζ - 1})) μ) = 1 := by 72 | obtain ⟨k, -, rfl⟩ := hζ.eq_pow_of_pow_eq_one hμ 73 | rw [map_pow, eq_one_mod_one_sub, one_pow] 74 | 75 | set_option synthInstance.maxHeartbeats 40000 in 76 | -- needed for `AddMonoidHomClass (𝓞 K →+* 𝓞 K ⧸ Ideal.span {↑hζ.unit' - 1}) ? ?` 77 | theorem aux {t} {l : 𝓞 K} {f : Fin t → ℤ} {μ : K} (hμ : IsPrimitiveRoot μ p) 78 | (h : ∑ x : Fin t, f x • (⟨μ, hμ.isIntegral (NeZero.pos p)⟩ : 𝓞 K) ^ (x : ℕ) = l) : 79 | algebraMap (𝓞 K) (𝓞 K ⧸ I) l = ∑ x : Fin t, (f x : 𝓞 K ⧸ I) := by 80 | apply_fun algebraMap (𝓞 K) (𝓞 K ⧸ I) at h 81 | simp only [map_sum, map_zsmul] at h 82 | convert h.symm using 1 83 | congr 84 | funext x 85 | have : (⟨μ, hμ.isIntegral (NeZero.pos p)⟩ : 𝓞 K) ^ p = 1 := by 86 | ext 87 | push_cast 88 | exact hμ.pow_eq_one 89 | have := hζ.unit'_coe.eq_one_mod_sub_of_pow this 90 | simp only [map_pow (algebraMap (𝓞 K) (𝓞 K ⧸ I)), this, one_pow, zsmul_one] 91 | 92 | theorem IsPrimitiveRoot.p_mem_one_sub_zeta [hp : Fact p.Prime] : (p : 𝓞 K) ∈ I := by 93 | classical 94 | have key : _ = (p : 𝓞 K) := @Polynomial.eval_one_cyclotomic_prime _ _ _ hp 95 | rw [cyclotomic_eq_prod_X_sub_primitiveRoots hζ.unit'_coe, eval_prod] at key 96 | simp only [eval_sub, eval_X, eval_C] at key 97 | have : {↑hζ.unit'} ⊆ primitiveRoots p (𝓞 K) := by simpa [NeZero.pos p] using hζ.unit'_coe 98 | rw [← Finset.prod_sdiff this, Finset.prod_singleton] at key 99 | rw [← key] 100 | have := (Ideal.neg_mem_iff I).mpr (Ideal.subset_span (Set.mem_singleton ζ1)) 101 | rw [neg_sub] at this 102 | exact Ideal.mul_mem_left _ _ this 103 | 104 | variable [IsCyclotomicExtension {p} ℚ K] 105 | 106 | theorem roots_of_unity_in_cyclo_aux {x : K} {l : ℕ} (hl : l ≠ 0) (hx : IsIntegral ℤ x) 107 | (hhl : (cyclotomic l R).IsRoot ⟨x, hx⟩) {ζ : K} (hζ : IsPrimitiveRoot ζ p) : l ∣ 2 * p := by 108 | by_contra h 109 | have hpl' : IsPrimitiveRoot (⟨x, hx⟩ : R) l := by 110 | have nezero : NeZero (l : 𝓞 K) := by 111 | refine ⟨fun hzero ↦ ?_⟩ 112 | simp only [Nat.cast_eq_zero, hl] at hzero 113 | rw [isRoot_cyclotomic_iff.symm] 114 | apply hhl 115 | have hpl : IsPrimitiveRoot x l := by 116 | have : (algebraMap R K) ⟨x, hx⟩ = x := by rfl 117 | have h4 := IsPrimitiveRoot.map_of_injective hpl' (f := algebraMap (𝓞 K) K); rw [← this] 118 | apply h4 119 | apply IsFractionRing.injective 120 | have hirr : Irreducible (cyclotomic p ℚ) := cyclotomic.irreducible_rat (NeZero.pos p) 121 | have KEY := IsPrimitiveRoot.lcm_totient_le_finrank hpl hζ <| 122 | cyclotomic.irreducible_rat <| Nat.lcm_pos (Nat.zero_lt_of_ne_zero hl) (NeZero.pos p) 123 | have hrank := IsCyclotomicExtension.finrank K hirr 124 | rw [hrank] at KEY 125 | have pdivlcm : p ∣ lcm l p := dvd_lcm_right l p 126 | rcases pdivlcm with ⟨pdivlcm_w, pdivlcm_h⟩ 127 | have ineq1 := Nat.totient_super_multiplicative p pdivlcm_w 128 | rw [← pdivlcm_h] at ineq1 129 | have KEY3 := (mul_le_iff_le_one_right (Nat.totient_pos.2 (NeZero.pos p))).mp (le_trans ineq1 KEY) 130 | have pdiv_ne_zero : 0 < pdivlcm_w := by 131 | by_contra h 132 | simp only [not_lt, le_zero_iff] at h 133 | rw [h] at pdivlcm_h 134 | simp only [MulZeroClass.mul_zero, lcm_eq_zero_iff, PNat.ne_zero, or_false] at pdivlcm_h 135 | aesop 136 | have K5 := (Nat.dvd_prime Nat.prime_two).1 (totient_le_one_dvd_two pdiv_ne_zero KEY3) 137 | rcases K5 with K5 | K5 138 | · rw [K5] at pdivlcm_h 139 | simp only [mul_one] at pdivlcm_h 140 | rw [lcm_eq_right_iff] at pdivlcm_h 141 | · have K6 : p ∣ 2 * p := dvd_mul_left p 2 142 | apply absurd (dvd_trans pdivlcm_h K6) h 143 | simp only [eq_self_iff_true, normalize_eq, PNat.coe_inj] 144 | · rw [K5] at pdivlcm_h 145 | rw [mul_comm] at pdivlcm_h 146 | have := dvd_lcm_left l p 147 | simp_rw [pdivlcm_h] at this 148 | apply absurd this h 149 | 150 | --do more generally 151 | theorem roots_of_unity_in_cyclo (hpo : Odd p) (x : K) 152 | (h : ∃ (n : ℕ) (_ : 0 < n), x ^ n = 1) : 153 | ∃ (m k : ℕ), x = (-1) ^ k * (hζ.unit' : K) ^ m := by 154 | obtain ⟨n, hn0, hn⟩ := h 155 | have hx : IsIntegral ℤ x := by 156 | refine ⟨X ^ n - 1, ⟨?_, ?_⟩⟩ 157 | · exact monic_X_pow_sub_C 1 (ne_of_lt hn0).symm 158 | · simp only [hn, eval₂_one, eval₂_X_pow, eval₂_sub, sub_self] 159 | have hxu : (⟨x, hx⟩ : R) ^ n = 1 := by ext; simp [hn] 160 | have H : ∃ (m k : ℕ), (⟨x, hx⟩ : R) = (-1) ^ k * (hζ.unit' : K) ^ m := by 161 | obtain ⟨l, hl, hhl⟩ := (_root_.isRoot_of_unity_iff hn0 _).1 hxu 162 | replace hl : l ≠ 0 := fun H ↦ by simp [H] at hl 163 | have hlp := roots_of_unity_in_cyclo_aux hl hx hhl hζ 164 | have isPrimRoot : IsPrimitiveRoot (hζ.unit' : R) p := hζ.unit'_coe 165 | have hxl : (⟨x, hx⟩ : R) ^ l = 1 := by 166 | apply isRoot_of_unity_of_root_cyclotomic _ hhl 167 | simp only [Nat.mem_divisors, dvd_refl, Ne, true_and] 168 | apply pos_iff_ne_zero.1 (Nat.pos_of_ne_zero hl) 169 | have hxp' : (⟨x, hx⟩ : R) ^ (2 * p) = 1 := by 170 | rcases hlp with ⟨hlp_w, hlp_h⟩ 171 | rw [hlp_h, pow_mul, hxl]; simp only [one_pow] 172 | have hxp'' : (⟨x, hx⟩ : R) ^ p = 1 ∨ (⟨x, hx⟩ : R) ^ p = -1 := by 173 | rw [mul_comm] at hxp'; rw [pow_mul] at hxp' 174 | suffices (⟨x, hx⟩ : 𝓞 K) ^ p = 1 ∨ (⟨x, hx⟩ : 𝓞 K) ^ p = -1 by 175 | · rcases this with h1 | h2 176 | · left; simp only [h1] 177 | · right; simp only [h2] 178 | apply eq_or_eq_neg_of_sq_eq_sq 179 | simp only [one_pow] 180 | apply hxp' 181 | rcases hxp'' with hxp'' | hxp'' 182 | · obtain ⟨i, _, Hi⟩ := IsPrimitiveRoot.eq_pow_of_pow_eq_one isPrimRoot hxp'' 183 | refine ⟨i, 2, ?_⟩ 184 | rw [← Subtype.val_inj] at Hi 185 | simp only [SubmonoidClass.coe_pow] at Hi 186 | simp only [PNat.val_ofNat, even_two, Even.neg_pow, one_pow, one_mul] 187 | rw [← Hi] 188 | rfl 189 | · have hone : (-1 : R) ^ p = (-1 : R) := by apply Odd.neg_one_pow hpo 190 | have hxp3 : (-1 * ⟨x, hx⟩ : R) ^ p = 1 := by 191 | rw [mul_pow, hone, hxp''] 192 | ring 193 | obtain ⟨i, _, Hi⟩ := IsPrimitiveRoot.eq_pow_of_pow_eq_one isPrimRoot hxp3 194 | refine ⟨i, 1, ?_⟩ 195 | simp only [PNat.one_coe, pow_one, neg_mul, one_mul, neg_neg] 196 | rw [← Subtype.val_inj] at Hi 197 | simp only [SubmonoidClass.coe_pow, Submonoid.coe_mul, 198 | Subsemiring.coe_toSubmonoid, Subalgebra.coe_toSubsemiring, InvMemClass.coe_inv, 199 | OneMemClass.coe_one, neg_mul, one_mul] at Hi 200 | exact Iff.mp neg_eq_iff_eq_neg (id (Eq.symm (by simpa using Hi))) 201 | obtain ⟨m, k, hmk⟩ := H 202 | refine ⟨m, k, ?_⟩ 203 | have eq : ((⟨x, hx⟩ : R) : K) = x := rfl 204 | rw [← eq, hmk] 205 | 206 | theorem IsPrimitiveRoot.isPrime_one_sub_zeta [hp : Fact p.Prime] : 207 | (I : Ideal (𝓞 K)).IsPrime := by 208 | rw [Ideal.span_singleton_prime] 209 | · exact hζ.zeta_sub_one_prime' 210 | apply_fun (fun x : 𝓞 K => (x : K)) 211 | push_cast 212 | intro h 213 | refine hp.1.ne_one (hζ.unique ?_) 214 | simp only [one_right_iff] 215 | simp only [map_sub, map_one, map_zero, sub_eq_zero] at h 216 | exact h 217 | 218 | theorem IsPrimitiveRoot.two_not_mem_one_sub_zeta [hp : Fact p.Prime] (h : p ≠ 2) : 219 | (2 : 𝓞 K) ∉ I := by 220 | have hpm := hζ.p_mem_one_sub_zeta 221 | obtain ⟨k, hk⟩ := hp.1.odd_of_ne_two h 222 | apply_fun (fun n : ℕ => (n : 𝓞 K)) at hk 223 | rw [Nat.cast_add, Nat.cast_mul, Nat.cast_two, Nat.cast_one, add_comm] at hk 224 | intro h2m 225 | have := Ideal.sub_mem I hpm (Ideal.mul_mem_right (↑k) I h2m) 226 | rw [sub_eq_of_eq_add hk] at this 227 | exact hζ.isPrime_one_sub_zeta.ne_top (Ideal.eq_top_of_isUnit_mem I this isUnit_one) 228 | 229 | lemma map_two {S T F: Type*} [NonAssocSemiring S] [NonAssocSemiring T] [FunLike F S T] 230 | [RingHomClass F S T] (f : F) : f 2 = 2 := by 231 | rw [← one_add_one_eq_two, map_add, map_one] 232 | exact one_add_one_eq_two 233 | 234 | lemma neg_one_eq_one_iff_two_eq_zero {M : Type*} [AddGroupWithOne M] : 235 | (-1 : M) = 1 ↔ (2 : M) = 0 := by rw [neg_eq_iff_add_eq_zero, one_add_one_eq_two] 236 | 237 | lemma Units.coe_map_inv' {M N F : Type*} [Monoid M] [Monoid N] [FunLike F M N] 238 | [MonoidHomClass F M N] (f : F) (m : Mˣ) : 239 | ↑((Units.map (f : M →* N) m)⁻¹) = f ↑(m⁻¹ : Mˣ) := 240 | m.coe_map_inv (f : M →* N) 241 | 242 | lemma unit_inv_conj_not_neg_zeta_runity_aux (u : Rˣ) (hp : p.Prime) : 243 | algebraMap (𝓞 K) (𝓞 K ⧸ I) ((u * (unitGalConj K p u)⁻¹) : _) = 1 := by 244 | have := Units.coe_map_inv' (N := 𝓞 K ⧸ I) (algebraMap (𝓞 K) (𝓞 K ⧸ I)) (unitGalConj K p u) 245 | rw [Units.val_mul, map_mul, ← this, Units.mul_inv_eq_one, Units.coe_map , MonoidHom.coe_coe] 246 | haveI := Fact.mk hp 247 | have hu := hζ.integralPowerBasis'.basis.sum_repr u 248 | let a := hζ.integralPowerBasis'.basis.repr 249 | let φn := hζ.integralPowerBasis'.dim 250 | simp_rw [PowerBasis.basis_eq_pow, IsPrimitiveRoot.integralPowerBasis'_gen] at hu 251 | have hu' := congr_arg (intGal ↑(galConj K p)) hu 252 | replace hu' : ∑ x : Fin φn, (a u) x • (intGal ↑(galConj K p)) 253 | (⟨ζ, hζ.isIntegral (NeZero.pos p)⟩ ^ (x : ℕ)) = unitGalConj K p u := by 254 | refine Eq.trans ?_ hu' 255 | rw [map_sum] 256 | congr 1 257 | ext x 258 | congr 1 259 | rw [map_zsmul] 260 | have : ∀ x : Fin φn, intGal ((galConj K p)) (⟨ζ, hζ.isIntegral (NeZero.pos p)⟩ ^ (x : ℕ)) = 261 | ⟨ζ⁻¹, hζ.inv.isIntegral (NeZero.pos p)⟩ ^ (x : ℕ) := by 262 | intro x 263 | ext 264 | simp only [map_pow, intGal_apply_coe, RingOfIntegers.map_mk, AlgHom.coe_coe, inv_pow] 265 | rw [← map_pow, galConj_zeta_runity_pow hζ, inv_pow] 266 | conv_lhs at hu' => 267 | congr 268 | congr 269 | ext a 270 | rw [this a] 271 | exact (aux hζ hζ hu).trans (aux hζ hζ.inv hu').symm 272 | 273 | set_option synthInstance.maxHeartbeats 40000 in 274 | -- Needed for `AddMonoidHomClass (𝓞 K →+* 𝓞 K ⧸ Ideal.span {↑hζ.unit' - 1}) ? ?` 275 | theorem unit_inv_conj_not_neg_zeta_runity (h : p ≠ 2) (u : Rˣ) (n : ℕ) (hp : p.Prime) : 276 | u * (unitGalConj K p u)⁻¹ ≠ -hζ.unit' ^ n := by 277 | by_contra H 278 | have hμ : algebraMap (𝓞 K) (𝓞 K ⧸ I) ((IsPrimitiveRoot.unit' hζ : 𝓞 K) ^ n) = 1 := by 279 | apply hζ.unit'_coe.eq_one_mod_sub_of_pow 280 | rw [← pow_mul, mul_comm, pow_mul, hζ.unit'_coe.pow_eq_one, one_pow] 281 | have hμ' : algebraMap (𝓞 K) (𝓞 K ⧸ I) ((IsPrimitiveRoot.unit' hζ : 𝓞 K) ^ n) = -1 := by 282 | rw [← neg_eq_iff_eq_neg, ← map_neg, ← Units.val_pow_eq_pow_val, ← Units.val_neg, ← H] 283 | apply unit_inv_conj_not_neg_zeta_runity_aux hζ u hp 284 | haveI := Fact.mk hp 285 | apply hζ.two_not_mem_one_sub_zeta h 286 | rw [← Ideal.Quotient.eq_zero_iff_mem, map_two, ← neg_one_eq_one_iff_two_eq_zero, ← hμ', hμ] 287 | 288 | -- this proof has mild coe annoyances rn 289 | theorem unit_inv_conj_is_root_of_unity (h : p ≠ 2) (hp : p.Prime) (u : Rˣ) : 290 | ∃ m : ℕ, u * (unitGalConj K p u)⁻¹ = (hζ.unit' ^ m) ^ 2 := by 291 | have hpo : Odd p := hp.odd_of_ne_two h 292 | haveI : NormedAlgebra ℚ ℂ := normedAlgebraRat 293 | have := 294 | @NumberField.Embeddings.pow_eq_one_of_norm_eq_one K _ _ ℂ _ _ _ (u * (unitGalConj K p u)⁻¹ : K) 295 | ?_ ?_ 296 | · have H := roots_of_unity_in_cyclo hζ hpo (u * (unitGalConj K p u)⁻¹ : K) this 297 | obtain ⟨n, k, hz⟩ := H 298 | simp_rw [← pow_mul] 299 | have hk := Nat.even_or_odd k 300 | rcases hk with hk | hk 301 | · simp only [hk.neg_one_pow, one_mul] at hz 302 | rw [← map_mul, ← Units.val_mul, ← map_pow, ← Units.val_pow_eq_pow_val] at hz 303 | norm_cast at hz 304 | rw [hz] 305 | refine (exists_congr fun a => ?_).mp (zeta_runity_pow_even hζ hpo n) 306 | · rw [mul_comm] 307 | · by_contra 308 | simp only [hk.neg_one_pow, neg_mul, one_mul] at hz 309 | rw [← map_mul, ← Units.val_mul, ← map_pow, ← Units.val_pow_eq_pow_val, ← map_neg] at hz 310 | norm_cast at hz 311 | simpa [hz] using unit_inv_conj_not_neg_zeta_runity hζ h u n hp 312 | · apply RingHom.IsIntegralElem.mul 313 | · exact NumberField.RingOfIntegers.isIntegral_coe _ 314 | · exact NumberField.RingOfIntegers.isIntegral_coe _ 315 | · exact unit_lemma_val_one p u 316 | 317 | lemma IsPrimitiveRoot.eq_one_mod_one_sub' {A : Type*} [CommRing A] [IsDomain A] 318 | {n : ℕ} [NeZero n] {ζ : A} (hζ : IsPrimitiveRoot ζ n) {η : A} (hη : η ∈ nthRootsFinset n 1) : 319 | Ideal.Quotient.mk (Ideal.span ({ζ - 1} : Set A)) η = 1 := by 320 | obtain ⟨i, ⟨_, rfl⟩⟩ := hζ.eq_pow_of_pow_eq_one ((Polynomial.mem_nthRootsFinset 321 | (NeZero.pos n) 1).1 hη) 322 | rw [map_pow, ← Ideal.Quotient.algebraMap_eq, eq_one_mod_one_sub, one_pow] 323 | -------------------------------------------------------------------------------- /FltRegular/NumberTheory/CyclotomicRing.lean: -------------------------------------------------------------------------------- 1 | import Mathlib.NumberTheory.Cyclotomic.Rat 2 | import Mathlib.NumberTheory.Cyclotomic.PrimitiveRoots 3 | import FltRegular.NumberTheory.Cyclotomic.MoreLemmas 4 | 5 | noncomputable section 6 | 7 | open Polynomial NumberField 8 | 9 | variable (p : ℕ) [hpri : Fact p.Prime] 10 | 11 | def CyclotomicIntegers : Type := AdjoinRoot (cyclotomic p ℤ) 12 | 13 | instance : CommRing (CyclotomicIntegers p) := by delta CyclotomicIntegers; infer_instance 14 | 15 | open Polynomial in 16 | lemma IsPrimitiveRoot.cyclotomic_eq_minpoly 17 | (x : 𝓞 (CyclotomicField p ℚ)) (hx : IsPrimitiveRoot x.1 p) : 18 | minpoly ℤ x = cyclotomic p ℤ := by 19 | apply Polynomial.map_injective (algebraMap ℤ ℚ) (RingHom.injective_int (algebraMap ℤ ℚ)) 20 | rw [← minpoly.isIntegrallyClosed_eq_field_fractions ℚ (CyclotomicField p ℚ), 21 | ← cyclotomic_eq_minpoly_rat (n := p), map_cyclotomic] 22 | · exact hx 23 | · exact hpri.out.pos 24 | · exact IsIntegralClosure.isIntegral _ (CyclotomicField p ℚ) _ 25 | 26 | lemma AdjoinRoot.aeval_root {R} [CommRing R] (P : R[X]) : aeval (root P) P = 0 := by simp 27 | 28 | def AdjoinRoot.equivOfMinpolyEq {R S} [CommRing R] [CommRing S] [Algebra R S] 29 | (P : R[X]) (pb : PowerBasis R S) (hpb : minpoly R pb.gen = P) : 30 | AdjoinRoot P ≃ₐ[R] S := AdjoinRoot.equiv' P pb (hpb ▸ aeval_root _) (hpb ▸ minpoly.aeval _ _) 31 | 32 | namespace CyclotomicIntegers 33 | 34 | @[simps! -isSimp] 35 | def equiv : 36 | CyclotomicIntegers p ≃+* 𝓞 (CyclotomicField p ℚ) := by 37 | have H := IsCyclotomicExtension.zeta_spec p ℚ (CyclotomicField p ℚ) 38 | exact (AdjoinRoot.equivOfMinpolyEq (cyclotomic p ℤ) H.integralPowerBasis' 39 | (H.integralPowerBasis'_gen ▸ IsPrimitiveRoot.cyclotomic_eq_minpoly p H.toInteger H)).toRingEquiv 40 | 41 | instance : IsDomain (CyclotomicIntegers p) := 42 | AdjoinRoot.isDomain_of_prime (UniqueFactorizationMonoid.irreducible_iff_prime.mp 43 | (cyclotomic.irreducible hpri.out.pos)) 44 | 45 | def zeta : CyclotomicIntegers p := AdjoinRoot.root _ 46 | 47 | lemma equiv_zeta : equiv p (zeta p) = (IsCyclotomicExtension.zeta_spec 48 | p ℚ (CyclotomicField p ℚ)).toInteger := by 49 | rw [equiv_apply, zeta] 50 | simp only [AdjoinRoot.liftHom_root, IsPrimitiveRoot.integralPowerBasis'_gen] 51 | 52 | lemma prime_one_sub_zeta : 53 | Prime (1 - zeta p) := by 54 | rw [← prime_units_mul (a := -1), Units.val_neg, Units.val_one, neg_mul, one_mul, neg_sub] 55 | apply (MulEquiv.prime_iff (equiv p)).1 56 | simp only [RingEquiv.toMulEquiv_eq_coe, RingEquiv.coe_toMulEquiv, 57 | (equiv p).map_sub, (equiv p).map_one, equiv_zeta] 58 | have H := IsCyclotomicExtension.zeta_spec p ℚ (CyclotomicField p ℚ) 59 | exact H.zeta_sub_one_prime' 60 | 61 | lemma one_sub_zeta_mem_nonZeroDivisors : 62 | 1 - zeta p ∈ nonZeroDivisors (CyclotomicIntegers p) := by 63 | rw [mem_nonZeroDivisors_iff_ne_zero] 64 | exact (prime_one_sub_zeta p).1 65 | 66 | lemma not_isUnit_one_sub_zeta : 67 | ¬ IsUnit (1 - zeta p) := (prime_one_sub_zeta p).irreducible.1 68 | 69 | lemma one_sub_zeta_dvd_int_iff (n : ℤ) : 1 - zeta p ∣ n ↔ ↑p ∣ n := by 70 | have H := IsCyclotomicExtension.zeta_spec p ℚ (CyclotomicField p ℚ) 71 | rw [← map_dvd_iff (equiv p), map_sub, map_one, equiv_zeta, map_intCast, 72 | ← neg_dvd, neg_sub] 73 | exact zeta_sub_one_dvd_Int_iff H 74 | 75 | lemma one_sub_zeta_dvd : 1 - zeta p ∣ p := by 76 | show 1 - zeta p ∣ (p : ℤ) 77 | rw [one_sub_zeta_dvd_int_iff] 78 | 79 | lemma isCoprime_one_sub_zeta (n : ℤ) (hn : ¬ (p : ℤ) ∣ n) : IsCoprime (1 - zeta p) n := by 80 | apply (((Nat.prime_iff_prime_int.mp hpri.out).coprime_iff_not_dvd.mpr hn).map 81 | (algebraMap ℤ <| CyclotomicIntegers p)).of_isCoprime_of_dvd_left 82 | exact one_sub_zeta_dvd p 83 | 84 | lemma exists_dvd_int (n : CyclotomicIntegers p) (hn : n ≠ 0) : ∃ m : ℤ, m ≠ 0 ∧ n ∣ m := by 85 | refine ⟨Algebra.norm ℤ ((equiv p) n), by simpa, ?_⟩ 86 | rw [← map_dvd_iff (equiv p), map_intCast] 87 | convert RingOfIntegers.dvd_norm ℚ (equiv p n) using 1 88 | ext1 89 | exact DFunLike.congr_arg (algebraMap ℚ _) (Algebra.coe_norm_int (equiv p n)) 90 | 91 | def powerBasis : PowerBasis ℤ (CyclotomicIntegers p) := 92 | AdjoinRoot.powerBasis' (cyclotomic.monic _ _) 93 | 94 | lemma powerBasis_dim : (powerBasis p).dim = p - 1 := by 95 | simp [powerBasis, Nat.totient_prime hpri.out, natDegree_cyclotomic] 96 | 97 | instance : Module.Free ℤ (CyclotomicIntegers p) := ⟨_, (powerBasis p).basis⟩ 98 | 99 | lemma nontrivial {p} (hp : p ≠ 0) : Nontrivial (CyclotomicIntegers p) := by 100 | apply Ideal.Quotient.nontrivial 101 | simp only [ne_eq, Ideal.span_singleton_eq_top] 102 | intro h 103 | have := natDegree_eq_zero_of_isUnit h 104 | rw [natDegree_cyclotomic] at this 105 | exact this.not_gt (Nat.totient_pos.2 <| Nat.zero_lt_of_ne_zero hp) 106 | 107 | lemma charZero {p} (hp : p ≠ 0) : CharZero (CyclotomicIntegers p) := 108 | have := nontrivial hp 109 | ⟨(FaithfulSMul.algebraMap_injective _ _).comp (algebraMap ℕ ℤ).injective_nat⟩ 110 | 111 | instance : CharZero (CyclotomicIntegers p) := charZero hpri.out.ne_zero 112 | 113 | end CyclotomicIntegers 114 | end 115 | -------------------------------------------------------------------------------- /FltRegular/NumberTheory/GaloisPrime.lean: -------------------------------------------------------------------------------- 1 | import Mathlib.NumberTheory.RamificationInertia.Galois 2 | 3 | /-! 4 | # Galois action on primes 5 | 6 | ## Main Result 7 | - `exists_comap_galRestrict_eq`: The galois action on `primesOver` is transitive. 8 | 9 | -/ 10 | open Ideal 11 | 12 | section primesOver 13 | variable {R S : Type*} [CommRing R] [CommRing S] [Algebra R S] 14 | 15 | lemma ne_bot_of_mem_primesOver [IsDedekindDomain S] [NoZeroSMulDivisors R S] {p : Ideal R} 16 | (hp : p ≠ ⊥) {P : Ideal S} (hP : P ∈ primesOver p S) : 17 | P ≠ ⊥ := by 18 | have : P.LiesOver p := hP.2 19 | exact Ideal.ne_bot_of_liesOver_of_ne_bot hp _ 20 | 21 | lemma isMaximal_of_mem_primesOver [IsDedekindDomain S] [NoZeroSMulDivisors R S] {p : Ideal R} 22 | (hp : p ≠ ⊥) {P : Ideal S} (hP : P ∈ primesOver p S) : P.IsMaximal := 23 | hP.1.isMaximal (ne_bot_of_mem_primesOver hp hP) 24 | 25 | lemma prime_of_mem_primesOver [IsDedekindDomain S] [NoZeroSMulDivisors R S] {p : Ideal R} 26 | (hp : p ≠ ⊥) {P : Ideal S} (hP : P ∈ primesOver p S) : Prime P := 27 | prime_of_isPrime (ne_bot_of_mem_primesOver hp hP) hP.1 28 | 29 | end primesOver 30 | 31 | variable (R K L S : Type*) [CommRing R] [CommRing S] [Algebra R S] [Field K] [Field L] 32 | [Algebra R K] [IsFractionRing R K] [Algebra S L] 33 | [Algebra K L] [Algebra R L] [IsScalarTower R S L] [IsScalarTower R K L] 34 | [IsIntegralClosure S R L] [FiniteDimensional K L] 35 | 36 | lemma exists_comap_galRestrict_eq [IsDedekindDomain R] [IsGalois K L] {p : Ideal R} 37 | {P₁ P₂ : Ideal S} (hP₁ : P₁ ∈ primesOver p S) (hP₂ : P₂ ∈ primesOver p S) : 38 | ∃ σ, P₁.comap (galRestrict R K L S σ) = P₂ := by 39 | have : IsDomain S := 40 | (IsIntegralClosure.equiv R S L (integralClosure R L)).toMulEquiv.isDomain (integralClosure R L) 41 | have := IsIntegralClosure.isDedekindDomain R K L S 42 | have : Module.Finite R S := IsIntegralClosure.finite R K L S 43 | have := hP₁.1 44 | have := hP₁.2 45 | have := hP₂.1 46 | have := hP₂.2 47 | have : IsFractionRing S L := IsIntegralClosure.isFractionRing_of_finite_extension R K L S 48 | obtain ⟨σ, hσ⟩ := exists_map_eq_of_isGalois p P₂ P₁ K L 49 | refine ⟨(galRestrict R K L S).symm σ, ?_⟩ 50 | simp only [MulEquiv.apply_symm_apply, ← hσ] 51 | refine comap_map_of_bijective _ (AlgEquiv.bijective σ) 52 | -------------------------------------------------------------------------------- /FltRegular/NumberTheory/Hilbert90.lean: -------------------------------------------------------------------------------- 1 | import Mathlib.RingTheory.IntegralClosure.IntegralRestrict 2 | import Mathlib.RepresentationTheory.GroupCohomology.Hilbert90 3 | 4 | open scoped nonZeroDivisors 5 | open FiniteDimensional Finset BigOperators Submodule groupCohomology Submonoid 6 | 7 | variable {K L : Type*} [Field K] [Field L] [Algebra K L] [FiniteDimensional K L] 8 | variable {σ : L ≃ₐ[K] L} (hσ : ∀ x, x ∈ Subgroup.zpowers σ) 9 | variable {η : Lˣ} (hη : Algebra.norm K η.1 = 1) 10 | 11 | local notation3 "φ" => (finEquivZPowers (isOfFinOrder_of_finite σ)).symm 12 | 13 | lemma hφ : ∀ (n : ℕ), φ ⟨σ ^ n, hσ _⟩ = n % (orderOf σ) := fun n ↦ by 14 | simpa [Fin.ext_iff] using finEquivZPowers_symm_apply (isOfFinOrder_of_finite σ) n 15 | 16 | variable (η) in 17 | noncomputable 18 | def cocycle : (L ≃ₐ[K] L) → Lˣ := fun τ ↦ ∏ i ∈ range (φ ⟨τ, hσ τ⟩), Units.map (σ ^ i) η 19 | 20 | include hσ hη in 21 | lemma aux1 [IsGalois K L] {a: ℕ} (h : a % orderOf σ = 0) : ∏ i ∈ range a, (σ ^ i) η = 1 := by 22 | obtain ⟨n, hn⟩ := Nat.dvd_iff_mod_eq_zero.2 h 23 | rw [hn] 24 | revert a 25 | induction n with 26 | | zero => simp 27 | | succ n ih => 28 | intro a _ _ 29 | rw [Nat.mul_succ, prod_range_add, ih (Nat.mul_mod_right (orderOf σ) n) rfl, one_mul] 30 | simp only [pow_add, pow_mul, pow_orderOf_eq_one, one_pow, one_mul] 31 | have := Algebra.norm_eq_prod_automorphisms K η.1 32 | simp only [hη, map_one] at this 33 | convert this.symm 34 | refine prod_bij (fun n (_ : n ∈ range (orderOf σ)) ↦ σ ^ n) (by simp) 35 | (fun a ha b hb hab ↦ ?_) (fun τ _ ↦ ?_) (fun _ _ ↦ by rfl) 36 | · rwa [pow_inj_mod, Nat.mod_eq_of_lt (mem_range.1 ha), 37 | Nat.mod_eq_of_lt (mem_range.1 hb)] at hab 38 | · refine ⟨(finEquivZPowers (isOfFinOrder_of_finite σ)).symm ⟨τ, hσ τ⟩, by simp, ?_⟩ 39 | have := Equiv.symm_apply_apply (finEquivZPowers (isOfFinOrder_of_finite σ)).symm ⟨τ, hσ τ⟩ 40 | simp only [SetLike.coe_sort_coe, Equiv.symm_symm, ← Subtype.coe_inj] at this ⊢ 41 | rw [← this] 42 | simp only [SetLike.coe_sort_coe, Subtype.coe_eta, Equiv.symm_apply_apply] 43 | rfl 44 | 45 | include hσ hη in 46 | lemma aux2 [IsGalois K L] {a b : ℕ} (h : a % orderOf σ = b % orderOf σ) : 47 | ∏ i ∈ range a, (σ ^ i) η = ∏ i ∈ range b, (σ ^ i) η := by 48 | wlog hab : b ≤ a generalizing a b 49 | · exact (this h.symm (not_le.1 hab).le).symm 50 | obtain ⟨c, hc⟩ := Nat.dvd_iff_mod_eq_zero.2 (Nat.sub_mod_eq_zero_of_mod_eq h) 51 | rw [Nat.sub_eq_iff_eq_add hab] at hc 52 | rw [hc, prod_range_add, aux1 hσ hη (Nat.mul_mod_right (orderOf σ) c), one_mul] 53 | simp [pow_add, pow_mul, pow_orderOf_eq_one] 54 | 55 | lemma cocycle_spec (hone : orderOf σ ≠ 1) : (cocycle hσ η) σ = η := by 56 | haveI nezero : NeZero (orderOf σ) := 57 | ⟨fun hzero ↦ orderOf_eq_zero_iff.1 hzero (isOfFinOrder_of_finite σ)⟩ 58 | conv => 59 | enter [1, 3] 60 | rw [← pow_one σ] 61 | have : 1 % orderOf σ = 1 := by 62 | suffices (orderOf σ).pred.pred + 2 = orderOf σ by 63 | rw [← this] 64 | exact Nat.one_mod _ 65 | rw [← Nat.succ_eq_add_one, ← Nat.succ_eq_add_one, Nat.succ_pred, Nat.succ_pred nezero.1] 66 | intro h 67 | rw [show 0 = Nat.pred 1 by rfl] at h 68 | apply hone 69 | exact Nat.pred_inj (Nat.pos_of_ne_zero nezero.1) zero_lt_one h 70 | simp [this] 71 | have horder := hφ hσ 1 72 | simp only [SetLike.coe_sort_coe, pow_one] at horder 73 | simp only [cocycle, SetLike.coe_sort_coe, horder, this, range_one, prod_singleton, pow_zero] 74 | rfl 75 | 76 | include hη in 77 | lemma is_cocycle_aux [IsGalois K L] : ∀ (α β : (L ≃ₐ[K] L)), (cocycle hσ η) (α * β) = 78 | α ((cocycle hσ η) β) * (cocycle hσ η) α := by 79 | intro α β 80 | have hσmon : ∀ x, x ∈ powers σ := by simpa [← mem_powers_iff_mem_zpowers] using hσ 81 | obtain ⟨a, ha⟩ := (mem_powers_iff _ _).1 (hσmon α) 82 | obtain ⟨b, hb⟩ := (mem_powers_iff _ _).1 (hσmon β) 83 | rw [← ha, ← hb, ← pow_add] 84 | have Hab := hφ hσ (a + b) 85 | have Ha := hφ hσ a 86 | have Hb := hφ hσ b 87 | simp only [SetLike.coe_sort_coe, Nat.cast_add, Fin.ext_iff, Fin.mod_val, Fin.coe_ofNat_eq_mod, 88 | Nat.mod_self, Nat.mod_zero, cocycle, Units.coe_prod, Units.coe_map, MonoidHom.coe_coe, 89 | map_prod] at Hab Ha Hb ⊢ 90 | rw [Hab, Ha, Hb, mul_comm] 91 | have H : ∀ n, σ ^ (a + n) = σ ^ (a % orderOf σ + n) := fun n ↦ by simp [pow_inj_mod] 92 | conv => 93 | enter [2, 2, 2, x] 94 | rw [← AlgEquiv.mul_apply, ← pow_add, H] 95 | rw [← prod_range_add (fun (n : ℕ) ↦ (σ ^ n) η) (a % orderOf σ) (b % orderOf σ)] 96 | simpa using aux2 hσ hη (by simp) 97 | 98 | include hη in 99 | lemma is_cocycle [IsGalois K L] : IsMulOneCocycle (cocycle hσ η) := by 100 | intro α β 101 | simp [← Units.eq_iff, is_cocycle_aux hσ hη α β] 102 | 103 | include hη hσ in 104 | lemma Hilbert90 [IsGalois K L] : ∃ ε : L, η = ε / σ ε := by 105 | by_cases hone : orderOf σ = 1 106 | · suffices Module.finrank K L = 1 by 107 | obtain ⟨a, ha⟩ := mem_span_singleton.1 <| (eq_top_iff'.1 <| 108 | (finrank_eq_one_iff_of_nonzero _ one_ne_zero).1 this) η 109 | rw [← Algebra.algebraMap_eq_smul_one] at ha 110 | rw [← ha, Algebra.norm_algebraMap, this, pow_one] at hη 111 | exact ⟨1, by simp [← ha, hη]⟩ 112 | rw [← IsGalois.card_aut_eq_finrank, Fintype.card_eq_one_iff] 113 | refine ⟨σ, fun τ ↦ ?_⟩ 114 | simp only [orderOf_eq_one_iff.1 hone, Subgroup.zpowers_one_eq_bot, Subgroup.mem_bot] at hσ 115 | rw [orderOf_eq_one_iff.1 hone, hσ τ] 116 | obtain ⟨ε, hε⟩ := isMulOneCoboundary_of_isMulOneCocycle_of_aut_to_units _ (is_cocycle hσ hη) 117 | use ε⁻¹ 118 | simp only [map_inv₀, div_inv_eq_mul] 119 | specialize hε σ 120 | nth_rewrite 2 [← inv_inv ε] at hε 121 | rw [div_inv_eq_mul, cocycle_spec hσ hone, mul_inv_eq_iff_eq_mul, mul_comm, 122 | ← Units.eq_iff] at hε 123 | simp only [AlgEquiv.smul_units_def, Units.coe_map, MonoidHom.coe_coe, Units.val_mul] at hε 124 | symm 125 | rw [inv_mul_eq_iff_eq_mul₀ ε.ne_zero, hε] 126 | 127 | variable {A B : Type*} [CommRing A] [CommRing B] [Algebra A B] [Algebra A L] [Algebra A K] 128 | variable [Algebra B L] [IsScalarTower A B L] [IsScalarTower A K L] [IsFractionRing A K] [IsDomain A] 129 | variable [IsIntegralClosure B A L] 130 | 131 | include hσ in 132 | lemma Hilbert90_integral [IsGalois K L] {η : B} (hη : Algebra.norm K (algebraMap B L η) = 1) : 133 | ∃ ε : B, ε ≠ 0 ∧ η * galRestrict A K L B σ ε = ε := by 134 | have : NoZeroSMulDivisors A L := by 135 | rw [NoZeroSMulDivisors.iff_algebraMap_injective, IsScalarTower.algebraMap_eq A K L] 136 | exact (algebraMap K L).injective.comp (IsFractionRing.injective A K) 137 | have : IsLocalization (Algebra.algebraMapSubmonoid B A⁰) L := 138 | IsIntegralClosure.isLocalization A K L B 139 | let η' : Lˣ := IsUnit.unit (a := (algebraMap B L η)) (isUnit_iff_ne_zero.2 140 | (fun h ↦ by simp [h] at hη)) 141 | replace hη : Algebra.norm K η'.1 = 1 := hη 142 | obtain ⟨ε, hε⟩ := Hilbert90 hσ hη 143 | obtain ⟨x, y, rfl⟩ := IsLocalization.mk'_surjective (Algebra.algebraMapSubmonoid B A⁰) ε 144 | obtain ⟨t, ht, ht'⟩ := y.prop 145 | have : t • IsLocalization.mk' L x y = algebraMap _ _ x := by 146 | rw [Algebra.smul_def, IsScalarTower.algebraMap_apply A B L, ht', IsLocalization.mk'_spec'] 147 | refine ⟨x, ?_, ?_⟩ 148 | · rintro rfl 149 | simp only [IsLocalization.mk'_zero, _root_.map_zero, ne_eq, not_true, div_zero] at hε 150 | rw [hε, Algebra.norm_zero] at hη 151 | exact zero_ne_one hη 152 | · rw [eq_div_iff_mul_eq] at hε 153 | · replace hε := congr_arg (t • ·) hε 154 | simp only at hε 155 | rw [Algebra.smul_def, mul_left_comm, ← Algebra.smul_def t] at hε 156 | change (algebraMap B L) η * t • σ.toAlgHom _ = _ at hε 157 | rw [← AlgHom.map_smul_of_tower, this] at hε 158 | apply IsIntegralClosure.algebraMap_injective B A L 159 | rw [map_mul, ← hε] 160 | congr 1 161 | · exact algebraMap_galRestrictHom_apply A K L B σ x 162 | · intro e 163 | rw [(map_eq_zero _).mp e, zero_div] at hε 164 | rw [hε, Algebra.norm_zero] at hη 165 | exact zero_ne_one hη 166 | -------------------------------------------------------------------------------- /FltRegular/NumberTheory/Hilbert94.lean: -------------------------------------------------------------------------------- 1 | import FltRegular.NumberTheory.Unramified 2 | import FltRegular.NumberTheory.Hilbert92 3 | import FltRegular.NumberTheory.Hilbert90 4 | import FltRegular.NumberTheory.IdealNorm 5 | import FltRegular.NumberTheory.RegularPrimes 6 | import Mathlib.NumberTheory.NumberField.ClassNumber 7 | 8 | open scoped NumberField 9 | 10 | variable {K : Type*} {p : ℕ} [hpri : Fact p.Prime] [Field K] 11 | 12 | open Polynomial Module 13 | 14 | variable {L} [Field L] [Algebra K L] [FiniteDimensional K L] 15 | (σ : L ≃ₐ[K] L) (hσ : ∀ x, x ∈ Subgroup.zpowers σ) (hKL : finrank K L = p) 16 | 17 | variable {A B : Type*} [CommRing A] [CommRing B] [Algebra A B] [Algebra A L] [Algebra A K] 18 | [Algebra B L] [IsScalarTower A B L] [IsScalarTower A K L] [IsFractionRing A K] 19 | [IsIntegralClosure B A L] 20 | 21 | instance : Algebra.IsAlgebraic K L := Algebra.IsAlgebraic.of_finite K L 22 | 23 | include hσ in 24 | lemma comap_span_galRestrict_eq_of_cyclic (β : B) (η : Bˣ) (hβ : η * (galRestrict A K L B σ) β = β) 25 | (σ' : L ≃ₐ[K] L) : 26 | (Ideal.span {β}).comap (galRestrict A K L B σ') = Ideal.span {β} := by 27 | suffices (Ideal.span {β}).map 28 | (galRestrict A K L B σ'⁻¹).toRingEquiv.toRingHom = Ideal.span {β} by 29 | rwa [RingEquiv.toRingHom_eq_coe, Ideal.map_comap_of_equiv, map_inv] at this 30 | apply_fun (Ideal.span {·}) at hβ 31 | rw [← Ideal.span_singleton_mul_span_singleton, Ideal.span_singleton_eq_top.mpr η.isUnit, 32 | ← Ideal.one_eq_top, one_mul, ← Set.image_singleton, ← Ideal.map_span] at hβ 33 | change Ideal.map (galRestrict A K L B σ : B →+* B) _ = _ at hβ 34 | generalize σ'⁻¹ = σ' 35 | obtain ⟨n, rfl : σ ^ n = σ'⟩ := mem_powers_iff_mem_zpowers.mpr (hσ σ') 36 | rw [map_pow] 37 | induction n with 38 | | zero => 39 | simp only [Nat.zero_eq, pow_zero, AlgEquiv.toRingEquiv_eq_coe, RingEquiv.toRingHom_eq_coe] 40 | exact Ideal.map_id _ 41 | | succ n IH => 42 | simp only [AlgEquiv.toRingEquiv_eq_coe, RingEquiv.toRingHom_eq_coe, pow_succ] at IH ⊢ 43 | conv_lhs at IH => rw [← hβ, Ideal.map_map] 44 | exact IH 45 | 46 | variable [IsGalois K L] 47 | 48 | include hσ in 49 | open FiniteDimensional in 50 | theorem exists_not_isPrincipal_and_isPrincipal_map_aux 51 | [IsDedekindDomain A] [IsUnramified A B] (η : Bˣ) 52 | (hη : Algebra.norm K (algebraMap B L η) = 1) 53 | (hη' : ¬∃ α : Bˣ, algebraMap B L η = (algebraMap B L α) / σ (algebraMap B L α)) : 54 | ∃ I : Ideal A, ¬I.IsPrincipal ∧ (I.map (algebraMap A B)).IsPrincipal := by 55 | obtain ⟨β, hβ_zero, hβ⟩ := Hilbert90_integral (A := A) (B := B) hσ hη 56 | haveI : IsDomain B := 57 | (IsIntegralClosure.equiv A B L (integralClosure A L)).toMulEquiv.isDomain (integralClosure A L) 58 | have hβ' := comap_map_eq_of_isUnramified K L _ 59 | (comap_span_galRestrict_eq_of_cyclic σ hσ (A := A) (B := B) β η hβ) 60 | refine ⟨(Ideal.span {β}).comap (algebraMap A B), ?_, ?_⟩ 61 | · rintro ⟨⟨γ, hγ : _ = Ideal.span _⟩⟩ 62 | rw [hγ, Ideal.map_span, Set.image_singleton, Ideal.span_singleton_eq_span_singleton] at hβ' 63 | obtain ⟨a, rfl⟩ := hβ' 64 | rw [map_mul, AlgEquiv.commutes, mul_left_comm, (mul_right_injective₀ _).eq_iff] at hβ 65 | · apply hη' 66 | use a 67 | conv_rhs => enter [1]; rw [← hβ] 68 | rw [map_mul, ← AlgHom.coe_coe σ, ← algebraMap_galRestrictHom_apply A K L B σ a] 69 | refine (mul_div_cancel_right₀ _ ?_).symm 70 | · rw [ne_eq, (injective_iff_map_eq_zero' _).mp (IsIntegralClosure.algebraMap_injective B A L), 71 | (injective_iff_map_eq_zero' _).mp (galRestrict A K L B σ).injective] 72 | exact a.isUnit.ne_zero 73 | · exact (mul_ne_zero_iff.mp hβ_zero).1 74 | · rw [hβ'] 75 | exact ⟨⟨_, rfl⟩⟩ 76 | 77 | theorem Ideal.isPrincipal_pow_finrank_of_isPrincipal_map [IsDedekindDomain A] {I : Ideal A} 78 | (hI : (I.map (algebraMap A B)).IsPrincipal) : (I ^ finrank K L).IsPrincipal := by 79 | haveI : IsDomain B := 80 | (IsIntegralClosure.equiv A B L (integralClosure A L)).toMulEquiv.isDomain (integralClosure A L) 81 | haveI := IsIntegralClosure.isNoetherian A K L B 82 | haveI := IsIntegralClosure.isDedekindDomain A K L B 83 | haveI := IsIntegralClosure.isFractionRing_of_finite_extension A K L B 84 | have hAB : Function.Injective (algebraMap A B) := by 85 | refine Function.Injective.of_comp (f := algebraMap B L) ?_ 86 | rw [← RingHom.coe_comp, ← IsScalarTower.algebraMap_eq, IsScalarTower.algebraMap_eq A K L] 87 | exact (algebraMap K L).injective.comp (IsFractionRing.injective _ _) 88 | rw [← NoZeroSMulDivisors.iff_algebraMap_injective] at hAB 89 | letI : Algebra (FractionRing A) (FractionRing B) := FractionRing.liftAlgebra _ _ 90 | have : IsScalarTower A (FractionRing A) (FractionRing B) := 91 | FractionRing.isScalarTower_liftAlgebra _ _ 92 | have H : RingHom.comp (algebraMap (FractionRing A) (FractionRing B)) 93 | (FractionRing.algEquiv A K).symm.toRingEquiv = 94 | RingHom.comp (FractionRing.algEquiv B L).symm.toRingEquiv (algebraMap K L) := by 95 | apply IsLocalization.ringHom_ext (nonZeroDivisors A) 96 | ext 97 | simp only [AlgEquiv.toRingEquiv_eq_coe, RingHom.coe_comp, RingHom.coe_coe, 98 | AlgEquiv.coe_ringEquiv, Function.comp_apply, AlgEquiv.commutes, 99 | ← IsScalarTower.algebraMap_apply] 100 | rw [IsScalarTower.algebraMap_apply A B L, AlgEquiv.commutes, ← IsScalarTower.algebraMap_apply] 101 | have : Algebra.IsSeparable (FractionRing A) (FractionRing B) := 102 | Algebra.IsSeparable.of_equiv_equiv _ _ H 103 | have hLK : finrank (FractionRing A) (FractionRing B) = finrank K L := by 104 | simpa only [Cardinal.toNat_lift] using congr_arg Cardinal.toNat 105 | (Algebra.lift_rank_eq_of_equiv_equiv (FractionRing.algEquiv A K).symm.toRingEquiv 106 | (FractionRing.algEquiv B L).symm.toRingEquiv H).symm 107 | rw [← hLK, ← Ideal.spanNorm_map, ← (I.map (algebraMap A B)).span_singleton_generator, 108 | Ideal.spanNorm_singleton] 109 | exact ⟨⟨_, rfl⟩⟩ 110 | 111 | /-- This is the first part of **Hilbert Theorem 94**, which states that if `L/K` is an unramified 112 | cyclic finite extension of number fields of odd prime degree, 113 | then there is an ideal that capitulates in `K`. -/ 114 | theorem exists_not_isPrincipal_and_isPrincipal_map (K L : Type*) 115 | [Field K] [Field L] [NumberField K] [NumberField L] [Algebra K L] 116 | [FiniteDimensional K L] [IsGalois K L] [IsUnramified (𝓞 K) (𝓞 L)] [h : IsCyclic (L ≃ₐ[K] L)] 117 | (hKL : Nat.Prime (finrank K L)) 118 | (hKL' : finrank K L ≠ 2) : 119 | ∃ I : Ideal (𝓞 K), ¬I.IsPrincipal ∧ (I.map (algebraMap (𝓞 K) (𝓞 L))).IsPrincipal := by 120 | obtain ⟨⟨σ, hσ⟩⟩ := h 121 | obtain ⟨η, hη, hη'⟩ := Hilbert92 hKL hKL' σ hσ 122 | exact exists_not_isPrincipal_and_isPrincipal_map_aux σ hσ η hη (not_exists.mpr hη') 123 | 124 | /-- This is the second part of **Hilbert Theorem 94**, which states that if `L/K` is an unramified 125 | cyclic finite extension of number fields of odd prime degree, 126 | then the degree divides the class number of `K`. -/ 127 | theorem dvd_card_classGroup_of_isUnramified_isCyclic (K L : Type*) 128 | [Field K] [Field L] [NumberField K] [NumberField L] [Algebra K L] 129 | [FiniteDimensional K L] [IsGalois K L] [IsUnramified (𝓞 K) (𝓞 L)] [IsCyclic (L ≃ₐ[K] L)] 130 | (hKL : Nat.Prime (finrank K L)) 131 | (hKL' : finrank K L ≠ 2) : 132 | finrank K L ∣ Fintype.card (ClassGroup (𝓞 K)) := by 133 | obtain ⟨I, hI, hI'⟩ := exists_not_isPrincipal_and_isPrincipal_map K L hKL hKL' 134 | have := Fact.mk hKL 135 | rw [hKL.dvd_iff_not_coprime] 136 | exact fun h ↦ hI (isPrincipal_of_isPrincipal_pow_of_coprime h 137 | (Ideal.isPrincipal_pow_finrank_of_isPrincipal_map hI')) 138 | -------------------------------------------------------------------------------- /FltRegular/NumberTheory/IdealNorm.lean: -------------------------------------------------------------------------------- 1 | import Mathlib.RingTheory.DedekindDomain.Ideal 2 | import Mathlib.RingTheory.DedekindDomain.Dvr 3 | import Mathlib.RingTheory.Localization.NormTrace 4 | import Mathlib.RingTheory.Localization.LocalizationLocalization 5 | import Mathlib.RingTheory.Nakayama 6 | import Mathlib.RingTheory.LocalProperties.IntegrallyClosed 7 | import Mathlib.RingTheory.IntegralClosure.IntegralRestrict 8 | import Mathlib.RingTheory.IntegralClosure.IntegrallyClosed 9 | import Mathlib.LinearAlgebra.FreeModule.PID 10 | import Mathlib.RingTheory.DedekindDomain.IntegralClosure 11 | import Mathlib.RingTheory.DedekindDomain.PID 12 | import Mathlib.FieldTheory.Separable 13 | import Mathlib.RingTheory.Ideal.Norm.RelNorm 14 | 15 | open scoped nonZeroDivisors 16 | 17 | attribute [local instance] FractionRing.liftAlgebra 18 | section intNorm 19 | 20 | variable (R S K L) [CommRing R] [CommRing S] [Field K] [Field L] 21 | [IsIntegrallyClosed R] [Algebra R S] [Algebra R K] [Algebra S L] [Algebra K L] 22 | [Algebra R L] [IsScalarTower R K L] [IsScalarTower R S L] [IsIntegralClosure S R L] 23 | [IsFractionRing R K] [FiniteDimensional K L] [Algebra.IsSeparable K L] 24 | [IsDomain R] [IsDomain S] [NoZeroSMulDivisors R S] [hRS : Module.Finite R S] 25 | [IsIntegrallyClosed S] [Algebra.IsSeparable (FractionRing R) (FractionRing S)] 26 | 27 | instance : FiniteDimensional (FractionRing R) (FractionRing S) := 28 | Module.Finite_of_isLocalization R S _ _ R⁰ 29 | 30 | open nonZeroDivisors 31 | 32 | end intNorm 33 | 34 | section SpanNorm 35 | 36 | namespace Ideal 37 | 38 | open Submodule 39 | 40 | variable (R : Type*) [CommRing R] {S : Type*} [CommRing S] [Algebra R S] 41 | variable [IsIntegrallyClosed R] [IsDomain R] [IsDomain S] [NoZeroSMulDivisors R S] 42 | [IsIntegrallyClosed S] [Algebra.IsSeparable (FractionRing R) (FractionRing S)] 43 | 44 | section 45 | 46 | variable [hRS : Module.Finite R S] 47 | 48 | variable [IsDedekindDomain R] [IsDedekindDomain S] [Module.Finite R S] [Module.Free R S] 49 | 50 | theorem IsLocalization.AtPrime.PID_of_dedekind_domain {A} [CommRing A] 51 | [IsDedekindDomain A] 52 | (P : Ideal A) [pP : P.IsPrime] (Aₘ : Type*) [CommRing Aₘ] [IsDomain Aₘ] 53 | [Algebra A Aₘ] [IsLocalization.AtPrime Aₘ P] : IsPrincipalIdealRing Aₘ := 54 | have : IsNoetherianRing Aₘ := 55 | IsLocalization.isNoetherianRing P.primeCompl _ IsDedekindRing.toIsNoetherian 56 | have : IsLocalRing Aₘ := IsLocalization.AtPrime.isLocalRing Aₘ P 57 | ((tfae_of_isNoetherianRing_of_isLocalRing_of_isDomain Aₘ).out 2 0).mp 58 | (IsLocalization.AtPrime.isDedekindDomain A P _) 59 | 60 | end 61 | 62 | variable [Module.Finite R S] [IsDedekindDomain R] [IsDedekindDomain S] 63 | 64 | /-- Multiplicativity of `Ideal.spanIntNorm`. simp-normal form is `map_mul (Ideal.relNorm R)`. -/ 65 | theorem spanNorm_map (I : Ideal R) : 66 | spanNorm R (I.map (algebraMap R S)) = 67 | I ^ (Module.finrank (FractionRing R) (FractionRing S)) := by 68 | nontriviality R 69 | nontriviality S 70 | refine eq_of_localization_maximal ?_ 71 | intro P hPd 72 | let P' := Algebra.algebraMapSubmonoid S P.primeCompl 73 | let Rₚ := Localization.AtPrime P 74 | let Sₚ := Localization P' 75 | let _ : Algebra Rₚ Sₚ := localizationAlgebra P.primeCompl S 76 | haveI : IsScalarTower R Rₚ Sₚ := 77 | IsScalarTower.of_algebraMap_eq (fun x => 78 | (IsLocalization.map_eq (T := P') (Q := Localization P') P.primeCompl.le_comap_map x).symm) 79 | have h : P' ≤ S⁰ := 80 | map_le_nonZeroDivisors_of_injective _ (FaithfulSMul.algebraMap_injective _ _) 81 | P.primeCompl_le_nonZeroDivisors 82 | have : IsDomain Sₚ := IsLocalization.isDomain_localization h 83 | have : IsDedekindDomain Sₚ := IsLocalization.isDedekindDomain S h _ 84 | have : IsPrincipalIdealRing Rₚ := IsLocalization.AtPrime.PID_of_dedekind_domain P Rₚ 85 | have := isIntegrallyClosed_of_isLocalization Rₚ P.primeCompl P.primeCompl_le_nonZeroDivisors 86 | have := NoZeroSMulDivisors_of_isLocalization R S Rₚ Sₚ P.primeCompl_le_nonZeroDivisors 87 | have := Module.Finite_of_isLocalization R S Rₚ Sₚ P.primeCompl 88 | let K := FractionRing R 89 | let f : Rₚ →+* K := IsLocalization.map _ (T := R⁰) (RingHom.id R) P.primeCompl_le_nonZeroDivisors 90 | letI := f.toAlgebra 91 | have : IsScalarTower R Rₚ K := IsScalarTower.of_algebraMap_eq' 92 | (by rw [RingHom.algebraMap_toAlgebra, IsLocalization.map_comp, RingHomCompTriple.comp_eq]) 93 | letI := IsFractionRing.isFractionRing_of_isDomain_of_isLocalization P.primeCompl Rₚ K 94 | let L := FractionRing S 95 | let g : Sₚ →+* L := IsLocalization.map _ (M := P') (T := S⁰) (RingHom.id S) h 96 | letI := g.toAlgebra 97 | haveI : IsScalarTower S Sₚ (FractionRing S) := IsScalarTower.of_algebraMap_eq' 98 | (by rw [RingHom.algebraMap_toAlgebra, IsLocalization.map_comp, RingHom.comp_id]) 99 | letI := IsFractionRing.isFractionRing_of_isDomain_of_isLocalization 100 | P' Sₚ (FractionRing S) 101 | let _ := ((algebraMap K L).comp f).toAlgebra 102 | have : IsScalarTower Rₚ K L := IsScalarTower.of_algebraMap_eq' rfl 103 | have : IsScalarTower Rₚ Sₚ L := by 104 | apply IsScalarTower.of_algebraMap_eq' 105 | apply IsLocalization.ringHom_ext P.primeCompl 106 | rw [RingHom.algebraMap_toAlgebra, RingHom.algebraMap_toAlgebra (R := Sₚ), RingHom.comp_assoc, 107 | RingHom.comp_assoc, ← IsScalarTower.algebraMap_eq, IsScalarTower.algebraMap_eq R S Sₚ, 108 | IsLocalization.map_comp, RingHom.comp_id, ← RingHom.comp_assoc, IsLocalization.map_comp, 109 | RingHom.comp_id, ← IsScalarTower.algebraMap_eq, ← IsScalarTower.algebraMap_eq] 110 | have : IsIntegralClosure Sₚ Rₚ L := IsIntegralClosure.of_isIntegrallyClosed _ _ _ 111 | have : Algebra.IsSeparable (FractionRing Rₚ) (FractionRing Sₚ) := by 112 | apply Algebra.IsSeparable.of_equiv_equiv 113 | (FractionRing.algEquiv Rₚ (FractionRing R)).symm.toRingEquiv 114 | (FractionRing.algEquiv Sₚ (FractionRing S)).symm.toRingEquiv 115 | apply IsLocalization.ringHom_ext R⁰ 116 | ext 117 | simp only [AlgEquiv.toRingEquiv_eq_coe, RingHom.coe_comp, 118 | RingHom.coe_coe, Function.comp_apply, ← IsScalarTower.algebraMap_apply] 119 | rw [IsScalarTower.algebraMap_apply R Rₚ (FractionRing R), AlgEquiv.coe_ringEquiv, 120 | AlgEquiv.commutes, IsScalarTower.algebraMap_apply R S L, 121 | IsScalarTower.algebraMap_apply S Sₚ L, AlgEquiv.coe_ringEquiv, AlgEquiv.commutes] 122 | simp only [← IsScalarTower.algebraMap_apply] 123 | have : Module.Free Rₚ Sₚ := Module.free_of_finite_type_torsion_free' 124 | simp only [Ideal.map_mul, ← spanIntNorm_localization (R := R) (S := S) 125 | (Rₘ := Localization.AtPrime P) (Sₘ := Localization P') _ _ P.primeCompl_le_nonZeroDivisors] 126 | rw [Ideal.map_pow, I.map_map, ← IsScalarTower.algebraMap_eq, IsScalarTower.algebraMap_eq R Rₚ Sₚ, 127 | ← I.map_map, ← (I.map _).span_singleton_generator, Ideal.map_span, Set.image_singleton, 128 | spanNorm_singleton, Ideal.span_singleton_pow] 129 | congr 2 130 | apply IsFractionRing.injective Rₚ K 131 | rw [Algebra.algebraMap_intNorm (L := L), ← IsScalarTower.algebraMap_apply, 132 | IsScalarTower.algebraMap_apply Rₚ K L, Algebra.norm_algebraMap, map_pow] 133 | 134 | end Ideal 135 | 136 | end SpanNorm 137 | -------------------------------------------------------------------------------- /FltRegular/NumberTheory/KummersLemma/Field.lean: -------------------------------------------------------------------------------- 1 | import Mathlib.NumberTheory.Cyclotomic.Rat 2 | import Mathlib.FieldTheory.KummerExtension 3 | import FltRegular.NumberTheory.Unramified 4 | import FltRegular.NumberTheory.Cyclotomic.MoreLemmas 5 | 6 | open scoped NumberField BigOperators 7 | 8 | variable {K : Type*} {p : ℕ} [hpri : Fact p.Prime] [Field K] [NumberField K] (hp : p ≠ 2) 9 | 10 | variable {ζ : K} (hζ : IsPrimitiveRoot ζ p) (u : (𝓞 K)ˣ) 11 | (hcong : (hζ.unit' - 1 : 𝓞 K) ^ p ∣ (↑u : 𝓞 K) - 1) (hu : ∀ v : K, v ^ p ≠ u) 12 | 13 | open Polynomial 14 | 15 | include hcong hp in 16 | lemma zeta_sub_one_pow_dvd_poly [IsCyclotomicExtension {p} ℚ K] : 17 | C ((hζ.unit' - 1 : 𝓞 K) ^ p) ∣ 18 | (C (hζ.unit' - 1 : 𝓞 K) * X - 1) ^ p + C (u : 𝓞 K) := by 19 | rw [← dvd_sub_left (_root_.map_dvd C hcong), add_sub_assoc, C.map_sub (u : 𝓞 K), ← sub_add, 20 | sub_self, map_one, zero_add] 21 | refine dvd_C_mul_X_sub_one_pow_add_one hpri.out hp _ _ dvd_rfl ?_ 22 | convert mul_dvd_mul_right (associated_zeta_sub_one_pow_prime hζ).dvd _ 23 | rw [← pow_succ, tsub_add_cancel_of_le (Nat.Prime.one_lt hpri.out).le] 24 | 25 | namespace KummersLemma 26 | 27 | lemma natDegree_poly_aux : 28 | natDegree ((C (hζ.unit' - 1 : 𝓞 K) * X - 1) ^ p + C (u : 𝓞 K)) = p := by 29 | haveI : Fact (Nat.Prime p) := hpri 30 | rw [natDegree_add_C, natDegree_pow, ← C.map_one, natDegree_sub_C, natDegree_mul_X, natDegree_C, 31 | zero_add, mul_one] 32 | exact C_ne_zero.mpr (hζ.unit'_coe.sub_one_ne_zero hpri.out.one_lt) 33 | 34 | lemma monic_poly_aux : 35 | leadingCoeff ((C (hζ.unit' - 1 : 𝓞 K) * X - 1) ^ p + C (u : 𝓞 K)) = 36 | (hζ.unit' - 1 : 𝓞 K) ^ p := by 37 | haveI : Fact (Nat.Prime p) := hpri 38 | trans leadingCoeff ((C (hζ.unit' - 1 : 𝓞 K) * X - 1) ^ p) 39 | · rw [leadingCoeff, leadingCoeff, coeff_add] 40 | nth_rewrite 1 [natDegree_add_C] 41 | convert add_zero _ using 2 42 | rw [natDegree_poly_aux hζ, coeff_C, if_neg (NeZero.pos p).ne.symm] 43 | · rw [leadingCoeff_pow, ← C.map_one, leadingCoeff, natDegree_sub_C, natDegree_mul_X] 44 | · simp only [map_one, natDegree_C, zero_add, coeff_sub, coeff_mul_X, coeff_C, ite_true, 45 | coeff_one, ite_false, sub_zero, one_ne_zero, ↓reduceIte] 46 | · exact C_ne_zero.mpr (hζ.unit'_coe.sub_one_ne_zero hpri.out.one_lt) 47 | 48 | 49 | variable [IsCyclotomicExtension {p} ℚ K] 50 | 51 | noncomputable def poly : (𝓞 K)[X] := (zeta_sub_one_pow_dvd_poly hp hζ u hcong).choose 52 | 53 | lemma poly_spec : 54 | C ((hζ.unit' - 1 : 𝓞 K) ^ p) * poly hp hζ u hcong = 55 | (C (hζ.unit' - 1 : 𝓞 K) * X - 1) ^ p + C (u : 𝓞 K) := 56 | (zeta_sub_one_pow_dvd_poly hp hζ u hcong).choose_spec.symm 57 | 58 | lemma monic_poly : Monic (poly hp hζ u hcong) := by 59 | haveI : Fact (Nat.Prime p) := hpri 60 | have := congr_arg leadingCoeff (poly_spec hp hζ u hcong) 61 | simp only [map_pow, leadingCoeff_mul, leadingCoeff_pow, leadingCoeff_C, ne_eq, PNat.pos, 62 | pow_eq_zero_iff, monic_poly_aux hζ] at this 63 | refine mul_right_injective₀ ?_ (this.trans (mul_one _).symm) 64 | exact pow_ne_zero _ (hζ.unit'_coe.sub_one_ne_zero hpri.out.one_lt) 65 | 66 | lemma natDegree_poly : natDegree (poly hp hζ u hcong) = p := by 67 | haveI : Fact (Nat.Prime p) := hpri 68 | have := congr_arg natDegree (poly_spec hp hζ u hcong) 69 | rwa [natDegree_C_mul, natDegree_poly_aux hζ] at this 70 | exact pow_ne_zero _ (hζ.unit'_coe.sub_one_ne_zero hpri.out.one_lt) 71 | 72 | lemma map_poly : (poly hp hζ u hcong).map (algebraMap (𝓞 K) K) = 73 | (X - C (1 / (ζ - 1))) ^ p + C (u / (ζ - 1) ^ p : K) := by 74 | ext i 75 | have := congr_arg (fun P : (𝓞 K)[X] ↦ (↑(coeff P i) : K)) (poly_spec hp hζ u hcong) 76 | change _ = algebraMap (𝓞 K) K _ at this 77 | rw [← coeff_map] at this 78 | replace this : (ζ - 1) ^ p * ↑((poly hp hζ u hcong).coeff i) = 79 | (((C ((algebraMap ((𝓞 K)) K) ↑hζ.unit') - 1) * X - 1) ^ p).coeff i + 80 | (C ((algebraMap ((𝓞 K)) K) ↑u)).coeff i := by 81 | simp only [map_pow, map_sub, map_one, Polynomial.map_add, Polynomial.map_pow, 82 | Polynomial.map_sub, Polynomial.map_mul, map_C, 83 | Polynomial.map_one, map_X, coeff_add] at this 84 | convert this 85 | simp only [NumberField.RingOfIntegers.coe_eq_algebraMap, ← Polynomial.coeff_map] 86 | simp only [coeff_map, Polynomial.map_mul, Polynomial.map_pow, Polynomial.map_sub, map_C, 87 | Polynomial.map_one] 88 | rw [← Polynomial.coeff_map, mul_comm, ← Polynomial.coeff_mul_C, mul_comm] 89 | simp [show hζ.unit'.1 = ζ from rfl] 90 | apply mul_right_injective₀ (pow_ne_zero p (hζ.sub_one_ne_zero hpri.out.one_lt)) 91 | simp only [Subalgebra.algebraMap_eq, Algebra.id.map_eq_id, RingHomCompTriple.comp_eq, coeff_map, 92 | RingHom.coe_coe, Subalgebra.coe_val, one_div, map_sub, map_one, coeff_add, coeff_sub, 93 | PNat.pos, pow_eq_zero_iff, this, mul_add] 94 | simp_rw [← smul_eq_mul (α := K), ← coeff_smul, show hζ.unit'.1 = ζ from rfl] 95 | rw [smul_C, smul_eq_mul, ← _root_.smul_pow, ← mul_div_assoc, mul_div_cancel_left₀, smul_sub, 96 | smul_C, smul_eq_mul, mul_inv_cancel₀, map_one, Algebra.smul_def, ← C_eq_algebraMap, map_sub, 97 | map_one] 98 | · exact hζ.sub_one_ne_zero hpri.out.one_lt 99 | · exact pow_ne_zero _ (hζ.sub_one_ne_zero hpri.out.one_lt) 100 | 101 | include hu in 102 | lemma irreducible_map_poly : 103 | Irreducible ((poly hp hζ u hcong).map (algebraMap (𝓞 K) K)) := by 104 | rw [map_poly] 105 | refine Irreducible.of_map (f := algEquivAevalXAddC (1 / (ζ - 1))) ?_ 106 | simp only [one_div, map_add, algEquivAevalXAddC_apply, map_pow, map_sub, aeval_X, aeval_C, 107 | algebraMap_eq, add_sub_cancel_right] 108 | rw [← sub_neg_eq_add, ← (C : K →+* _).map_neg] 109 | apply X_pow_sub_C_irreducible_of_prime hpri.out 110 | intro b hb 111 | apply hu (- b * (ζ - 1)) 112 | rw [mul_pow, (hpri.out.odd_of_ne_two hp).neg_pow, hb, neg_neg, 113 | div_mul_cancel₀ _ (pow_ne_zero _ (hζ.sub_one_ne_zero hpri.out.one_lt))] 114 | 115 | theorem aeval_poly {L : Type*} [Field L] [Algebra K L] (α : L) 116 | (e : α ^ p = algebraMap K L u) (m : ℕ) : 117 | aeval (((1 : L) - ζ ^ m • α) / (algebraMap K L (ζ - 1))) (poly hp hζ u hcong) = 0 := by 118 | have hζ' : algebraMap K L ζ - 1 ≠ 0 := by 119 | simpa using (algebraMap K L).injective.ne (hζ.sub_one_ne_zero hpri.out.one_lt) 120 | rw [map_sub, map_one] 121 | have := congr_arg (aeval ((1 - ζ ^ m • α) / (algebraMap K L (ζ - 1)))) 122 | (poly_spec hp hζ u hcong) 123 | -- also add to mathlib 124 | have hcoe : (algebraMap (𝓞 K) L) (↑hζ.unit') = algebraMap K L ζ := rfl 125 | have hcoe1 : (algebraMap (𝓞 K) L) ↑u = algebraMap K L ↑↑u := rfl 126 | simp only [map_sub, map_one, map_pow, map_mul, aeval_C, Subalgebra.algebraMap_eq, _root_.smul_pow, 127 | hcoe, RingHom.coe_comp, RingHom.coe_coe, Subalgebra.coe_val, Function.comp_apply, e, hcoe1, 128 | map_add, aeval_X, ← mul_div_assoc, mul_div_cancel_left₀ _ hζ', 129 | sub_sub_cancel_left, (hpri.out.odd_of_ne_two hp).neg_pow] at this 130 | rw [← pow_mul, mul_comm m, pow_mul, hζ.pow_eq_one, one_pow, one_smul, neg_add_cancel, 131 | mul_eq_zero] at this 132 | exact this.resolve_left (pow_ne_zero _ hζ') 133 | 134 | def polyRoot {L : Type*} [Field L] [Algebra K L] (α : L) 135 | (e : α ^ p = algebraMap K L u) (m : ℕ) : 𝓞 L := 136 | ⟨((1 : L) - ζ ^ m • α) / (algebraMap K L (ζ - 1)), isIntegral_trans _ 137 | ⟨poly hp hζ u hcong, monic_poly hp hζ u hcong, aeval_poly hp hζ u hcong α e m⟩⟩ 138 | 139 | theorem roots_poly {L : Type*} [Field L] [Algebra K L] (α : L) 140 | (e : α ^ p = algebraMap K L u) : 141 | roots ((poly hp hζ u hcong).map (algebraMap (𝓞 K) L)) = 142 | (Finset.range p).val.map 143 | (fun i ↦ ((1 : L) - ζ ^ i • α) / (algebraMap K L (ζ - 1))) := by 144 | by_cases hα : α = 0 145 | · rw [hα, zero_pow (NeZero.ne p)] at e 146 | exact (((algebraMap (𝓞 K) L).isUnit_map u.isUnit).ne_zero e.symm).elim 147 | have hζ' : algebraMap K L ζ - 1 ≠ 0 := by 148 | simpa using (algebraMap K L).injective.ne (hζ.sub_one_ne_zero hpri.out.one_lt) 149 | classical 150 | symm; apply Multiset.eq_of_le_of_card_le 151 | · rw [← Finset.image_val_of_injOn, Finset.val_le_iff_val_subset] 152 | · intro x hx 153 | simp only [Finset.image_val, Finset.range_val, Multiset.mem_dedup, Multiset.mem_map, 154 | Multiset.mem_range] at hx 155 | obtain ⟨m, _, rfl⟩ := hx 156 | rw [mem_roots, IsRoot.def, eval_map, ← aeval_def, aeval_poly hp hζ u hcong α e] 157 | exact ((monic_poly hp hζ u hcong).map (algebraMap (𝓞 K) L)).ne_zero 158 | · intros i hi j hj e 159 | apply (hζ.map_of_injective (algebraMap K L).injective).injOn_pow_mul hα hi hj 160 | apply_fun (1 - · * (algebraMap K L ζ - 1)) at e 161 | dsimp only at e 162 | simpa only [Nat.cast_one, map_sub, map_one, Algebra.smul_def, map_pow, 163 | div_mul_cancel₀ _ hζ', sub_sub_cancel] using e 164 | · simp only [Finset.range_val, Multiset.card_map, Multiset.card_range] 165 | refine (Polynomial.card_roots' _).trans ?_ 166 | rw [(monic_poly hp hζ u hcong).natDegree_map, natDegree_poly hp hζ] 167 | 168 | theorem splits_poly {L : Type*} [Field L] [Algebra K L] (α : L) 169 | (e : α ^ p = algebraMap K L u) : 170 | (poly hp hζ u hcong).Splits (algebraMap (𝓞 K) L) := by 171 | rw [← splits_id_iff_splits, splits_iff_card_roots, roots_poly hp hζ u hcong α e, 172 | (monic_poly hp hζ u hcong).natDegree_map, natDegree_poly hp hζ, 173 | Finset.range_val, Multiset.card_map, Multiset.card_range] 174 | 175 | theorem map_poly_eq_prod {L : Type*} [Field L] [Algebra K L] (α : L) 176 | (e : α ^ p = algebraMap K L u) : 177 | (poly hp hζ u hcong).map (algebraMap (𝓞 K) (𝓞 L)) = 178 | ∏ i ∈ Finset.range p, (X - C (polyRoot hp hζ u hcong α e i)) := by 179 | apply map_injective (algebraMap (𝓞 L) L) Subtype.coe_injective 180 | have : (algebraMap (𝓞 L) L).comp (algebraMap (𝓞 K) (𝓞 L)) = algebraMap (𝓞 K) L := by 181 | ext; rfl 182 | rw [← coe_mapRingHom, map_prod, coe_mapRingHom, map_map, this] 183 | rw [eq_prod_roots_of_monic_of_splits_id ((monic_poly hp hζ u hcong).map _) 184 | ((splits_id_iff_splits _).mpr (splits_poly hp hζ u hcong α e)), roots_poly hp hζ u hcong α e, 185 | Multiset.map_map] 186 | simp only [Polynomial.map_sub, map_X, map_C] 187 | rfl 188 | 189 | lemma isIntegralClosure_of_isScalarTower (R A K L B) [CommRing R] [CommRing A] [CommRing K] 190 | [CommRing L] [CommRing B] [Algebra R K] [Algebra A K] [Algebra R L] [Algebra B L] 191 | [Algebra A L] [Algebra R A] [IsScalarTower R A K] [IsScalarTower R A L] 192 | [IsIntegralClosure A R K] [IsIntegralClosure B R L] : 193 | IsIntegralClosure B A L where 194 | algebraMap_injective' := IsIntegralClosure.algebraMap_injective B R L 195 | isIntegral_iff := fun {x} ↦ by 196 | refine Iff.trans ?_ (IsIntegralClosure.isIntegral_iff (R := R) (A := B) (B := L)) 197 | have := (IsIntegralClosure.isIntegral_algebra R (A := A) K) 198 | exact ⟨isIntegral_trans x, IsIntegral.tower_top⟩ 199 | 200 | instance {K L} [Field K] [Field L] [Algebra K L] : 201 | IsIntegralClosure (𝓞 L) (𝓞 K) L := isIntegralClosure_of_isScalarTower ℤ _ K _ _ 202 | 203 | instance {K L} [Field K] [Field L] [Algebra K L] : 204 | IsScalarTower (𝓞 K) (𝓞 L) L := IsScalarTower.of_algebraMap_eq (fun _ ↦ rfl) 205 | 206 | include hu in 207 | lemma minpoly_polyRoot'' {L : Type*} [Field L] [Algebra K L] (α : L) 208 | (e : α ^ p = algebraMap K L u) (i) : 209 | minpoly K (polyRoot hp hζ u hcong α e i : L) = 210 | (poly hp hζ u hcong).map (algebraMap (𝓞 K) K) := by 211 | have : IsIntegral K (polyRoot hp hζ u hcong α e i : L) := 212 | IsIntegral.tower_top (polyRoot hp hζ u hcong α e i).prop 213 | apply eq_of_monic_of_associated (minpoly.monic this) ((monic_poly hp hζ u hcong).map _) 214 | refine Irreducible.associated_of_dvd (minpoly.irreducible this) 215 | (irreducible_map_poly hp hζ u hcong hu) (minpoly.dvd _ _ ?_) 216 | rw [aeval_def, eval₂_map, ← IsScalarTower.algebraMap_eq, ← aeval_def] 217 | exact aeval_poly hp hζ u hcong α e i 218 | 219 | include hu in 220 | lemma minpoly_polyRoot' {L : Type*} [Field L] [Algebra K L] (α : L) 221 | (e : α ^ p = algebraMap K L u) (i) : 222 | minpoly (𝓞 K) (polyRoot hp hζ u hcong α e i : L) = (poly hp hζ u hcong) := by 223 | apply map_injective (algebraMap (𝓞 K) K) Subtype.coe_injective 224 | rw [← minpoly.isIntegrallyClosed_eq_field_fractions' K] 225 | · exact minpoly_polyRoot'' hp hζ u hcong hu α e i 226 | · exact IsIntegral.tower_top (polyRoot hp hζ u hcong α e i).prop 227 | 228 | lemma separable_poly_aux {L : Type*} [Field L] [Algebra K L] (α : L) 229 | (e : α ^ p = algebraMap K L u) : Separable ((poly hp hζ u hcong).map 230 | (algebraMap (𝓞 K) (𝓞 L))) := by 231 | have hζ' : algebraMap K L ζ - 1 ≠ 0 := by 232 | simpa using (algebraMap K L).injective.ne (hζ.sub_one_ne_zero hpri.out.one_lt) 233 | rw [map_poly_eq_prod (e := e)] 234 | refine separable_prod' ?_ (fun _ _ => separable_X_sub_C) 235 | intros i hi j hj hij 236 | apply isCoprime_X_sub_C_of_isUnit_sub 237 | obtain ⟨v, hv⟩ : 238 | Associated (hζ.unit' - 1 : 𝓞 K) ((hζ.unit' : 𝓞 K) ^ j - (hζ.unit' : 𝓞 K) ^ i) := by 239 | refine hζ.unit'_coe.associated_sub_one hpri.out ?_ ?_ ?_ 240 | · rw [mem_nthRootsFinset (NeZero.pos p), ← pow_mul, mul_comm, pow_mul, hζ.unit'_coe.pow_eq_one, 241 | one_pow] 242 | · rw [mem_nthRootsFinset (NeZero.pos p), ← pow_mul, mul_comm, pow_mul, hζ.unit'_coe.pow_eq_one, 243 | one_pow] 244 | · exact mt (hζ.unit'_coe.injOn_pow hj hi) hij.symm 245 | rw [NumberField.RingOfIntegers.ext_iff] at hv 246 | have hcoe : (algebraMap (𝓞 K) K) (↑hζ.unit') = ζ := rfl 247 | simp only [map_mul, map_sub, map_one, map_pow, hcoe] at hv 248 | have hα : IsIntegral (𝓞 K) α := by 249 | apply IsIntegral.of_pow (NeZero.pos p); rw [e]; exact isIntegral_algebraMap 250 | have : IsUnit (⟨α, isIntegral_trans _ hα⟩ : 𝓞 L) := by 251 | rw [← isUnit_pow_iff (NeZero.pos p).ne.symm] 252 | convert (algebraMap (𝓞 K) (𝓞 L)).isUnit_map u.isUnit 253 | ext; simp only [SubmonoidClass.coe_pow, e]; rfl 254 | convert ((algebraMap (𝓞 K) (𝓞 L)).isUnit_map v.isUnit).mul this using 1 255 | ext 256 | simp only [polyRoot, map_sub, map_one, sub_div, one_div, map_sub, 257 | sub_sub_sub_cancel_left, map_mul, NumberField.RingOfIntegers.map_mk] 258 | rw [← sub_div, ← sub_smul, ← hv, Algebra.smul_def, map_mul, map_sub, map_one, mul_assoc, 259 | mul_div_cancel_left₀ _ hζ'] 260 | rfl 261 | 262 | include hu in 263 | open scoped KummerExtension in 264 | attribute [local instance] Ideal.Quotient.field in 265 | lemma separable_poly (I : Ideal (𝓞 K)) [I.IsMaximal] : 266 | Separable ((poly hp hζ u hcong).map (Ideal.Quotient.mk I)) := by 267 | let L := K[p√(u : K)] 268 | have := Fact.mk (X_pow_sub_C_irreducible_of_prime hpri.out hu) 269 | let J := I.map (algebraMap (𝓞 K) (𝓞 L)) 270 | let i : 𝓞 K ⧸ I →+* 𝓞 L ⧸ J := Ideal.quotientMap _ 271 | (algebraMap (𝓞 K) (𝓞 L)) Ideal.le_comap_map 272 | haveI : Nontrivial (𝓞 L ⧸ J) := by 273 | apply Ideal.Quotient.nontrivial 274 | rw [ne_eq, Ideal.map_eq_top_iff]; exact Ideal.IsMaximal.ne_top ‹_› 275 | · intros x y e; ext; exact (algebraMap K L).injective (congr_arg Subtype.val e) 276 | · intros x; exact IsIntegral.tower_top (IsIntegralClosure.isIntegral ℤ L x) 277 | rw [← Polynomial.separable_map i, map_map, Ideal.quotientMap_comp_mk, ← map_map] 278 | apply Separable.map 279 | apply separable_poly_aux hp hζ u hcong 280 | exact root_X_pow_sub_C_pow _ _ 281 | 282 | lemma polyRoot_spec {L : Type*} [Field L] [Algebra K L] (α : L) 283 | (e : α ^ p = algebraMap K L u) (i) : 284 | α = (ζ ^ i)⁻¹ • (1 - (ζ - 1) • (polyRoot hp hζ u hcong α e i : L)) := by 285 | apply smul_right_injective (M := L) (c := ζ ^ i) (pow_ne_zero _ <| hζ.ne_zero 286 | (NeZero.pos p).ne.symm) 287 | simp only [polyRoot, map_sub, map_one, NumberField.RingOfIntegers.map_mk, 288 | Algebra.smul_def (ζ - 1), ← mul_div_assoc, 289 | mul_div_cancel_left₀ _ 290 | ((hζ.map_of_injective (algebraMap K L).injective).sub_one_ne_zero hpri.out.one_lt), 291 | sub_sub_cancel, smul_smul, inv_mul_cancel₀ (pow_ne_zero _ <| hζ.ne_zero (NeZero.pos p).ne.symm), 292 | one_smul] 293 | 294 | lemma mem_adjoin_polyRoot {L : Type*} [Field L] [Algebra K L] (α : L) 295 | (e : α ^ p = algebraMap K L u) (i) : 296 | α ∈ Algebra.adjoin K {(polyRoot hp hζ u hcong α e i : L)} := by 297 | conv => enter [2]; rw [polyRoot_spec hp hζ u hcong α e i] 298 | exact Subalgebra.smul_mem _ (sub_mem (one_mem _) 299 | (Subalgebra.smul_mem _ (Algebra.self_mem_adjoin_singleton K _) _)) _ 300 | 301 | include hu hp hζ hcong in 302 | attribute [local instance] Ideal.Quotient.field in 303 | lemma isUnramified (L) [Field L] [Algebra K L] [IsSplittingField K L (X ^ p - C (u : K))] : 304 | IsUnramified (𝓞 K) (𝓞 L) := by 305 | let α := polyRoot hp hζ u hcong _ (rootOfSplitsXPowSubC_pow _ L) 0 306 | haveI := Polynomial.IsSplittingField.finiteDimensional L (X ^ p - C (u : K)) 307 | have hα : Algebra.adjoin K {(α : L)} = ⊤ := by 308 | rw [eq_top_iff, ← Algebra.adjoin_root_eq_top_of_isSplittingField 309 | ⟨ζ, (mem_primitiveRoots (NeZero.pos p)).mpr hζ⟩ (X_pow_sub_C_irreducible_of_prime hpri.out hu) 310 | (rootOfSplitsXPowSubC_pow (u : K) L), Algebra.adjoin_le_iff, Set.singleton_subset_iff] 311 | exact mem_adjoin_polyRoot hp hζ u hcong _ _ 0 312 | constructor 313 | intros I hI hIbot 314 | refine isUnramifiedAt_of_Separable_minpoly (R := 𝓞 K) K (S := 𝓞 L) L I hIbot α ?_ hα ?_ 315 | · exact IsIntegral.tower_top α.prop 316 | · rw [minpoly_polyRoot' hp hζ u hcong hu] 317 | haveI := hI.isMaximal hIbot 318 | exact separable_poly hp hζ u hcong hu I 319 | 320 | end KummersLemma 321 | -------------------------------------------------------------------------------- /FltRegular/NumberTheory/KummersLemma/KummersLemma.lean: -------------------------------------------------------------------------------- 1 | import FltRegular.NumberTheory.KummersLemma.Field 2 | import FltRegular.NumberTheory.Hilbert94 3 | 4 | open Polynomial 5 | open scoped NumberField 6 | 7 | variable {K : Type*} {p : ℕ} [hpri : Fact p.Prime] [Field K] [NumberField K] 8 | [IsCyclotomicExtension {p} ℚ K] (hp : p ≠ 2) [Fintype (ClassGroup (𝓞 K))] 9 | (hreg : p.Coprime <| Fintype.card <| ClassGroup (𝓞 K)) 10 | {ζ : K} (hζ : IsPrimitiveRoot ζ p) 11 | 12 | include hp hreg 13 | 14 | theorem exists_pow_eq_of_zeta_sub_one_pow_dvd_sub_one {u : (𝓞 K)ˣ} 15 | (hcong : (hζ.unit' - 1 : 𝓞 K) ^ p ∣ (u : 𝓞 K) - 1) : ∃ v : K, v ^ p = u := by 16 | by_contra! hu 17 | have hirr := X_pow_sub_C_irreducible_of_prime hpri.out hu 18 | have := Fact.mk hirr 19 | let L := AdjoinRoot (X ^ p - C (u : K)) 20 | have := isSplittingField_AdjoinRoot_X_pow_sub_C ⟨ζ, (mem_primitiveRoots (NeZero.pos p)).mpr hζ⟩ 21 | hirr 22 | have := isGalois_of_isSplittingField_X_pow_sub_C ⟨ζ, (mem_primitiveRoots (NeZero.pos p)).mpr hζ⟩ 23 | hirr L 24 | have := IsSplittingField.finiteDimensional L (X ^ p - C (u : K)) 25 | have := isCyclic_of_isSplittingField_X_pow_sub_C ⟨ζ, (mem_primitiveRoots (NeZero.pos p)).mpr hζ⟩ 26 | hirr L 27 | have : CharZero L := charZero_of_injective_algebraMap (algebraMap K L).injective 28 | have : FiniteDimensional ℚ L := Module.Finite.trans K L 29 | have : NumberField L := ⟨⟩ 30 | have hKL : Module.finrank K L = p := 31 | finrank_of_isSplittingField_X_pow_sub_C ⟨ζ, (mem_primitiveRoots (NeZero.pos p)).mpr hζ⟩ hirr L 32 | have := KummersLemma.isUnramified hp hζ u hcong hu L 33 | have := dvd_card_classGroup_of_isUnramified_isCyclic K L (hKL.symm ▸ hpri.out) (hKL.symm ▸ hp) 34 | rw [hKL, hpri.out.dvd_iff_not_coprime] at this 35 | exact this (by convert hreg) 36 | 37 | -- Let 𝑝 be a regular prime (i.e. an odd prime which does not divide the class number off 38 | -- the 𝑝-th cyclotomic field) and 𝜉 a primitive 𝑝-th root of unity; 39 | -- if a unit 𝑢∈𝐐(𝜉) is congruent to an integer modulo 𝑝, then 𝑢 is a 𝑝-th power in 𝐐(𝜉). 40 | theorem eq_pow_prime_of_unit_of_congruent (u : (𝓞 K)ˣ) 41 | (hcong : ∃ n : ℤ, (p : 𝓞 K) ∣ (u - n : 𝓞 K)) : 42 | ∃ v, u = v ^ p := by 43 | obtain ⟨ζ, hζ⟩ := IsCyclotomicExtension.exists_isPrimitiveRoot ℚ (B := K) (Set.mem_singleton p) 44 | (NeZero.ne p) 45 | obtain ⟨x, hx⟩ : (p : 𝓞 K) ∣ (↑(u ^ (p - 1)) : 𝓞 K) - 1 := by 46 | obtain ⟨n, hn⟩ := hcong 47 | have hn' : (p : ℤ) ∣ n ^ (p - 1) - 1 := by 48 | refine Int.modEq_iff_dvd.mp (Int.ModEq.pow_card_sub_one_eq_one hpri.out ?_).symm 49 | rw [isCoprime_comm, (Nat.prime_iff_prime_int.mp hpri.out).coprime_iff_not_dvd] 50 | intro h 51 | replace h := Int.cast_dvd_cast (α := 𝓞 K) _ _ h 52 | simp only [Int.cast_natCast, ← dvd_iff_dvd_of_dvd_sub hn] at h 53 | refine hζ.zeta_sub_one_prime'.not_unit ((isUnit_pow_iff ?_).mp 54 | (isUnit_of_dvd_unit ((associated_zeta_sub_one_pow_prime hζ).dvd.trans h) u.isUnit)) 55 | simpa only [ne_eq, tsub_eq_zero_iff_le, not_le] using hpri.out.one_lt 56 | replace hn' := Int.cast_dvd_cast (α := 𝓞 K) _ _ hn' 57 | simp only [Int.cast_natCast, Int.cast_sub, Int.cast_pow, Int.cast_one] at hn' 58 | rw [← Ideal.mem_span_singleton, ← Ideal.Quotient.eq_zero_iff_mem, 59 | RingHom.map_sub, sub_eq_zero] at hn hn' ⊢ 60 | rw [Units.val_pow_eq_pow_val, RingHom.map_pow, hn, ← RingHom.map_pow, hn'] 61 | have : (hζ.unit' - 1 : 𝓞 K) ^ p ∣ (↑(u ^ (p - 1)) : 𝓞 K) - 1 := by 62 | rw [hx] 63 | rw [sub_eq_iff_eq_add, add_comm] at hx 64 | have H : Algebra.norm ℤ (1 + p • x) = 1 := norm_add_one_smul_of_isUnit hpri.out 65 | hp x (by rw [nsmul_eq_mul, ← hx]; exact Units.isUnit _) 66 | have := H ▸ zeta_sub_one_pow_dvd_norm_sub_pow hζ x 67 | simpa [ge_iff_le, Int.cast_one, sub_self, nsmul_eq_mul, Nat.cast_mul, PNat.pos, 68 | Nat.cast_pred, zero_sub, IsUnit.mul_iff, ne_eq, tsub_eq_zero_iff_le, not_le, dvd_neg, 69 | Units.isUnit, and_true, zero_add] using this 70 | obtain ⟨v, hv⟩ := exists_pow_eq_of_zeta_sub_one_pow_dvd_sub_one hp hreg hζ this 71 | have hv' : IsIntegral ℤ v := by 72 | apply IsIntegral.of_pow (NeZero.pos p); rw [hv] 73 | exact NumberField.RingOfIntegers.isIntegral_coe _ 74 | set w : 𝓞 K := ⟨v, hv'⟩ 75 | have : IsUnit w := by 76 | rw [← isUnit_pow_iff (NeZero.pos p).ne.symm]; convert (u ^ (p - 1) : (𝓞 K)ˣ).isUnit; ext 77 | exact hv 78 | have hv'' : this.unit ^ p = u ^ (p - 1) := by 79 | ext; simpa only [Units.val_pow_eq_pow_val, IsUnit.unit_spec, SubmonoidClass.coe_pow] using hv 80 | use u / this.unit 81 | rw [div_pow, hv'', div_eq_mul_inv, ← pow_sub _ tsub_le_self, 82 | tsub_tsub_cancel_of_le (Nat.Prime.one_lt hpri.out).le, pow_one] 83 | -------------------------------------------------------------------------------- /FltRegular/NumberTheory/RegularPrimes.lean: -------------------------------------------------------------------------------- 1 | import Mathlib.FieldTheory.SplittingField.Construction 2 | import Mathlib.NumberTheory.ClassNumber.Finite 3 | import Mathlib.NumberTheory.ClassNumber.AdmissibleAbs 4 | import FltRegular.NumberTheory.Cyclotomic.CyclRat 5 | import Mathlib.NumberTheory.Cyclotomic.PID 6 | 7 | /-! 8 | # Regular primes 9 | 10 | ## Main definitions 11 | 12 | * `IsRegularNumber`: a natural number `n` is regular if `n` is coprime with the cardinal of the 13 | class group. 14 | 15 | -/ 16 | 17 | noncomputable section 18 | 19 | open Nat Polynomial NumberField 20 | 21 | open scoped NumberField 22 | 23 | variable (n p : ℕ) [hp : Fact p.Prime] 24 | 25 | /-- A natural number `n` is regular if `n` is coprime with the cardinal of the class group -/ 26 | def IsRegularNumber [NeZero n] : Prop := 27 | n.Coprime <| Fintype.card <| ClassGroup (𝓞 <| CyclotomicField n ℚ) 28 | 29 | /-- The definition of regular primes. -/ 30 | def IsRegularPrime : Prop := 31 | IsRegularNumber p 32 | 33 | section TwoRegular 34 | 35 | variable (A K : Type*) [CommRing A] [IsDomain A] [Field K] [Algebra A K] [IsFractionRing A K] 36 | 37 | variable (L : Type*) [Field L] [Algebra K L] 38 | 39 | /-- The second cyclotomic field is equivalent to the base field. -/ 40 | def cyclotomicFieldTwoEquiv [IsCyclotomicExtension {2} K L] : L ≃ₐ[K] K := by 41 | suffices IsSplittingField K K (cyclotomic 2 K) by 42 | have : IsSplittingField K L (cyclotomic 2 K) := 43 | IsCyclotomicExtension.splitting_field_cyclotomic 2 K L 44 | exact 45 | (IsSplittingField.algEquiv L (cyclotomic 2 K)).trans 46 | (IsSplittingField.algEquiv K <| cyclotomic 2 K).symm 47 | exact ⟨by simpa using @splits_X_sub_C _ _ _ _ (RingHom.id K) (-1), 48 | by simp [eq_iff_true_of_subsingleton]⟩ 49 | 50 | instance IsPrincipalIdealRing_of_IsCyclotomicExtension_two 51 | (L : Type*) [Field L] [CharZero L] [IsCyclotomicExtension {2} ℚ L] : 52 | IsPrincipalIdealRing (𝓞 L) := by 53 | haveI : IsIntegralClosure ℤ ℤ L := 54 | { algebraMap_injective' := (algebraMap ℤ L).injective_int 55 | isIntegral_iff := fun {x} => by 56 | let f := cyclotomicFieldTwoEquiv ℚ L 57 | refine 58 | ⟨fun hx => ⟨IsIntegralClosure.mk' ℤ (f x) (map_isIntegral_int f hx), f.injective ?_⟩, ?_⟩ 59 | · convert IsIntegralClosure.algebraMap_mk' ℤ (f x) (map_isIntegral_int f hx) 60 | simp 61 | · rintro ⟨y, hy⟩ 62 | simpa [← hy] using isIntegral_algebraMap } 63 | let F : 𝓞 L ≃+* ℤ := NumberField.RingOfIntegers.equiv _ 64 | exact IsPrincipalIdealRing.of_surjective F.symm.toRingHom F.symm.surjective 65 | 66 | instance : IsCyclotomicExtension {2} ℚ (CyclotomicField 2 ℚ) := 67 | CyclotomicField.isCyclotomicExtension 2 ℚ 68 | 69 | instance : IsPrincipalIdealRing (𝓞 (CyclotomicField 2 ℚ)) := 70 | IsPrincipalIdealRing_of_IsCyclotomicExtension_two _ 71 | 72 | theorem isRegularPrime_two : IsRegularPrime 2 := by 73 | rw [IsRegularPrime, IsRegularNumber] 74 | convert coprime_one_right _ 75 | dsimp 76 | apply (card_classGroup_eq_one_iff (R := 𝓞 (CyclotomicField 2 ℚ))).2 77 | infer_instance 78 | 79 | theorem isRegularPrime_three : 80 | haveI : Fact (Nat.Prime 3) := ⟨Nat.prime_three⟩ 81 | IsRegularPrime 3 := by 82 | rw [IsRegularPrime, IsRegularNumber] 83 | convert coprime_one_right _ 84 | exact classNumber_eq_one_iff.2 85 | (IsCyclotomicExtension.Rat.three_pid (CyclotomicField _ ℚ)) 86 | 87 | end TwoRegular 88 | 89 | theorem isPrincipal_of_isPrincipal_pow_of_coprime 90 | {A : Type*} [CommRing A] [IsDedekindDomain A] [Fintype (ClassGroup A)] 91 | {p : ℕ} [Fact p.Prime] 92 | (H : p.Coprime <| Fintype.card <| ClassGroup A) {I : Ideal A} 93 | (hI : (I ^ p).IsPrincipal) : I.IsPrincipal := by 94 | by_cases Izero : I = 0 95 | · rw [Izero] 96 | exact bot_isPrincipal 97 | rw [← ClassGroup.mk0_eq_one_iff (mem_nonZeroDivisors_of_ne_zero _)] at hI ⊢ 98 | swap 99 | · exact Izero 100 | swap 101 | · exact pow_ne_zero p Izero 102 | · rw [← orderOf_eq_one_iff, ← Nat.dvd_one, ← H, Nat.dvd_gcd_iff] 103 | refine ⟨?_, orderOf_dvd_card⟩ 104 | rwa [orderOf_dvd_iff_pow_eq_one, ← map_pow, SubmonoidClass.mk_pow] 105 | -------------------------------------------------------------------------------- /FltRegular/NumberTheory/SystemOfUnits.lean: -------------------------------------------------------------------------------- 1 | import FltRegular.NumberTheory.Cyclotomic.UnitLemmas 2 | import FltRegular.NumberTheory.CyclotomicRing 3 | import Mathlib.LinearAlgebra.FreeModule.StrongRankCondition 4 | import Mathlib.LinearAlgebra.Dimension.Localization 5 | 6 | open FiniteDimensional 7 | open NumberField 8 | 9 | variable (p : ℕ) {K : Type*} [Field K] [NumberField K] [IsCyclotomicExtension {p} ℚ K] 10 | variable {k : Type*} [Field k] [NumberField k] (hp : Nat.Prime p) 11 | 12 | open Module BigOperators Finset 13 | open CyclotomicIntegers(zeta) 14 | 15 | variable 16 | (G : Type*) {H : Type*} [AddCommGroup G] (s : ℕ) 17 | (hf : finrank ℤ G = s * (p - 1)) 18 | 19 | local notation "A" => (CyclotomicIntegers p) 20 | 21 | section 22 | 23 | variable [Module (CyclotomicIntegers p) G] 24 | 25 | structure systemOfUnits (s : ℕ) 26 | where 27 | units : Fin s → G 28 | linearIndependent : LinearIndependent A units 29 | 30 | namespace systemOfUnits 31 | 32 | lemma existence0 : Nonempty (systemOfUnits p G 0) := by 33 | exact ⟨⟨fun _ => 0, linearIndependent_empty_type⟩⟩ 34 | 35 | theorem _root_.PowerBasis.finrank' {R S} [CommRing R] [Nontrivial R] [CommRing S] [Algebra R S] 36 | (pb : PowerBasis R S) : finrank R S = pb.dim := by 37 | rw [finrank_eq_card_basis pb.basis, Fintype.card_fin] 38 | 39 | open Cardinal Submodule in 40 | theorem _root_.finrank_span_set_eq_card' {R M} [CommRing R] [AddCommGroup M] [Module R M] 41 | [Nontrivial R] (s : Set M) [Fintype s] (hs : LinearIndependent R ((↑) : s → M)) : 42 | finrank R (span R s) = s.toFinset.card := 43 | finrank_eq_of_rank_eq 44 | (by 45 | have : Module.rank R (span R s) = #s := rank_span_set hs 46 | rwa [Cardinal.mk_fintype, ← Set.toFinset_card] at this) 47 | 48 | include hp 49 | 50 | lemma finrank_spanA {R : ℕ} (f : Fin R → G) (hf : LinearIndependent A f) : 51 | finrank ℤ (Submodule.span A (Set.range f)) = (p - 1) * R := by 52 | classical 53 | have := Fact.mk hp 54 | have := finrank_span_set_eq_card' (R := A) (Set.range f) 55 | ((linearIndepOn_id_range_iff hf.injective).mpr hf) 56 | simp only [Set.toFinset_range, Finset.card_image_of_injective _ hf.injective, card_fin] at this 57 | rw [← CyclotomicIntegers.powerBasis_dim, ← PowerBasis.finrank'] 58 | conv_rhs => rw [← this] 59 | have := Module.Free.of_basis (Basis.span hf) 60 | have := Module.Finite.of_basis (Basis.span hf) 61 | rw [finrank_mul_finrank] 62 | 63 | include hf 64 | 65 | lemma ex_not_mem [Module.Free ℤ G] {R : ℕ} (S : systemOfUnits p G R) (hR : R < s) : 66 | ∃ g, ∀ (k : ℤ), k ≠ 0 → ¬(k • g ∈ Submodule.span A (Set.range S.units)) := by 67 | have := Fact.mk hp 68 | have : Module.Finite ℤ G := Module.finite_of_finrank_pos 69 | (by simp [hf, R.zero_le.trans_lt hR, hp.one_lt]) 70 | refine Submodule.exists_of_finrank_lt 71 | ((Submodule.span A (Set.range S.units)).restrictScalars ℤ) ?_ 72 | show finrank ℤ (Submodule.span A _) < _ 73 | rw [finrank_spanA p hp G S.units S.linearIndependent, hf, mul_comm] 74 | exact Nat.mul_lt_mul_of_lt_of_le hR rfl.le hp.pred_pos 75 | 76 | end systemOfUnits 77 | 78 | end 79 | 80 | namespace systemOfUnits 81 | 82 | include hp hf 83 | 84 | variable [Module.Free ℤ G] 85 | 86 | lemma existence' [Module A G] {R : ℕ} (S : systemOfUnits p G R) (hR : R < s) : 87 | Nonempty (systemOfUnits p G (R + 1)) := by 88 | obtain ⟨g, hg⟩ := ex_not_mem p hp G s hf S hR 89 | refine ⟨⟨Fin.cases g S.units, ?_⟩⟩ 90 | refine LinearIndependent.fin_cons' g S.units S.linearIndependent (fun a y hy ↦ ?_) 91 | by_contra! ha 92 | have := Fact.mk hp 93 | obtain ⟨n, h0, f, Hf⟩ := CyclotomicIntegers.exists_dvd_int p _ ha 94 | replace hy := congr_arg (f • ·) hy 95 | simp only at hy 96 | rw [smul_zero, smul_add, smul_smul, mul_comm f, 97 | ← Hf, ← eq_neg_iff_add_eq_zero, Int.cast_smul_eq_zsmul] at hy 98 | apply hg _ h0 99 | rw [hy] 100 | exact Submodule.neg_mem _ (Submodule.smul_mem _ _ y.2) 101 | 102 | lemma existence'' [Module A G] {R : ℕ} (hR : R ≤ s) : Nonempty (systemOfUnits p G R) := by 103 | induction R with 104 | | zero => exact existence0 p G 105 | | succ n ih => 106 | obtain ⟨S⟩ := ih (le_trans (Nat.le_succ n) hR) 107 | exact existence' p hp G s hf S (lt_of_lt_of_le (Nat.lt.base n) hR) 108 | 109 | lemma existence [Module A G] : Nonempty (systemOfUnits p G s) := existence'' p hp G s hf rfl.le 110 | 111 | end systemOfUnits 112 | -------------------------------------------------------------------------------- /FltRegular/NumberTheory/Unramified.lean: -------------------------------------------------------------------------------- 1 | import FltRegular.NumberTheory.GaloisPrime 2 | import Mathlib.NumberTheory.KummerDedekind 3 | import Mathlib.RingTheory.DedekindDomain.Different 4 | import Mathlib.Order.CompletePartialOrder 5 | import Mathlib.NumberTheory.RamificationInertia.Basic 6 | 7 | /-! 8 | # Unramified extensions 9 | 10 | ## Main definition 11 | - `IsUnramifiedAt`: 12 | `IsUnramifiedAt S p` states that every prime in `S` over `p` has ramification index `1`. 13 | - `IsUnramified`: 14 | An extension is unramified if it is unramified at every (finite) primes. 15 | 16 | ## Main results 17 | - `comap_map_eq_of_isUnramified`: If `K/L` is galois, `S/R` is unramified, then any 18 | ideal `I` fixed by `Gal(L/K) satisfies `(I ∩ R)S = I`. 19 | - `isUnramifiedAt_of_Separable_minpoly`: If `L = K[α]` with `α` integral over `R`, and `f'(α) mod p` 20 | is separable for some ideal `p` of `R` (with `f` being the minpoly of `α` over `R`), then `S/R` is 21 | unramified at `p`. 22 | -/ 23 | open BigOperators UniqueFactorizationMonoid Ideal 24 | 25 | attribute [local instance] FractionRing.liftAlgebra 26 | 27 | variable (R K L S : Type*) [CommRing R] [CommRing S] [Algebra R S] [Field K] [Field L] 28 | [IsDedekindDomain R] [Algebra R K] [IsFractionRing R K] [Algebra S L] 29 | [Algebra K L] [Algebra R L] [IsScalarTower R S L] [IsScalarTower R K L] 30 | [IsIntegralClosure S R L] [FiniteDimensional K L] 31 | 32 | def IsUnramifiedAt {R} (S : Type*) [CommRing R] [CommRing S] [Algebra R S] (p : Ideal R) : Prop := 33 | ∀ P ∈ primesOver p S, Ideal.ramificationIdx (algebraMap R S) p P = 1 34 | 35 | /-- TODO: Link this to `FormallyUnramified`. -/ 36 | -- Should we name this `IsUnramifiedAtFinitePrimes`? 37 | class IsUnramified (R S : Type*) [CommRing R] [CommRing S] [Algebra R S] : Prop where 38 | isUnramifiedAt : ∀ (p : Ideal R) [p.IsPrime] (_ : p ≠ ⊥), IsUnramifiedAt S p 39 | 40 | variable {R} {S} 41 | 42 | lemma prod_primesOverFinset_of_isUnramified [IsUnramified R S] [IsDedekindDomain S] 43 | [NoZeroSMulDivisors R S] (p : Ideal R) [p.IsPrime] (hp : p ≠ ⊥) : 44 | ∏ P ∈ primesOverFinset p S, P = p.map (algebraMap R S) := by 45 | classical 46 | have hpbot' : p.map (algebraMap R S) ≠ ⊥ := (Ideal.map_eq_bot_iff_of_injective 47 | (NoZeroSMulDivisors.iff_algebraMap_injective.mp ‹_›)).not.mpr hp 48 | rw [← associated_iff_eq.mp (factors_pow_count_prod hpbot')] 49 | apply Finset.prod_congr rfl 50 | intros P hP 51 | convert (pow_one _).symm 52 | have : p.IsMaximal := Ring.DimensionLEOne.maximalOfPrime hp ‹_› 53 | rw [← Finset.mem_coe, coe_primesOverFinset hp] at hP 54 | rw [← Ideal.IsDedekindDomain.ramificationIdx_eq_factors_count hpbot' hP.1 55 | (ne_bot_of_mem_primesOver hp hP)] 56 | exact IsUnramified.isUnramifiedAt _ hp _ hP 57 | 58 | lemma comap_map_eq_of_isUnramified [IsGalois K L] [IsUnramified R S] (I : Ideal S) 59 | (hI : ∀ σ : L ≃ₐ[K] L, I.comap (galRestrict R K L S σ) = I) : 60 | (I.comap (algebraMap R S)).map (algebraMap R S) = I := by 61 | classical 62 | have : IsDomain S := 63 | (IsIntegralClosure.equiv R S L (integralClosure R L)).toMulEquiv.isDomain (integralClosure R L) 64 | have := IsIntegralClosure.isDedekindDomain R K L S 65 | have hRS : Function.Injective (algebraMap R S) := by 66 | refine Function.Injective.of_comp (f := algebraMap S L) ?_ 67 | rw [← RingHom.coe_comp, ← IsScalarTower.algebraMap_eq, IsScalarTower.algebraMap_eq R K L] 68 | exact (algebraMap K L).injective.comp (IsFractionRing.injective _ _) 69 | have := NoZeroSMulDivisors.iff_algebraMap_injective.mpr hRS 70 | by_cases hIbot : I = ⊥ 71 | · rw [hIbot, Ideal.comap_bot_of_injective _ hRS, Ideal.map_bot] 72 | have h1 : Algebra.IsIntegral R S := IsIntegralClosure.isIntegral_algebra R L 73 | have hIbot' : I.comap (algebraMap R S) ≠ ⊥ := mt Ideal.eq_bot_of_comap_eq_bot hIbot 74 | have : ∀ p, (p.IsPrime ∧ I.comap (algebraMap R S) ≤ p) → ∃ P ≥ I, P ∈ primesOver p S := by 75 | intro p ⟨hp₁, hp₂⟩ 76 | obtain ⟨P, hP1, hP2, hP3⟩ := Ideal.exists_ideal_over_prime_of_isIntegral _ _ hp₂ 77 | exact ⟨P, hP1, hP2, ⟨hP3.symm⟩⟩ 78 | choose 𝔓 h𝔓 h𝔓' using this 79 | suffices I = ∏ p ∈ (factors (I.comap <| algebraMap R S)).toFinset, 80 | (p.map (algebraMap R S)) ^ (if h : _ then (factors I).count (𝔓 p h) else 0) by 81 | simp_rw [← Ideal.mapHom_apply, ← map_pow, ← map_prod, Ideal.mapHom_apply] at this 82 | rw [this, Ideal.map_comap_map] 83 | conv_lhs => rw [← associated_iff_eq.mp (factors_pow_count_prod hIbot)] 84 | rw [← Finset.prod_fiberwise_of_maps_to (g := (Ideal.comap (algebraMap R S) : Ideal S → Ideal R)) 85 | (t := (factors (I.comap (algebraMap R S))).toFinset)] 86 | · apply Finset.prod_congr rfl 87 | intros p hp 88 | simp only [factors_eq_normalizedFactors, Multiset.mem_toFinset, 89 | Ideal.mem_normalizedFactors_iff hIbot'] at hp 90 | have hpbot : p ≠ ⊥ := fun hp' ↦ hIbot' (eq_bot_iff.mpr (hp.2.trans_eq hp')) 91 | have hpbot' : p.map (algebraMap R S) ≠ ⊥ := (Ideal.map_eq_bot_iff_of_injective hRS).not.mpr 92 | hpbot 93 | have := hp.1 94 | rw [← prod_primesOverFinset_of_isUnramified p hpbot, ← Finset.prod_pow] 95 | have : p.IsMaximal := Ring.DimensionLEOne.maximalOfPrime hpbot this 96 | apply Finset.prod_congr 97 | · ext P 98 | rw [factors_eq_normalizedFactors, Finset.mem_filter, Multiset.mem_toFinset, 99 | Ideal.mem_normalizedFactors_iff hIbot, ← Finset.mem_coe, coe_primesOverFinset hpbot S] 100 | refine ⟨fun H ↦ ⟨H.1.1, ⟨H.2.symm⟩⟩, fun H ↦ ⟨⟨H.1, ?_⟩, ?_⟩⟩ 101 | · have ⟨σ, hσ⟩ := exists_comap_galRestrict_eq R K L S (h𝔓' _ hp) H 102 | rw [← hσ, ← hI σ] 103 | exact Ideal.comap_mono (h𝔓 _ hp) 104 | · have := H.2.1 105 | rw [Ideal.under_def] at this 106 | exact this.symm 107 | · intro P hP 108 | rw [← Finset.mem_coe, coe_primesOverFinset hpbot S] at hP 109 | congr 110 | rw [dif_pos hp, ← Nat.cast_inj (R := ENat), ← normalize_eq P, factors_eq_normalizedFactors, 111 | ← emultiplicity_eq_count_normalizedFactors 112 | (prime_of_mem_primesOver hpbot hP).irreducible hIbot, 113 | ← normalize_eq (𝔓 p hp), ← emultiplicity_eq_count_normalizedFactors 114 | (prime_of_mem_primesOver hpbot <| h𝔓' p hp).irreducible hIbot, 115 | emultiplicity_eq_emultiplicity_iff] 116 | intro n 117 | have ⟨σ, hσ⟩ := exists_comap_galRestrict_eq R K L S (h𝔓' _ hp) hP 118 | rw [Ideal.dvd_iff_le, Ideal.dvd_iff_le] 119 | conv_lhs => rw [← hI σ, ← hσ, 120 | Ideal.comap_le_iff_le_map _ (AlgEquiv.bijective _), Ideal.map_pow, 121 | Ideal.map_comap_of_surjective _ (AlgEquiv.surjective _)] 122 | · intro P hP 123 | simp only [factors_eq_normalizedFactors, Multiset.mem_toFinset, 124 | Ideal.mem_normalizedFactors_iff hIbot] at hP 125 | simp only [factors_eq_normalizedFactors, Multiset.mem_toFinset, 126 | Ideal.mem_normalizedFactors_iff hIbot'] 127 | exact ⟨hP.1.comap _, Ideal.comap_mono hP.2⟩ 128 | 129 | section KummerDedekind 130 | 131 | end KummerDedekind 132 | 133 | open nonZeroDivisors Polynomial 134 | 135 | attribute [local instance] Ideal.Quotient.field in 136 | lemma isUnramifiedAt_of_Separable_minpoly' [Algebra.IsSeparable K L] 137 | (p : Ideal R) [hp : p.IsPrime] (hpbot : p ≠ ⊥) (x : S) 138 | (hx' : Algebra.adjoin K {algebraMap S L x} = ⊤) 139 | (h : Polynomial.Separable ((minpoly R x).map (Ideal.Quotient.mk p))) : 140 | IsUnramifiedAt S p := by 141 | classical 142 | have : IsDomain S := 143 | (IsIntegralClosure.equiv R S L (integralClosure R L)).toMulEquiv.isDomain (integralClosure R L) 144 | have hRS : Function.Injective (algebraMap R S) := by 145 | refine Function.Injective.of_comp (f := algebraMap S L) ?_ 146 | rw [← RingHom.coe_comp, ← IsScalarTower.algebraMap_eq, IsScalarTower.algebraMap_eq R K L] 147 | exact (algebraMap K L).injective.comp (IsFractionRing.injective _ _) 148 | have := NoZeroSMulDivisors.iff_algebraMap_injective.mpr hRS 149 | have := IsIntegralClosure.isNoetherian R K L S 150 | have := IsIntegralClosure.isDedekindDomain R K L S 151 | have := IsIntegralClosure.isFractionRing_of_finite_extension R K L S 152 | have := aeval_derivative_mem_differentIdeal R K L x hx' 153 | have H : RingHom.comp (algebraMap (FractionRing R) (FractionRing S)) 154 | (FractionRing.algEquiv R K).symm.toRingEquiv = 155 | RingHom.comp (FractionRing.algEquiv S L).symm.toRingEquiv (algebraMap K L) := by 156 | apply IsLocalization.ringHom_ext R⁰ 157 | ext 158 | simp only [AlgEquiv.toRingEquiv_eq_coe, RingHom.coe_comp, RingHom.coe_coe, 159 | AlgEquiv.coe_ringEquiv, Function.comp_apply, AlgEquiv.commutes, 160 | ← IsScalarTower.algebraMap_apply] 161 | rw [IsScalarTower.algebraMap_apply R S L, AlgEquiv.commutes, ← IsScalarTower.algebraMap_apply] 162 | have : Algebra.IsSeparable (FractionRing R) (FractionRing S) := 163 | Algebra.IsSeparable.of_equiv_equiv _ _ H 164 | have := hp.isMaximal hpbot 165 | intro P hP 166 | letI : IsScalarTower S (S ⧸ P) (S ⧸ P) := IsScalarTower.right 167 | have := isMaximal_of_mem_primesOver hpbot hP 168 | apply le_antisymm 169 | · rw [← tsub_eq_zero_iff_le] 170 | by_contra H 171 | have hxP : aeval x (derivative (minpoly R x)) ∈ P := by 172 | have : differentIdeal R S ≤ P := by 173 | rw [← Ideal.dvd_iff_le] 174 | exact (dvd_pow_self _ H).trans (pow_sub_one_dvd_differentIdeal R P _ hpbot <| 175 | Ideal.dvd_iff_le.mpr <| Ideal.le_pow_ramificationIdx) 176 | exact this (aeval_derivative_mem_differentIdeal R K L _ hx') 177 | rw [← Ideal.Quotient.eq_zero_iff_mem, ← Ideal.Quotient.algebraMap_eq] at hxP 178 | have := hP.2.1 179 | rw [Ideal.under_def] at this 180 | have := (separable_map (Ideal.quotientMap P (algebraMap R S) this.symm.ge)).mpr h 181 | rw [Polynomial.map_map, Ideal.quotientMap_comp_mk] at this 182 | obtain ⟨a, b, e⟩ := this 183 | apply_fun (aeval (Ideal.Quotient.mk P x)) at e 184 | simp_rw [← Ideal.Quotient.algebraMap_eq, ← Polynomial.map_map, derivative_map, map_add, 185 | _root_.map_mul, aeval_map_algebraMap, aeval_algebraMap_apply, minpoly.aeval, hxP, map_zero, 186 | mul_zero, zero_add, map_one, zero_ne_one] at e 187 | · rwa [Ideal.IsDedekindDomain.ramificationIdx_eq_factors_count _ 188 | (isMaximal_of_mem_primesOver hpbot hP).isPrime (ne_bot_of_mem_primesOver hpbot hP), 189 | Multiset.one_le_count_iff_mem, ← Multiset.mem_toFinset, ← primesOverFinset, 190 | ← Finset.mem_coe, coe_primesOverFinset hpbot] 191 | rwa [ne_eq, Ideal.map_eq_bot_iff_of_injective hRS] 192 | 193 | lemma isUnramifiedAt_of_Separable_minpoly [Algebra.IsSeparable K L] 194 | (p : Ideal R) [hp : p.IsPrime] (hpbot : p ≠ ⊥) (x : L) (hx : IsIntegral R x) 195 | (hx' : Algebra.adjoin K {x} = ⊤) 196 | (h : Polynomial.Separable ((minpoly R x).map (Ideal.Quotient.mk p))) : 197 | IsUnramifiedAt S p := by 198 | rw [← IsIntegralClosure.algebraMap_mk' S x hx, minpoly.algebraMap_eq 199 | (IsIntegralClosure.algebraMap_injective S R L)] at h 200 | exact isUnramifiedAt_of_Separable_minpoly' K L p hpbot (IsIntegralClosure.mk' S x hx) 201 | (by rwa [IsIntegralClosure.algebraMap_mk']) h 202 | -------------------------------------------------------------------------------- /LICENSE: -------------------------------------------------------------------------------- 1 | Apache License 2 | Version 2.0, January 2004 3 | http://www.apache.org/licenses/ 4 | 5 | TERMS AND CONDITIONS FOR USE, REPRODUCTION, AND DISTRIBUTION 6 | 7 | 1. Definitions. 8 | 9 | "License" shall mean the terms and conditions for use, reproduction, 10 | and distribution as defined by Sections 1 through 9 of this document. 11 | 12 | "Licensor" shall mean the copyright owner or entity authorized by 13 | the copyright owner that is granting the License. 14 | 15 | "Legal Entity" shall mean the union of the acting entity and all 16 | other entities that control, are controlled by, or are under common 17 | control with that entity. For the purposes of this definition, 18 | "control" means (i) the power, direct or indirect, to cause the 19 | direction or management of such entity, whether by contract or 20 | otherwise, or (ii) ownership of fifty percent (50%) or more of the 21 | outstanding shares, or (iii) beneficial ownership of such entity. 22 | 23 | "You" (or "Your") shall mean an individual or Legal Entity 24 | exercising permissions granted by this License. 25 | 26 | "Source" form shall mean the preferred form for making modifications, 27 | including but not limited to software source code, documentation 28 | source, and configuration files. 29 | 30 | "Object" form shall mean any form resulting from mechanical 31 | transformation or translation of a Source form, including but 32 | not limited to compiled object code, generated documentation, 33 | and conversions to other media types. 34 | 35 | "Work" shall mean the work of authorship, whether in Source or 36 | Object form, made available under the License, as indicated by a 37 | copyright notice that is included in or attached to the work 38 | (an example is provided in the Appendix below). 39 | 40 | "Derivative Works" shall mean any work, whether in Source or Object 41 | form, that is based on (or derived from) the Work and for which the 42 | editorial revisions, annotations, elaborations, or other modifications 43 | represent, as a whole, an original work of authorship. For the purposes 44 | of this License, Derivative Works shall not include works that remain 45 | separable from, or merely link (or bind by name) to the interfaces of, 46 | the Work and Derivative Works thereof. 47 | 48 | "Contribution" shall mean any work of authorship, including 49 | the original version of the Work and any modifications or additions 50 | to that Work or Derivative Works thereof, that is intentionally 51 | submitted to Licensor for inclusion in the Work by the copyright owner 52 | or by an individual or Legal Entity authorized to submit on behalf of 53 | the copyright owner. For the purposes of this definition, "submitted" 54 | means any form of electronic, verbal, or written communication sent 55 | to the Licensor or its representatives, including but not limited to 56 | communication on electronic mailing lists, source code control systems, 57 | and issue tracking systems that are managed by, or on behalf of, the 58 | Licensor for the purpose of discussing and improving the Work, but 59 | excluding communication that is conspicuously marked or otherwise 60 | designated in writing by the copyright owner as "Not a Contribution." 61 | 62 | "Contributor" shall mean Licensor and any individual or Legal Entity 63 | on behalf of whom a Contribution has been received by Licensor and 64 | subsequently incorporated within the Work. 65 | 66 | 2. Grant of Copyright License. Subject to the terms and conditions of 67 | this License, each Contributor hereby grants to You a perpetual, 68 | worldwide, non-exclusive, no-charge, royalty-free, irrevocable 69 | copyright license to reproduce, prepare Derivative Works of, 70 | publicly display, publicly perform, sublicense, and distribute the 71 | Work and such Derivative Works in Source or Object form. 72 | 73 | 3. Grant of Patent License. Subject to the terms and conditions of 74 | this License, each Contributor hereby grants to You a perpetual, 75 | worldwide, non-exclusive, no-charge, royalty-free, irrevocable 76 | (except as stated in this section) patent license to make, have made, 77 | use, offer to sell, sell, import, and otherwise transfer the Work, 78 | where such license applies only to those patent claims licensable 79 | by such Contributor that are necessarily infringed by their 80 | Contribution(s) alone or by combination of their Contribution(s) 81 | with the Work to which such Contribution(s) was submitted. If You 82 | institute patent litigation against any entity (including a 83 | cross-claim or counterclaim in a lawsuit) alleging that the Work 84 | or a Contribution incorporated within the Work constitutes direct 85 | or contributory patent infringement, then any patent licenses 86 | granted to You under this License for that Work shall terminate 87 | as of the date such litigation is filed. 88 | 89 | 4. Redistribution. You may reproduce and distribute copies of the 90 | Work or Derivative Works thereof in any medium, with or without 91 | modifications, and in Source or Object form, provided that You 92 | meet the following conditions: 93 | 94 | (a) You must give any other recipients of the Work or 95 | Derivative Works a copy of this License; and 96 | 97 | (b) You must cause any modified files to carry prominent notices 98 | stating that You changed the files; and 99 | 100 | (c) You must retain, in the Source form of any Derivative Works 101 | that You distribute, all copyright, patent, trademark, and 102 | attribution notices from the Source form of the Work, 103 | excluding those notices that do not pertain to any part of 104 | the Derivative Works; and 105 | 106 | (d) If the Work includes a "NOTICE" text file as part of its 107 | distribution, then any Derivative Works that You distribute must 108 | include a readable copy of the attribution notices contained 109 | within such NOTICE file, excluding those notices that do not 110 | pertain to any part of the Derivative Works, in at least one 111 | of the following places: within a NOTICE text file distributed 112 | as part of the Derivative Works; within the Source form or 113 | documentation, if provided along with the Derivative Works; or, 114 | within a display generated by the Derivative Works, if and 115 | wherever such third-party notices normally appear. The contents 116 | of the NOTICE file are for informational purposes only and 117 | do not modify the License. You may add Your own attribution 118 | notices within Derivative Works that You distribute, alongside 119 | or as an addendum to the NOTICE text from the Work, provided 120 | that such additional attribution notices cannot be construed 121 | as modifying the License. 122 | 123 | You may add Your own copyright statement to Your modifications and 124 | may provide additional or different license terms and conditions 125 | for use, reproduction, or distribution of Your modifications, or 126 | for any such Derivative Works as a whole, provided Your use, 127 | reproduction, and distribution of the Work otherwise complies with 128 | the conditions stated in this License. 129 | 130 | 5. Submission of Contributions. Unless You explicitly state otherwise, 131 | any Contribution intentionally submitted for inclusion in the Work 132 | by You to the Licensor shall be under the terms and conditions of 133 | this License, without any additional terms or conditions. 134 | Notwithstanding the above, nothing herein shall supersede or modify 135 | the terms of any separate license agreement you may have executed 136 | with Licensor regarding such Contributions. 137 | 138 | 6. Trademarks. This License does not grant permission to use the trade 139 | names, trademarks, service marks, or product names of the Licensor, 140 | except as required for reasonable and customary use in describing the 141 | origin of the Work and reproducing the content of the NOTICE file. 142 | 143 | 7. Disclaimer of Warranty. Unless required by applicable law or 144 | agreed to in writing, Licensor provides the Work (and each 145 | Contributor provides its Contributions) on an "AS IS" BASIS, 146 | WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or 147 | implied, including, without limitation, any warranties or conditions 148 | of TITLE, NON-INFRINGEMENT, MERCHANTABILITY, or FITNESS FOR A 149 | PARTICULAR PURPOSE. You are solely responsible for determining the 150 | appropriateness of using or redistributing the Work and assume any 151 | risks associated with Your exercise of permissions under this License. 152 | 153 | 8. Limitation of Liability. In no event and under no legal theory, 154 | whether in tort (including negligence), contract, or otherwise, 155 | unless required by applicable law (such as deliberate and grossly 156 | negligent acts) or agreed to in writing, shall any Contributor be 157 | liable to You for damages, including any direct, indirect, special, 158 | incidental, or consequential damages of any character arising as a 159 | result of this License or out of the use or inability to use the 160 | Work (including but not limited to damages for loss of goodwill, 161 | work stoppage, computer failure or malfunction, or any and all 162 | other commercial damages or losses), even if such Contributor 163 | has been advised of the possibility of such damages. 164 | 165 | 9. Accepting Warranty or Additional Liability. While redistributing 166 | the Work or Derivative Works thereof, You may choose to offer, 167 | and charge a fee for, acceptance of support, warranty, indemnity, 168 | or other liability obligations and/or rights consistent with this 169 | License. However, in accepting such obligations, You may act only 170 | on Your own behalf and on Your sole responsibility, not on behalf 171 | of any other Contributor, and only if You agree to indemnify, 172 | defend, and hold each Contributor harmless for any liability 173 | incurred by, or claims asserted against, such Contributor by reason 174 | of your accepting any such warranty or additional liability. 175 | 176 | END OF TERMS AND CONDITIONS 177 | 178 | APPENDIX: How to apply the Apache License to your work. 179 | 180 | To apply the Apache License to your work, attach the following 181 | boilerplate notice, with the fields enclosed by brackets "{}" 182 | replaced with your own identifying information. (Don't include 183 | the brackets!) The text should be enclosed in the appropriate 184 | comment syntax for the file format. We also recommend that a 185 | file or class name and description of purpose be included on the 186 | same "printed page" as the copyright notice for easier 187 | identification within third-party archives. 188 | 189 | Copyright {yyyy} {name of copyright owner} 190 | 191 | Licensed under the Apache License, Version 2.0 (the "License"); 192 | you may not use this file except in compliance with the License. 193 | You may obtain a copy of the License at 194 | 195 | http://www.apache.org/licenses/LICENSE-2.0 196 | 197 | Unless required by applicable law or agreed to in writing, software 198 | distributed under the License is distributed on an "AS IS" BASIS, 199 | WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. 200 | See the License for the specific language governing permissions and 201 | limitations under the License. 202 | -------------------------------------------------------------------------------- /README.md: -------------------------------------------------------------------------------- 1 | # Fermat's Last Theorem for regular primes 2 | The goal of this project is to prove Fermat's Last Theorem for [regular primes](https://en.wikipedia.org/wiki/Regular_prime) 3 | in [Lean](https://leanprover-community.github.io/). 4 | 5 | The following readme has been shamelessly copied from the [Liquid Tensor Experiment](https://github.com/leanprover-community/lean-liquid/). 6 | 7 | ## How to browse this repository 8 | 9 | ### Blueprint 10 | 11 | Here are a draft [blueprint](https://leanprover-community.github.io/flt-regular/blueprint) and [dependency graph](https://leanprover-community.github.io/flt-regular/blueprint/dep_graph_document.html). 12 | 13 | ### Getting the project 14 | 15 | The recommended way of browsing this repository is by using a Lean development environment. 16 | Crucially, this will allow you to inspect Lean's "Goal state" during proofs, 17 | and easily jump to definitions or otherwise follow paths through the code. Please use the 18 | [installation instructions](https://leanprover-community.github.io/get_started.html#regular-install) 19 | to install Lean and a supporting toolchain. 20 | After that, download and open a copy of the repository 21 | by executing the following command in a terminal: 22 | ``` 23 | git clone https://github.com/leanprover-community/flt-regular.git 24 | cd flt-regular 25 | lake exe cache get 26 | lake build 27 | code . 28 | ``` 29 | For detailed instructions on how to work with Lean projects, 30 | see [this](https://leanprover-community.github.io/install/project.html). 31 | 32 | You can also use gitpod and do everything directly in your browser, without installing anything. 33 | Just click on [![Gitpod](https://img.shields.io/badge/Gitpod-ready--to--code-908a85?logo=gitpod)](https://gitpod.io/#https://github.com/leanprover-community/flt-regular), but beware that everything will be slower than on your computer. 34 | 35 | ### Reading the project 36 | 37 | With the project opened in VScode, 38 | you are all set to start exploring the code. 39 | There are two pieces of functionality that help a lot when browsing through Lean code: 40 | 41 | * "Go to definition": if you right-click on a name of a definition or lemma 42 | (such as `is_regular_number`, or `flt_regular_case_one`), then you can choose "Go to definition" from the menu, 43 | and you will be taken to the relevant location in the source files. 44 | This also works by `Ctrl`-clicking on the name. 45 | * "Goal view": in the event that you would like to read a *proof*, 46 | you can step through the proof line-by-line, 47 | and see the internals of Lean's "brain" in the Goal window. 48 | If the Goal window is not open, 49 | you can open it by clicking on one of the icons in the top right hand corner. 50 | 51 | ### Organization of the project 52 | 53 | * All the Lean code (the juicy stuff) is contained in the directory `FltRegular/`. 54 | * The file `FltRegular/FltRegular.lean` contains the statement of Fermat's Last Theorem for 55 | regular primes. 56 | * The ingredients that go into the theorem statement are defined in several other files. 57 | The most important pieces are: 58 | - `NumberTheory/RegularPrimes.lean` we give the definition of what a regular number is. 59 | - `NumberTheory/*` are the files we are actively working on. 60 | - `ReadyForMathlib/*` are the files that are (almost) ready to be PRed to mathlib. 61 | 62 | Note that we are trying to move all our results to mathlib as fast as possible, so the 63 | folder `ReadyForMathlib` changes rapidly. You should also check `Mathlib.NumberTheory.Cyclotomic.*`. 64 | 65 | ## Brief note on type theory 66 | 67 | Lean is based on type theory, 68 | which means that some things work slightly differently from set theory. 69 | We highlight two syntactical differences. 70 | 71 | * Firstly, the element-of relation (`∈`) plays no fundamental role. 72 | Instead, there is a typing judgment (`:`). 73 | 74 | This means that we write `x : X` to say that "`x` is a term of type `X`" 75 | instead of "`x` is an element of the set `X`". 76 | Conveniently, we can write `f : X → Y` to mean "`f` has type `X → Y`", 77 | in other words "`f` is a function from `X` to `Y`". 78 | 79 | * Secondly, type theorists use lambda-notation. 80 | This means that we can define the square function on the integers via 81 | `fun x ↦ x^2`, which translates to `x ↦ x^2` in set-theoretic notation. 82 | For more information about `λ` (called `fun` in Lean 4), see the Wikipedia page on 83 | [lambda calculus](https://en.wikipedia.org/wiki/Lambda_calculus). 84 | 85 | For a more extensive discussion of type theory, 86 | see the dedicated 87 | [page](https://leanprover-community.github.io/lean-perfectoid-spaces/type_theory.html) 88 | on the perfectoid project website. 89 | 90 | [![Gitpod ready-to-code](https://img.shields.io/badge/Gitpod-ready--to--code-908a85?logo=gitpod)](https://gitpod.io/#https://github.com/leanprover-community/flt-regular) 91 | -------------------------------------------------------------------------------- /blueprint/print/print.aux: -------------------------------------------------------------------------------- 1 | \relax 2 | \providecommand\hyper@newdestlabel[2]{} 3 | \providecommand\babel@aux[2]{} 4 | \@nameuse{bbl@beforestart} 5 | \providecommand\HyperFirstAtBeginDocument{\AtBeginDocument} 6 | \HyperFirstAtBeginDocument{\ifx\hyper@anchor\@undefined 7 | \global\let\oldcontentsline\contentsline 8 | \gdef\contentsline#1#2#3#4{\oldcontentsline{#1}{#2}{#3}} 9 | \global\let\oldnewlabel\newlabel 10 | \gdef\newlabel#1#2{\newlabelxx{#1}#2} 11 | \gdef\newlabelxx#1#2#3#4#5#6{\oldnewlabel{#1}{{#2}{#3}}} 12 | \AtEndDocument{\ifx\hyper@anchor\@undefined 13 | \let\contentsline\oldcontentsline 14 | \let\newlabel\oldnewlabel 15 | \fi} 16 | \fi} 17 | \global\let\hyper@last\relax 18 | \gdef\HyperFirstAtBeginDocument#1{#1} 19 | \providecommand*\HyPL@Entry[1]{} 20 | \citation{Samuel} 21 | \citation{marcus} 22 | \citation{washington} 23 | \HyPL@Entry{0<>} 24 | \babel@aux{english}{} 25 | \@writefile{toc}{\contentsline {section}{\tocsection {}{1}{Introduction}}{1}{section.1}\protected@file@percent } 26 | \@writefile{toc}{\contentsline {section}{\tocsection {}{2}{Discriminants of number fields}}{1}{section.2}\protected@file@percent } 27 | \newlabel{lemma:alt_definition_of_norm}{{2.1}{1}{}{subsection.2.1}{}} 28 | \newlabel{lemma:alt_definition_of_trace}{{2.2}{1}{}{subsection.2.2}{}} 29 | \newlabel{defn_of_disc}{{2.3}{1}{}{subsection.2.3}{}} 30 | \newlabel{lem:lin_indep_iff_disc_ne_zero}{{2.4}{2}{}{subsection.2.4}{}} 31 | \newlabel{lem:disc_change_of_basis}{{2.5}{2}{}{subsection.2.5}{}} 32 | \newlabel{lemma:disc_via_embs}{{2.6}{2}{}{subsection.2.6}{}} 33 | \newlabel{lemma:disc_of_prim_elt_basis}{{2.7}{2}{}{subsection.2.7}{}} 34 | \newlabel{lemma:diff_of_irr_pol}{{2.8}{2}{}{subsection.2.8}{}} 35 | \newlabel{lemma:num_field_disc_in_terms_of_norm}{{2.9}{3}{}{subsection.2.9}{}} 36 | \newlabel{lemma:norm_of_alg_int_is_int}{{2.10}{3}{}{subsection.2.10}{}} 37 | \newlabel{lemma:trace_of_alg_int_is_int}{{2.11}{3}{}{subsection.2.11}{}} 38 | \newlabel{lemma:int_basis_int_disc}{{2.12}{3}{}{subsection.2.12}{}} 39 | \newlabel{lemma:disc_int_basis}{{2.13}{3}{}{subsection.2.13}{}} 40 | \newlabel{lemma:eis_crit_and_alg_ints}{{2.14}{3}{}{subsection.2.14}{}} 41 | \@writefile{toc}{\contentsline {section}{\tocsection {}{3}{Cyclotomic fields}}{4}{section.3}\protected@file@percent } 42 | \newlabel{lemma:cyclo_poly_deg}{{3.1}{4}{}{subsection.3.1}{}} 43 | \newlabel{lemma:cyclo_poly_irr}{{3.2}{4}{}{subsection.3.2}{}} 44 | \newlabel{lem:discr_of_cyclo}{{3.3}{4}{}{subsection.3.3}{}} 45 | \newlabel{theorem:ring_of_ints_of_cyclo}{{3.4}{4}{}{subsection.3.4}{}} 46 | \citation{washington} 47 | \citation{marcus} 48 | \newlabel{lemma:alg_int_abs_val_one}{{3.5}{5}{}{subsection.3.5}{}} 49 | \newlabel{lem:roots_of_unity_in_cyclo}{{3.6}{5}{}{subsection.3.6}{}} 50 | \newlabel{lemma:unit_lemma}{{3.7}{5}{}{subsection.3.7}{}} 51 | \newlabel{lemma:zeta_pow_sub_eq_unit_zeta_sub_one}{{3.8}{5}{}{subsection.3.8}{}} 52 | \newlabel{lemma:ideals_mult_to_power}{{3.9}{5}{}{subsection.3.9}{}} 53 | \@writefile{toc}{\contentsline {section}{\tocsection {}{4}{Fermat's Last Theorem for regular primes}}{5}{section.4}\protected@file@percent } 54 | \newlabel{lem:flt_ideals_coprime}{{4.1}{5}{}{subsection.4.1}{}} 55 | \newlabel{lem:exists_int_sub_pow_prime_dvd}{{4.2}{6}{}{subsection.4.2}{}} 56 | \newlabel{lem:dvd_coeff_cycl_integer}{{4.3}{6}{}{subsection.4.3}{}} 57 | \newlabel{lem:exists_int_sum_eq_zero}{{4.4}{6}{}{subsection.4.4}{}} 58 | \newlabel{lemma:may_assume_coprime}{{4.5}{6}{}{subsection.4.5}{}} 59 | \citation{Borevich_Shafarevich} 60 | \newlabel{defn:is_regular_number}{{4.6}{7}{}{subsection.4.6}{}} 61 | \newlabel{theorem:FLT_case_one}{{4.7}{7}{}{subsection.4.7}{}} 62 | \newlabel{thm:Kummers_lemma}{{4.8}{7}{}{subsection.4.8}{}} 63 | \newlabel{lem:gen_dvd_by_frak_p}{{4.9}{7}{}{subsection.4.9}{}} 64 | \bibstyle{amsalpha} 65 | \bibdata{references} 66 | \bibcite{Borevich_Shafarevich}{BS66} 67 | \bibcite{marcus}{Mar18} 68 | \bibcite{Samuel}{Sam70} 69 | \bibcite{washington}{Was82} 70 | \newlabel{tocindent-1}{0pt} 71 | \newlabel{tocindent0}{13.9941pt} 72 | \newlabel{tocindent1}{19.4691pt} 73 | \newlabel{tocindent2}{0pt} 74 | \newlabel{tocindent3}{0pt} 75 | \newlabel{lem:gen_ideal_coprimality}{{4.10}{8}{}{subsection.4.10}{}} 76 | \newlabel{thm:gen_flt_eqn}{{4.11}{8}{}{subsection.4.11}{}} 77 | \newlabel{theorem:FLT_case_two}{{4.12}{8}{}{subsection.4.12}{}} 78 | \newlabel{FLT_regular}{{4.13}{8}{}{subsection.4.13}{}} 79 | \@writefile{toc}{\contentsline {section}{\tocsection {}{}{References}}{8}{section*.2}\protected@file@percent } 80 | \gdef \@abspage@last{8} 81 | -------------------------------------------------------------------------------- /blueprint/print/print.bbl: -------------------------------------------------------------------------------- 1 | \providecommand{\bysame}{\leavevmode\hbox to3em{\hrulefill}\thinspace} 2 | \providecommand{\MR}{\relax\ifhmode\unskip\space\fi MR } 3 | % \MRhref is called by the amsart/book/proc definition of \MR. 4 | \providecommand{\MRhref}[2]{% 5 | \href{http://www.ams.org/mathscinet-getitem?mr=#1}{#2} 6 | } 7 | \providecommand{\href}[2]{#2} 8 | \begin{thebibliography}{Was82} 9 | 10 | \bibitem[BS66]{Borevich_Shafarevich} 11 | A.~I. Borevich and I.~R. Shafarevich, \emph{Number theory}, Pure and Applied 12 | Mathematics, Vol. 20, Academic Press, New York-London, 1966, Translated from 13 | the Russian by Newcomb Greenleaf. \MR{0195803} 14 | 15 | \bibitem[Mar18]{marcus} 16 | Daniel~A. Marcus, \emph{Number fields}, Universitext, Springer, Cham, 2018, 17 | Second edition of [ MR0457396], With a foreword by Barry Mazur. \MR{3822326} 18 | 19 | \bibitem[Sam70]{Samuel} 20 | Pierre Samuel, \emph{Algebraic theory of numbers}, Translated from the French 21 | by Allan J. Silberger, Houghton Mifflin Co., Boston, Mass., 1970. 22 | \MR{0265266} 23 | 24 | \bibitem[Was82]{washington} 25 | Lawrence~C. Washington, \emph{Introduction to cyclotomic fields}, Graduate 26 | Texts in Mathematics, vol.~83, Springer-Verlag, New York, 1982. \MR{718674 27 | (85g:11001)} 28 | 29 | \end{thebibliography} 30 | -------------------------------------------------------------------------------- /blueprint/print/print.blg: -------------------------------------------------------------------------------- 1 | This is BibTeX, Version 0.99d 2 | Capacity: max_strings=200000, hash_size=200000, hash_prime=170003 3 | The top-level auxiliary file: print.aux 4 | Reallocating 'name_of_file' (item size: 1) to 9 items. 5 | The style file: amsalpha.bst 6 | Reallocating 'name_of_file' (item size: 1) to 11 items. 7 | Database file #1: references.bib 8 | You've used 4 entries, 9 | 2624 wiz_defined-function locations, 10 | 595 strings with 5027 characters, 11 | and the built_in function-call counts, 1198 in all, are: 12 | = -- 89 13 | > -- 43 14 | < -- 4 15 | + -- 12 16 | - -- 12 17 | * -- 91 18 | := -- 203 19 | add.period$ -- 4 20 | call.type$ -- 4 21 | change.case$ -- 21 22 | chr.to.int$ -- 4 23 | cite$ -- 4 24 | duplicate$ -- 64 25 | empty$ -- 109 26 | format.name$ -- 19 27 | if$ -- 237 28 | int.to.chr$ -- 1 29 | int.to.str$ -- 0 30 | missing$ -- 12 31 | newline$ -- 22 32 | num.names$ -- 12 33 | pop$ -- 24 34 | preamble$ -- 1 35 | purify$ -- 25 36 | quote$ -- 0 37 | skip$ -- 46 38 | stack$ -- 0 39 | substring$ -- 28 40 | swap$ -- 10 41 | text.length$ -- 4 42 | text.prefix$ -- 3 43 | top$ -- 0 44 | type$ -- 16 45 | warning$ -- 0 46 | while$ -- 9 47 | width$ -- 5 48 | write$ -- 60 49 | -------------------------------------------------------------------------------- /blueprint/print/print.out: -------------------------------------------------------------------------------- 1 | \BOOKMARK [1][-]{section.1}{\376\377\0001\000.\000\040\000I\000n\000t\000r\000o\000d\000u\000c\000t\000i\000o\000n}{}% 1 2 | \BOOKMARK [1][-]{section.2}{\376\377\0002\000.\000\040\000D\000i\000s\000c\000r\000i\000m\000i\000n\000a\000n\000t\000s\000\040\000o\000f\000\040\000n\000u\000m\000b\000e\000r\000\040\000f\000i\000e\000l\000d\000s}{}% 2 3 | \BOOKMARK [1][-]{section.3}{\376\377\0003\000.\000\040\000C\000y\000c\000l\000o\000t\000o\000m\000i\000c\000\040\000f\000i\000e\000l\000d\000s}{}% 3 4 | \BOOKMARK [1][-]{section.4}{\376\377\0004\000.\000\040\000F\000e\000r\000m\000a\000t\000'\000s\000\040\000L\000a\000s\000t\000\040\000T\000h\000e\000o\000r\000e\000m\000\040\000f\000o\000r\000\040\000r\000e\000g\000u\000l\000a\000r\000\040\000p\000r\000i\000m\000e\000s}{}% 4 5 | \BOOKMARK [1][-]{section*.2}{\376\377\000R\000e\000f\000e\000r\000e\000n\000c\000e\000s}{}% 5 6 | -------------------------------------------------------------------------------- /blueprint/print/print.pdf: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/leanprover-community/flt-regular/05b30b586026461e1d87c57b924e37a632fa4ba9/blueprint/print/print.pdf -------------------------------------------------------------------------------- /blueprint/print/print.toc: -------------------------------------------------------------------------------- 1 | \babel@toc {english}{}\relax 2 | \contentsline {section}{\tocsection {}{1}{Introduction}}{1}{section.1}% 3 | \contentsline {section}{\tocsection {}{2}{Discriminants of number fields}}{1}{section.2}% 4 | \contentsline {section}{\tocsection {}{3}{Cyclotomic fields}}{4}{section.3}% 5 | \contentsline {section}{\tocsection {}{4}{Fermat's Last Theorem for regular primes}}{5}{section.4}% 6 | \contentsline {section}{\tocsection {}{}{References}}{8}{section*.2}% 7 | -------------------------------------------------------------------------------- /blueprint/requirements.txt: -------------------------------------------------------------------------------- 1 | invoke==1.7.1 2 | plasTeX @ git+https://github.com/plastex/plastex.git 3 | leanblueprint @ git+https://github.com/utensil/leanblueprint.git@lean4-only 4 | watchfiles==0.16.1 5 | -------------------------------------------------------------------------------- /blueprint/src/blueprint.sty: -------------------------------------------------------------------------------- 1 | \DeclareOption*{} 2 | \ProcessOptions 3 | 4 | \newcommand{\uses}[1]{} 5 | \newcommand{\proves}[1]{} 6 | \newcommand{\lean}[1]{} 7 | \newcommand{\leanfour}[1]{} 8 | \newcommand{\leanok}{} 9 | \newcommand{\alsoIn}[1]{} 10 | \newcommand{\home}[1]{} 11 | \newcommand{\github}[1]{} 12 | \newcommand{\dochome}[1]{} 13 | 14 | -------------------------------------------------------------------------------- /blueprint/src/content.tex: -------------------------------------------------------------------------------- 1 | \title{Fermat's Last Theorem for regular primes} 2 | 3 | 4 | 5 | 6 | \home{https://leanprover-community.github.io/flt-regular/} 7 | \github{https://github.com/leanprover-community/flt-regular/} 8 | \dochome{https://cbirkbeck.github.io/FltRegulartest/docs/} 9 | 10 | \maketitle 11 | 12 | 13 | \input{demo} 14 | -------------------------------------------------------------------------------- /blueprint/src/plastex.cfg: -------------------------------------------------------------------------------- 1 | [general] 2 | plugins=leanblueprint 3 | 4 | [document] 5 | toc-depth=2 6 | toc-non-files=True 7 | 8 | [files] 9 | directory=../web/ 10 | split-level=3 11 | 12 | [html5] 13 | localtoc-level=0 14 | mathjax-dollars=True 15 | extra-css=style.css 16 | -------------------------------------------------------------------------------- /blueprint/src/preamble.tex: -------------------------------------------------------------------------------- 1 | 2 | 3 | \usepackage[english]{babel} 4 | \usepackage{enumerate} 5 | \usepackage{setspace} 6 | \usepackage[project=..]{blueprint} 7 | \usepackage{amsfonts,amsthm,amsmath,amssymb,comment,url} 8 | 9 | % Set up blueprint options 10 | 11 | 12 | \home{https://leanprover-community.github.io/flt-regular/} 13 | \github{https://github.com/leanprover-community/flt-regular/} 14 | 15 | 16 | \usepackage[ocgcolorlinks]{hyperref} 17 | \hypersetup{citecolor=blue, 18 | linkcolor=red} 19 | 20 | %%============================================================================== 21 | % 22 | % GEOMETRY 23 | % 24 | %%============================================================================== 25 | 26 | 27 | 28 | 29 | %\addbibresource{bibliog.bib} 30 | %\usepackage{fullpage} 31 | % \bibliography{bibliog} 32 | 33 | \renewcommand{\baselinestretch}{1} 34 | 35 | 36 | 37 | 38 | %%============================================================================== 39 | % 40 | % ADDRESSES 41 | % 42 | %%============================================================================== 43 | 44 | %\author{Christopher Birkbeck} 45 | %\address{Department of Mathematics, University College London, Gower street, London, WC1E 6BT} 46 | %\email{c.birkbeck@ucl.ac.uk} 47 | %\urladdr{orcid.org/0000-0002-7546-9028} 48 | 49 | 50 | 51 | 52 | 53 | %%============================================================================== 54 | % 55 | % COMMENTS 56 | % 57 | %%============================================================================== 58 | 59 | 60 | 61 | %%============================================================================== 62 | % 63 | % THEOREM STYLES 64 | % 65 | %%============================================================================== 66 | \theoremstyle{plain} 67 | \newtheorem{theorem}[subsection]{Theorem} 68 | \newtheorem*{theorem*}{Theorem} 69 | \newtheorem{lemma}[subsection]{Lemma} 70 | \newtheorem{cor}[subsection]{Corollary} 71 | \newtheorem{con}[subsection]{Conjecture} 72 | \newtheorem*{con*}{Conjecture} 73 | \newtheorem{prop}[subsection]{Proposition} 74 | \newtheorem*{prop*}{Proposition} 75 | \theoremstyle{definition} 76 | \newtheorem{definition}[subsection]{Definition} 77 | \newtheorem{exmp}[subsection]{Example} 78 | \newtheorem{nexmp}[subsection]{Non-example} 79 | \theoremstyle{definition} 80 | \newtheorem{rmrk}[subsection]{Remark} 81 | \newtheorem*{rmrk*}{Remark} 82 | \newtheorem{warn}[subsection]{Warning} 83 | \newtheorem*{warn*}{Warning} 84 | \newtheorem{exmps}[subsection]{Examples} 85 | \newtheorem*{exmp*}{Example} 86 | \newtheorem{obs}[subsection]{Observation} 87 | \newtheorem*{obs*}{Observation} 88 | \newtheorem{nota}[subsection]{Notation} 89 | \newtheorem{question}[subsection]{Exercise} 90 | \newtheorem{num}[subsection]{\unskip} 91 | \newtheorem{alg}[subsection]{Algorithm} 92 | 93 | 94 | 95 | 96 | %%============================================================================== 97 | % 98 | % MACROS 99 | % 100 | %%============================================================================== 101 | 102 | %Shortcuts: mathbb 103 | \newcommand{\Q}{\mathbb{Q}} 104 | \renewcommand{\O}{\mathcal{O}} 105 | \newcommand{\Tr}{\operatorname{Tr}} 106 | \newcommand{\Ql}{\mathbb{Q}_\ell} 107 | \newcommand{\Qp}{\mathbb{Q}_p} 108 | \renewcommand{\c}{\mathfrak c} 109 | 110 | \newcommand{\Z}{\mathbb{Z}} 111 | \newcommand{\ZZ}{\mathbb{Z}} 112 | \newcommand{\Zp}{\mathbb{Z}_p} 113 | \newcommand{\Zpthick}{\ZZ_p^\times} 114 | 115 | \newcommand{\R}{\mathbb{R}} 116 | \newcommand{\RR}{\mathbb{R}} 117 | 118 | \newcommand{\A}{\mathbb{A}} 119 | \newcommand{\D}{\mathbb{D}} 120 | \newcommand{\C}{\mathbb{C}} 121 | \newcommand{\Cp}{\mathbb{C}_p} 122 | \newcommand{\N}{\mathbb{N}} 123 | \newcommand{\F}{\mathbb{F}} 124 | 125 | \newcommand{\TT}{\mathbb{T}} 126 | \DeclareMathOperator{\bA}{\mathbb{A}} 127 | \DeclareMathOperator{\bG}{\mathbb{G}} 128 | \DeclareMathOperator{\bH}{\mathbb{H}} 129 | \DeclareMathOperator{\bT}{\mathbb{T}} 130 | \DeclareMathOperator{\bV}{\mathbb{V}} 131 | \DeclareMathOperator{\bW}{\mathbb{W}} 132 | \newcommand{\bD}{\mathbb{D}} 133 | \def\AA{\mathbb{A}} 134 | \def\CC{\mathbb{C}} 135 | \def\FF{\mathbb{F}} 136 | \def\GG{\mathbb{G}} 137 | \def\NN{\mathbb{N}} 138 | \renewcommand{\P}{\mathbb{P}} 139 | \renewcommand{\A}{\mathbb{A}} 140 | \newcommand{\G}{\mathbb{G}} 141 | \def\PP{\mathbb{P}} 142 | \def\QQ{\mathbb{Q}} 143 | \def\TT{\mathbb{T}} 144 | \def \Nm {N_{K/\QQ}} 145 | 146 | 147 | % Shortcuts: mathcal 148 | \newcommand{\f}{\mathcal{F}} 149 | \newcommand{\T}{\mathcal{T}} 150 | \newcommand{\uhp}{\mathcal{H}} 151 | \newcommand{\roi}{\mathcal{O}} 152 | \newcommand{\DD}{\mathcal{D}} 153 | \newcommand{\cA}{\mathcal{A}} 154 | \newcommand{\cU}{\mathcal{U}} 155 | \newcommand{\cB}{\mathcal{B}} 156 | \newcommand{\cC}{\mathcal{C}} 157 | \newcommand{\cD}{\mathcal{D}} 158 | \newcommand{\cE}{\mathcal{E}} 159 | \newcommand{\cF}{\mathcal{F}} 160 | \newcommand{\cH}{\mathcal{H}} 161 | \newcommand{\cM}{\mathcal{M}} 162 | \newcommand{\cP}{\mathcal{P}} 163 | \newcommand{\cR}{\mathcal{R}} 164 | \newcommand{\cV}{\mathcal{V}} 165 | \newcommand{\cW}{\mathcal{W}} 166 | \newcommand{\cX}{\mathcal{X}} 167 | \newcommand{\cO}{\mathcal{O}} 168 | \newcommand{\cG}{\mathcal{G}} 169 | \newcommand{\cY}{\mathcal{Y}} 170 | \newcommand{\cI}{\mathcal{I}} 171 | \newcommand{\cZ}{\mathcal{Z}} 172 | \newcommand{\UU}{\mathcal{U}} 173 | \newcommand{\W}{\mathcal{W}} 174 | \def \D {\mathcal{D}} 175 | \def \E {\mathcal{E}} 176 | \def \B {\mathcal{B}} 177 | \def \M {\mathcal{M}} 178 | \def \Y {\mathcal{Y}} 179 | \def \XX {\mathcal{X}} 180 | \def \XXt {\mathcal{X}^{\ast}} 181 | \newcommand*\pp{{\rlap{\('\)}}} 182 | \newcommand*\mA{\mathcal A} 183 | \newcommand*\mB{\mathcal B} 184 | \newcommand*\mC{\mathcal C} 185 | 186 | % Shortcuts: mathscr 187 | 188 | 189 | \newcommand{\fc}{\mathfrak{c}} 190 | \newcommand{\ff}{\mathfrak{f}} 191 | \newcommand{\fg}{\mathfrak{g}} 192 | \newcommand{\fH}{\mathfrak{H}} 193 | \newcommand{\fn}{\mathfrak{n}} 194 | \newcommand{\fm}{\mathfrak{m}} 195 | \newcommand{\m}{\mathfrak{m}} 196 | \newcommand{\n}{N} 197 | \newcommand{\fp}{\mathfrak{p}} 198 | \newcommand{\pri}{\mathfrak{p}} 199 | \newcommand{\fW}{\mathfrak{W}} 200 | \newcommand{\fX}{\mathfrak{X}} 201 | \def\gothb{\mathfrak{b}} 202 | \def\gothc{\mathfrak{c}} 203 | \def\ee{\mathfrak{e}} 204 | \def\gothm{\mathfrak{m}} 205 | \def\gothd{\mathfrak{d}} 206 | \def\gotho{\mathfrak{o}} 207 | \def\goths{\mathfrak{s}} 208 | \def\gothz{\mathfrak{z}} 209 | \def \gothP{\mathfrak{P}} 210 | \def \gothQ{\mathfrak{Q}} 211 | \def \nn {N} 212 | \def \zz {\mathfrak{z}} 213 | \def \MM {\mathfrak{M}} 214 | \def\gothp{\mathfrak{p}} 215 | \def\ps{\mathfrak{p}} 216 | \def\gothq{\mathfrak{q}} 217 | 218 | 219 | 220 | 221 | \def \a{\alpha} 222 | \def \e {\epsilon} 223 | \def \aa {\pmb{a}} 224 | \def \g {\gamma} 225 | \def \GA {\Gamma} 226 | \def \w {\omega} 227 | \def \LL {\mathcal{L}} 228 | \def \W {\mathcal{W}} 229 | \def \p {\partial} 230 | \def \ll {\lambda} 231 | \def \L {\Lambda} 232 | \def \st {\star} 233 | \def \bu {\bullet} 234 | 235 | \def \s {\sigma} 236 | \def \Si {\Sigma} 237 | \def \lam {\lambda} 238 | \def \ii {\iota} 239 | \def \d {\mathfrak{d}} 240 | \def \OO {\mathcal{O}} 241 | \def \OF {\mathcal{O}_F} 242 | \def \OFU {\mathcal{O}_F^{\times,+}} 243 | \def \OOO {\mathscr{O}} 244 | 245 | \def \CD {\mathfrak{D}} 246 | \def \CX {\mathscr{X}} 247 | \def \CZ {\mathscr{Z}} 248 | \def \CM {\mathscr{M}} 249 | \def \Z {\ZZ} 250 | \def \s {\sigma} 251 | \def \SS {S_{k}} 252 | \def \PGI {\mathrm{P}\Gamma_{\!\scaleto{0}{3pt}}(p^n)} 253 | \def \GI {\Gamma_{\!\scaleto{0}{3pt}}(p^n)} 254 | \def \GIst {\Gamma_{\!\scaleto{0}{3pt}}^*(p^n)} 255 | \def \un {\underline} 256 | \def \ov{\overline} 257 | \def \wt{\widetilde} 258 | \def \wh{\widehat} 259 | \def \H{\widehat{H}} 260 | \def \HHH {\mathcal{H}} 261 | \def \U {\mathcal{U}} 262 | 263 | 264 | \def\gotha{\mathfrak{a}} 265 | \def \tr {\ttr_{F/\QQ}} 266 | \def \eU {\e_\k} 267 | 268 | 269 | \def \k {\kappa} 270 | \def \Norm {N_{F/\QQ}} 271 | 272 | 273 | % Shortcuts: Operators 274 | \DeclareMathOperator{\coker}{coker} 275 | \DeclareMathOperator{\Sch}{Sch} 276 | \DeclareMathOperator{\Set}{Set} 277 | \DeclareMathOperator{\hH}{H} 278 | \DeclareMathOperator{\tor}{tor} 279 | \DeclareMathOperator{\et}{et} 280 | \DeclareMathOperator{\qproet}{qproet} 281 | \DeclareMathOperator{\End}{End} 282 | \DeclareMathOperator{\Lie}{Lie} 283 | \DeclareMathOperator{\Sym}{Sym} 284 | \DeclareMathOperator{\Ker}{Ker} 285 | \DeclareMathOperator{\Sp}{Sp} 286 | \DeclareMathOperator{\Ext}{Ext} 287 | \DeclareMathOperator{\val}{val} 288 | \DeclareMathOperator{\im}{Im} 289 | \DeclareMathOperator{\lcm}{lcm} 290 | 291 | \DeclareMathOperator{\unn}{univ} 292 | \DeclareMathOperator{\Spa}{Spa} 293 | \DeclareMathOperator{\Spf}{Spf} 294 | \DeclareMathOperator{\Spec}{Spec} 295 | \DeclareMathOperator{\NP}{NP} 296 | \DeclareMathOperator{\Inv}{Inv} 297 | \newcommand{\an}{\mathrm{an}} 298 | \newcommand{\ad}{\mathrm{ad}} 299 | \newcommand{\Ig}{\mathrm{Ig}} 300 | \DeclareMathOperator{\ab}{ab} 301 | \DeclareMathOperator{\fs}{fs} 302 | \newcommand{\id}{\operatorname{id}} 303 | \DeclareMathOperator{\Stab}{Stab} 304 | \DeclareMathOperator{\coinf}{Coinf} 305 | \DeclareMathOperator{\trr}{tr} 306 | \DeclareMathOperator{\Det}{Det} 307 | \DeclareMathOperator{\nrd}{nrd} 308 | \DeclareMathOperator{\ttr}{Tr} 309 | \DeclareMathOperator{\suup}{sup} 310 | \DeclareMathOperator{\Tor}{Tor} 311 | 312 | \newcommand{\AIP}{\mathrm{AIP}} 313 | \DeclareMathOperator{\Fred}{Fred} 314 | \DeclareMathOperator{\Frob}{Frob} 315 | \DeclareMathOperator{\Hdg}{Hdg} 316 | \DeclareMathOperator{\Coeq}{Coeq} 317 | \DeclareMathOperator{\Cor}{Cor} 318 | \newcommand{\cyc}{{\text{cyc}}} 319 | \newcommand{\cts}{{\text{cts}}} 320 | \DeclareMathOperator{\disc}{Disc} 321 | \DeclareMathOperator{\Gal}{Gal} 322 | \DeclareMathOperator{\SL}{SL} 323 | \DeclareMathOperator{\Sh}{Sh} 324 | \DeclareMathOperator{\Gr}{Gr} 325 | \DeclareMathOperator{\Vol}{Vol} 326 | \DeclareMathOperator{\GL}{GL} 327 | \DeclareMathOperator{\GSp}{GSp} 328 | \DeclareMathOperator{\Isom}{Isom} 329 | \DeclareMathOperator{\Hom}{Hom} 330 | \DeclareMathOperator{\Ind}{Ind} 331 | \DeclareMathOperator{\inv}{inv} 332 | \DeclareMathOperator{\Res}{Res} 333 | \DeclareMathOperator{\smcy}{smcy} 334 | \DeclareMathOperator{\Trace}{Trace} 335 | \DeclareMathOperator{\Cl}{Cl} 336 | \DeclareMathOperator{\Inf}{Inf} 337 | \DeclareMathOperator{\Aut}{Aut} 338 | \DeclareMathOperator{\Mapc}{\operatorname{Map}_{\cts}} 339 | \newcommand{\HTf}{\HT_{\mathrm{fin}}} 340 | \newcommand{\hcg}{\mathrm{H}_{\mathrm{c}}^g} 341 | 342 | 343 | \usepackage{amsfonts, amsthm, amssymb, amsmath} 344 | 345 | -------------------------------------------------------------------------------- /blueprint/src/print.tex: -------------------------------------------------------------------------------- 1 | % This file is the entry point for generating the PDF by LaTex. 2 | \documentclass[11pt]{amsart} 3 | 4 | % Define `ifplastex` so we can have complicated macros for PDF 5 | % and simplified macros for the website. 6 | % See http://plastex.github.io/plastex/plastex/sect0006.html 7 | \newif\ifplastex 8 | \plastexfalse 9 | 10 | \input{preamble} 11 | 12 | \usepackage{blueprint} 13 | 14 | % The rest of this file is identical to web.tex 15 | 16 | \begin{document} 17 | 18 | \input{content} 19 | 20 | \bibliographystyle{amsalpha} 21 | \bibliography{references} 22 | 23 | \end{document} 24 | 25 | % vim: ts=2 et sw=2 sts=2 26 | -------------------------------------------------------------------------------- /blueprint/src/references.bib: -------------------------------------------------------------------------------- 1 | 2 | 3 | 4 | @book {marcus, 5 | AUTHOR = {Marcus, Daniel A.}, 6 | TITLE = {Number fields}, 7 | SERIES = {Universitext}, 8 | NOTE = {Second edition of [ MR0457396], 9 | With a foreword by Barry Mazur}, 10 | PUBLISHER = {Springer, Cham}, 11 | YEAR = {2018}, 12 | PAGES = {xviii+203}, 13 | ISBN = {978-3-319-90232-6; 978-3-319-90233-3}, 14 | MRCLASS = {11-01 (11Rxx 11Txx 12-01)}, 15 | MRNUMBER = {3822326}, 16 | DOI = {10.1007/978-3-319-90233-3}, 17 | URL = {https://doi.org/10.1007/978-3-319-90233-3}, 18 | } 19 | 20 | @book {Samuel, 21 | AUTHOR = {Samuel, Pierre}, 22 | TITLE = {Algebraic theory of numbers}, 23 | SERIES = {Translated from the French by Allan J. Silberger}, 24 | PUBLISHER = {Houghton Mifflin Co., Boston, Mass.}, 25 | YEAR = {1970}, 26 | PAGES = {109}, 27 | MRCLASS = {10.01}, 28 | MRNUMBER = {0265266}, 29 | } 30 | 31 | @book {washington, 32 | 33 | AUTHOR = {Washington, Lawrence C.}, 34 | 35 | TITLE = {Introduction to cyclotomic fields}, 36 | 37 | SERIES = {Graduate Texts in Mathematics}, 38 | 39 | VOLUME = {83}, 40 | 41 | PUBLISHER = {Springer-Verlag}, 42 | 43 | ADDRESS = {New York}, 44 | 45 | YEAR = {1982}, 46 | 47 | PAGES = {xi+389}, 48 | 49 | ISBN = {0-387-90622-3}, 50 | 51 | MRCLASS = {11-01 (11R18 11R23)}, 52 | 53 | MRNUMBER = {718674 (85g:11001)}, 54 | 55 | MRREVIEWER = {T. Mets{\"a}nkyl{\"a}}, 56 | 57 | DOI = {10.1007/978-1-4684-0133-2}, 58 | 59 | URL = {http://dx.doi.org/10.1007/978-1-4684-0133-2}, 60 | 61 | BOEKCODE = {12-01}, 62 | 63 | } 64 | 65 | @book {Borevich_Shafarevich, 66 | AUTHOR = {Borevich, A. I. and Shafarevich, I. R.}, 67 | TITLE = {Number theory}, 68 | SERIES = {Pure and Applied Mathematics, Vol. 20}, 69 | NOTE = {Translated from the Russian by Newcomb Greenleaf}, 70 | PUBLISHER = {Academic Press, New York-London}, 71 | YEAR = {1966}, 72 | PAGES = {x+435}, 73 | MRCLASS = {10.00}, 74 | MRNUMBER = {0195803}, 75 | } 76 | 77 | @book {SD, 78 | place={Cambridge}, 79 | series={London Mathematical Society Student Texts}, 80 | title={A Brief Guide to Algebraic Number Theory}, 81 | DOI={10.1017/CBO9781139173360}, 82 | publisher={Cambridge University Press}, 83 | author={Swinnerton-Dyer, H. P. F.}, 84 | year={2001}, 85 | collection={London Mathematical Society Student Texts} 86 | } 87 | 88 | @book{Hilbert, 89 | title={The Theory of Algebraic Number Fields}, 90 | author={Hilbert, D. and Lemmermeyer, F. and Adamson, I.T. and Schappacher, N. and Schoof, R.}, 91 | isbn={9783540627791}, 92 | lccn={98042185}, 93 | url={https://books.google.co.uk/books?id=_Q2h83Bm94cC}, 94 | year={1998}, 95 | publisher={Springer Berlin Heidelberg} 96 | } -------------------------------------------------------------------------------- /blueprint/src/style.css: -------------------------------------------------------------------------------- 1 | /* 2 | red RawSienna #8C2700 3 | light red Salmon!5 #fef9f8 4 | dark green ForestGreen!70!black #2f6023 5 | green ForestGreen #458933 6 | light green ForestGreen!5 #f5f9f4 7 | dark blue MidnightBlue #19196B 8 | */ 9 | .example_thmwrapper { 10 | border: 1px solid #8C2700; 11 | padding: 1em 1em 1em 1em; 12 | background-color: #fef9f8; 13 | } 14 | 15 | .example_thmheading { 16 | color: #8C2700; 17 | } 18 | 19 | .remark_thmwrapper { 20 | border-left: 5px solid #458933; 21 | padding: 1em 1em 1em 1em; 22 | background-color: #f5f9f4; 23 | } 24 | 25 | .remark_thmheading { 26 | color: #2f6023; /* #399D3C; */ 27 | } -------------------------------------------------------------------------------- /blueprint/src/web.tex: -------------------------------------------------------------------------------- 1 | % This file is the entry point for generating the website by Plastex. 2 | \documentclass{amsart} 3 | 4 | \input{preamble} 5 | 6 | \usepackage[dep_graph, coverage, project=../../]{blueprint} 7 | % \dochome{https://math.iisc.ac.in/~gadgil/proofs-and-programs-2023/doc/} 8 | 9 | % The rest of this file is identical to print.tex 10 | 11 | \begin{document} 12 | 13 | \input{content} 14 | 15 | \bibliographystyle{amsalpha} 16 | \bibliography{references} 17 | 18 | \end{document} 19 | 20 | % vim: ts=2 et sw=2 sts=2 21 | -------------------------------------------------------------------------------- /blueprint/tasks.py: -------------------------------------------------------------------------------- 1 | """ 2 | This file specifies tasks mostly for generating the blueprint for formalizing Geometric Algebra. 3 | 4 | Check out blueprint/README.md for more information. 5 | """ 6 | 7 | import os 8 | import warnings 9 | import shutil 10 | from pathlib import Path 11 | from invoke import run, task 12 | 13 | import http.server 14 | import socketserver 15 | 16 | ROOT = Path(__file__).parent 17 | BP_DIR = ROOT 18 | 19 | @task 20 | def bp(ctx): 21 | """ 22 | Build the blueprint PDF file and prepare src/web.bbl for task `web` using TeXLive 23 | """ 24 | 25 | cwd = os.getcwd() 26 | os.chdir(BP_DIR) 27 | run('mkdir -p print && cd src && xelatex -output-directory=../print print.tex') 28 | run('cd print && bibtex print.aux', env={'BIBINPUTS': '../src'}) 29 | run('cd src && xelatex -output-directory=../print print.tex') 30 | run('cd src && xelatex -output-directory=../print print.tex') 31 | run('cp print/print.bbl src/web.bbl') 32 | os.chdir(cwd) 33 | 34 | @task 35 | def web(ctx): 36 | """ 37 | Build the blueprint website 38 | """ 39 | 40 | cwd = os.getcwd() 41 | os.chdir(BP_DIR/'src') 42 | run('plastex -c plastex.cfg web.tex') 43 | os.chdir(cwd) 44 | 45 | try: 46 | shutil.copy2(BP_DIR/'print'/'print.pdf', BP_DIR/'web'/'blueprint.pdf') 47 | except Exception as e: 48 | print(e) 49 | 50 | @task 51 | def serve(ctx, port=8080): 52 | """ 53 | Serve the blueprint website assuming the blueprint website is already built 54 | """ 55 | 56 | class MyTCPServer(socketserver.TCPServer): 57 | def server_bind(self): 58 | import socket 59 | self.socket.setsockopt(socket.SOL_SOCKET, socket.SO_REUSEADDR, 1) 60 | self.socket.bind(self.server_address) 61 | 62 | cwd = os.getcwd() 63 | os.chdir(BP_DIR/'web') 64 | 65 | Handler = http.server.SimpleHTTPRequestHandler 66 | httpd = MyTCPServer(('', port), Handler) 67 | 68 | try: 69 | (ip, port) = httpd.server_address 70 | ip = ip or 'localhost' 71 | print(f'Serving http://{ip}:{port}/ ...') 72 | httpd.serve_forever() 73 | except KeyboardInterrupt: 74 | pass 75 | httpd.server_close() 76 | 77 | os.chdir(cwd) 78 | 79 | @task(bp, web) 80 | def dev(ctx): 81 | """ 82 | Serve the blueprint website, rebuild PDF and the website on file changes 83 | 84 | Note: this will not run/rerun task `decls` 85 | """ 86 | 87 | from watchfiles import run_process, DefaultFilter 88 | 89 | def callback(changes): 90 | print('Changes detected:', changes) 91 | bp(ctx) 92 | web(ctx) 93 | 94 | run_process(BP_DIR/'src', target='inv serve', callback=callback, 95 | watch_filter=DefaultFilter( 96 | ignore_entity_patterns=( 97 | '.*\.aux$', 98 | '.*\.log$', 99 | '.*\.fls$', 100 | '.*\.fdb_latexmk$', 101 | '.*\.bbl$', 102 | '.*\.paux$', 103 | '.*\.pdf$', 104 | '.*\.out$', 105 | '.*\.blg$', 106 | '.*\.synctex.*$', 107 | ) 108 | )) 109 | 110 | @task(bp, web) 111 | def all(ctx): 112 | """ 113 | Run all tasks: `bp`, and `web` 114 | """ 115 | 116 | pass -------------------------------------------------------------------------------- /lake-manifest.json: -------------------------------------------------------------------------------- 1 | {"version": "1.1.0", 2 | "packagesDir": ".lake/packages", 3 | "packages": 4 | [{"url": "https://github.com/leanprover/doc-gen4", 5 | "type": "git", 6 | "subDir": null, 7 | "scope": "", 8 | "rev": "6dff9d050137b425784b1ad17980160df3235478", 9 | "name": "«doc-gen4»", 10 | "manifestFile": "lake-manifest.json", 11 | "inputRev": "main", 12 | "inherited": false, 13 | "configFile": "lakefile.lean"}, 14 | {"url": "https://github.com/PatrickMassot/checkdecls.git", 15 | "type": "git", 16 | "subDir": null, 17 | "scope": "", 18 | "rev": "3d425859e73fcfbef85b9638c2a91708ef4a22d4", 19 | "name": "checkdecls", 20 | "manifestFile": "lake-manifest.json", 21 | "inputRev": "master", 22 | "inherited": false, 23 | "configFile": "lakefile.lean"}, 24 | {"url": "https://github.com/leanprover-community/mathlib4.git", 25 | "type": "git", 26 | "subDir": null, 27 | "scope": "", 28 | "rev": "b5368c7d48612974c32bab5cebc280f2517b33be", 29 | "name": "mathlib", 30 | "manifestFile": "lake-manifest.json", 31 | "inputRev": null, 32 | "inherited": false, 33 | "configFile": "lakefile.lean"}, 34 | {"url": "https://github.com/mhuisi/lean4-cli", 35 | "type": "git", 36 | "subDir": null, 37 | "scope": "", 38 | "rev": "a0abd472348dd725adbb26732e79b26e7e220913", 39 | "name": "Cli", 40 | "manifestFile": "lake-manifest.json", 41 | "inputRev": "main", 42 | "inherited": true, 43 | "configFile": "lakefile.toml"}, 44 | {"url": "https://github.com/fgdorais/lean4-unicode-basic", 45 | "type": "git", 46 | "subDir": null, 47 | "scope": "", 48 | "rev": "9f94839235c03d3e04aaed60d277a287f9c84873", 49 | "name": "UnicodeBasic", 50 | "manifestFile": "lake-manifest.json", 51 | "inputRev": "main", 52 | "inherited": true, 53 | "configFile": "lakefile.lean"}, 54 | {"url": "https://github.com/dupuisf/BibtexQuery", 55 | "type": "git", 56 | "subDir": null, 57 | "scope": "", 58 | "rev": "dbfe2b7630c5f7c5c1cf71e7747ffc0a30337f69", 59 | "name": "BibtexQuery", 60 | "manifestFile": "lake-manifest.json", 61 | "inputRev": "master", 62 | "inherited": true, 63 | "configFile": "lakefile.toml"}, 64 | {"url": "https://github.com/acmepjz/md4lean", 65 | "type": "git", 66 | "subDir": null, 67 | "scope": "", 68 | "rev": "8ba0ef10d178ab95a5d6fe3cfbd586c6ecef2717", 69 | "name": "MD4Lean", 70 | "manifestFile": "lake-manifest.json", 71 | "inputRev": "main", 72 | "inherited": true, 73 | "configFile": "lakefile.lean"}, 74 | {"url": "https://github.com/leanprover-community/plausible", 75 | "type": "git", 76 | "subDir": null, 77 | "scope": "leanprover-community", 78 | "rev": "1603151ac0db4e822908e18094f16acc250acaff", 79 | "name": "plausible", 80 | "manifestFile": "lake-manifest.json", 81 | "inputRev": "main", 82 | "inherited": true, 83 | "configFile": "lakefile.toml"}, 84 | {"url": "https://github.com/leanprover-community/LeanSearchClient", 85 | "type": "git", 86 | "subDir": null, 87 | "scope": "leanprover-community", 88 | "rev": "6c62474116f525d2814f0157bb468bf3a4f9f120", 89 | "name": "LeanSearchClient", 90 | "manifestFile": "lake-manifest.json", 91 | "inputRev": "main", 92 | "inherited": true, 93 | "configFile": "lakefile.toml"}, 94 | {"url": "https://github.com/leanprover-community/import-graph", 95 | "type": "git", 96 | "subDir": null, 97 | "scope": "leanprover-community", 98 | "rev": "e25fe66cf13e902ba550533ef681cc35a9f18dc2", 99 | "name": "importGraph", 100 | "manifestFile": "lake-manifest.json", 101 | "inputRev": "main", 102 | "inherited": true, 103 | "configFile": "lakefile.toml"}, 104 | {"url": "https://github.com/leanprover-community/ProofWidgets4", 105 | "type": "git", 106 | "subDir": null, 107 | "scope": "leanprover-community", 108 | "rev": "6980f6ca164de593cb77cd03d8eac549cc444156", 109 | "name": "proofwidgets", 110 | "manifestFile": "lake-manifest.json", 111 | "inputRev": "v0.0.62", 112 | "inherited": true, 113 | "configFile": "lakefile.lean"}, 114 | {"url": "https://github.com/leanprover-community/aesop", 115 | "type": "git", 116 | "subDir": null, 117 | "scope": "leanprover-community", 118 | "rev": "c3a19fa17982c5c1413fea335f371869b8b12e1d", 119 | "name": "aesop", 120 | "manifestFile": "lake-manifest.json", 121 | "inputRev": "master", 122 | "inherited": true, 123 | "configFile": "lakefile.toml"}, 124 | {"url": "https://github.com/leanprover-community/quote4", 125 | "type": "git", 126 | "subDir": null, 127 | "scope": "leanprover-community", 128 | "rev": "e1d2994e0acdee2f0c03c9d84d28a5df34aa0020", 129 | "name": "Qq", 130 | "manifestFile": "lake-manifest.json", 131 | "inputRev": "master", 132 | "inherited": true, 133 | "configFile": "lakefile.toml"}, 134 | {"url": "https://github.com/leanprover-community/batteries", 135 | "type": "git", 136 | "subDir": null, 137 | "scope": "leanprover-community", 138 | "rev": "4765a4ee7ab89f407e4bc1d2aa1dbae83bd95e16", 139 | "name": "batteries", 140 | "manifestFile": "lake-manifest.json", 141 | "inputRev": "main", 142 | "inherited": true, 143 | "configFile": "lakefile.toml"}], 144 | "name": "«flt-regular»", 145 | "lakeDir": ".lake"} 146 | -------------------------------------------------------------------------------- /lakefile.toml: -------------------------------------------------------------------------------- 1 | name = "«flt-regular»" 2 | defaultTargets = ["FltRegular"] 3 | 4 | [leanOptions] 5 | pp.unicode.fun = true # pretty-prints `fun a ↦ b` 6 | autoImplicit = false # no "assume a typo is a new variable" 7 | relaxedAutoImplicit = false # no "assume a typo is a new variable" 8 | maxSynthPendingDepth = 3 # same as mathlib, changes behaviour of typeclass inference 9 | linter.flexible = true # no rigid tactic (e.g. `exact`) after a flexible tactic (e.g. `simp`) 10 | # Enable all mathlib linters: automatically matches what mathlib uses. 11 | linter.mathlibStandardSet = true 12 | 13 | [[require]] 14 | name = "mathlib" 15 | git = "https://github.com/leanprover-community/mathlib4.git" 16 | 17 | [[require]] 18 | name = "checkdecls" 19 | git = "https://github.com/PatrickMassot/checkdecls.git" 20 | rev = "master" 21 | 22 | [[require]] 23 | name = "«doc-gen4»" 24 | git = "https://github.com/leanprover/doc-gen4" 25 | rev = "main" 26 | 27 | [[lean_lib]] 28 | name = "FltRegular" 29 | globs = ["FltRegular"] 30 | -------------------------------------------------------------------------------- /lean-toolchain: -------------------------------------------------------------------------------- 1 | leanprover/lean4:v4.21.0-rc3 -------------------------------------------------------------------------------- /scripts/noshake.json: -------------------------------------------------------------------------------- 1 | {} 2 | -------------------------------------------------------------------------------- /unseen.lean: -------------------------------------------------------------------------------- 1 | import Lean 2 | import FltRegular 3 | 4 | open Lean 5 | 6 | def nameCode (n : Name) : String := 7 | if n = .anonymous then 8 | "anonymous" 9 | else 10 | n.toString.replace "." "_" 11 | 12 | def nameDisplay (n : Name) : String := 13 | n.components.getLast!.toString 14 | 15 | def printDeps₁ (k : Name) (_v : Array Name) (print : String → IO Unit) : IO Unit := do 16 | let n := k.componentsRev[1]! 17 | print (nameCode k ++ " [label=\"" ++ nameDisplay k ++ "\"" ++ 18 | " group=\"" ++ n.toString ++ "\"]" ++ ";\n") 19 | 20 | def printDeps₂ (k : Name) (v : Array Name) (print : String → IO Unit) : IO Unit := do 21 | for val in v do 22 | if (`FltRegular).isPrefixOf val then 23 | print (nameCode val ++ " -> " ++ nameCode k ++ ";\n") 24 | 25 | def group (name : Name) : Name := 26 | (name.eraseSuffix? name.componentsRev.head!).get! 27 | 28 | def groups (imports : NameMap (Array Name)) : NameMap Unit := 29 | RBMap.fromList (imports.fold (fun xs k _ => 30 | if (`FltRegular).isPrefixOf k then (group k, ()) :: xs else xs) []) _ 31 | 32 | /-- `#deptree` outputs a graphviz dependency graph to `depgraph.dot`. Build it with 33 | `dot -Tpdf -Gnewrank=true -Goverlaps=false -Gsplines=ortho depgraph.dot > depgraph.pdf`. -/ 34 | elab "#deptree " : command => do 35 | let env ← getEnv 36 | let imports := env.importGraph 37 | IO.FS.withFile "docs/depgraph.dot" IO.FS.Mode.write fun h => do 38 | h.write "digraph {\n".toUTF8 39 | h.write "compound=true;\n".toUTF8 40 | for (gp, _) in groups imports do 41 | h.write ("subgraph cluster_" ++ nameCode gp ++ " {\n").toUTF8 42 | for (k, v) in imports do 43 | if (`FltRegular).isPrefixOf k && group k = gp then do 44 | printDeps₁ k v (fun s => h.write s.toUTF8) 45 | h.write ("label = \"" ++ gp.toString ++ "\";\n").toUTF8 46 | h.write ("margin = 32;\n").toUTF8 47 | h.write ("pad = 32;\n").toUTF8 48 | h.write ("penwidth = 5;\n").toUTF8 49 | h.write ("color = cyan4;\n").toUTF8 50 | h.write "}\n".toUTF8 51 | for (k, v) in imports do 52 | if (`FltRegular).isPrefixOf k then do 53 | printDeps₂ k v (fun s => h.write s.toUTF8) 54 | h.write "}\n".toUTF8 55 | 56 | /-- Extracts the names of the declarations in `env` on which `decl` depends. -/ 57 | -- source: 58 | -- https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/Counting.20prerequisites.20of.20a.20theorem/near/425370265 59 | def getVisited (env : Environment) (decl : Name) := 60 | let (_, { visited, .. }) := Lean.CollectAxioms.collect decl |>.run env |>.run {} 61 | visited.erase decl 62 | 63 | partial def allDeclsIn (module : Name) : Elab.Command.CommandElabM (Array Name) := do 64 | let mFile ← findOLean module 65 | unless (← mFile.pathExists) do 66 | logError m!"object file '{mFile}' of module {module} does not exist" 67 | let (md, _) ← readModuleData mFile 68 | let decls ← md.constNames.filterM fun d => 69 | return !(← d.isBlackListed) && !(`injEq).isSuffixOf d && !(`sizeOf_spec).isSuffixOf d 70 | return decls 71 | 72 | def allFiles (env : Environment) : List Name := 73 | (env.importGraph.fold (fun xs k _ => if (`FltRegular).isPrefixOf k then k :: xs else xs) []).mergeSort 74 | (toString · < toString ·) 75 | 76 | def allDecls (env : Environment) : Elab.Command.CommandElabM NameSet := 77 | (fun l => RBTree.ofList (l.map (fun a => a.toList)).flatten) <$> 78 | (List.mapM allDeclsIn (allFiles env)) 79 | 80 | /-- `#index` computes an index of the declations in the project and saves it to `index.csv`. -/ 81 | elab "#index " : command => do 82 | let env ← getEnv 83 | let allDecls ← allDecls env 84 | let result ← List.mapM (fun decl => do 85 | let ranges ← findDeclarationRanges? decl 86 | let mod ← findModuleOf? decl 87 | match (ranges, mod) with 88 | | (some ranges, some mod) => pure (some (decl, ranges, mod)) 89 | | _ => pure none) 90 | (allDecls.toList.mergeSort (toString · < toString ·)) 91 | let result' := result.filterMap id 92 | IO.FS.withFile "docs/index.csv" IO.FS.Mode.write (fun h => do 93 | for (decl, ranges, mod) in result' do 94 | h.write (decl.toString ++ ", " ++ mod.toString ++ ", " ++ 95 | ranges.range.pos.line.repr ++ ", " ++ ranges.range.pos.column.repr ++ "\n").toUTF8) 96 | 97 | def seenIn (env : Environment) (allDecls : NameSet) (decl : Name) : NameSet := 98 | (getVisited env decl).fold 99 | (fun decls x => if allDecls.contains x then decls.insert x else decls) RBTree.empty 100 | 101 | /-- `#unseen` computes a list of the declarations in the project that are 102 | defined but not used in the current file. The list is stored in `unseen_defs.txt`. 103 | The average runtime on my computer is 1 minute, with 16 threads. -/ 104 | elab "#unseen " : command => do 105 | let env ← getEnv 106 | let allDecls ← allDecls env 107 | let mut unseen := allDecls 108 | let timeStart ← IO.monoMsNow 109 | let tasks := (fun l => Task.spawn (fun _ => seenIn env allDecls l)) <$> 110 | allDecls.toList.mergeSort (toString · < toString ·) 111 | for task in tasks do 112 | for v in task.get do 113 | unseen := unseen.erase v 114 | IO.FS.withFile "docs/unseen_defs.txt" IO.FS.Mode.write (fun h => do 115 | for v in unseen.toList.mergeSort (toString · < toString ·) do 116 | h.write (v.toString ++ "\n").toUTF8) 117 | let timeEnd ← IO.monoMsNow 118 | logInfo m!"operation took {(timeEnd - timeStart) / 1000}s" 119 | 120 | #unseen 121 | --------------------------------------------------------------------------------