PhysLean/HepLean/SpaceTime/Basic.lean

105 lines
3.4 KiB
Text
Raw Normal View History

2024-05-14 08:25:03 -04:00
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
2024-07-12 16:39:44 -04:00
Released under Apache 2.0 license as described in the file LICENSE.
2024-05-14 08:25:03 -04:00
Authors: Joseph Tooby-Smith
-/
2024-06-25 07:06:32 -04:00
import Mathlib.Geometry.Manifold.SmoothManifoldWithCorners
import Mathlib.Analysis.InnerProductSpace.PiL2
2024-05-14 08:25:03 -04:00
/-!
# Space time
This file introduce 4d Minkowski spacetime.
-/
noncomputable section
2024-11-12 06:33:58 +00:00
/-! TODO: SpaceTime should be refactored into a structure, to prevent casting. -/
2024-05-14 08:25:03 -04:00
/-- The space-time -/
def SpaceTime : Type := Fin 4 →
2024-05-14 08:25:03 -04:00
/-- Give spacetime the structure of an additive commutative monoid. -/
instance : AddCommMonoid SpaceTime := Pi.addCommMonoid
2024-05-14 08:25:03 -04:00
/-- Give spacetime the structure of a module over the reals. -/
instance : Module SpaceTime := Pi.module _ _ _
2024-05-14 08:25:03 -04:00
2024-11-11 16:55:15 +00:00
/-- The instance of a normed group on spacetime defined via the Euclidean norm. -/
instance euclideanNormedAddCommGroup : NormedAddCommGroup SpaceTime := Pi.normedAddCommGroup
2024-05-14 08:25:03 -04:00
2024-11-11 16:55:15 +00:00
/-- The Euclidean norm-structure on space time. This is used to give it a smooth structure. -/
instance euclideanNormedSpace : NormedSpace SpaceTime := Pi.normedSpace
2024-11-11 16:55:15 +00:00
namespace SpaceTime
2024-05-14 08:25:03 -04:00
open Manifold
open Matrix
open Complex
open ComplexConjugate
/-- The space part of spacetime. -/
@[simp]
def space (x : SpaceTime) : EuclideanSpace (Fin 3) := ![x 1, x 2, x 3]
2024-05-14 08:25:03 -04:00
/-- The structure of a smooth manifold on spacetime. -/
def asSmoothManifold : ModelWithCorners SpaceTime SpaceTime := 𝓘(, SpaceTime)
2024-05-14 08:25:03 -04:00
2024-11-11 16:55:15 +00:00
/-- The instance of a `ChartedSpace` on `SpaceTime`. -/
instance : ChartedSpace SpaceTime SpaceTime := chartedSpaceSelf SpaceTime
2024-05-14 08:27:14 -04:00
2024-05-14 08:25:03 -04:00
/-- The standard basis for spacetime. -/
def stdBasis : Basis (Fin 4) SpaceTime := Pi.basisFun (Fin 4)
2024-05-14 08:25:03 -04:00
lemma stdBasis_apply (μ ν : Fin 4) : stdBasis μ ν = if μ = ν then 1 else 0 := by
2024-09-04 06:28:46 -04:00
erw [stdBasis, Pi.basisFun_apply, Pi.single_apply]
refine ite_congr ?h₁ (congrFun rfl) (congrFun rfl)
exact Eq.propIntro (fun a => id (Eq.symm a)) fun a => id (Eq.symm a)
2024-05-14 08:25:03 -04:00
lemma stdBasis_not_eq {μ ν : Fin 4} (h : μ ≠ ν) : stdBasis μ ν = 0 := by
rw [stdBasis_apply]
exact if_neg h
2024-11-12 06:33:58 +00:00
/-- For space-time,`stdBasis 0` is equal to `![1, 0, 0, 0] `. -/
2024-05-14 08:25:03 -04:00
lemma stdBasis_0 : stdBasis 0 = ![1, 0, 0, 0] := by
funext i
2024-08-20 14:04:52 +02:00
fin_cases i <;> simp [stdBasis_apply]
2024-05-14 08:25:03 -04:00
2024-11-12 06:33:58 +00:00
/-- For space-time,`stdBasis 1` is equal to `![0, 1, 0, 0] `. -/
2024-05-14 08:25:03 -04:00
lemma stdBasis_1 : stdBasis 1 = ![0, 1, 0, 0] := by
funext i
2024-08-20 14:04:52 +02:00
fin_cases i <;> simp [stdBasis_apply]
2024-05-14 08:25:03 -04:00
2024-11-12 06:33:58 +00:00
/-- For space-time,`stdBasis 2` is equal to `![0, 0, 1, 0] `. -/
2024-05-14 08:25:03 -04:00
lemma stdBasis_2 : stdBasis 2 = ![0, 0, 1, 0] := by
funext i
2024-08-20 14:04:52 +02:00
fin_cases i <;> simp [stdBasis_apply]
2024-05-14 08:25:03 -04:00
2024-11-12 06:33:58 +00:00
/-- For space-time,`stdBasis 3` is equal to `![0, 0, 0, 1] `. -/
2024-05-14 08:25:03 -04:00
lemma stdBasis_3 : stdBasis 3 = ![0, 0, 0, 1] := by
funext i
2024-08-20 14:04:52 +02:00
fin_cases i <;> simp [stdBasis_apply]
2024-05-14 08:25:03 -04:00
lemma stdBasis_mulVec (μ ν : Fin 4) (Λ : Matrix (Fin 4) (Fin 4) ) :
(Λ *ᵥ stdBasis μ) ν = Λ ν μ := by
rw [mulVec, dotProduct, Fintype.sum_eq_single μ, stdBasis_apply]
· simp only [↓reduceIte, mul_one]
· intro x h
rw [stdBasis_apply, if_neg (Ne.symm h)]
exact CommMonoidWithZero.mul_zero (Λ ν x)
2024-05-14 08:25:03 -04:00
2024-11-12 06:33:58 +00:00
/-- The explicit expansion of a point in spacetime as `![x 0, x 1, x 2, x 3]`. -/
lemma explicit (x : SpaceTime) : x = ![x 0, x 1, x 2, x 3] := by
2024-05-14 08:25:03 -04:00
funext i
fin_cases i <;> rfl
@[simp]
lemma add_apply (x y : SpaceTime) (i : Fin 4) : (x + y) i = x i + y i := rfl
@[simp]
lemma smul_apply (x : SpaceTime) (a : ) (i : Fin 4) : (a • x) i = a * x i := rfl
2024-05-14 08:25:03 -04:00
end SpaceTime
2024-05-14 08:25:03 -04:00
end