diff --git a/HepLean/SpaceTime/LorentzGroup/Basic.lean b/HepLean/SpaceTime/LorentzGroup/Basic.lean index 30d2b82..5c737b5 100644 --- a/HepLean/SpaceTime/LorentzGroup/Basic.lean +++ b/HepLean/SpaceTime/LorentzGroup/Basic.lean @@ -89,7 +89,6 @@ lemma mem_iff_self_mul_dual : Λ ∈ LorentzGroup d ↔ Λ * dual Λ = 1 := by rw [mem_iff_dual_mul_self] exact mul_eq_one_comm - lemma mem_iff_transpose : Λ ∈ LorentzGroup d ↔ Λᵀ ∈ LorentzGroup d := by apply Iff.intro · intro h @@ -236,7 +235,6 @@ lemma toGL_embedding : Embedding (@toGL d).toFun where rw [isOpen_induced_iff, isOpen_induced_iff] exact exists_exists_and_eq_and - instance : TopologicalGroup (LorentzGroup d) := Inducing.topologicalGroup toGL toGL_embedding.toInducing @@ -276,9 +274,5 @@ lemma timeComp_mul (Λ Λ' : LorentzGroup d) : timeComp (Λ * Λ') = erw [Pi.basisFun_apply, mulVec_stdBasis] simp - - - - end end LorentzGroup diff --git a/HepLean/SpaceTime/LorentzVector/Basic.lean b/HepLean/SpaceTime/LorentzVector/Basic.lean index 45543c2..c3d2a5a 100644 --- a/HepLean/SpaceTime/LorentzVector/Basic.lean +++ b/HepLean/SpaceTime/LorentzVector/Basic.lean @@ -65,7 +65,7 @@ scoped[LorentzVector] notation "e" => stdBasis lemma stdBasis_apply (μ ν : Fin 1 ⊕ Fin d) : e μ ν = if μ = ν then 1 else 0 := by rw [stdBasis] erw [Pi.basisFun_apply] - simp + exact LinearMap.stdBasis_apply' ℝ μ ν /-- The standard unit time vector. -/ 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 funext i rcases i with i | i - · simp only [Sum.elim_inl] - apply Eq.refl + · rfl · simp only [Sum.elim_inr, Pi.neg_apply] apply neg_add map_smul' c x := by funext i rcases i with i | i - · simp only [Sum.elim_inl, Pi.smul_apply] - apply smul_eq_mul - · simp [ HSMul.hSMul, SMul.smul] + · rfl + · simp [HSMul.hSMul, SMul.smul] /-- The reflection of space. -/ @[simp] diff --git a/HepLean/SpaceTime/LorentzVector/NormOne.lean b/HepLean/SpaceTime/LorentzVector/NormOne.lean index 532d8dd..2866976 100644 --- a/HepLean/SpaceTime/LorentzVector/NormOne.lean +++ b/HepLean/SpaceTime/LorentzVector/NormOne.lean @@ -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) : ⟪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 apply le_of_le_of_eq h1 ?_ simp [neg] @@ -207,7 +207,7 @@ open LorentzVector @[simps!] noncomputable def timeVecNormOneFuture : FuturePointing d := ⟨⟨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. -/ @@ -222,15 +222,12 @@ noncomputable def pathFromTime (u : FuturePointing d) : Path timeVecNormOneFutur conj_trivial] rw [Real.mul_self_sqrt, ← @real_inner_self_eq_norm_sq, @PiLp.inner_apply] 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)) = - t ^ 2 * ∑ x : Fin d, u.1.1 (Sum.inr x) * u.1.1 (Sum.inr x) := by - rw [Finset.mul_sum] - apply Finset.sum_congr rfl - intro i _ - ring - rw [h1] + refine Eq.symm (eq_sub_of_add_eq (congrArg (HAdd.hAdd 1) ?_)) + rw [Finset.mul_sum] + apply Finset.sum_congr rfl + intro i _ 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 simp only [space, Function.comp_apply, mem_iff_time_nonneg, time, Real.sqrt_pos] exact Real.sqrt_nonneg _⟩