refactor: Minor golfing
This commit is contained in:
parent
304c3542b5
commit
a4afeba3cd
3 changed files with 11 additions and 22 deletions
|
@ -89,7 +89,6 @@ lemma mem_iff_self_mul_dual : Λ ∈ LorentzGroup d ↔ Λ * dual Λ = 1 := by
|
||||||
rw [mem_iff_dual_mul_self]
|
rw [mem_iff_dual_mul_self]
|
||||||
exact mul_eq_one_comm
|
exact mul_eq_one_comm
|
||||||
|
|
||||||
|
|
||||||
lemma mem_iff_transpose : Λ ∈ LorentzGroup d ↔ Λᵀ ∈ LorentzGroup d := by
|
lemma mem_iff_transpose : Λ ∈ LorentzGroup d ↔ Λᵀ ∈ LorentzGroup d := by
|
||||||
apply Iff.intro
|
apply Iff.intro
|
||||||
· intro h
|
· intro h
|
||||||
|
@ -236,7 +235,6 @@ lemma toGL_embedding : Embedding (@toGL d).toFun where
|
||||||
rw [isOpen_induced_iff, isOpen_induced_iff]
|
rw [isOpen_induced_iff, isOpen_induced_iff]
|
||||||
exact exists_exists_and_eq_and
|
exact exists_exists_and_eq_and
|
||||||
|
|
||||||
|
|
||||||
instance : TopologicalGroup (LorentzGroup d) :=
|
instance : TopologicalGroup (LorentzGroup d) :=
|
||||||
Inducing.topologicalGroup toGL toGL_embedding.toInducing
|
Inducing.topologicalGroup toGL toGL_embedding.toInducing
|
||||||
|
|
||||||
|
@ -276,9 +274,5 @@ lemma timeComp_mul (Λ Λ' : LorentzGroup d) : timeComp (Λ * Λ') =
|
||||||
erw [Pi.basisFun_apply, mulVec_stdBasis]
|
erw [Pi.basisFun_apply, mulVec_stdBasis]
|
||||||
simp
|
simp
|
||||||
|
|
||||||
|
|
||||||
|
|
||||||
|
|
||||||
|
|
||||||
end
|
end
|
||||||
end LorentzGroup
|
end LorentzGroup
|
||||||
|
|
|
@ -65,7 +65,7 @@ scoped[LorentzVector] notation "e" => stdBasis
|
||||||
lemma stdBasis_apply (μ ν : Fin 1 ⊕ Fin d) : e μ ν = if μ = ν then 1 else 0 := by
|
lemma stdBasis_apply (μ ν : Fin 1 ⊕ Fin d) : e μ ν = if μ = ν then 1 else 0 := by
|
||||||
rw [stdBasis]
|
rw [stdBasis]
|
||||||
erw [Pi.basisFun_apply]
|
erw [Pi.basisFun_apply]
|
||||||
simp
|
exact LinearMap.stdBasis_apply' ℝ μ ν
|
||||||
|
|
||||||
/-- The standard unit time vector. -/
|
/-- The standard unit time vector. -/
|
||||||
noncomputable abbrev timeVec : (LorentzVector d) := e (Sum.inl 0)
|
noncomputable abbrev timeVec : (LorentzVector d) := e (Sum.inl 0)
|
||||||
|
@ -91,16 +91,14 @@ def spaceReflectionLin : LorentzVector d →ₗ[ℝ] LorentzVector d where
|
||||||
map_add' x y := by
|
map_add' x y := by
|
||||||
funext i
|
funext i
|
||||||
rcases i with i | i
|
rcases i with i | i
|
||||||
· simp only [Sum.elim_inl]
|
· rfl
|
||||||
apply Eq.refl
|
|
||||||
· simp only [Sum.elim_inr, Pi.neg_apply]
|
· simp only [Sum.elim_inr, Pi.neg_apply]
|
||||||
apply neg_add
|
apply neg_add
|
||||||
map_smul' c x := by
|
map_smul' c x := by
|
||||||
funext i
|
funext i
|
||||||
rcases i with i | i
|
rcases i with i | i
|
||||||
· simp only [Sum.elim_inl, Pi.smul_apply]
|
· rfl
|
||||||
apply smul_eq_mul
|
· simp [HSMul.hSMul, SMul.smul]
|
||||||
· simp [ HSMul.hSMul, SMul.smul]
|
|
||||||
|
|
||||||
/-- The reflection of space. -/
|
/-- The reflection of space. -/
|
||||||
@[simp]
|
@[simp]
|
||||||
|
|
|
@ -184,7 +184,7 @@ lemma metric_reflect_mem_not_mem (h : v ∈ FuturePointing d) (hw : w ∉ Future
|
||||||
|
|
||||||
lemma metric_reflect_not_mem_mem (h : v ∉ FuturePointing d) (hw : w ∈ FuturePointing d) :
|
lemma metric_reflect_not_mem_mem (h : v ∉ FuturePointing d) (hw : w ∈ FuturePointing d) :
|
||||||
⟪v.1, w.1.spaceReflection⟫ₘ ≤ 0 := by
|
⟪v.1, w.1.spaceReflection⟫ₘ ≤ 0 := by
|
||||||
rw [show (0 : ℝ) = - 0 by simp, le_neg]
|
rw [show (0 : ℝ) = - 0 from zero_eq_neg.mpr rfl, le_neg]
|
||||||
have h1 := metric_reflect_mem_mem ((not_mem_iff_neg v).mp h) hw
|
have h1 := metric_reflect_mem_mem ((not_mem_iff_neg v).mp h) hw
|
||||||
apply le_of_le_of_eq h1 ?_
|
apply le_of_le_of_eq h1 ?_
|
||||||
simp [neg]
|
simp [neg]
|
||||||
|
@ -207,7 +207,7 @@ open LorentzVector
|
||||||
@[simps!]
|
@[simps!]
|
||||||
noncomputable def timeVecNormOneFuture : FuturePointing d := ⟨⟨timeVec, by
|
noncomputable def timeVecNormOneFuture : FuturePointing d := ⟨⟨timeVec, by
|
||||||
rw [NormOneLorentzVector.mem_iff, on_timeVec]⟩, by
|
rw [NormOneLorentzVector.mem_iff, on_timeVec]⟩, by
|
||||||
rw [mem_iff, timeVec_time]; simp⟩
|
rw [mem_iff, timeVec_time]; exact Real.zero_lt_one⟩
|
||||||
|
|
||||||
|
|
||||||
/-- A continuous path from `timeVecNormOneFuture` to any other. -/
|
/-- A continuous path from `timeVecNormOneFuture` to any other. -/
|
||||||
|
@ -222,15 +222,12 @@ noncomputable def pathFromTime (u : FuturePointing d) : Path timeVecNormOneFutur
|
||||||
conj_trivial]
|
conj_trivial]
|
||||||
rw [Real.mul_self_sqrt, ← @real_inner_self_eq_norm_sq, @PiLp.inner_apply]
|
rw [Real.mul_self_sqrt, ← @real_inner_self_eq_norm_sq, @PiLp.inner_apply]
|
||||||
simp only [Function.comp_apply, RCLike.inner_apply, conj_trivial]
|
simp only [Function.comp_apply, RCLike.inner_apply, conj_trivial]
|
||||||
have h1 : ∑ x : Fin d, t.1 * u.1.1 (Sum.inr x) * (↑t.1 * u.1.1 (Sum.inr x)) =
|
refine Eq.symm (eq_sub_of_add_eq (congrArg (HAdd.hAdd 1) ?_))
|
||||||
t ^ 2 * ∑ x : Fin d, u.1.1 (Sum.inr x) * u.1.1 (Sum.inr x) := by
|
rw [Finset.mul_sum]
|
||||||
rw [Finset.mul_sum]
|
apply Finset.sum_congr rfl
|
||||||
apply Finset.sum_congr rfl
|
intro i _
|
||||||
intro i _
|
|
||||||
ring
|
|
||||||
rw [h1]
|
|
||||||
ring
|
ring
|
||||||
refine Right.add_nonneg (zero_le_one' ℝ) $ mul_nonneg (sq_nonneg _) (sq_nonneg _) ⟩,
|
exact Right.add_nonneg (zero_le_one' ℝ) $ mul_nonneg (sq_nonneg _) (sq_nonneg _) ⟩,
|
||||||
by
|
by
|
||||||
simp only [space, Function.comp_apply, mem_iff_time_nonneg, time, Real.sqrt_pos]
|
simp only [space, Function.comp_apply, mem_iff_time_nonneg, time, Real.sqrt_pos]
|
||||||
exact Real.sqrt_nonneg _⟩
|
exact Real.sqrt_nonneg _⟩
|
||||||
|
|
Loading…
Add table
Add a link
Reference in a new issue