Skip to content

Commit fbe1ae5

Browse files
bump
1 parent c0d3922 commit fbe1ae5

File tree

2 files changed

+8
-8
lines changed

2 files changed

+8
-8
lines changed

FltRegular/NumberTheory/Hilbert92.lean

+4-4
Original file line numberDiff line numberDiff line change
@@ -130,7 +130,7 @@ lemma lemma2 [Module A G] (S : systemOfUnits p G s) (hs : S.IsFundamental)
130130
let a' := a.comapDomain (Fin.succAbove i) Fin.succAbove_right_injective.injOn
131131
have hS' : S'.units ∘ Fin.succAbove i = S.units ∘ Fin.succAbove i := by
132132
ext; simp only [Function.comp_apply, ne_eq, Fin.succAbove_ne, not_false_eq_true,
133-
Function.update_noteq]
133+
Function.update_of_ne]
134134
have ha' :
135135
Finsupp.linearCombination A (S'.units ∘ Fin.succAbove i) a' + S.units i = (1 - zeta p) • g := by
136136
rw [hS', Finsupp.linearCombination_comp, LinearMap.comp_apply, Finsupp.lmapDomain_apply,
@@ -158,11 +158,11 @@ lemma lemma2 [Module A G] (S : systemOfUnits p G s) (hs : S.IsFundamental)
158158
· rw [add_comm, ← eq_sub_iff_add_eq] at ha'
159159
rw [← hij, ha']
160160
apply sub_mem
161-
· exact Submodule.smul_mem _ _ (Submodule.subset_span ⟨i, Function.update_same _ _ _⟩)
161+
· exact Submodule.smul_mem _ _ (Submodule.subset_span ⟨i, Function.update_self _ _ _⟩)
162162
· rw [← Finsupp.range_linearCombination, Finsupp.linearCombination_comp, LinearMap.comp_apply]
163163
exact ⟨_, rfl⟩
164-
· exact Submodule.subset_span ⟨j, Function.update_noteq (Ne.symm hij) _ _⟩
165-
· refine ⟨g, Submodule.subset_span ⟨i, Function.update_same _ _ _⟩, ?_⟩
164+
· exact Submodule.subset_span ⟨j, Function.update_of_ne (Ne.symm hij) _ _⟩
165+
· refine ⟨g, Submodule.subset_span ⟨i, Function.update_self _ _ _⟩, ?_⟩
166166
rw [← Finsupp.range_linearCombination]
167167
rintro ⟨l, rfl⟩
168168
letI := (Algebra.id A).toModule

lake-manifest.json

+4-4
Original file line numberDiff line numberDiff line change
@@ -5,7 +5,7 @@
55
"type": "git",
66
"subDir": null,
77
"scope": "",
8-
"rev": "fd52917a1ac853bcb73cd089ce43e6a4257c98e0",
8+
"rev": "43d30441447b3b8da2ae24e2091751feb743640c",
99
"name": "«doc-gen4»",
1010
"manifestFile": "lake-manifest.json",
1111
"inputRev": "main",
@@ -15,7 +15,7 @@
1515
"type": "git",
1616
"subDir": null,
1717
"scope": "",
18-
"rev": "7edf946a4217aa3aa911290811204096e8464ada",
18+
"rev": "75d7aea6c81efeb68701164edaaa9116f06b91ba",
1919
"name": "checkdecls",
2020
"manifestFile": "lake-manifest.json",
2121
"inputRev": "master",
@@ -25,7 +25,7 @@
2525
"type": "git",
2626
"subDir": null,
2727
"scope": "",
28-
"rev": "d4736a1fb2613cc1c6eac9321e0471f744ff3551",
28+
"rev": "58dc043d6bdac239bfbdcb8de5257893aac92baa",
2929
"name": "mathlib",
3030
"manifestFile": "lake-manifest.json",
3131
"inputRev": null,
@@ -85,7 +85,7 @@
8585
"type": "git",
8686
"subDir": null,
8787
"scope": "leanprover-community",
88-
"rev": "d7caecce0d0f003fd5e9cce9a61f1dd6ba83142b",
88+
"rev": "003ff459cdd85de551f4dcf95cdfeefe10f20531",
8989
"name": "LeanSearchClient",
9090
"manifestFile": "lake-manifest.json",
9191
"inputRev": "main",

0 commit comments

Comments
 (0)