refactor: Lint

This commit is contained in:
jstoobysmith 2024-09-06 05:43:52 -04:00
parent 27721bc476
commit 5a613b4de6
4 changed files with 13 additions and 10 deletions

View file

@ -59,6 +59,8 @@ lemma smooth_toFin2 : Smooth 𝓘(, HiggsVec) 𝓘(, Fin 2 → ) toF
def orthonormBasis : OrthonormalBasis (Fin 2) HiggsVec :=
EuclideanSpace.basisFun (Fin 2)
/-- Generating a Higgs vector from a real number, such that the norm-squared of that Higgs vector
is the given real number. -/
def ofReal (a : ) : HiggsVec :=
![Real.sqrt a, 0]
@ -164,6 +166,8 @@ lemma isConst_of_higgsVec (φ : HiggsVec) : φ.toField.IsConst := fun _ => congr
lemma isConst_iff_of_higgsVec (Φ : HiggsField) : Φ.IsConst ↔ ∃ (φ : HiggsVec), Φ = φ.toField :=
Iff.intro (fun h ↦ ⟨Φ 0, by ext x y; rw [← h x 0]; rfl⟩) (fun ⟨φ, hφ⟩ x y ↦ by subst hφ; rfl)
/-- Generating a constant Higgs field from a real number, such that the norm-squared of that Higgs
vector is the given real number. -/
def ofReal (a : ) : HiggsField := (HiggsVec.ofReal a).toField
end HiggsField