diff --git a/FltRegular/NumberTheory/KummersLemma/KummersLemma.lean b/FltRegular/NumberTheory/KummersLemma/KummersLemma.lean index 3d667dd7..084d2903 100644 --- a/FltRegular/NumberTheory/KummersLemma/KummersLemma.lean +++ b/FltRegular/NumberTheory/KummersLemma/KummersLemma.lean @@ -47,6 +47,7 @@ theorem not_for_all_zeta_sub_one_pow_dvd_sub_one_of_pow_ne (u : (𝓞 K)ˣ) -- the 𝑝-th cyclotomic field) and 𝜉 a primitive 𝑝-th root of unity; -- if a unit 𝑢∈𝐐(𝜉) is congruent to an integer modulo 𝑝, then 𝑢 is a 𝑝-th power in 𝐐(𝜉). set_option synthInstance.maxHeartbeats 40000 in +set_option maxHeartbeats 400000 in theorem eq_pow_prime_of_unit_of_congruent (u : (𝓞 K)ˣ) (hcong : ∃ n : ℤ, (p : 𝓞 K) ∣ (u - n : 𝓞 K)) : ∃ v, u = v ^ (p : ℕ) := by @@ -55,7 +56,7 @@ theorem eq_pow_prime_of_unit_of_congruent (u : (𝓞 K)ˣ) obtain ⟨x, hx⟩ : (p : 𝓞 K) ∣ (↑(u ^ (p - 1 : ℕ)) : 𝓞 K) - 1 := by obtain ⟨n, hn⟩ := hcong have hn' : (p : ℤ) ∣ n ^ (p - 1 : ℕ) - 1 := by - refine Int.modEq_iff_dvd.mp (Int.ModEq.pow_card_sub_one_eq_one hpri.out (n := n) ?_).symm + refine Int.modEq_iff_dvd.mp (Int.ModEq.pow_card_sub_one_eq_one hpri.out ?_).symm rw [isCoprime_comm, (Nat.prime_iff_prime_int.mp hpri.out).coprime_iff_not_dvd] intro h replace h := _root_.map_dvd (Int.castRingHom (𝓞 K)) h diff --git a/lake-manifest.json b/lake-manifest.json index e620906d..062e0939 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -15,7 +15,7 @@ "type": "git", "subDir": null, "scope": "", - "rev": "8e459c606cbf19248c6a59956570f051f05e6067", + "rev": "7edf946a4217aa3aa911290811204096e8464ada", "name": "checkdecls", "manifestFile": "lake-manifest.json", "inputRev": "master", @@ -25,7 +25,7 @@ "type": "git", "subDir": null, "scope": "", - "rev": "1ec33a404152f54bbabe1162c3dfa356a6f8b0ab", + "rev": "67d092e380faf850fd6a1febf81c27fb33c6c5c8", "name": "mathlib", "manifestFile": "lake-manifest.json", "inputRev": null, @@ -135,7 +135,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "7805acf1864ba1a2e359f86a8f092ccf1438ad83", + "rev": "c016aa9938c4cedc9b7066099f99bcae1b1af625", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "main",