refactor: Change case of type and props

This commit is contained in:
jstoobysmith 2024-06-26 11:54:02 -04:00
parent 18b83f582e
commit f7a638d32e
58 changed files with 695 additions and 696 deletions

View file

@ -14,7 +14,7 @@ and the vector space of 2×2-complex self-adjoint matrices.
In this file we define this linear equivalence in `toSelfAdjointMatrix`.
-/
namespace spaceTime
namespace SpaceTime
open Matrix
open MatrixGroups
@ -22,11 +22,11 @@ open Complex
/-- A 2×2-complex matrix formed from a space-time point. -/
@[simp]
def toMatrix (x : spaceTime) : Matrix (Fin 2) (Fin 2) :=
def toMatrix (x : SpaceTime) : Matrix (Fin 2) (Fin 2) :=
!![x 0 + x 3, x 1 - x 2 * I; x 1 + x 2 * I, x 0 - x 3]
/-- The matrix `x.toMatrix` for `x ∈ spaceTime` is self adjoint. -/
lemma toMatrix_isSelfAdjoint (x : spaceTime) : IsSelfAdjoint x.toMatrix := by
lemma toMatrix_isSelfAdjoint (x : SpaceTime) : IsSelfAdjoint x.toMatrix := by
rw [isSelfAdjoint_iff, star_eq_conjTranspose, ← Matrix.ext_iff]
intro i j
fin_cases i <;> fin_cases j <;>
@ -35,17 +35,17 @@ lemma toMatrix_isSelfAdjoint (x : spaceTime) : IsSelfAdjoint x.toMatrix := by
/-- A self-adjoint matrix formed from a space-time point. -/
@[simps!]
def toSelfAdjointMatrix' (x : spaceTime) : selfAdjoint (Matrix (Fin 2) (Fin 2) ) :=
def toSelfAdjointMatrix' (x : SpaceTime) : selfAdjoint (Matrix (Fin 2) (Fin 2) ) :=
⟨x.toMatrix, toMatrix_isSelfAdjoint x⟩
/-- A self-adjoint matrix formed from a space-time point. -/
@[simp]
noncomputable def fromSelfAdjointMatrix' (x : selfAdjoint (Matrix (Fin 2) (Fin 2) )) : spaceTime :=
noncomputable def fromSelfAdjointMatrix' (x : selfAdjoint (Matrix (Fin 2) (Fin 2) )) : SpaceTime :=
![1/2 * (x.1 0 0 + x.1 1 1).re, (x.1 1 0).re, (x.1 1 0).im , (x.1 0 0 - x.1 1 1).re/2]
/-- The linear equivalence between the vector-space `spaceTime` and self-adjoint
2×2-complex matrices. -/
noncomputable def toSelfAdjointMatrix : spaceTime ≃ₗ[] selfAdjoint (Matrix (Fin 2) (Fin 2) ) where
noncomputable def toSelfAdjointMatrix : SpaceTime ≃ₗ[] selfAdjoint (Matrix (Fin 2) (Fin 2) ) where
toFun := toSelfAdjointMatrix'
invFun := fromSelfAdjointMatrix'
left_inv x := by
@ -53,7 +53,7 @@ noncomputable def toSelfAdjointMatrix : spaceTime ≃ₗ[] selfAdjoint (Matri
cons_val_zero, empty_val', cons_val_fin_one, cons_val_one, head_cons, head_fin_const,
add_add_sub_cancel, add_re, ofReal_re, mul_re, I_re, mul_zero, ofReal_im, I_im, mul_one,
sub_self, add_zero, add_im, mul_im, zero_add, add_sub_sub_cancel, half_add_self]
field_simp [spaceTime]
field_simp [SpaceTime]
ext1 x
fin_cases x <;> rfl
right_inv x := by
@ -83,10 +83,10 @@ noncomputable def toSelfAdjointMatrix : spaceTime ≃ₗ[] selfAdjoint (Matri
field_simp [fromSelfAdjointMatrix', toMatrix, conj_ofReal, smul_apply]
<;> ring
lemma det_eq_ηLin (x : spaceTime) : det (toSelfAdjointMatrix x).1 = ηLin x x := by
lemma det_eq_ηLin (x : SpaceTime) : det (toSelfAdjointMatrix x).1 = ηLin x x := by
simp [toSelfAdjointMatrix, ηLin_expand]
ring_nf
simp only [Fin.isValue, I_sq, mul_neg, mul_one]
ring
end spaceTime
end SpaceTime

View file

@ -17,21 +17,21 @@ This file introduce 4d Minkowski spacetime.
noncomputable section
/-- The space-time -/
def spaceTime : Type := Fin 4 →
def SpaceTime : Type := Fin 4 →
/-- Give spacetime the structure of an additive commutative monoid. -/
instance : AddCommMonoid spaceTime := Pi.addCommMonoid
instance : AddCommMonoid SpaceTime := Pi.addCommMonoid
/-- Give spacetime the structure of a module over the reals. -/
instance : Module spaceTime := Pi.module _ _ _
instance : Module SpaceTime := Pi.module _ _ _
instance euclideanNormedAddCommGroup : NormedAddCommGroup spaceTime := Pi.normedAddCommGroup
instance euclideanNormedAddCommGroup : NormedAddCommGroup SpaceTime := Pi.normedAddCommGroup
instance euclideanNormedSpace : NormedSpace spaceTime := Pi.normedSpace
instance euclideanNormedSpace : NormedSpace SpaceTime := Pi.normedSpace
namespace spaceTime
namespace SpaceTime
open Manifold
open Matrix
@ -40,15 +40,15 @@ open ComplexConjugate
/-- The space part of spacetime. -/
@[simp]
def space (x : spaceTime) : EuclideanSpace (Fin 3) := ![x 1, x 2, x 3]
def space (x : SpaceTime) : EuclideanSpace (Fin 3) := ![x 1, x 2, x 3]
/-- The structure of a smooth manifold on spacetime. -/
def asSmoothManifold : ModelWithCorners spaceTime spaceTime := 𝓘(, spaceTime)
def asSmoothManifold : ModelWithCorners SpaceTime SpaceTime := 𝓘(, SpaceTime)
instance : ChartedSpace spaceTime spaceTime := chartedSpaceSelf spaceTime
instance : ChartedSpace SpaceTime SpaceTime := chartedSpaceSelf SpaceTime
/-- The standard basis for spacetime. -/
def stdBasis : Basis (Fin 4) spaceTime := Pi.basisFun (Fin 4)
def stdBasis : Basis (Fin 4) SpaceTime := Pi.basisFun (Fin 4)
lemma stdBasis_apply (μ ν : Fin 4) : stdBasis μ ν = if μ = ν then 1 else 0 := by
erw [stdBasis, Pi.basisFun_apply, LinearMap.stdBasis_apply']
@ -85,16 +85,16 @@ lemma stdBasis_mulVec (μ ν : Fin 4) (Λ : Matrix (Fin 4) (Fin 4) ) :
rw [stdBasis_apply, if_neg (Ne.symm h)]
exact CommMonoidWithZero.mul_zero (Λ ν x)
lemma explicit (x : spaceTime) : x = ![x 0, x 1, x 2, x 3] := by
lemma explicit (x : SpaceTime) : x = ![x 0, x 1, x 2, x 3] := by
funext i
fin_cases i <;> rfl
@[simp]
lemma add_apply (x y : spaceTime) (i : Fin 4) : (x + y) i = x i + y i := rfl
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
lemma smul_apply (x : SpaceTime) (a : ) (i : Fin 4) : (a • x) i = a * x i := rfl
end spaceTime
end SpaceTime
end

View file

@ -13,18 +13,18 @@ We define
- `FourVelocity` : as a space-time element with norm 1 and future pointing.
-/
open spaceTime
open SpaceTime
/-- A `spaceTime` vector for which `ηLin x x = 1`. -/
@[simp]
def PreFourVelocity : Set spaceTime :=
def PreFourVelocity : Set SpaceTime :=
fun x => ηLin x x = 1
instance : TopologicalSpace PreFourVelocity := by
exact instTopologicalSpaceSubtype
namespace PreFourVelocity
lemma mem_PreFourVelocity_iff {x : spaceTime} : x ∈ PreFourVelocity ↔ ηLin x x = 1 := by
lemma mem_PreFourVelocity_iff {x : SpaceTime} : x ∈ PreFourVelocity ↔ ηLin x x = 1 := by
rfl
/-- The negative of a `PreFourVelocity` as a `PreFourVelocity`. -/
@ -228,7 +228,7 @@ lemma η_pos (u v : FourVelocity) : 0 < ηLin u v := by
lemma one_plus_ne_zero (u v : FourVelocity) : 1 + ηLin u v ≠ 0 := by
linarith [η_pos u v]
lemma η_continuous (u : spaceTime) : Continuous (fun (a : FourVelocity) => ηLin u a) := by
lemma η_continuous (u : SpaceTime) : Continuous (fun (a : FourVelocity) => ηLin u a) := by
simp only [ηLin_expand]
refine Continuous.add ?_ ?_
refine Continuous.add ?_ ?_

View file

@ -19,7 +19,7 @@ We define
-/
namespace spaceTime
namespace SpaceTime
open Matrix
open TensorProduct
@ -96,7 +96,7 @@ lemma space_comps (Λ : lorentzAlgebra) (i j : Fin 3) :
end lorentzAlgebra
@[simps!]
instance spaceTimeAsLieRingModule : LieRingModule lorentzAlgebra spaceTime where
instance spaceTimeAsLieRingModule : LieRingModule lorentzAlgebra SpaceTime where
bracket Λ x := Λ.1.mulVec x
add_lie Λ1 Λ2 x := by
simp [add_mulVec]
@ -107,7 +107,7 @@ instance spaceTimeAsLieRingModule : LieRingModule lorentzAlgebra spaceTime where
simp [mulVec_add, Bracket.bracket, sub_mulVec]
@[simps!]
instance spaceTimeAsLieModule : LieModule lorentzAlgebra spaceTime where
instance spaceTimeAsLieModule : LieModule lorentzAlgebra SpaceTime where
smul_lie r Λ x := by
simp [Bracket.bracket, smul_mulVec_assoc]
lie_smul r Λ x := by
@ -115,4 +115,4 @@ instance spaceTimeAsLieModule : LieModule lorentzAlgebra spaceTime where
rw [mulVec_smul]
end spaceTime
end SpaceTime

View file

@ -11,7 +11,7 @@ We define the standard basis of the Lorentz group.
-/
namespace spaceTime
namespace SpaceTime
namespace lorentzAlgebra
open Matrix
@ -145,4 +145,4 @@ theorem finrank_eq_six : FiniteDimensional.finrank lorentzAlgebra = 6 := by
end lorentzAlgebra
end spaceTime
end SpaceTime

View file

@ -26,7 +26,7 @@ identity.
noncomputable section
namespace spaceTime
namespace SpaceTime
open Manifold
open Matrix
@ -44,14 +44,14 @@ These matrices form the Lorentz group, which we will define in the next section
/-- We say a matrix `Λ` preserves `ηLin` if for all `x` and `y`,
`ηLin (Λ *ᵥ x) (Λ *ᵥ y) = ηLin x y`. -/
def PreservesηLin (Λ : Matrix (Fin 4) (Fin 4) ) : Prop :=
∀ (x y : spaceTime), ηLin (Λ *ᵥ x) (Λ *ᵥ y) = ηLin x y
∀ (x y : SpaceTime), ηLin (Λ *ᵥ x) (Λ *ᵥ y) = ηLin x y
namespace PreservesηLin
variable (Λ : Matrix (Fin 4) (Fin 4) )
lemma iff_norm : PreservesηLin Λ ↔
∀ (x : spaceTime), ηLin (Λ *ᵥ x) (Λ *ᵥ x) = ηLin x x := by
∀ (x : SpaceTime), ηLin (Λ *ᵥ x) (Λ *ᵥ x) = ηLin x x := by
refine Iff.intro (fun h x => h x x) (fun h x y => ?_)
have hp := h (x + y)
have hn := h (x - y)
@ -75,7 +75,7 @@ lemma iff_det_selfAdjoint : PreservesηLin Λ ↔
simpa [det_eq_ηLin] using h1
lemma iff_on_right : PreservesηLin Λ ↔
∀ (x y : spaceTime), ηLin x ((η * Λᵀ * η * Λ) *ᵥ y) = ηLin x y := by
∀ (x y : SpaceTime), ηLin x ((η * Λᵀ * η * Λ) *ᵥ y) = ηLin x y := by
apply Iff.intro
intro h x y
have h1 := h x y
@ -138,10 +138,10 @@ We show that the Lorentz group is indeed a group.
-/
/-- The Lorentz group is the subset of matrices which preserve ηLin. -/
def lorentzGroup : Type := {Λ : Matrix (Fin 4) (Fin 4) // PreservesηLin Λ}
def LorentzGroup : Type := {Λ : Matrix (Fin 4) (Fin 4) // PreservesηLin Λ}
@[simps mul_coe one_coe inv div]
instance lorentzGroupIsGroup : Group lorentzGroup where
instance lorentzGroupIsGroup : Group LorentzGroup where
mul A B := ⟨A.1 * B.1, PreservesηLin.mul A.2 B.2⟩
mul_assoc A B C := by
apply Subtype.eq
@ -160,20 +160,20 @@ instance lorentzGroupIsGroup : Group lorentzGroup where
exact (PreservesηLin.iff_matrix A.1).mp A.2
/-- Notation for the Lorentz group. -/
scoped[spaceTime] notation (name := lorentzGroup_notation) "𝓛" => lorentzGroup
scoped[SpaceTime] notation (name := lorentzGroup_notation) "𝓛" => LorentzGroup
/-- `lorentzGroup` has the subtype topology. -/
instance : TopologicalSpace lorentzGroup := instTopologicalSpaceSubtype
instance : TopologicalSpace LorentzGroup := instTopologicalSpaceSubtype
namespace lorentzGroup
namespace LorentzGroup
lemma coe_inv (A : 𝓛) : (A⁻¹).1 = A.1⁻¹:= by
lemma coe_inv (A : LorentzGroup) : (A⁻¹).1 = A.1⁻¹:= by
refine (inv_eq_left_inv ?h).symm
exact (PreservesηLin.iff_matrix A.1).mp A.2
/-- The transpose of an matrix in the Lorentz group is an element of the Lorentz group. -/
def transpose (Λ : lorentzGroup) : lorentzGroup := ⟨Λ.1ᵀ, (PreservesηLin.iff_transpose Λ.1).mp Λ.2⟩
def transpose (Λ : LorentzGroup) : LorentzGroup := ⟨Λ.1ᵀ, (PreservesηLin.iff_transpose Λ.1).mp Λ.2⟩
/-!
@ -186,7 +186,7 @@ embedding.
-/
/-- The homomorphism of the Lorentz group into `GL (Fin 4) `. -/
def toGL : 𝓛 →* GL (Fin 4) where
def toGL : LorentzGroup →* GL (Fin 4) where
toFun A := ⟨A.1, (A⁻¹).1, mul_eq_one_comm.mpr $ (PreservesηLin.iff_matrix A.1).mp A.2,
(PreservesηLin.iff_matrix A.1).mp A.2⟩
map_one' := by
@ -206,7 +206,7 @@ lemma toGL_injective : Function.Injective toGL := by
/-- The homomorphism from the Lorentz Group into the monoid of matrices times the opposite of
the monoid of matrices. -/
@[simps!]
def toProd : 𝓛 →* (Matrix (Fin 4) (Fin 4) ) × (Matrix (Fin 4) (Fin 4) )ᵐᵒᵖ :=
def toProd : LorentzGroup →* (Matrix (Fin 4) (Fin 4) ) × (Matrix (Fin 4) (Fin 4) )ᵐᵒᵖ :=
MonoidHom.comp (Units.embedProduct _) toGL
lemma toProd_eq_transpose_η : toProd A = (A.1, ⟨η * A.1ᵀ * η⟩) := rfl
@ -248,11 +248,11 @@ lemma toGL_embedding : Embedding toGL.toFun where
exact exists_exists_and_eq_and
instance : TopologicalGroup 𝓛 := Inducing.topologicalGroup toGL toGL_embedding.toInducing
instance : TopologicalGroup LorentzGroup := Inducing.topologicalGroup toGL toGL_embedding.toInducing
end lorentzGroup
end LorentzGroup
end spaceTime
end SpaceTime

View file

@ -28,14 +28,14 @@ A boost is the special case of a generalised boost when `u = stdBasis 0`.
-/
noncomputable section
namespace spaceTime
namespace SpaceTime
namespace lorentzGroup
namespace LorentzGroup
open FourVelocity
/-- An auxillary linear map used in the definition of a generalised boost. -/
def genBoostAux₁ (u v : FourVelocity) : spaceTime →ₗ[] spaceTime where
def genBoostAux₁ (u v : FourVelocity) : SpaceTime →ₗ[] SpaceTime where
toFun x := (2 * ηLin x u) • v.1.1
map_add' x y := by
simp only [map_add, LinearMap.add_apply]
@ -46,7 +46,7 @@ def genBoostAux₁ (u v : FourVelocity) : spaceTime →ₗ[] spaceTime where
rw [← mul_assoc, mul_comm 2 c, mul_assoc, mul_smul]
/-- An auxillary linear map used in the definition of a genearlised boost. -/
def genBoostAux₂ (u v : FourVelocity) : spaceTime →ₗ[] spaceTime where
def genBoostAux₂ (u v : FourVelocity) : SpaceTime →ₗ[] SpaceTime where
toFun x := - (ηLin x (u + v) / (1 + ηLin u v)) • (u + v)
map_add' x y := by
simp only
@ -62,7 +62,7 @@ def genBoostAux₂ (u v : FourVelocity) : spaceTime →ₗ[] spaceTime where
/-- An generalised boost. This is a Lorentz transformation which takes the four velocity `u`
to `v`. -/
def genBoost (u v : FourVelocity) : spaceTime →ₗ[] spaceTime :=
def genBoost (u v : FourVelocity) : SpaceTime →ₗ[] SpaceTime :=
LinearMap.id + genBoostAux₁ u v + genBoostAux₂ u v
namespace genBoost
@ -91,7 +91,7 @@ lemma self (u : FourVelocity) : genBoost u u = LinearMap.id := by
def toMatrix (u v : FourVelocity) : Matrix (Fin 4) (Fin 4) :=
LinearMap.toMatrix stdBasis stdBasis (genBoost u v)
lemma toMatrix_mulVec (u v : FourVelocity) (x : spaceTime) :
lemma toMatrix_mulVec (u v : FourVelocity) (x : SpaceTime) :
(toMatrix u v).mulVec x = genBoost u v x :=
LinearMap.toMatrix_mulVec_repr stdBasis stdBasis (genBoost u v) x
@ -143,7 +143,7 @@ lemma toMatrix_PreservesηLin (u v : FourVelocity) : PreservesηLin (toMatrix u
ring
/-- A generalised boost as an element of the Lorentz Group. -/
def toLorentz (u v : FourVelocity) : lorentzGroup :=
def toLorentz (u v : FourVelocity) : LorentzGroup :=
⟨toMatrix u v, toMatrix_PreservesηLin u v⟩
lemma toLorentz_continuous (u : FourVelocity) : Continuous (toLorentz u) := by
@ -171,8 +171,8 @@ lemma isProper (u v : FourVelocity) : IsProper (toLorentz u v) :=
end genBoost
end lorentzGroup
end LorentzGroup
end spaceTime
end SpaceTime
end

View file

@ -20,19 +20,19 @@ matrices.
noncomputable section
namespace spaceTime
namespace SpaceTime
open Manifold
open Matrix
open Complex
open ComplexConjugate
namespace lorentzGroup
namespace LorentzGroup
open PreFourVelocity
/-- The first column of a lorentz matrix as a `PreFourVelocity`. -/
@[simp]
def fstCol (Λ : lorentzGroup) : PreFourVelocity := ⟨Λ.1 *ᵥ stdBasis 0, by
def fstCol (Λ : LorentzGroup) : PreFourVelocity := ⟨Λ.1 *ᵥ stdBasis 0, by
rw [mem_PreFourVelocity_iff, ηLin_expand]
simp only [Fin.isValue, stdBasis_mulVec]
have h00 := congrFun (congrFun ((PreservesηLin.iff_matrix Λ.1).mp Λ.2) 0) 0
@ -46,18 +46,18 @@ def fstCol (Λ : lorentzGroup) : PreFourVelocity := ⟨Λ.1 *ᵥ stdBasis 0, by
exact h00⟩
/-- A Lorentz transformation is `orthochronous` if its `0 0` element is non-negative. -/
def IsOrthochronous (Λ : lorentzGroup) : Prop := 0 ≤ Λ.1 0 0
def IsOrthochronous (Λ : LorentzGroup) : Prop := 0 ≤ Λ.1 0 0
lemma IsOrthochronous_iff_transpose (Λ : lorentzGroup) :
lemma IsOrthochronous_iff_transpose (Λ : LorentzGroup) :
IsOrthochronous Λ ↔ IsOrthochronous (transpose Λ) := by rfl
lemma IsOrthochronous_iff_fstCol_IsFourVelocity (Λ : lorentzGroup) :
lemma IsOrthochronous_iff_fstCol_IsFourVelocity (Λ : LorentzGroup) :
IsOrthochronous Λ ↔ IsFourVelocity (fstCol Λ) := by
simp [IsOrthochronous, IsFourVelocity]
rw [stdBasis_mulVec]
/-- The continuous map taking a Lorentz transformation to its `0 0` element. -/
def mapZeroZeroComp : C(lorentzGroup, ) := ⟨fun Λ => Λ.1 0 0,
def mapZeroZeroComp : C(LorentzGroup, ) := ⟨fun Λ => Λ.1 0 0,
Continuous.matrix_elem (continuous_iff_le_induced.mpr fun _ a => a) 0 0⟩
/-- An auxillary function used in the definition of `orthchroMapReal`. -/
@ -78,10 +78,10 @@ lemma stepFunction_continuous : Continuous stepFunction := by
/-- The continuous map from `lorentzGroup` to `` wh
taking Orthochronous elements to `1` and non-orthochronous to `-1`. -/
def orthchroMapReal : C(lorentzGroup, ) := ContinuousMap.comp
def orthchroMapReal : C(LorentzGroup, ) := ContinuousMap.comp
⟨stepFunction, stepFunction_continuous⟩ mapZeroZeroComp
lemma orthchroMapReal_on_IsOrthochronous {Λ : lorentzGroup} (h : IsOrthochronous Λ) :
lemma orthchroMapReal_on_IsOrthochronous {Λ : LorentzGroup} (h : IsOrthochronous Λ) :
orthchroMapReal Λ = 1 := by
rw [IsOrthochronous_iff_fstCol_IsFourVelocity] at h
simp only [IsFourVelocity] at h
@ -92,7 +92,7 @@ lemma orthchroMapReal_on_IsOrthochronous {Λ : lorentzGroup} (h : IsOrthochronou
rw [stepFunction, if_neg h1, if_pos h]
lemma orthchroMapReal_on_not_IsOrthochronous {Λ : lorentzGroup} (h : ¬ IsOrthochronous Λ) :
lemma orthchroMapReal_on_not_IsOrthochronous {Λ : LorentzGroup} (h : ¬ IsOrthochronous Λ) :
orthchroMapReal Λ = - 1 := by
rw [IsOrthochronous_iff_fstCol_IsFourVelocity] at h
rw [not_IsFourVelocity_iff, zero_nonpos_iff] at h
@ -100,7 +100,7 @@ lemma orthchroMapReal_on_not_IsOrthochronous {Λ : lorentzGroup} (h : ¬ IsOrtho
change stepFunction (Λ.1 0 0) = - 1
rw [stepFunction, if_pos h]
lemma orthchroMapReal_minus_one_or_one (Λ : lorentzGroup) :
lemma orthchroMapReal_minus_one_or_one (Λ : LorentzGroup) :
orthchroMapReal Λ = -1 orthchroMapReal Λ = 1 := by
by_cases h : IsOrthochronous Λ
apply Or.inr $ orthchroMapReal_on_IsOrthochronous h
@ -109,21 +109,21 @@ lemma orthchroMapReal_minus_one_or_one (Λ : lorentzGroup) :
local notation "ℤ₂" => Multiplicative (ZMod 2)
/-- A continuous map from `lorentzGroup` to `ℤ₂` whose kernal are the Orthochronous elements. -/
def orthchroMap : C(lorentzGroup, ℤ₂) :=
def orthchroMap : C(LorentzGroup, ℤ₂) :=
ContinuousMap.comp coeFor₂ {
toFun := fun Λ => ⟨orthchroMapReal Λ, orthchroMapReal_minus_one_or_one Λ⟩,
continuous_toFun := Continuous.subtype_mk (ContinuousMap.continuous orthchroMapReal) _}
lemma orthchroMap_IsOrthochronous {Λ : lorentzGroup} (h : IsOrthochronous Λ) :
lemma orthchroMap_IsOrthochronous {Λ : LorentzGroup} (h : IsOrthochronous Λ) :
orthchroMap Λ = 1 := by
simp [orthchroMap, orthchroMapReal_on_IsOrthochronous h]
lemma orthchroMap_not_IsOrthochronous {Λ : lorentzGroup} (h : ¬ IsOrthochronous Λ) :
lemma orthchroMap_not_IsOrthochronous {Λ : LorentzGroup} (h : ¬ IsOrthochronous Λ) :
orthchroMap Λ = Additive.toMul (1 : ZMod 2) := by
simp [orthchroMap, orthchroMapReal_on_not_IsOrthochronous h]
rfl
lemma zero_zero_mul (Λ Λ' : lorentzGroup) :
lemma zero_zero_mul (Λ Λ' : LorentzGroup) :
(Λ * Λ').1 0 0 = (fstCol (transpose Λ)).1 0 * (fstCol Λ').1 0 +
⟪(fstCol (transpose Λ)).1.space, (fstCol Λ').1.space⟫_ := by
simp only [Fin.isValue, lorentzGroupIsGroup_mul_coe, mul_apply, Fin.sum_univ_four, fstCol,
@ -132,21 +132,21 @@ lemma zero_zero_mul (Λ Λ' : lorentzGroup) :
cons_val_one, head_cons, cons_val_two, tail_cons]
ring
lemma mul_othchron_of_othchron_othchron {Λ Λ' : lorentzGroup} (h : IsOrthochronous Λ)
lemma mul_othchron_of_othchron_othchron {Λ Λ' : LorentzGroup} (h : IsOrthochronous Λ)
(h' : IsOrthochronous Λ') : IsOrthochronous (Λ * Λ') := by
rw [IsOrthochronous_iff_transpose] at h
rw [IsOrthochronous_iff_fstCol_IsFourVelocity] at h h'
rw [IsOrthochronous, zero_zero_mul]
exact euclid_norm_IsFourVelocity_IsFourVelocity h h'
lemma mul_othchron_of_not_othchron_not_othchron {Λ Λ' : lorentzGroup} (h : ¬ IsOrthochronous Λ)
lemma mul_othchron_of_not_othchron_not_othchron {Λ Λ' : LorentzGroup} (h : ¬ IsOrthochronous Λ)
(h' : ¬ IsOrthochronous Λ') : IsOrthochronous (Λ * Λ') := by
rw [IsOrthochronous_iff_transpose] at h
rw [IsOrthochronous_iff_fstCol_IsFourVelocity] at h h'
rw [IsOrthochronous, zero_zero_mul]
exact euclid_norm_not_IsFourVelocity_not_IsFourVelocity h h'
lemma mul_not_othchron_of_othchron_not_othchron {Λ Λ' : lorentzGroup} (h : IsOrthochronous Λ)
lemma mul_not_othchron_of_othchron_not_othchron {Λ Λ' : LorentzGroup} (h : IsOrthochronous Λ)
(h' : ¬ IsOrthochronous Λ') : ¬ IsOrthochronous (Λ * Λ') := by
rw [IsOrthochronous_iff_transpose] at h
rw [IsOrthochronous_iff_fstCol_IsFourVelocity] at h h'
@ -156,7 +156,7 @@ lemma mul_not_othchron_of_othchron_not_othchron {Λ Λ' : lorentzGroup} (h : Is
rw [zero_zero_mul]
exact euclid_norm_IsFourVelocity_not_IsFourVelocity h h'
lemma mul_not_othchron_of_not_othchron_othchron {Λ Λ' : lorentzGroup} (h : ¬ IsOrthochronous Λ)
lemma mul_not_othchron_of_not_othchron_othchron {Λ Λ' : LorentzGroup} (h : ¬ IsOrthochronous Λ)
(h' : IsOrthochronous Λ') : ¬ IsOrthochronous (Λ * Λ') := by
rw [IsOrthochronous_iff_transpose] at h
rw [IsOrthochronous_iff_fstCol_IsFourVelocity] at h h'
@ -167,7 +167,7 @@ lemma mul_not_othchron_of_not_othchron_othchron {Λ Λ' : lorentzGroup} (h : ¬
exact euclid_norm_not_IsFourVelocity_IsFourVelocity h h'
/-- The homomorphism from `lorentzGroup` to `ℤ₂` whose kernel are the Orthochronous elements. -/
def orthchroRep : lorentzGroup →* ℤ₂ where
def orthchroRep : LorentzGroup →* ℤ₂ where
toFun := orthchroMap
map_one' := orthchroMap_IsOrthochronous (by simp [IsOrthochronous])
map_mul' Λ Λ' := by
@ -187,7 +187,7 @@ def orthchroRep : lorentzGroup →* ℤ₂ where
orthchroMap_IsOrthochronous (mul_othchron_of_not_othchron_not_othchron h h')]
rfl
end lorentzGroup
end LorentzGroup
end spaceTime
end SpaceTime
end

View file

@ -14,14 +14,14 @@ We define the give a series of lemmas related to the determinant of the lorentz
noncomputable section
namespace spaceTime
namespace SpaceTime
open Manifold
open Matrix
open Complex
open ComplexConjugate
namespace lorentzGroup
namespace LorentzGroup
/-- The determinant of a member of the lorentz group is `1` or `-1`. -/
lemma det_eq_one_or_neg_one (Λ : 𝓛) : Λ.1.det = 1 Λ.1.det = -1 := by
@ -49,14 +49,14 @@ def coeFor₂ : C(({-1, 1} : Set ), ℤ₂) where
/-- The continuous map taking a lorentz matrix to its determinant. -/
def detContinuous : C(𝓛, ℤ₂) :=
ContinuousMap.comp coeFor₂ {
toFun := fun Λ => ⟨Λ.1.det, Or.symm (lorentzGroup.det_eq_one_or_neg_one _)⟩,
toFun := fun Λ => ⟨Λ.1.det, Or.symm (LorentzGroup.det_eq_one_or_neg_one _)⟩,
continuous_toFun := by
refine Continuous.subtype_mk ?_ _
apply Continuous.matrix_det $
Continuous.comp' (continuous_iff_le_induced.mpr fun U a => a) continuous_id'
}
lemma detContinuous_eq_iff_det_eq (Λ Λ' : lorentzGroup) :
lemma detContinuous_eq_iff_det_eq (Λ Λ' : LorentzGroup) :
detContinuous Λ = detContinuous Λ' ↔ Λ.1.det = Λ'.1.det := by
apply Iff.intro
intro h
@ -90,7 +90,7 @@ def detRep : 𝓛 →* ℤ₂ where
lemma detRep_continuous : Continuous detRep := detContinuous.2
lemma det_on_connected_component {Λ Λ' : lorentzGroup} (h : Λ' ∈ connectedComponent Λ) :
lemma det_on_connected_component {Λ Λ' : LorentzGroup} (h : Λ' ∈ connectedComponent Λ) :
Λ.1.det = Λ'.1.det := by
obtain ⟨s, hs, hΛ'⟩ := h
let f : ContinuousMap s ℤ₂ := ContinuousMap.restrict s detContinuous
@ -99,23 +99,23 @@ lemma det_on_connected_component {Λ Λ' : lorentzGroup} (h : Λ' ∈ connected
(@IsPreconnected.subsingleton ℤ₂ _ _ _ (isPreconnected_range f.2))
(Set.mem_range_self ⟨Λ, hs.2⟩) (Set.mem_range_self ⟨Λ', hΛ'⟩)
lemma detRep_on_connected_component {Λ Λ' : lorentzGroup} (h : Λ' ∈ connectedComponent Λ) :
lemma detRep_on_connected_component {Λ Λ' : LorentzGroup} (h : Λ' ∈ connectedComponent Λ) :
detRep Λ = detRep Λ' := by
simp [detRep_apply, detRep_apply, detContinuous]
rw [det_on_connected_component h]
lemma det_of_joined {Λ Λ' : lorentzGroup} (h : Joined Λ Λ') : Λ.1.det = Λ'.1.det :=
lemma det_of_joined {Λ Λ' : LorentzGroup} (h : Joined Λ Λ') : Λ.1.det = Λ'.1.det :=
det_on_connected_component $ pathComponent_subset_component _ h
/-- A Lorentz Matrix is proper if its determinant is 1. -/
@[simp]
def IsProper (Λ : lorentzGroup) : Prop := Λ.1.det = 1
def IsProper (Λ : LorentzGroup) : Prop := Λ.1.det = 1
instance : DecidablePred IsProper := by
intro Λ
apply Real.decidableEq
lemma IsProper_iff (Λ : lorentzGroup) : IsProper Λ ↔ detRep Λ = 1 := by
lemma IsProper_iff (Λ : LorentzGroup) : IsProper Λ ↔ detRep Λ = 1 := by
rw [show 1 = detRep 1 from Eq.symm (MonoidHom.map_one detRep)]
rw [detRep_apply, detRep_apply, detContinuous_eq_iff_det_eq]
simp only [IsProper, lorentzGroupIsGroup_one_coe, det_one]
@ -123,12 +123,12 @@ lemma IsProper_iff (Λ : lorentzGroup) : IsProper Λ ↔ detRep Λ = 1 := by
lemma id_IsProper : IsProper 1 := by
simp [IsProper]
lemma isProper_on_connected_component {Λ Λ' : lorentzGroup} (h : Λ' ∈ connectedComponent Λ) :
lemma isProper_on_connected_component {Λ Λ' : LorentzGroup} (h : Λ' ∈ connectedComponent Λ) :
IsProper Λ ↔ IsProper Λ' := by
simp [detRep_apply, detRep_apply, detContinuous]
rw [det_on_connected_component h]
end lorentzGroup
end LorentzGroup
end spaceTime
end SpaceTime
end

View file

@ -11,7 +11,7 @@ import Mathlib.Topology.Constructions
-/
noncomputable section
namespace spaceTime
namespace SpaceTime
namespace lorentzGroup
open GroupTheory
@ -54,5 +54,5 @@ def SO3ToLorentz : SO(3) →* 𝓛 where
end lorentzGroup
end spaceTime
end SpaceTime
end

View file

@ -15,7 +15,7 @@ This file introduces the metric on spacetime in the (+, -, -, -) signature.
noncomputable section
namespace spaceTime
namespace SpaceTime
open Manifold
open Matrix
@ -108,7 +108,7 @@ lemma η_sq : η * η = 1 := by
lemma η_diag_mul_self (μ : Fin 4) : η μ μ * η μ μ = 1 := by
fin_cases μ <;> simp [η_explicit]
lemma η_mulVec (x : spaceTime) : η *ᵥ x = ![x 0, -x 1, -x 2, -x 3] := by
lemma η_mulVec (x : SpaceTime) : η *ᵥ x = ![x 0, -x 1, -x 2, -x 3] := by
rw [explicit x, η_explicit]
funext i
fin_cases i <;>
@ -143,7 +143,7 @@ lemma η_contract_self' (μ ν : Fin 4) : ∑ x, (η^[x]_[μ] * η_[ν]_[x]) =
/-- Given a point in spaceTime `x` the linear map `y → x ⬝ᵥ (η *ᵥ y)`. -/
@[simps!]
def linearMapForSpaceTime (x : spaceTime) : spaceTime →ₗ[] where
def linearMapForSpaceTime (x : SpaceTime) : SpaceTime →ₗ[] where
toFun y := x ⬝ᵥ (η *ᵥ y)
map_add' y z := by
simp only
@ -154,7 +154,7 @@ def linearMapForSpaceTime (x : spaceTime) : spaceTime →ₗ[] where
rfl
/-- The metric as a bilinear map from `spaceTime` to ``. -/
def ηLin : LinearMap.BilinForm spaceTime where
def ηLin : LinearMap.BilinForm SpaceTime where
toFun x := linearMapForSpaceTime x
map_add' x y := by
apply LinearMap.ext
@ -168,7 +168,7 @@ def ηLin : LinearMap.BilinForm spaceTime where
rw [smul_dotProduct]
rfl
lemma ηLin_expand (x y : spaceTime) : ηLin x y = x 0 * y 0 - x 1 * y 1 - x 2 * y 2 - x 3 * y 3 := by
lemma ηLin_expand (x y : SpaceTime) : ηLin x y = x 0 * y 0 - x 1 * y 1 - x 2 * y 2 - x 3 * y 3 := by
rw [ηLin]
simp only [LinearMap.coe_mk, AddHom.coe_mk, linearMapForSpaceTime_apply, Fin.isValue]
erw [η_mulVec]
@ -177,40 +177,40 @@ lemma ηLin_expand (x y : spaceTime) : ηLin x y = x 0 * y 0 - x 1 * y 1 - x 2 *
cons_val_zero, cons_val_one, head_cons, mul_neg, cons_val_two, tail_cons, cons_val_three]
ring
lemma ηLin_expand_self (x : spaceTime) : ηLin x x = x 0 ^ 2 - ‖x.space‖ ^ 2 := by
lemma ηLin_expand_self (x : SpaceTime) : ηLin x x = x 0 ^ 2 - ‖x.space‖ ^ 2 := by
rw [← @real_inner_self_eq_norm_sq, @PiLp.inner_apply, Fin.sum_univ_three, ηLin_expand]
noncomm_ring
lemma time_elm_sq_of_ηLin (x : spaceTime) : x 0 ^ 2 = ηLin x x + ‖x.space‖ ^ 2 := by
lemma time_elm_sq_of_ηLin (x : SpaceTime) : x 0 ^ 2 = ηLin x x + ‖x.space‖ ^ 2 := by
rw [ηLin_expand_self]
ring
lemma ηLin_leq_time_sq (x : spaceTime) : ηLin x x ≤ x 0 ^ 2 := by
lemma ηLin_leq_time_sq (x : SpaceTime) : ηLin x x ≤ x 0 ^ 2 := by
rw [time_elm_sq_of_ηLin]
exact (le_add_iff_nonneg_right _).mpr $ sq_nonneg ‖x.space‖
lemma ηLin_space_inner_product (x y : spaceTime) :
lemma ηLin_space_inner_product (x y : SpaceTime) :
ηLin x y = x 0 * y 0 - ⟪x.space, y.space⟫_ := by
rw [ηLin_expand, @PiLp.inner_apply, Fin.sum_univ_three]
noncomm_ring
lemma ηLin_ge_abs_inner_product (x y : spaceTime) :
lemma ηLin_ge_abs_inner_product (x y : SpaceTime) :
x 0 * y 0 - ‖⟪x.space, y.space⟫_‖ ≤ (ηLin x y) := by
rw [ηLin_space_inner_product, sub_le_sub_iff_left]
exact Real.le_norm_self ⟪x.space, y.space⟫_
lemma ηLin_ge_sub_norm (x y : spaceTime) :
lemma ηLin_ge_sub_norm (x y : SpaceTime) :
x 0 * y 0 - ‖x.space‖ * ‖y.space‖ ≤ (ηLin x y) := by
apply le_trans ?_ (ηLin_ge_abs_inner_product x y)
rw [sub_le_sub_iff_left]
exact norm_inner_le_norm x.space y.space
lemma ηLin_symm (x y : spaceTime) : ηLin x y = ηLin y x := by
lemma ηLin_symm (x y : SpaceTime) : ηLin x y = ηLin y x := by
rw [ηLin_expand, ηLin_expand]
ring
lemma ηLin_stdBasis_apply (μ : Fin 4) (x : spaceTime) : ηLin (stdBasis μ) x = η μ μ * x μ := by
lemma ηLin_stdBasis_apply (μ : Fin 4) (x : SpaceTime) : ηLin (stdBasis μ) x = η μ μ * x μ := by
rw [ηLin_expand]
fin_cases μ
<;> simp [stdBasis_0, stdBasis_1, stdBasis_2, stdBasis_3, η_explicit]
@ -227,13 +227,13 @@ lemma ηLin_η_stdBasis (μ ν : Fin 4) : ηLin (stdBasis μ) (stdBasis ν) = η
exact fun a ↦ h (id a.symm)
set_option maxHeartbeats 0
lemma ηLin_mulVec_left (x y : spaceTime) (Λ : Matrix (Fin 4) (Fin 4) ) :
lemma ηLin_mulVec_left (x y : SpaceTime) (Λ : Matrix (Fin 4) (Fin 4) ) :
ηLin (Λ *ᵥ x) y = ηLin x ((η * Λᵀ * η) *ᵥ y) := by
simp [ηLin, LinearMap.coe_mk, AddHom.coe_mk, linearMapForSpaceTime_apply,
mulVec_mulVec, (vecMul_transpose Λ x).symm, ← dotProduct_mulVec, mulVec_mulVec,
← mul_assoc, ← mul_assoc, η_sq, one_mul]
lemma ηLin_mulVec_right (x y : spaceTime) (Λ : Matrix (Fin 4) (Fin 4) ) :
lemma ηLin_mulVec_right (x y : SpaceTime) (Λ : Matrix (Fin 4) (Fin 4) ) :
ηLin x (Λ *ᵥ y) = ηLin ((η * Λᵀ * η) *ᵥ x) y := by
rw [ηLin_symm, ηLin_symm ((η * Λᵀ * η) *ᵥ x) _ ]
exact ηLin_mulVec_left y x Λ
@ -247,7 +247,7 @@ lemma ηLin_matrix_stdBasis' (μ ν : Fin 4) (Λ : Matrix (Fin 4) (Fin 4) ) :
rw [ηLin_matrix_stdBasis, ← mul_assoc, η_diag_mul_self, one_mul]
lemma ηLin_matrix_eq_identity_iff (Λ : Matrix (Fin 4) (Fin 4) ) :
Λ = 1 ↔ ∀ (x y : spaceTime), ηLin x y = ηLin x (Λ *ᵥ y) := by
Λ = 1 ↔ ∀ (x y : SpaceTime), ηLin x y = ηLin x (Λ *ᵥ y) := by
apply Iff.intro
· intro h
subst h
@ -260,8 +260,8 @@ lemma ηLin_matrix_eq_identity_iff (Λ : Matrix (Fin 4) (Fin 4) ) :
simp_all [η_explicit, Fin.mk_one, Fin.mk_one, vecHead, vecTail]
/-- The metric as a quadratic form on `spaceTime`. -/
def quadraticForm : QuadraticForm spaceTime := ηLin.toQuadraticForm
def quadraticForm : QuadraticForm SpaceTime := ηLin.toQuadraticForm
end spaceTime
end SpaceTime
end

View file

@ -11,7 +11,7 @@ import Mathlib.RepresentationTheory.Basic
The aim of this file is to give the relationship between `SL(2, )` and the Lorentz group.
-/
namespace spaceTime
namespace SpaceTime
open Matrix
open MatrixGroups
@ -19,7 +19,7 @@ open Complex
namespace SL2C
open spaceTime
open SpaceTime
noncomputable section
@ -64,7 +64,7 @@ def repSelfAdjointMatrix : Representation SL(2, ) $ selfAdjoint (Matrix (
/-- The representation of `SL(2, )` on `spaceTime` obtained from `toSelfAdjointMatrix` and
`repSelfAdjointMatrix`. -/
def repSpaceTime : Representation SL(2, ) spaceTime where
def repSpaceTime : Representation SL(2, ) SpaceTime where
toFun M := toSelfAdjointMatrix.symm.comp ((repSelfAdjointMatrix M).comp
toSelfAdjointMatrix.toLinearMap)
map_one' := by
@ -120,4 +120,4 @@ Complete this section.
end
end SL2C
end spaceTime
end SpaceTime