Skip to content

Commit 36e783c

Browse files
bump
1 parent 359ab56 commit 36e783c

File tree

9 files changed

+18
-17
lines changed

9 files changed

+18
-17
lines changed

FltRegular/CaseI/Statement.lean

+1-1
Original file line numberDiff line numberDiff line change
@@ -205,7 +205,7 @@ theorem auxf' (hp5 : 5 ≤ p) (a b : ℤ) (k₁ k₂ : Fin p) :
205205
rw [← Finset.card_pos, hcard, card_range]
206206
exact Nat.sub_pos_of_lt (lt_of_lt_of_le this hp5)
207207
obtain ⟨i, hi⟩ := hcard
208-
refine' ⟨i, sdiff_subset _ _ hi, _⟩
208+
refine' ⟨i, sdiff_subset hi, _⟩
209209
have hi0 : i ≠ 0 := fun h => by simp [h, s] at hi
210210
have hi1 : i ≠ 1 := fun h => by simp [h, s] at hi
211211
have hik₁ : i ≠ k₁ := fun h => by simp [h, s] at hi

FltRegular/FltThree/Primes.lean

-1
Original file line numberDiff line numberDiff line change
@@ -2,7 +2,6 @@
22
Copyright (c) 2020 Ruben Van de Velde. All rights reserved.
33
Released under Apache 2.0 license as described in the file LICENSE.
44
-/
5-
import Mathlib.Data.Int.Parity
65
import Mathlib.RingTheory.Int.Basic
76

87
section

FltRegular/NumberTheory/Cyclotomic/MoreLemmas.lean

