Skip to content

Commit

Permalink
Prove Mbar_deg, Mbar_monic (#151)
Browse files Browse the repository at this point in the history
  • Loading branch information
Ruben-VandeVelde authored Oct 2, 2024
1 parent 078603f commit ea954b8
Showing 1 changed file with 14 additions and 4 deletions.
18 changes: 14 additions & 4 deletions FLT/MathlibExperiments/FrobeniusRiou.lean
Original file line number Diff line number Diff line change
Expand Up @@ -494,10 +494,20 @@ noncomputable def Mbar

variable (hFull' : ∀ (b : B), (∀ (g : G), g • b = b) → ∃ a : A, b = a)

theorem Mbar_deg (bbar : B ⧸ Q) : degree (Mbar P hFull' bbar) = Nat.card G := sorry

theorem Mbar_monic (bbar : B ⧸ Q) : (Mbar P hFull' bbar).Monic := by
sorry
omit [SMulCommClass G A B] [Q.IsPrime] [P.IsPrime] [Algebra (A ⧸ P) (B ⧸ Q)]
[IsScalarTower A (A ⧸ P) (B ⧸ Q)] in
theorem Mbar_monic [Nontrivial B] (bbar : B ⧸ Q) : (Mbar P hFull' bbar).Monic := by
have := M_monic hFull'
simp [Mbar, (M_monic hFull' _).map]

omit [SMulCommClass G A B] [Q.IsPrime] [Algebra (A ⧸ P) (B ⧸ Q)] [IsScalarTower A (A ⧸ P) (B ⧸ Q)] in
theorem Mbar_deg [Nontrivial A] [Nontrivial B] (bbar : B ⧸ Q) :
degree (Mbar P hFull' bbar) = Nat.card G := by
simp only [Mbar]
rw [degree_map_eq_of_leadingCoeff_ne_zero]
· exact M_deg hFull' _
· rw [(M_monic hFull' _).leadingCoeff]
simp only [map_one, ne_eq, one_ne_zero, not_false_eq_true]

theorem Mbar_eval_eq_zero (bbar : B ⧸ Q) : eval₂ (algebraMap (A ⧸ P) (B ⧸ Q)) bbar (Mbar P hFull' bbar) = 0 := by
sorry
Expand Down

0 comments on commit ea954b8

Please sign in to comment.