refactor: Lint

This commit is contained in:
jstoobysmith 2025-03-04 09:08:21 +00:00
parent e22483c780
commit 75d864df77
16 changed files with 34 additions and 42 deletions

View file

@ -10,7 +10,6 @@ Authors: Joseph Tooby-Smith
This directory is currently a place holder.
Please feel free to contribute!
Some directories which are NOT currently place holders are:
- Mathematics
- Meta
@ -19,5 +18,4 @@ Some directories which are NOT currently place holders are:
- Quantum Mechanics
- Relativity
-/

View file

@ -18,5 +18,4 @@ Some directories which are NOT currently place holders are:
- Quantum Mechanics
- Relativity
-/

View file

@ -113,7 +113,7 @@ lemma smooth_innerProd (φ1 φ2 : HiggsField) : ContMDiff 𝓘(, SpaceTime)
the function `SpaceTime → ` obtained by taking the square norm of the
pointwise Higgs vector. In other words, `normSq φ x = ‖φ x‖ ^ 2`.
The notation `‖φ‖_H^2` is used for the `normSq φ` -/
The notation `‖φ‖_H^2` is used for the `normSq φ`. -/
@[simp]
def normSq (φ : HiggsField) : SpaceTime → := fun x => ‖φ x‖ ^ 2

View file

@ -134,7 +134,7 @@ lemma one_over_ξ : 1/Q.ξ = √(Q.m * Q.ω / Q.ℏ) := by
have := Q.hℏ
field_simp [ξ]
lemma ξ_inverse : Q.ξ⁻¹ = √(Q.m * Q.ω / Q.ℏ):= by
lemma ξ_inverse : Q.ξ⁻¹ = √(Q.m * Q.ω / Q.ℏ) := by
rw [inv_eq_one_div]
exact one_over_ξ Q
@ -182,7 +182,7 @@ lemma schrodingerOperator_eq (ψ : ) :
/-- The schrodinger operator written in terms of `ξ`. -/
lemma schrodingerOperator_eq_ξ (ψ : ) : Q.schrodingerOperator ψ =
fun x => (Q.ℏ ^2 / (2 * Q.m)) * (- (deriv (deriv ψ) x) + (1/Q.ξ^2 )^2 * x^2 * ψ x) := by
fun x => (Q.ℏ ^2 / (2 * Q.m)) * (- (deriv (deriv ψ) x) + (1/Q.ξ^2)^2 * x^2 * ψ x) := by
funext x
simp [schrodingerOperator_eq, ξ_sq, ξ_inverse, ξ_ne_zero, ξ_pos, ξ_abs, ← Complex.ofReal_pow]
have hm' := Complex.ofReal_ne_zero.mpr Q.m_ne_zero

View file