+2-1
Original file line numberDiff line numberDiff line change
@@ -31,7 +31,7 @@ lemma norm_Int_zeta_sub_one : Algebra.norm ℤ (↑(IsPrimitiveRoot.unit' hζ) -
3131
letI := IsCyclotomicExtension.numberField {p} ℚ K
3232
haveI : Fact (Nat.Prime p) := hpri
3333
apply RingHom.injective_int (algebraMap ℤ ℚ)
34-
simp [Algebra.coe_norm_int, hζ.sub_one_norm_prime (cyclotomic.irreducible_rat p.2) hp]
34+
simp [Algebra.coe_norm_int, hζ.norm_sub_one_of_prime_ne_two' (cyclotomic.irreducible_rat p.2) hp]
3535

3636
@[simp]
3737
lemma PNat.coe_two : (2 : ℕ+) = (2 : ℕ) := rfl
@@ -179,6 +179,7 @@ lemma quotient_zero_sub_one_comp_aut (σ : 𝓞 K →+* 𝓞 K) :
179179
· rw [mem_nthRootsFinset p.pos, ← map_pow, hζ.unit'_coe.pow_eq_one, map_one]
180180
· rw [mem_nthRootsFinset p.pos, hζ.unit'_coe.pow_eq_one]
181181

182+
set_option maxHeartbeats 400000 in
182183
set_option synthInstance.maxHeartbeats 80000 in
183184
lemma zeta_sub_one_dvd_trace_sub_smul (x : 𝓞 K) :
184185
(hζ.unit' - 1 : 𝓞 K) ∣ Algebra.trace ℤ _ x - (p - 1 : ℕ) • x := by

FltRegular/NumberTheory/Cyclotomic/UnitLemmas.lean

+2-1
Original file line numberDiff line numberDiff line change
@@ -148,7 +148,7 @@ theorem IsPrimitiveRoot.eq_one_mod_sub_of_pow {A : Type _} [CommRing A] [IsDomai
148148
obtain ⟨k, -, rfl⟩ := hζ.eq_pow_of_pow_eq_one hμ p.pos
149149
rw [map_pow, eq_one_mod_one_sub, one_pow]
150150

151-
set_option synthInstance.maxHeartbeats 40000 in
151+
set_option synthInstance.maxHeartbeats 80000 in
152152
set_option maxHeartbeats 400000 in
153153
theorem aux {t} {l : 𝓞 K} {f : Fin t → ℤ} {μ : K} (hμ : IsPrimitiveRoot μ p)
154154
(h : ∑ x : Fin t, f x • (⟨μ, hμ.isIntegral p.pos⟩ : 𝓞 K) ^ (x : ℕ) = l) :
@@ -359,6 +359,7 @@ lemma unit_inv_conj_not_neg_zeta_runity_aux (u : Rˣ) (hp : (p : ℕ).Prime) :
359359
rw [this a]
360360
exact (aux hζ hζ hu).trans (aux hζ hζ.inv hu').symm
361361

362+
set_option maxHeartbeats 400000 in
362363
set_option synthInstance.maxHeartbeats 80000 in
363364
theorem unit_inv_conj_not_neg_zeta_runity (h : p ≠ 2) (u : Rˣ) (n : ℕ) (hp : (p : ℕ).Prime) :
364365
u * (unitGalConj K p u)⁻¹ ≠ -hζ.unit' ^ n := by

FltRegular/NumberTheory/Finrank.lean

+2-2
Original file line numberDiff line numberDiff line change
@@ -54,9 +54,9 @@ lemma FiniteDimensional.finrank_add_finrank_quotient_le (N : Submodule R M) :
5454
apply LinearIndependent.finset_card_le_finrank
5555
· rw [← LinearIndependent.finset_toSet, Finset.coe_union, Finset.coe_image, Finset.coe_image]
5656
refine LinearIndependent.union ?_ ?_ H
57-
· rw [← linearIndependent_image (Subtype.val_injective.injOn _)]
57+
· rw [← linearIndependent_image Subtype.val_injective.injOn]
5858
exact hs'.map' N.subtype N.ker_subtype
59-
· rw [← linearIndependent_image (hf.injective.injOn _)]
59+
· rw [← linearIndependent_image hf.injective.injOn]
6060
apply LinearIndependent.of_comp N.mkQ
6161
convert ht'
6262
exact funext fun x => hf _

FltRegular/NumberTheory/Hilbert92.lean

+3-3
Original file line numberDiff line numberDiff line change
@@ -56,8 +56,8 @@ theorem Finsupp.prod_congr' {α M N} [Zero M] [CommMonoid N] {f₁ f₂ : α →
5656
(h : ∀ x, g1 x (f₁ x) = g2 x (f₂ x)) (hg1 : ∀ i, g1 i 0 = 1) (hg2 : ∀ i, g2 i 0 = 1) :
5757
f₁.prod g1 = f₂.prod g2 := by
5858
classical
59-
rw [f₁.prod_of_support_subset (Finset.subset_union_left _ f₂.support) _ (fun i _ ↦ hg1 i),
60-
f₂.prod_of_support_subset (Finset.subset_union_right f₁.support _) _ (fun i _ ↦ hg2 i)]
59+
rw [f₁.prod_of_support_subset Finset.subset_union_left _ (fun i _ ↦ hg1 i),
60+
f₂.prod_of_support_subset Finset.subset_union_right _ (fun i _ ↦ hg2 i)]
6161
exact Finset.prod_congr rfl (fun x _ ↦ h x)
6262

6363
lemma LinearIndependent.update {ι} [DecidableEq ι] {R} [CommRing R] [Module R G]
@@ -160,7 +160,7 @@ lemma lemma2 [Module A G] (S : systemOfUnits p G r) (hs : S.IsFundamental) (i :
160160
let S' : systemOfUnits p G (r + 1) := ⟨Function.update S.units i g,
161161
LinearIndependent.update' _ _ _ _ _ _ (CyclotomicIntegers.one_sub_zeta_mem_nonZeroDivisors p)
162162
hg (ha ▸ one_mem A⁰) S.linearIndependent⟩
163-
let a' := a.comapDomain (Fin.succAbove i) (Fin.succAbove_right_injective.injOn _)
163+
let a' := a.comapDomain (Fin.succAbove i) Fin.succAbove_right_injective.injOn
164164
have hS' : S'.units ∘ Fin.succAbove i = S.units ∘ Fin.succAbove i
165165
· ext; simp only [Function.comp_apply, ne_eq, Fin.succAbove_ne, not_false_eq_true,
166166
Function.update_noteq]

FltRegular/ReadyForMathlib/Homogenization.lean

+1-1
Original file line numberDiff line numberDiff line change
@@ -808,7 +808,7 @@ theorem support_sum_monomial_subset' [DecidableEq ι] {α : Type _} (S : Finset
808808
apply Finset.union_subset
809809
· apply Finset.Subset.trans support_monomial_subset _
810810
rw [Finset.image_insert]
811-
convert Finset.subset_union_left _ (Finset.image g S)
811+
exact union_subset_left fun ⦃a⦄ a => a
812812
· apply Finset.Subset.trans hsi _
813813
rw [Finset.image_insert]
814814
exact Finset.subset_insert (g s) (Finset.image g S)

lake-manifest.json

+6-6
Original file line numberDiff line numberDiff line change
@@ -1,10 +1,10 @@
1-
{"version": 7,
1+
{"version": "1.0.0",
22
"packagesDir": ".lake/packages",
33
"packages":
44
[{"url": "https://github.com/leanprover-community/batteries",
55
"type": "git",
66
"subDir": null,
7-
"rev": "60d622c124cebcecc000853cdae93f4251f4beb5",
7+
"rev": "6a63eb6a326181df29d95a84ce1f16c1145e66d8",
88
"name": "batteries",
99
"manifestFile": "lake-manifest.json",
1010
"inputRev": "main",
@@ -13,7 +13,7 @@
1313
{"url": "https://github.com/leanprover-community/quote4",
1414
"type": "git",
1515
"subDir": null,
16-
"rev": "53156671405fbbd5402ed17a79bd129b961bd8d6",
16+
"rev": "a7bfa63f5dddbcab2d4e0569c4cac74b2585e2c6",
1717
"name": "Qq",
1818
"manifestFile": "lake-manifest.json",
1919
"inputRev": "master",
@@ -22,7 +22,7 @@
2222
{"url": "https://github.com/leanprover-community/aesop",
2323
"type": "git",
2424
"subDir": null,
25-
"rev": "70ec1d99be1e1b835d831f39c01b0d14921d2118",
25+
"rev": "7e3bd939c6badfcb1e607c0fddb509548baafd05",
2626
"name": "aesop",
2727
"manifestFile": "lake-manifest.json",
2828
"inputRev": "master",
@@ -49,7 +49,7 @@
4949
{"url": "https://github.com/leanprover-community/import-graph.git",
5050
"type": "git",
5151
"subDir": null,
52-
"rev": "b167323652ab59a5d1b91e906ca4172d1c0474b7",
52+
"rev": "7983e959f8f4a79313215720de3ef1eca2d6d474",
5353
"name": "importGraph",
5454
"manifestFile": "lake-manifest.json",
5555
"inputRev": "main",
@@ -58,7 +58,7 @@
5858
{"url": "https://github.com/leanprover-community/mathlib4.git",
5959
"type": "git",
6060
"subDir": null,
61-
"rev": "9220d10df329407f9ddb9911812afacd2a7e1af2",
61+
"rev": "303e71911e1c040ac1278de13fc2032102888a6e",
6262
"name": "mathlib",
6363
"manifestFile": "lake-manifest.json",
6464
"inputRev": null,

lean-toolchain

+1-1
Original file line numberDiff line numberDiff line change
@@ -1 +1 @@
1-
leanprover/lean4:v4.8.0-rc2
1+
leanprover/lean4:v4.9.0-rc1

0 commit comments

Comments
 (0)