PhysLean/HepLean/PerturbationTheory/Wick/Algebra.lean

148 lines
6 KiB
Text
Raw Normal View History

2024-11-25 05:41:46 +00:00
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
2024-12-02 16:26:19 +00:00
import HepLean.PerturbationTheory.Wick.Species
2024-12-03 15:08:14 +00:00
import HepLean.Mathematics.SuperAlgebra.Basic
2024-12-04 13:37:23 +00:00
import HepLean.Meta.Notes.Basic
2024-11-25 05:41:46 +00:00
/-!
# Operator algebra
2024-11-25 05:41:46 +00:00
Currently this file is only for an example of Wick strings, correpsonding to a
theory with two complex scalar fields. The concepts will however generalize.
We will formally define the operator ring, in terms of the fields present in the theory.
## Futher reading
- https://physics.stackexchange.com/questions/258718/ and links therein
- Ryan Thorngren (https://physics.stackexchange.com/users/10336/ryan-thorngren), Fermions,
2024-11-29 06:38:52 +00:00
different species and (anti-)commutation rules, URL (version: 2019-02-20) :
https://physics.stackexchange.com/q/461929
2024-12-03 07:03:14 +00:00
- Tong, https://www.damtp.cam.ac.uk/user/tong/qft/qft.pdf
2024-11-25 05:41:46 +00:00
-/
2024-12-04 13:37:23 +00:00
note "
2024-12-06 06:52:44 +00:00
<h2>Operator algebra</h2>
The operator algebra is a super-algebra over the complex numbers, which acts on
the Hilbert space of the theory. A super-algebra is an algebra with a Z/2 grading.
To do pertubation theory in a QFT we need a need some basic properties of the operator algebra,
$A$.
"
2024-12-04 13:37:23 +00:00
2024-11-29 15:02:07 +00:00
namespace Wick
2024-11-25 05:41:46 +00:00
2024-12-04 13:37:23 +00:00
informal_definition_note WickAlgebra where
math :≈ "
2024-12-03 07:03:14 +00:00
Modifications of this may be needed.
A structure with the following data:
2024-12-03 15:08:14 +00:00
- A super algebra A.
2024-12-03 15:17:52 +00:00
- A map from `ψ : S.𝓯 × SpaceTime → A` where S.𝓯 are field colors.
- A map `ψc : S.𝓯 × SpaceTime → A`.
- A map `ψd : S.𝓯 × SpaceTime → A`.
Subject to the conditions:
2024-12-02 06:00:59 +00:00
- The sum of `ψc` and `ψd` is `ψ`.
2024-12-03 07:03:14 +00:00
- All maps land on homogeneous elements.
- Two fields super-commute if there colors are not dual to each other.
- The super-commutator of two fields is always in the
2024-12-03 07:03:14 +00:00
center of the algebra.
Asympotic states:
2024-12-03 15:17:52 +00:00
- `φc : S.𝓯 × SpaceTime → A`. The creation asympotic state (incoming).
- `φd : S.𝓯 × SpaceTime → A`. The destruction asympotic state (outgoing).
2024-12-03 07:51:10 +00:00
Subject to the conditions:
...
2024-12-03 07:03:14 +00:00
"
physics :≈ "This is defined to be an
abstraction of the notion of an operator algebra."
ref :≈ "https://physics.stackexchange.com/questions/24157/"
2024-12-03 15:30:47 +00:00
deps :≈ [``SuperAlgebra, ``SuperAlgebra.superCommuator]
2024-12-04 13:37:23 +00:00
informal_definition_note WickMonomial where
math :≈ "The type of elements of the Wick algebra which is a product of fields."
deps :≈ [``WickAlgebra]
namespace WickMonomial
informal_definition toWickAlgebra where
math :≈ "A function from WickMonomial to WickAlgebra which takes a monomial and
returns the product of the fields in the monomial."
deps :≈ [``WickAlgebra, ``WickMonomial]
informal_definition timeOrder where
math :≈ "A function from WickMonomial to WickAlgebra which takes a monomial and
returns the monomial with the fields time ordered, with the correct sign
2024-12-02 07:18:42 +00:00
determined by the Koszul sign factor.
If two fields have the same time, then their order is preserved e.g.
T(ψ₁(t)ψ₂(t)) = ψ₁(t)ψ₂(t)
and
T(ψ₂(t)ψ₁(t)) = ψ₂(t)ψ₁(t).
This allows us to make sense of the construction in e.g.
https://www.physics.purdue.edu/~clarkt/Courses/Physics662/ps/qftch32.pdf
which permits normal-ordering within time-ordering.
"
deps :≈ [``WickAlgebra, ``WickMonomial]
informal_definition normalOrder where
math :≈ "A function from WickMonomial to WickAlgebra which takes a monomial and
returns the element in `WickAlgebra` defined as follows
2024-12-02 06:00:59 +00:00
- The ψd fields are move to the right.
- The ψc fields are moved to the left.
- Othewise the order of the fields is preserved."
ref :≈ "https://www.imperial.ac.uk/media/imperial-college/research-centres-and-groups/theoretical-physics/msc/current/qft/handouts/qftwickstheorem.pdf"
deps :≈ [``WickAlgebra, ``WickMonomial]
end WickMonomial
2024-12-03 07:51:10 +00:00
informal_definition asymptoicContract where
2024-12-03 15:17:52 +00:00
math :≈ "Given two `i j : S.𝓯 × SpaceTime`, the super-commutator [φd(i), ψ(j)]."
2024-12-03 07:51:10 +00:00
ref :≈ "See e.g. http://www.dylanjtemples.com:82/solutions/QFT_Solution_I-6.pdf"
informal_definition contractAsymptotic where
2024-12-03 15:17:52 +00:00
math :≈ "Given two `i j : S.𝓯 × SpaceTime`, the super-commutator [ψ(i), φc(j)]."
2024-12-03 07:51:10 +00:00
informal_definition asymptoicContractAsymptotic where
2024-12-03 15:17:52 +00:00
math :≈ "Given two `i j : S.𝓯 × SpaceTime`, the super-commutator
2024-12-03 07:51:10 +00:00
[φd(i), φc(j)]."
informal_definition contraction where
2024-12-03 15:17:52 +00:00
math :≈ "Given two `i j : S.𝓯 × SpaceTime`, the element of WickAlgebra
defined by subtracting the normal ordering of `ψ i ψ j` from the time-ordering of
`ψ i ψ j`."
deps :≈ [``WickAlgebra, ``WickMonomial]
informal_lemma contraction_in_center where
math :≈ "The contraction of two fields is in the center of the algebra."
deps :≈ [``WickAlgebra, ``contraction]
informal_lemma contraction_non_dual_is_zero where
math :≈ "The contraction of two fields is zero if the fields are not dual to each other."
deps :≈ [``WickAlgebra, ``contraction]
informal_lemma timeOrder_single where
math :≈ "The time ordering of a single field is the normal ordering of that field."
2024-11-29 10:53:20 +00:00
proof :≈ "Follows from the definitions."
deps :≈ [``WickAlgebra, ``WickMonomial.timeOrder, ``WickMonomial.normalOrder]
informal_lemma timeOrder_pair where
math :≈ "The time ordering of two fields is the normal ordering of the fields plus the
contraction of the fields."
2024-11-29 10:53:20 +00:00
proof :≈ "Follows from the definition of contraction."
deps :≈ [``WickAlgebra, ``WickMonomial.timeOrder, ``WickMonomial.normalOrder,
``contraction]
2024-12-02 07:33:53 +00:00
informal_definition WickMap where
math :≈ "A linear map `vev` from the Wick algebra `A` to the underlying field such that
2024-12-04 16:24:23 +00:00
`vev(...ψd(t)) = 0` and `vev(ψc(t)...) = 0`."
2024-12-02 07:33:53 +00:00
physics :≈ "An abstraction of the notion of a vacuum expectation value, containing
the necessary properties for lots of theorems to hold."
deps :≈ [``WickAlgebra, ``WickMonomial]
2024-11-29 11:49:55 +00:00
2024-12-02 07:44:44 +00:00
informal_lemma normalOrder_wickMap where
math :≈ "Any normal ordering maps to zero under a Wick map."
deps :≈ [``WickMap, ``WickMonomial.normalOrder]
2024-11-29 15:02:07 +00:00
end Wick