@ -175,7 +175,7 @@ lemma orthogonal_physHermite_of_mem_orthogonal (f : ) (hf : MemHS f)
or_false, Real.sqrt_nonneg] at h1
have h0 : n ! ≠ 0 := by
exact factorial_ne_zero n
have h0' : ¬ (√Q.ξ = 0 √Real.pi = 0):= by
have h0' : ¬ (√Q.ξ = 0 √Real.pi = 0) := by
simpa using And.intro (Real.sqrt_ne_zero'.mpr Q.ξ_pos) (Real.sqrt_ne_zero'.mpr Real.pi_pos)
simp only [h0, h0', or_self, false_or] at h1
rw [← h1]

View file

@ -43,7 +43,7 @@ lemma eigenfunction_eq (n : ) :
ring
lemma eigenfunction_zero : Q.eigenfunction 0 = fun (x : ) =>
(1/ √(√Real.pi * Q.ξ)) * Complex.exp (- x^2 / (2 * Q.ξ^2)):= by
(1/ √(√Real.pi * Q.ξ)) * Complex.exp (- x^2 / (2 * Q.ξ^2)) := by
funext x
simp [eigenfunction]

View file

@ -158,7 +158,6 @@ lemma deriv_deriv_eigenfunction_zero (x : ) : deriv (deriv (Q.eigenfunction 0
simp only [Complex.ofReal_one, Complex.ofReal_pow, one_mul, one_pow, inv_pow]
ring
lemma deriv_deriv_eigenfunction_succ (n : ) (x : ) :
deriv (fun x => deriv (Q.eigenfunction (n + 1)) x) x =
Complex.ofReal (1/√(2 ^ (n + 1) * (n + 1) !) * (1/Q.ξ)) *
@ -195,7 +194,7 @@ lemma deriv_deriv_eigenfunction (n : ) (x : ) :
match n with
| 0 => simpa using Q.deriv_deriv_eigenfunction_zero x
| n + 1 =>
trans Complex.ofReal (1/Real.sqrt (2 ^ (n + 1) * (n + 1) !) ) *
trans Complex.ofReal (1/Real.sqrt (2 ^ (n + 1) * (n + 1) !)) *
(((- 1 / Q.ξ ^ 2) * (2 * (n + 1)
+ (1 + (- 1/ Q.ξ ^ 2) * x ^ 2)) *
(physHermite (n + 1) (x/Q.ξ))) * Q.eigenfunction 0 x)
@ -223,7 +222,7 @@ lemma deriv_deriv_eigenfunction (n : ) (x : ) :
trans (- 1/ Q.ξ^2) * (2 * (n + 1) *
(2 * ((1 / Q.ξ) * x) * (physHermite n (x/Q.ξ)) -
2 * n * (physHermite (n - 1) (x/Q.ξ)))
+ (1 + (- 1 / Q.ξ^2) * x ^ 2) * (physHermite (n + 1) ( x/Q.ξ)))
+ (1 + (- 1 / Q.ξ^2) * x ^ 2) * (physHermite (n + 1) (x/Q.ξ)))
· ring_nf
trans (- 1 / Q.ξ^2) * (2 * (n + 1) * (physHermite (n + 1) (x/Q.ξ))
+ (1 + (- 1/ Q.ξ^2) * x ^ 2) * (physHermite (n + 1) (x/Q.ξ)))

View file

@ -335,9 +335,7 @@ lemma _root_.LorentzGroup.mem_iff_norm : Λ ∈ LorentzGroup d ↔
apply e.injective
have hp' := e.injective.eq_iff.mpr hp
have hn' := e.injective.eq_iff.mpr hn
simp [Action.instMonoidalCategory_tensorUnit_V, Action.instMonoidalCategory_tensorObj_V,
Equivalence.symm_inverse, Action.functorCategoryEquivalence_functor,
Action.FunctorCategoryEquivalence.functor_obj_obj, map_add, map_sub] at hp' hn'
simp only [Action.instMonoidalCategory_tensorUnit_V, map_add, map_sub] at hp' hn'
linear_combination (norm := ring_nf) (1 / 4) * hp' + (-1/ 4) * hn'
rw [symm (Λ *ᵥ y) (Λ *ᵥ x), symm y x]
simp only [Action.instMonoidalCategory_tensorUnit_V]

View file

@ -72,8 +72,7 @@ $$\begin{align}
\begin{vmatrix}
\operatorname{Re}(x\bar{y}) & -\operatorname{Im}(x\bar{y}) \\
\operatorname{Im}(x\bar{y}) & \operatorname{Re}(x\bar{y})
\end{vmatrix} \det\left(
\begin{bmatrix} \lvert x\rvert^2 & □ \\ 0 & \lvert y\rvert^2 \end{bmatrix} -
\end{vmatrix} \det\left(\begin{bmatrix} \lvert x\rvert^2 & □ \\ 0 & \lvert y\rvert^2 \end{bmatrix} -
\begin{bmatrix} □ & □ \\ 0 & 0 \end{bmatrix}
\begin{bmatrix} □ & □ \\ □ & □ \end{bmatrix}
\begin{bmatrix} 0 & □ \\ 0 & □ \end{bmatrix}

View file

@ -190,8 +190,7 @@ def altRightRightUnit : 𝟙_ (Rep SL(2,)) ⟶ altRightHanded ⊗ rightHa
refine ModuleCat.hom_ext ?_
refine LinearMap.ext fun x : => ?_
simp only [Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V,
Action.tensorUnit_ρ
, CategoryTheory.Category.id_comp, Action.tensor_ρ, ModuleCat.hom_comp,
Action.tensorUnit_ρ, CategoryTheory.Category.id_comp, Action.tensor_ρ, ModuleCat.hom_comp,
Function.comp_apply]
change x • altRightRightUnitVal =
(TensorProduct.map (altRightHanded.ρ M) (rightHanded.ρ M)) (x • altRightRightUnitVal)

View file

@ -31,9 +31,9 @@ def PhysLeanTextLinter : Type := Array String → Array (String × × )
/-- Checks if there are two consecutive empty lines. -/
def doubleEmptyLineLinter : PhysLeanTextLinter := fun lines ↦ Id.run do
let enumLines := (lines.toList.enumFrom 1)
let enumLines := (lines.toList.zipIdx 1)
let pairLines := List.zip enumLines (List.tail! enumLines)
let errors := pairLines.filterMap (fun ((lno1, l1), _, l2) ↦
let errors := pairLines.filterMap (fun ((l1, lno1), l2, _) ↦
if l1.length == 0 && l2.length == 0 then
some (s!" Double empty line. ", lno1, 1)
else none)
@ -41,8 +41,8 @@ def doubleEmptyLineLinter : PhysLeanTextLinter := fun lines ↦ Id.run do
/-- Checks if there is a double space in the line, which is not at the start. -/
def doubleSpaceLinter : PhysLeanTextLinter := fun lines ↦ Id.run do
let enumLines := (lines.toList.enumFrom 1)
let errors := enumLines.filterMap (fun (lno, l) ↦
let enumLines := (lines.toList.zipIdx 1)
let errors := enumLines.filterMap (fun (l, lno) ↦
if String.containsSubstr l.trimLeft " " then
let k := (Substring.findAllSubstr l " ").toList.getLast?
let col := match k with
@ -53,8 +53,8 @@ def doubleSpaceLinter : PhysLeanTextLinter := fun lines ↦ Id.run do
errors.toArray
def longLineLinter : PhysLeanTextLinter := fun lines ↦ Id.run do
let enumLines := (lines.toList.enumFrom 1)
let errors := enumLines.filterMap (fun (lno, l) ↦
let enumLines := (lines.toList.zipIdx 1)
let errors := enumLines.filterMap (fun (l, lno) ↦
if l.length > 100 ∧ ¬ String.containsSubstr l "http" then
some (s!" Line is too long.", lno, 100)
else none)
@ -62,8 +62,8 @@ def longLineLinter : PhysLeanTextLinter := fun lines ↦ Id.run do
/-- Substring linter. -/
def substringLinter (s : String) : PhysLeanTextLinter := fun lines ↦ Id.run do
let enumLines := (lines.toList.enumFrom 1)
let errors := enumLines.filterMap (fun (lno, l) ↦
let enumLines := (lines.toList.zipIdx 1)
let errors := enumLines.filterMap (fun (l, lno) ↦
if String.containsSubstr l s then
let k := (Substring.findAllSubstr l s).toList.getLast?
let col := match k with
@ -74,8 +74,8 @@ def substringLinter (s : String) : PhysLeanTextLinter := fun lines ↦ Id.run do
errors.toArray
def endLineLinter (s : String) : PhysLeanTextLinter := fun lines ↦ Id.run do
let enumLines := (lines.toList.enumFrom 1)
let errors := enumLines.filterMap (fun (lno, l) ↦
let enumLines := (lines.toList.zipIdx 1)
let errors := enumLines.filterMap (fun (l, lno) ↦
if l.endsWith s then
some (s!" Line ends with `{s}`.", lno, l.length)
else none)
@ -83,8 +83,8 @@ def endLineLinter (s : String) : PhysLeanTextLinter := fun lines ↦ Id.run do
/-- Number of space at new line must be even. -/
def numInitialSpacesEven : PhysLeanTextLinter := fun lines ↦ Id.run do
let enumLines := (lines.toList.enumFrom 1)
let errors := enumLines.filterMap (fun (lno, l) ↦
let enumLines := (lines.toList.zipIdx 1)
let errors := enumLines.filterMap (fun (l, lno) ↦
let numSpaces := (l.takeWhile (· == ' ')).length
if numSpaces % 2 != 0 then
some (s!"Number of initial spaces is not even.", lno, 1)