feat: Start adding expansions in terms of basis
This commit is contained in:
parent
5a0e52e726
commit
1148234929
5 changed files with 196 additions and 2 deletions
|
@ -42,8 +42,14 @@ def basisVector {n : ℕ} (c : Fin n → complexLorentzTensor.C)
|
|||
complexLorentzTensor.F.obj (OverColor.mk c) :=
|
||||
PiTensorProduct.tprod ℂ (fun i => complexLorentzTensor.basis (c i) (b i))
|
||||
|
||||
/-!
|
||||
|
||||
## Useful expansions.
|
||||
|
||||
-/
|
||||
|
||||
/-- The expansion of the Lorentz covariant metric in terms of basis vectors. -/
|
||||
lemma coMetric_expand : {Lorentz.coMetric | μ ν}ᵀ.tensor =
|
||||
lemma coMetric_basis_expand : {Lorentz.coMetric | μ ν}ᵀ.tensor =
|
||||
basisVector ![Color.down, Color.down] (fun _ => 0)
|
||||
- basisVector ![Color.down, Color.down] (fun _ => 1)
|
||||
- basisVector ![Color.down, Color.down] (fun _ => 2)
|
||||
|
@ -68,6 +74,32 @@ lemma coMetric_expand : {Lorentz.coMetric | μ ν}ᵀ.tensor =
|
|||
simp only [Fin.isValue, Lorentz.complexCoBasisFin4, Basis.coe_reindex, Function.comp_apply]
|
||||
rfl
|
||||
|
||||
/-- The expansion of the Lorentz contrvariant metric in terms of basis vectors. -/
|
||||
lemma contrMatrix_basis_expand : {Lorentz.contrMetric | μ ν}ᵀ.tensor =
|
||||
basisVector ![Color.up, Color.up] (fun _ => 0)
|
||||
- basisVector ![Color.up, Color.up] (fun _ => 1)
|
||||
- basisVector ![Color.up, Color.up] (fun _ => 2)
|
||||
- basisVector ![Color.up, Color.up] (fun _ => 3) := by
|
||||
simp only [Nat.succ_eq_add_one, Nat.reduceAdd, constTwoNode_tensor,
|
||||
Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V,
|
||||
Functor.id_obj, Fin.isValue]
|
||||
erw [Lorentz.contrMetric_apply_one, Lorentz.contrMetricVal_expand_tmul]
|
||||
simp only [Fin.isValue, map_sub]
|
||||
congr 1
|
||||
congr 1
|
||||
congr 1
|
||||
all_goals
|
||||
erw [pairIsoSep_tmul, basisVector]
|
||||
apply congrArg
|
||||
funext i
|
||||
fin_cases i
|
||||
all_goals
|
||||
simp only [Nat.succ_eq_add_one, Nat.reduceAdd, Fin.zero_eta, Fin.isValue, OverColor.mk_hom,
|
||||
cons_val_zero, Fin.cases_zero]
|
||||
change _ = Lorentz.complexContrBasisFin4 _
|
||||
simp only [Fin.isValue, Lorentz.complexContrBasisFin4, Basis.coe_reindex, Function.comp_apply]
|
||||
rfl
|
||||
|
||||
end complexLorentzTensor
|
||||
end Fermion
|
||||
end
|
||||
|
|
|
@ -203,6 +203,11 @@ def specialTypes : List (String × (Term → Term)) := [
|
|||
mkIdent ``Fermion.complexLorentzTensor,
|
||||
mkIdent ``Fermion.Color.down,
|
||||
mkIdent ``Fermion.Color.down, T]),
|
||||
("𝟙_ (Rep ℂ SL(2, ℂ)) ⟶ Lorentz.complexContr ⊗ Lorentz.complexContr", fun T =>
|
||||
Syntax.mkApp (mkIdent ``TensorTree.constTwoNodeE) #[
|
||||
mkIdent ``Fermion.complexLorentzTensor,
|
||||
mkIdent ``Fermion.Color.up,
|
||||
mkIdent ``Fermion.Color.up, T]),
|
||||
("𝟙_ (Rep ℂ SL(2, ℂ)) ⟶ Lorentz.complexContr ⊗ Fermion.leftHanded ⊗ Fermion.rightHanded", fun T =>
|
||||
Syntax.mkApp (mkIdent ``TensorTree.constThreeNodeE) #[
|
||||
mkIdent ``Fermion.complexLorentzTensor, mkIdent ``Fermion.Color.up,
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue