chore: Bump to 4.14.0
This commit is contained in:
parent
c91ca06272
commit
5dfd29ab8d
32 changed files with 376 additions and 334 deletions
|
@ -334,7 +334,7 @@ def pauliCoMap := ((Sum.elim ![Color.down, Color.down] ![Color.up, Color.upL, Co
|
|||
⇑finSumFinEquiv.symm) ∘ Fin.succAbove 1 ∘ Fin.succAbove 1)
|
||||
|
||||
lemma pauliMatrix_contr_down_0 :
|
||||
(contr 1 1 rfl (((tensorNode (basisVector ![Color.down, Color.down] fun x => 0)).prod
|
||||
(contr 1 1 rfl (((tensorNode (basisVector ![Color.down, Color.down] fun _ => 0)).prod
|
||||
(tensorNode pauliContr)))).tensor
|
||||
= basisVector pauliCoMap (fun | 0 => 0 | 1 => 0 | 2 => 0)
|
||||
+ basisVector pauliCoMap (fun | 0 => 0 | 1 => 1 | 2 => 1) := by
|
||||
|
@ -379,7 +379,7 @@ lemma pauliMatrix_contr_down_0 :
|
|||
fin_cases k <;> rfl
|
||||
|
||||
lemma pauliMatrix_contr_down_1 :
|
||||
{(basisVector ![Color.down, Color.down] fun x => 1) | ν μ ⊗
|
||||
{(basisVector ![Color.down, Color.down] fun _ => 1) | ν μ ⊗
|
||||
pauliContr | μ α β}ᵀ.tensor
|
||||
= basisVector pauliCoMap (fun | 0 => 1 | 1 => 0 | 2 => 1)
|
||||
+ basisVector pauliCoMap (fun | 0 => 1 | 1 => 1 | 2 => 0) := by
|
||||
|
@ -424,7 +424,7 @@ lemma pauliMatrix_contr_down_1 :
|
|||
fin_cases k <;> rfl
|
||||
|
||||
lemma pauliMatrix_contr_down_2 :
|
||||
{(basisVector ![Color.down, Color.down] fun x => 2) | μ ν ⊗
|
||||
{(basisVector ![Color.down, Color.down] fun _ => 2) | μ ν ⊗
|
||||
pauliContr | ν α β}ᵀ.tensor
|
||||
= (- I) • basisVector pauliCoMap (fun | 0 => 2 | 1 => 0 | 2 => 1)
|
||||
+ (I) • basisVector pauliCoMap (fun | 0 => 2 | 1 => 1 | 2 => 0) := by
|
||||
|
@ -464,7 +464,7 @@ lemma pauliMatrix_contr_down_2 :
|
|||
· decide
|
||||
|
||||
lemma pauliMatrix_contr_down_3 :
|
||||
{(basisVector ![Color.down, Color.down] fun x => 3) | μ ν ⊗
|
||||
{(basisVector ![Color.down, Color.down] fun _ => 3) | μ ν ⊗
|
||||
pauliContr | ν α β}ᵀ.tensor
|
||||
= basisVector pauliCoMap (fun | 0 => 3 | 1 => 0 | 2 => 0)
|
||||
+ (- 1 : ℂ) • basisVector pauliCoMap (fun | 0 => 3 | 1 => 1 | 2 => 1) := by
|
||||
|
|
|
@ -241,17 +241,19 @@ lemma toProd_continuous : Continuous (@toProd d) := by
|
|||
MulOpposite.continuous_op.comp' ((continuous_const.matrix_mul (continuous_iff_le_induced.mpr
|
||||
fun U a => a).matrix_transpose).matrix_mul continuous_const)⟩
|
||||
|
||||
open Topology
|
||||
|
||||
/-- The embedding from the Lorentz Group into the monoid of matrices times the opposite of
|
||||
the monoid of matrices. -/
|
||||
lemma toProd_embedding : IsEmbedding (@toProd d) where
|
||||
inj := toProd_injective
|
||||
injective := toProd_injective
|
||||
eq_induced :=
|
||||
(isInducing_iff ⇑toProd).mp (IsInducing.of_comp toProd_continuous continuous_fst
|
||||
((isInducing_iff (Prod.fst ∘ ⇑toProd)).mpr rfl))
|
||||
|
||||
/-- The embedding from the Lorentz Group into `GL (Fin 4) ℝ`. -/
|
||||
lemma toGL_embedding : IsEmbedding (@toGL d).toFun where
|
||||
inj := toGL_injective
|
||||
injective := toGL_injective
|
||||
eq_induced := by
|
||||
refine ((fun {X} {t t'} => TopologicalSpace.ext_iff.mpr) fun _ ↦ ?_).symm
|
||||
rw [TopologicalSpace.ext_iff.mp toProd_embedding.eq_induced _, isOpen_induced_iff,
|
||||
|
|
|
@ -347,9 +347,7 @@ lemma _root_.LorentzGroup.mem_iff_norm : Λ ∈ LorentzGroup d ↔
|
|||
Action.FunctorCategoryEquivalence.functor_obj_obj, 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, Action.instMonoidalCategory_tensorObj_V,
|
||||
Equivalence.symm_inverse, Action.functorCategoryEquivalence_functor,
|
||||
Action.FunctorCategoryEquivalence.functor_obj_obj, add_sub_cancel, neg_add_cancel, e]
|
||||
ring
|
||||
|
||||
/-!
|
||||
|
||||
|
|
|
@ -197,7 +197,7 @@ def toSpace (v : ContrMod d) : EuclideanSpace ℝ (Fin d) := v.val ∘ Sum.inr
|
|||
/-- The representation of the Lorentz group acting on `ContrℝModule d`. -/
|
||||
def rep : Representation ℝ (LorentzGroup d) (ContrMod d) where
|
||||
toFun g := Matrix.toLinAlgEquiv stdBasis g
|
||||
map_one' := (MulEquivClass.map_eq_one_iff (Matrix.toLinAlgEquiv stdBasis)).mpr rfl
|
||||
map_one' := EmbeddingLike.map_eq_one_iff.mpr rfl
|
||||
map_mul' x y := by
|
||||
simp only [lorentzGroupIsGroup_mul_coe, _root_.map_mul]
|
||||
|
||||
|
@ -283,6 +283,8 @@ lemma toSelfAdjoint_symm_basis (i : Fin 1 ⊕ Fin 3) :
|
|||
instance : TopologicalSpace (ContrMod d) := TopologicalSpace.induced
|
||||
ContrMod.toFin1dℝEquiv (Pi.topologicalSpace)
|
||||
|
||||
open Topology
|
||||
|
||||
lemma toFin1dℝEquiv_isInducing : IsInducing (@ContrMod.toFin1dℝEquiv d) := by
|
||||
exact { eq_induced := rfl }
|
||||
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue