Skip to content

Commit 21b1db8

Browse files
committed
Merge branch 'pitmonticone-golf-proofs' into main
2 parents 10e7f80 + 720ebfd commit 21b1db8

File tree

5 files changed

+48
-90
lines changed

5 files changed

+48
-90
lines changed

FLT/AutomorphicForm/QuaternionAlgebra.lean

+1-1
Original file line numberDiff line numberDiff line change
@@ -60,7 +60,7 @@ instance : TopologicalRing (D ⊗[F] (FiniteAdeleRing (𝓞 F) F)) := moobar (Fi
6060

6161
namespace TotallyDefiniteQuaternionAlgebra
6262

63-
noncomputable example : D →+* (D ⊗[F] FiniteAdeleRing (𝓞 F) F) := by exact
63+
noncomputable example : D →+* (D ⊗[F] FiniteAdeleRing (𝓞 F) F) :=
6464
Algebra.TensorProduct.includeLeftRingHom
6565

6666
abbrev Dfx := (D ⊗[F] (FiniteAdeleRing (𝓞 F) F))ˣ

FLT/Basic/Reductions.lean

+8-26
Original file line numberDiff line numberDiff line change
@@ -103,20 +103,17 @@ lemma gcdab_eq_gcdac {a b c : ℤ} {p : ℕ} (hp : 0 < p) (h : a ^ p + b ^ p = c
103103
have foo : gcd a b ∣ gcd a c := by
104104
apply dvd_gcd (gcd_dvd_left a b)
105105
rw [← Int.pow_dvd_pow_iff hp.ne', ← h]
106-
apply dvd_add
107-
· rw [Int.pow_dvd_pow_iff hp.ne']
108-
exact gcd_dvd_left a b
109-
· rw [Int.pow_dvd_pow_iff hp.ne']
110-
exact gcd_dvd_right a b
106+
apply dvd_add <;> rw [Int.pow_dvd_pow_iff hp.ne']
107+
· exact gcd_dvd_left a b
108+
· exact gcd_dvd_right a b
111109
have bar : gcd a c ∣ gcd a b := by
112110
apply dvd_gcd (gcd_dvd_left a c)
113111
have h2 : b ^ p = c ^ p - a ^ p := eq_sub_of_add_eq' h
114112
rw [← Int.pow_dvd_pow_iff hp.ne', h2]
115113
apply dvd_add
116114
· rw [Int.pow_dvd_pow_iff hp.ne']
117115
exact gcd_dvd_right a c
118-
· rw [dvd_neg]
119-
rw [Int.pow_dvd_pow_iff hp.ne']
116+
· rw [dvd_neg, Int.pow_dvd_pow_iff hp.ne']
120117
exact gcd_dvd_left a c
121118
change _ ∣ (Int.gcd a c : ℤ) at foo
122119
apply Int.ofNat_dvd.1 at bar
@@ -221,45 +218,35 @@ def FreyCurve (P : FreyPackage) : EllipticCurve ℚ := {
221218
simp_rw [← mul_pow]
222219
refine pow_ne_zero 2 <| pow_ne_zero P.p <| (mul_ne_zero (mul_ne_zero P.ha0 P.hb0) P.hc0)
223220
coe_Δ' := by
224-
simp only [Units.val_mk0]
225-
rw [← Int.cast_pow P.c, ← P.hFLT]
221+
simp only [Units.val_mk0, ← Int.cast_pow P.c, ← P.hFLT]
226222
field_simp [EllipticCurve.Δ', WeierstrassCurve.Δ, WeierstrassCurve.b₂, WeierstrassCurve.b₄,
227223
WeierstrassCurve.b₆, WeierstrassCurve.b₈]
228224
ring }
229225

230226
lemma FreyCurve.b₂ (P : FreyPackage) :
231227
P.FreyCurve.b₂ = P.b ^ P.p - P.a ^ P.p := by
232228
simp [FreyCurve, WeierstrassCurve.b₂]
233-
field_simp
234-
norm_cast
235229
ring
236230

237231
lemma FreyCurve.b₄ (P : FreyPackage) :
238232
P.FreyCurve.b₄ = - (P.a * P.b) ^ P.p / 8 := by
239233
simp [FreyCurve, WeierstrassCurve.b₄]
240-
field_simp
241-
norm_cast
242234
ring
243235

244236
lemma FreyCurve.c₄ (P : FreyPackage) :
245237
P.FreyCurve.c₄ = (P.a ^ P.p) ^ 2 + P.a ^ P.p * P.b ^ P.p + (P.b ^ P.p) ^ 2 := by
246238
simp [FreyCurve.b₂, FreyCurve.b₄, WeierstrassCurve.c₄]
247-
field_simp
248-
norm_cast
249239
ring
250240

251241
lemma FreyCurve.c₄' (P : FreyPackage) :
252242
P.FreyCurve.c₄ = P.c ^ (2 * P.p) - (P.a * P.b) ^ P.p := by
253243
rw [FreyCurve.c₄]
254-
norm_cast
255-
rw [pow_mul', ← hFLT]
244+
rw_mod_cast [pow_mul', ← hFLT]
256245
ring
257246

258247
lemma FreyCurve.Δ'inv (P : FreyPackage) :
259248
(↑(P.FreyCurve.Δ'⁻¹) : ℚ) = 2 ^ 8 / (P.a*P.b*P.c)^(2*P.p) := by
260249
simp [FreyCurve]
261-
congr 1
262-
norm_cast
263250
ring
264251

265252
lemma FreyCurve.j (P : FreyPackage) :
@@ -277,7 +264,7 @@ private lemma j_pos_aux (a b : ℤ) (hb : b ≠ 0) : 0 < (a + b) ^ 2 - a * b :=
277264
apply mul_self_pos.mpr hb
278265
| inr h =>
279266
rw [sub_pos]
280-
apply h.trans_le (sq_nonneg _)
267+
exact h.trans_le (sq_nonneg _)
281268

282269
/-- The q-adic valuation of the j-invariant of the Frey curve is a multiple of p if 2 < q is
283270
a prime of bad reduction. -/
@@ -287,8 +274,7 @@ lemma FreyCurve.j_valuation_of_bad_prime (P : FreyPackage) {q : ℕ} (hqPrime :
287274
have := Fact.mk hqPrime
288275
have hqPrime' := Nat.prime_iff_prime_int.mp hqPrime
289276
have h₀ : ((P.c ^ (2 * P.p) - (P.a * P.b) ^ P.p) ^ 3 : ℚ) ≠ 0 := by
290-
norm_cast
291-
rw [pow_mul', ← P.hFLT, mul_pow]
277+
rw_mod_cast [pow_mul', ← P.hFLT, mul_pow]
292278
exact pow_ne_zero _ <| ne_of_gt <| j_pos_aux _ _ (pow_ne_zero _ P.hb0)
293279
have h₁ : P.a * P.b * P.c ≠ 0 := mul_ne_zero (mul_ne_zero P.ha0 P.hb0) P.hc0
294280
rw [FreyCurve.j, padicValRat.div (mul_ne_zero (by norm_num) h₀) (pow_ne_zero _ (mod_cast h₁)),
@@ -321,10 +307,6 @@ lemma FreyCurve.j_valuation_of_bad_prime (P : FreyPackage) {q : ℕ} (hqPrime :
321307

322308
end FreyPackage
323309

324-
325-
326-
327-
328310
/-!
329311
330312
Given an elliptic curve over `ℚ`, the p-torsion points defined over an algebraic

FLT/mathlibExperiments/Frobenius.lean

+36-56
Original file line numberDiff line numberDiff line change
@@ -384,7 +384,7 @@ noncomputable instance : Field (A ⧸ P) := Ideal.Quotient.field P
384384
lemma is_primepow_char_A_quot_P : IsPrimePow q := Fintype.isPrimePow_card_of_field
385385

386386
lemma ex_primepow_char_A_quot_P : ∃ p n : ℕ , Prime p ∧ 0 < n ∧ p ^ n = q := by
387-
apply is_primepow_char_A_quot_P
387+
exact is_primepow_char_A_quot_P _
388388

389389
local notation "p" => Classical.choose (CharP.exists (A ⧸ P))
390390
local notation "p'" => Classical.choose (ex_primepow_char_A_quot_P P)
@@ -407,7 +407,7 @@ lemma p_is_p' : p = p' := by
407407
have h1 : p ∣ q := charP_iff (A ⧸ P) p |>.1 (p_is_char P) q |>.1 eq0
408408
have eq1 : p' ^ n = q := q_is_p'_pow_n P
409409
rw [← eq1] at h1
410-
refine Nat.dvd_prime (p'_is_prime P) |>.1
410+
exact Nat.dvd_prime (p'_is_prime P) |>.1
411411
(Nat.Prime.dvd_of_dvd_pow (p_is_prime P) h1) |>.resolve_left <| Nat.Prime.ne_one <| p_is_prime P
412412

413413
lemma q_is_p_pow_n : p ^ n = q := by
@@ -421,7 +421,7 @@ theorem pow_eq_expand (a : (A ⧸ P)[X]) :
421421
have pprime : Nat.Prime p := p_is_prime P
422422
have factprime : Fact (Nat.Prime p) := { out := pprime }
423423
rw [← q_is_p_pow_n, ← map_expand_pow_char, map_expand, FiniteField.frobenius_pow]
424-
· simp only [exists_and_left, RingHom.one_def, map_id]
424+
· simp [exists_and_left, RingHom.one_def, map_id]
425425
· exact (q_is_p_pow_n P).symm
426426

427427
end FreshmansDream
@@ -488,11 +488,7 @@ theorem gal_smul_F_eq (σ : L ≃ₐ[K] L) :
488488
∏ τ : L ≃ₐ[K] L,
489489
(X - C (galRestrict A K L B σ
490490
(galRestrict A K L B τ α))):= by
491-
rw [Finset.smul_prod]
492-
simp_rw [smul_sub]
493-
ext
494-
congr
495-
simp only [smul_X, smul_C, AlgEquiv.smul_def]
491+
simp [Finset.smul_prod, smul_sub, smul_X, smul_C, AlgEquiv.smul_def]
496492

497493
/-- use `Finset.prod_bij` to show
498494
`(galRestrict A K L B σ • (∏ τ : L ≃ₐ[K] L, (X - C (galRestrict A K L B τ α))) =
@@ -507,7 +503,7 @@ lemma F_invariant_under_finite_aut (σ : L ≃ₐ[K] L) :
507503
have i_inj : ∀ (τ₁ : L ≃ₐ[K] L) (hτ₁ : τ₁ ∈ Finset.univ) (τ₂ : L ≃ₐ[K] L)
508504
(hτ₂ : τ₂ ∈ Finset.univ), i τ₁ hτ₁ = i τ₂ hτ₂ → τ₁ = τ₂ := by
509505
intros τ₁ _ τ₂ _ h
510-
simpa only [i, mul_right_inj] using h
506+
simpa [i, mul_right_inj] using h
511507
have i_surj : ∀ σ ∈ Finset.univ, ∃ (τ : L ≃ₐ[K] L) (hτ : τ ∈ Finset.univ), i τ hτ = σ := by
512508
intro τ'
513509
simp only [Finset.mem_univ, exists_true_left, forall_true_left, i]
@@ -520,7 +516,7 @@ lemma F_invariant_under_finite_aut (σ : L ≃ₐ[K] L) :
520516
simp only [i, map_mul, sub_right_inj, C_inj]
521517
rw [AlgEquiv.aut_mul]
522518
rfl
523-
apply Finset.prod_bij i hi i_inj i_surj h
519+
exact Finset.prod_bij i hi i_inj i_surj h
524520

525521
/-- ` L ≃ₐ[K] L` fixes `F`. -/
526522
theorem gal_smul_F_eq_self (σ : L ≃ₐ[K] L) :
@@ -532,9 +528,8 @@ theorem gal_smul_F_eq_self (σ : L ≃ₐ[K] L) :
532528

533529
theorem gal_smul_coeff_eq (n : ℕ) (h : ∀ σ : L ≃ₐ[K] L, galRestrict A K L B σ • F = F)
534530
(σ : L ≃ₐ[K] L) : galRestrict A K L B σ • (coeff F n) = coeff F n := by
535-
simp only [AlgEquiv.smul_def]
536531
nth_rewrite 2 [← h σ]
537-
simp only [coeff_smul, AlgEquiv.smul_def]
532+
simp [coeff_smul, AlgEquiv.smul_def]
538533

539534
variable {A K L B Q}
540535

@@ -580,7 +575,7 @@ theorem coeff_lives_in_A (n : ℕ) : ∃ a : A, algebraMap B L (coeff F n) = (al
580575
exact NoZeroSMulDivisors.algebraMap_injective K L
581576
cases' h1.1 h2 with a' ha'
582577
use a'
583-
simpa only [eq0, RingHom.coe_comp, Function.comp_apply, ha']
578+
simpa [eq0, RingHom.coe_comp, Function.comp_apply, ha']
584579

585580
/-- `α` is a root of `F`. -/
586581
lemma isRoot_α : eval α F = 0 := by
@@ -603,14 +598,12 @@ lemma isRoot_α : eval α F = 0 := by
603598
apply eq0 at evalid
604599
rwa [← eqF]
605600

606-
lemma quotient_isRoot_α : (eval α F) ∈ Q := by
607-
rw [isRoot_α Q_ne_bot]
608-
apply Ideal.zero_mem
601+
lemma quotient_isRoot_α : (eval α F) ∈ Q := (isRoot_α Q_ne_bot) ▸ Ideal.zero_mem _
609602

610603
lemma conjugate_isRoot_α (σ : L ≃ₐ[K] L) : (eval (galRestrict A K L B σ α) F) = 0 := by
611604
have evalσ : eval (galRestrict A K L B σ α)
612605
(X - C (galRestrict A K L B σ α)) = 0 := by
613-
simp only [eval_sub, eval_X, eval_C, sub_self]
606+
simp [eval_sub, eval_X, eval_C, sub_self]
614607
have eqF : (eval (galRestrict A K L B σ α) (∏ τ : L ≃ₐ[K] L,
615608
(X - C (galRestrict A K L B τ α)))) =
616609
eval (galRestrict A K L B σ α) F := rfl
@@ -630,9 +623,7 @@ lemma conjugate_isRoot_α (σ : L ≃ₐ[K] L) : (eval (galRestrict A K L B σ
630623
rw [← eqF, evalσ]
631624

632625
lemma conjugate_quotient_isRoot_α (σ : L ≃ₐ[K] L) :
633-
(eval (galRestrict A K L B σ α) F) ∈ Q := by
634-
rw [conjugate_isRoot_α Q_ne_bot]
635-
apply Ideal.zero_mem
626+
(eval (galRestrict A K L B σ α) F) ∈ Q := (conjugate_isRoot_α Q_ne_bot) _ ▸ Ideal.zero_mem _
636627

637628
lemma F_is_root_iff_is_conjugate {x : B} :
638629
IsRoot F x ↔ (∃ σ : L ≃ₐ[K] L, x = (galRestrict A K L B σ α)) := by
@@ -656,7 +647,7 @@ lemma F_is_root_iff_is_conjugate {x : B} :
656647
· intros h
657648
rcases h with ⟨σ, hσ⟩
658649
rw [Polynomial.IsRoot.def, hσ]
659-
apply conjugate_isRoot_α Q_ne_bot
650+
exact conjugate_isRoot_α Q_ne_bot _
660651

661652
lemma F_eval_zero_is_conjugate {x : B} (h : eval x F = 0) : ∃ σ : L ≃ₐ[K] L,
662653
x = ((galRestrict A K L B σ) α) := by
@@ -678,17 +669,15 @@ lemma ex_poly_in_A : ∃ m : A[X], Polynomial.map (algebraMap A B) m = F := by
678669
apply Iff.not
679670
rw [← _root_.map_eq_zero_iff (algebraMap B L), this, map_eq_zero_iff]
680671
have I : NoZeroSMulDivisors A L := NoZeroSMulDivisors.trans A K L
681-
· apply NoZeroSMulDivisors.algebraMap_injective
682-
· apply NoZeroSMulDivisors.algebraMap_injective
672+
· exact NoZeroSMulDivisors.algebraMap_injective _ _
673+
· exact NoZeroSMulDivisors.algebraMap_injective _ _
683674
}}
684675
use m
685676
ext n
686-
have := Classical.choose_spec (h n)
687677
simp only [coeff_map, coeff_ofFinsupp, Finsupp.coe_mk]
688678
set s := Classical.choose (h n)
689679
apply NoZeroSMulDivisors.algebraMap_injective B L
690-
rw [this]
691-
exact (IsScalarTower.algebraMap_apply A B L _).symm
680+
exact (Classical.choose_spec (h n)) ▸ (IsScalarTower.algebraMap_apply A B L _).symm
692681

693682
/--`m' : A[X]` such that `Polynomial.map (algebraMap A B) m = F`. -/
694683
noncomputable def m' : A[X] := Classical.choose (ex_poly_in_A Q_ne_bot)
@@ -703,11 +692,11 @@ lemma m_eq_F_in_B_quot_Q :
703692
Polynomial.map (algebraMap B (B ⧸ Q)) F := by
704693
suffices h : Polynomial.map (algebraMap A B) m = F by
705694
exact congrArg (map (algebraMap B (B ⧸ Q))) h
706-
apply m_mapsto_F
695+
exact m_mapsto_F _
707696

708697
lemma m_expand_char_q : (Polynomial.map (algebraMap A (A ⧸ P)) m) ^ q =
709698
(expand _ q (Polynomial.map (algebraMap A (A ⧸ P)) m)) := by
710-
apply pow_eq_expand
699+
exact pow_eq_expand _ _
711700

712701
lemma B_m_expand_char_q : (Polynomial.map (algebraMap A (B ⧸ Q)) m) ^ q =
713702
(expand _ q (Polynomial.map (algebraMap A (B ⧸ Q)) m)) := by
@@ -741,23 +730,23 @@ lemma pow_expand_A_B_scalar_tower_F :
741730
lemma F_expand_eval_eq_eval_pow :
742731
(eval₂ (Ideal.Quotient.mk Q) (Ideal.Quotient.mk Q α) F) ^ q =
743732
(eval₂ (Ideal.Quotient.mk Q) (Ideal.Quotient.mk Q (α ^ q)) F) := by
744-
simp only [← Polynomial.eval_map, ← Ideal.Quotient.algebraMap_eq, ← Polynomial.coe_evalRingHom,
745-
← map_pow, pow_expand_A_B_scalar_tower_F]
746-
simp only [Ideal.Quotient.algebraMap_eq, coe_evalRingHom, expand_eval, map_pow]
733+
simp_rw [← Polynomial.eval_map, ← Ideal.Quotient.algebraMap_eq, ← Polynomial.coe_evalRingHom,
734+
← map_pow, pow_expand_A_B_scalar_tower_F, Ideal.Quotient.algebraMap_eq, coe_evalRingHom,
735+
expand_eval, map_pow]
747736

748737
lemma quotient_F_is_product_of_quot :
749738
(Polynomial.map (Ideal.Quotient.mk Q) F) =
750739
∏ τ : L ≃ₐ[K] L, (X - C ((Ideal.Quotient.mk Q) ((galRestrict A K L B τ) α))) := by
751740
rw [← Polynomial.coe_mapRingHom]
752741
erw [map_prod]
753-
simp only [map_sub, coe_mapRingHom, map_X, map_C]
742+
simp [map_sub, coe_mapRingHom, map_X, map_C]
754743

755744
lemma quotient_F_is_root_iff_is_conjugate (x : (B ⧸ Q)) :
756745
IsRoot (Polynomial.map (Ideal.Quotient.mk Q) F) x ↔
757746
(∃ σ : L ≃ₐ[K] L, x = ((Ideal.Quotient.mk Q) ((galRestrict A K L B σ) α))) := by
758747
rw [quotient_F_is_product_of_quot, Polynomial.isRoot_prod]
759748
simp only [Finset.mem_univ, eval_sub, eval_X, eval_C, true_and, Polynomial.root_X_sub_C]
760-
simp only [eq_comm (a := x)]
749+
simp [eq_comm]
761750

762751
lemma pow_eval_root_in_Q : ((eval α F) ^ q) ∈ Q := by
763752
have h : (eval α F) ∈ Q := quotient_isRoot_α Q_ne_bot
@@ -767,16 +756,15 @@ lemma pow_eval_root_in_Q : ((eval α F) ^ q) ∈ Q := by
767756

768757
lemma expand_eval_root_eq_zero :
769758
(eval₂ (Ideal.Quotient.mk Q) (Ideal.Quotient.mk Q (α ^ q)) F) = 0 := by
770-
rw [← F_expand_eval_eq_eval_pow P Q_ne_bot]
771-
simp only [eval₂_at_apply, ne_eq, Fintype.card_ne_zero, not_false_eq_true, pow_eq_zero_iff]
759+
simp only [← F_expand_eval_eq_eval_pow P Q_ne_bot, eval₂_at_apply, ne_eq, Fintype.card_ne_zero,
760+
not_false_eq_true, pow_eq_zero_iff]
772761
have h : eval α F ∈ Q := quotient_isRoot_α Q_ne_bot
773762
rwa [← Ideal.Quotient.eq_zero_iff_mem] at h
774763

775764
-- now, want `∃ σ, α ^ q ≡ σ α mod Q`
776765
lemma pow_q_is_conjugate : ∃ σ : L ≃ₐ[K] L, (Ideal.Quotient.mk Q (α ^ q)) =
777766
(Ideal.Quotient.mk Q ((((galRestrict A K L B) σ)) α)) := by
778-
rw [← quotient_F_is_root_iff_is_conjugate]
779-
simp only [map_pow, IsRoot.def, Polynomial.eval_map]
767+
rw [← quotient_F_is_root_iff_is_conjugate, map_pow, IsRoot.def, Polynomial.eval_map]
780768
exact expand_eval_root_eq_zero P Q_ne_bot
781769

782770
-- following lemma suggested by Amelia
@@ -785,7 +773,7 @@ lemma pow_quotient_IsRoot_α : (eval (α ^ q) F) ∈ Q := by
785773
have h2 : (eval₂ (Ideal.Quotient.mk Q) (Ideal.Quotient.mk Q (α ^ q)) F) = 0 :=
786774
expand_eval_root_eq_zero P Q_ne_bot
787775
convert h2
788-
simp only [eval₂_at_apply]
776+
rw [eval₂_at_apply]
789777

790778
/--`α ^ q ≡ σ • α mod Q` for some `σ : L ≃ₐ[K] L` -/
791779
lemma pow_q_conjugate :
@@ -808,9 +796,8 @@ lemma inv_aut_not_mem_decomp (h : Frob ∉ decomposition_subgroup_Ideal' Q) : (F
808796
lemma gen_zero_mod_inv_aut (h1 : Frob ∉ decomposition_subgroup_Ideal' Q) :
809797
α ∈ (Frob⁻¹ • Q) := by
810798
have inv : Frob⁻¹ ∉ decomposition_subgroup_Ideal' Q := by
811-
simpa only [inv_mem_iff]
812-
apply generator_well_defined
813-
exact inv
799+
simpa [inv_mem_iff]
800+
exact generator_well_defined _ _ inv
814801

815802
lemma prop_Frob : (α ^ q - (galRestrict A K L B Frob) α) ∈ Q :=
816803
Classical.choose_spec (pow_q_conjugate P Q_ne_bot)
@@ -830,11 +817,9 @@ lemma is_zero_pow_gen_mod_Q (h : α ∈ (Frob⁻¹ • Q)) :
830817
rw [← Ideal.neg_mem_iff] at h1
831818
have h3 : ((α ^ q - (galRestrict A K L B Frob) α) -
832819
(-(galRestrict A K L B Frob) α)) ∈ Q := by
833-
apply Ideal.sub_mem
834-
· exact h2
835-
· exact h1
820+
exact Ideal.sub_mem Q h2 h1
836821
convert h3
837-
simp only [sub_neg_eq_add, sub_add_cancel]
822+
simp [sub_neg_eq_add, sub_add_cancel]
838823

839824
/-- `Frob ∈ decomposition_subgroup_Ideal' A K L B Q`. -/
840825
theorem Frob_is_in_decompositionSubgroup :
@@ -852,13 +837,12 @@ lemma γ_not_in_Q_is_pow_gen {γ : B} (h : γ ∉ Q) : ∃ (i : ℕ), γ - (α
852837
let g := Units.mk0 (((Ideal.Quotient.mk Q γ))) <| by
853838
intro h1
854839
rw [Ideal.Quotient.eq_zero_iff_mem] at h1
855-
apply h
856-
exact h1
840+
exact h h1
857841
rcases generator_mem_submonoid_powers Q_ne_bot g with ⟨i, hi⟩
858842
use i
859843
rw [← Ideal.Quotient.mk_eq_mk_iff_sub_mem]
860-
simp only [g, Units.ext_iff, Units.val_pow_eq_pow_val, IsUnit.unit_spec, Units.val_mk0] at hi
861-
simp only [g, map_pow, hi]
844+
simp [g, Units.ext_iff, Units.val_pow_eq_pow_val, IsUnit.unit_spec, Units.val_mk0] at hi
845+
simp [g, map_pow, hi]
862846

863847
/--`i' : ℕ` such that, for `(γ : B) (h : γ ∉ Q)`, `γ - (α ^ i) ∈ Q`. -/
864848
noncomputable def i' {γ : B} (h : γ ∉ Q) : ℕ :=
@@ -874,12 +858,11 @@ lemma eq_pow_gen_apply {γ : B} (h: γ ∉ Q) : (galRestrict A K L B Frob) γ -
874858
galRestrict A K L B Frob (α ^ (i h)) ∈ Q := by
875859
rw [← Ideal.Quotient.mk_eq_mk_iff_sub_mem]
876860
have h1 : γ - (α ^ (i h)) ∈ Q := prop_γ_not_in_Q_is_pow_gen Q_ne_bot h
877-
rw [← Ideal.Quotient.mk_eq_mk_iff_sub_mem] at h1
861+
rw [← Ideal.Quotient.mk_eq_mk_iff_sub_mem, Ideal.Quotient.eq] at h1
878862
rw [Ideal.Quotient.eq, ← map_sub]
879-
rw [Ideal.Quotient.eq] at h1
880863
have := Frob_is_in_decompositionSubgroup P Q_ne_bot
881864
rw [mem_decomposition_iff] at this
882-
apply (this _).1 h1
865+
exact (this _).1 h1
883866

884867
-- γ ∈ B \ Q is α^i mod Q
885868
/-- `Frob • (α ^ i) ≡ α ^ (i * q) mod Q` -/
@@ -916,10 +899,7 @@ theorem Frob_γ_not_in_Q_is_pow {γ : B} (h : γ ∉ Q) :
916899
· exact h3
917900
simp only [map_pow, sub_sub_sub_cancel_right] at h5
918901
have h6 : (( (((galRestrict A K L B) Frob)) γ - α ^ (i h * q)) +
919-
(((α ^ ((i h) * q)) - (γ ^ q)))) ∈ Q := by
920-
apply Ideal.add_mem
921-
· exact h5
922-
· exact h4
902+
(((α ^ ((i h) * q)) - (γ ^ q)))) ∈ Q := Ideal.add_mem _ h5 h4
923903
simp only [sub_add_sub_cancel] at h6
924904
rw [← Ideal.neg_mem_iff] at h6
925905
simp only [neg_sub] at h6

0 commit comments

Comments
 (0)