feat: replace FLT sorry with assumptions
This commit is contained in:
parent
388c824463
commit
70cafbf41e
2 changed files with 14 additions and 10 deletions
|
@ -54,7 +54,8 @@ lemma accGrav_Q_zero {S : (SMNoGrav 1).Sols} (hQ : Q S.val (0 : Fin 1) = 0) :
|
|||
simp_all
|
||||
linear_combination 3 * h2
|
||||
|
||||
lemma accGrav_Q_neq_zero {S : (SMNoGrav 1).Sols} (hQ : Q S.val (0 : Fin 1) ≠ 0) :
|
||||
lemma accGrav_Q_neq_zero {S : (SMNoGrav 1).Sols} (hQ : Q S.val (0 : Fin 1) ≠ 0)
|
||||
(FLTThree : FermatLastTheoremWith ℚ 3) :
|
||||
accGrav S.val = 0 := by
|
||||
have hE := E_zero_iff_Q_zero.mpr.mt hQ
|
||||
let S' := linearParametersQENeqZero.bijection.symm ⟨S.1.1, And.intro hQ hE⟩
|
||||
|
@ -64,13 +65,14 @@ lemma accGrav_Q_neq_zero {S : (SMNoGrav 1).Sols} (hQ : Q S.val (0 : Fin 1) ≠ 0
|
|||
change _ = S.val at hS'
|
||||
rw [← hS'] at hC
|
||||
rw [← hS']
|
||||
exact S'.grav_of_cubic hC
|
||||
exact S'.grav_of_cubic hC FLTThree
|
||||
|
||||
/-- Any solution to the ACCs without gravity satifies the gravitational ACC. -/
|
||||
theorem accGravSatisfied {S : (SMNoGrav 1).Sols} : accGrav S.val = 0 := by
|
||||
theorem accGravSatisfied {S : (SMNoGrav 1).Sols} (FLTThree : FermatLastTheoremWith ℚ 3) :
|
||||
accGrav S.val = 0 := by
|
||||
by_cases hQ : Q S.val (0 : Fin 1)= 0
|
||||
exact accGrav_Q_zero hQ
|
||||
exact accGrav_Q_neq_zero hQ
|
||||
exact accGrav_Q_neq_zero hQ FLTThree
|
||||
|
||||
|
||||
end One
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue