docs: Field specification docs

This commit is contained in:
jstoobysmith 2025-02-06 10:06:05 +00:00
parent 3f2593b5ff
commit 8cc273fe38
13 changed files with 234 additions and 166 deletions

View file

@ -10,7 +10,9 @@ import Mathlib.Algebra.BigOperators.Group.Finset
-/
/-- The type specifing whether an operator is a creation or annihilation operator. -/
/-- The type `CreateAnnihilate` is the type containing two elements `create` and `annihilate`.
This type is used to specify if an operator is a creation or annihilation operator
or the sum thereof or intergral thereover etc. -/
inductive CreateAnnihilate where
| create : CreateAnnihilate
| annihilate : CreateAnnihilate

View file

@ -512,18 +512,18 @@ def anPart (φ : 𝓕.FieldOp) : 𝓕.FieldOpAlgebra := ι (anPartF φ)
lemma anPart_eq_ι_anPartF (φ : 𝓕.FieldOp) : anPart φ = ι (anPartF φ) := rfl
@[simp]
lemma anPart_negAsymp (φ : 𝓕.asymptoticDOF × (Fin 3 → )) :
lemma anPart_negAsymp (φ : (Σ f, 𝓕.AsymptoticLabel f) × (Fin 3 → )) :
anPart (FieldOp.inAsymp φ) = 0 := by
simp [anPart, anPartF]
@[simp]
lemma anPart_position (φ : 𝓕.positionDOF × SpaceTime) :
lemma anPart_position (φ : (Σ f, 𝓕.PositionLabel f) × SpaceTime) :
anPart (FieldOp.position φ) =
ofCrAnFieldOp ⟨FieldOp.position φ, CreateAnnihilate.annihilate⟩ := by
simp [anPart, ofCrAnFieldOp]
@[simp]
lemma anPart_posAsymp (φ : 𝓕.asymptoticDOF × (Fin 3 → )) :
lemma anPart_posAsymp (φ : (Σ f, 𝓕.AsymptoticLabel f) × (Fin 3 → )) :
anPart (FieldOp.outAsymp φ) = ofCrAnFieldOp ⟨FieldOp.outAsymp φ, ()⟩ := by
simp [anPart, ofCrAnFieldOp]
@ -533,18 +533,18 @@ def crPart (φ : 𝓕.FieldOp) : 𝓕.FieldOpAlgebra := ι (crPartF φ)
lemma crPart_eq_ι_crPartF (φ : 𝓕.FieldOp) : crPart φ = ι (crPartF φ) := rfl
@[simp]
lemma crPart_negAsymp (φ : 𝓕.asymptoticDOF × (Fin 3 → )) :
lemma crPart_negAsymp (φ : (Σ f, 𝓕.AsymptoticLabel f) × (Fin 3 → )) :
crPart (FieldOp.inAsymp φ) = ofCrAnFieldOp ⟨FieldOp.inAsymp φ, ()⟩ := by
simp [crPart, ofCrAnFieldOp]
@[simp]
lemma crPart_position (φ : 𝓕.positionDOF × SpaceTime) :
lemma crPart_position (φ : (Σ f, 𝓕.PositionLabel f) × SpaceTime) :
crPart (FieldOp.position φ) =
ofCrAnFieldOp ⟨FieldOp.position φ, CreateAnnihilate.create⟩ := by
simp [crPart, ofCrAnFieldOp]
@[simp]
lemma crPart_posAsymp (φ : 𝓕.asymptoticDOF × (Fin 3 → )) :
lemma crPart_posAsymp (φ : (Σ f, 𝓕.AsymptoticLabel f) × (Fin 3 → )) :
crPart (FieldOp.outAsymp φ) = 0 := by
simp [crPart]

View file

@ -131,18 +131,18 @@ def crPartF : 𝓕.FieldOp → 𝓕.FieldOpFreeAlgebra := fun φ =>
| FieldOp.outAsymp _ => 0
@[simp]
lemma crPartF_negAsymp (φ : 𝓕.asymptoticDOF × (Fin 3 → )) :
lemma crPartF_negAsymp (φ : (Σ f, 𝓕.AsymptoticLabel f) × (Fin 3 → )) :
crPartF (FieldOp.inAsymp φ) = ofCrAnOpF ⟨FieldOp.inAsymp φ, ()⟩ := by
simp [crPartF]
@[simp]
lemma crPartF_position (φ : 𝓕.positionDOF × SpaceTime) :
lemma crPartF_position (φ : (Σ f, 𝓕.PositionLabel f) × SpaceTime) :
crPartF (FieldOp.position φ) =
ofCrAnOpF ⟨FieldOp.position φ, CreateAnnihilate.create⟩ := by
simp [crPartF]
@[simp]
lemma crPartF_posAsymp (φ : 𝓕.asymptoticDOF × (Fin 3 → )) :
lemma crPartF_posAsymp (φ : (Σ f, 𝓕.AsymptoticLabel f) × (Fin 3 → )) :
crPartF (FieldOp.outAsymp φ) = 0 := by
simp [crPartF]
@ -156,18 +156,18 @@ def anPartF : 𝓕.FieldOp → 𝓕.FieldOpFreeAlgebra := fun φ =>
| FieldOp.outAsymp φ => ofCrAnOpF ⟨FieldOp.outAsymp φ, ()⟩
@[simp]
lemma anPartF_negAsymp (φ : 𝓕.asymptoticDOF × (Fin 3 → )) :
lemma anPartF_negAsymp (φ : (Σ f, 𝓕.AsymptoticLabel f) × (Fin 3 → )) :
anPartF (FieldOp.inAsymp φ) = 0 := by
simp [anPartF]
@[simp]
lemma anPartF_position (φ : 𝓕.positionDOF × SpaceTime) :
lemma anPartF_position (φ : (Σ f, 𝓕.PositionLabel f) × SpaceTime) :
anPartF (FieldOp.position φ) =
ofCrAnOpF ⟨FieldOp.position φ, CreateAnnihilate.annihilate⟩ := by
simp [anPartF]
@[simp]
lemma anPartF_posAsymp (φ : 𝓕.asymptoticDOF × (Fin 3 → )) :
lemma anPartF_posAsymp (φ : (Σ f, 𝓕.AsymptoticLabel f) × (Fin 3 → )) :
anPartF (FieldOp.outAsymp φ) = ofCrAnOpF ⟨FieldOp.outAsymp φ, ()⟩ := by
simp [anPartF]

View file

@ -29,50 +29,63 @@ These states carry the same field statistic as the field they are derived from.
-/
/-- A field specification is defined as a structure containing the basic data needed to write down
position and asymptotic field operators for a theory. It contains:
- A type `positionDOF` containing the degree-of-freedom in position-based field
operators (excluding space-time position). Thus a sutible (but not unique) choice
- Real-scalar fields correspond to a single element of `positionDOF`.
- Complex-scalar fields correspond to two elements of `positionDOF`, one for the field and one
for its conjugate.
- Dirac fermions correspond to eight elements of `positionDOF`. One for each Lorentz index of
the field and its conjugate. (These are not all independent)
- Weyl fermions correspond to four elements of `positionDOF`. One for each Lorentz index of the
field. (These are not all independent)
- A type `asymptoticDOF` containing the degree-of-freedom in asymptotic field operators. Thus a
sutible (but not unique) choice is
- Real-scalar fields correspond to a single element of `asymptoticDOF`.
- Complex-scalar fields correspond to two elements of `asymptoticDOF`, one for the field and one
for its conjugate.
- Dirac fermions correspond to four elements of `asymptoticDOF`, two for each type of spin.
- Weyl fermions correspond to two elements of `asymptoticDOF`, one for each spin.
- A specification `statisticsPos` on a `positionDOF` is Fermionic or Bosonic.
- A specification `statisticsAsym` on a `asymptoticDOF` is Fermionic or Bosonic.
/--
The structure `FieldSpecification` is defined to have the following content:
- A type `Field` whose elements are the constituent fields of the theory.
- For every field `f` in `Field`, a type `PositionLabel f` whose elements label the different
position operators associated with the field `f`. For example,
- For `f` a *real-scalar field*, `PositionLabel f` will have a unique element.
- For `f` a *complex-scalar field*, `PositionLabel f` will have two elements, one for the field
operator and one for its conjugate.
- For `f` a *Dirac fermion*, `PositionLabel f` will have eight elements, one for each Lorentz
index of the field and its conjugate.
- For `f` a *Weyl fermion*, `PositionLabel f` will have four elements, one for each Lorentz
index of the field and its conjugate.
- For every field `f` in `Field`, a type `AsymptoticLabel f` whose elements label the different
asymptotic based field operators associated with the field `f`. For example,
- For `f` a *real-scalar field*, `AsymptoticLabel f` will have a unique element.
- For `f` a *complex-scalar field*, `AsymptoticLabel f` will have two elements, one for the
field operator and one for its conjugate.
- For `f` a *Dirac fermion*, `AsymptoticLabel f` will have four elements, two for each spin.
- For `f` a *Weyl fermion*, `AsymptoticLabel f` will have two elements, one for each spin.
- For each field `f` in `Field`, a field statistic `statistic f` which classifying `f` as either
`bosonic` or `fermionic`.
-/
structure FieldSpecification where
/-- Degrees of freedom for position based field operators. -/
positionDOF : Type
/-- Degrees of freedom for asymptotic based field operators. -/
asymptoticDOF : Type
/-- The specification if the `positionDOF` are Fermionic or Bosonic. -/
statisticsPos : positionDOF → FieldStatistic
/-- The specification if the `asymptoticDOF` are Fermionic or Bosonic. -/
statisticsAsym : asymptoticDOF → FieldStatistic
/-- A type whose elements are the constituent fields of the theory. -/
Field : Type
/-- For every field `f` in `Field`, the type `PositionLabel f` has elements that label the
different position operators associated with the field `f`. -/
PositionLabel : Field → Type
/-- For every field `f` in `Field`, the type `AsymptoticLabel f` has elements that label
the different asymptotic based field operators associated with the field `f`. -/
AsymptoticLabel : Field → Type
/-- For every field `f` in `Field`, the field statistic `statistic f` classifies `f` as either
`bosonic` or `fermionic`. -/
statistic : Field → FieldStatistic
namespace FieldSpecification
variable (𝓕 : FieldSpecification)
/-- For a field specification `𝓕`, the type `𝓕.FieldOp` is defined such that every element of
`FieldOp` corresponds either to:
- an incoming asymptotic field operator `.inAsymp` specified by a field and a `3`-momentum.
- an position operator `.position` specified by a field and a point in spacetime.
- an outgoing asymptotic field operator `.outAsymp` specified by a field and a `3`-momentum.
- an incoming asymptotic field operator `.inAsymp` which is specified by
a field `f` in `𝓕.Field`, an element of `AsymptoticLabel f` (which specifies exactly
which asymptotic field operator associated with `f`) and a `3`-momentum.
- an position operator `.position` which is specified by
a field `f` in `𝓕.Field`, an element of `PositionLabel f` (which specifies exactly
which position field operator associated with `f`) and a element of `SpaceTime`.
- an outgoing asymptotic field operator `.outAsymp` which is specified by
a field `f` in `𝓕.Field`, an element of `AsymptoticLabel f` (which specifies exactly
which asymptotic field operator associated with `f`) and a `3`-momentum.
Note the use of `3`-momentum here rather then `4`-momentum. This is because the asymptotic states
have on-shell momenta.
-/
inductive FieldOp (𝓕 : FieldSpecification) where
| inAsymp : 𝓕.asymptoticDOF × (Fin 3 → ) → 𝓕.FieldOp
| position : 𝓕.positionDOF × SpaceTime → 𝓕.FieldOp
| outAsymp : 𝓕.asymptoticDOF × (Fin 3 → ) → 𝓕.FieldOp
| inAsymp : (Σ f, 𝓕.AsymptoticLabel f) × (Fin 3 → ) → 𝓕.FieldOp
| position : (Σ f, 𝓕.PositionLabel f) × SpaceTime → 𝓕.FieldOp
| outAsymp : (Σ f, 𝓕.AsymptoticLabel f) × (Fin 3 → ) → 𝓕.FieldOp
/-- The bool on `FieldOp` which is true only for position field operator. -/
@ -80,22 +93,34 @@ def statesIsPosition : 𝓕.FieldOp → Bool
| FieldOp.position _ => true
| _ => false
/-- The statistics associated to a field operator. -/
def statesStatistic : 𝓕.FieldOp → FieldStatistic := fun f =>
match f with
| FieldOp.inAsymp (a, _) => 𝓕.statisticsAsym a
| FieldOp.position (a, _) => 𝓕.statisticsPos a
| FieldOp.outAsymp (a, _) => 𝓕.statisticsAsym a
/-- For a field specification `𝓕`, `𝓕.fieldOpToField` is defined to take field operators
to their underlying field. -/
def fieldOpToField : 𝓕.FieldOp → 𝓕.Field
| FieldOp.inAsymp (f, _) => f.1
| FieldOp.position (f, _) => f.1
| FieldOp.outAsymp (f, _) => f.1
/-- The field statistics associated with a field operator. -/
scoped[FieldSpecification] notation 𝓕 "|>ₛ" φ => statesStatistic 𝓕 φ
/-- For a field specification `𝓕`, and an element `φ` of `𝓕.FieldOp`.
The field statistic `fieldOpStatistic φ` is defined to be the statistic associated with
the field underlying `φ`.
/-- The field statistics associated with a list field operators. -/
The following notation is used in relation to `fieldOpStatistic`:
- For `φ` an element of `𝓕.FieldOp`, `𝓕 |>ₛ φ` is `fieldOpStatistic φ`.
- For `φs` a list of `𝓕.FieldOp`, `𝓕 |>ₛ φs` is the product of `fieldOpStatistic φ` over
the list `φs`.
- For a function `f : Fin n → 𝓕.FieldOp` and a finset `a` of `Fin n`, `𝓕 |>ₛ ⟨f, a⟩` is the
product of `fieldOpStatistic (f i)` for all `i ∈ a`. -/
def fieldOpStatistic : 𝓕.FieldOp → FieldStatistic := 𝓕.statistic ∘ 𝓕.fieldOpToField
@[inherit_doc fieldOpStatistic]
scoped[FieldSpecification] notation 𝓕 "|>ₛ" φ => fieldOpStatistic 𝓕 φ
@[inherit_doc fieldOpStatistic]
scoped[FieldSpecification] notation 𝓕 "|>ₛ" φ => FieldStatistic.ofList
(statesStatistic 𝓕) φ
(fieldOpStatistic 𝓕) φ
/-- The field statistic associated with a finite set-/
@[inherit_doc fieldOpStatistic]
scoped[FieldSpecification] notation 𝓕 "|>ₛ" "⟨" f ","a "⟩"=> FieldStatistic.ofFinset
(statesStatistic 𝓕) f a
(fieldOpStatistic 𝓕) f a
end FieldSpecification

View file

@ -63,14 +63,15 @@ def fieldOpToCreateAnnihilateTypeCongr : {i j : 𝓕.FieldOp} → i = j →
| _, _, rfl => Equiv.refl _
/--
For a field specification `𝓕`, the type `𝓕.CrAnFieldOp` is defined such that every element
corresponds to
- an incoming asymptotic field operator `.inAsymp` and the unique element of `Unit`,
corresponding to the statement that an `inAsymp` state is a creation operator.
- a position operator `.position` and an element of `CreateAnnihilate`,
corresponding to either the creation part or the annihilation part of a position operator.
- an outgoing asymptotic field operator `.outAsymp` and the unique element of `Unit`,
corresponding to the statement that an `outAsymp` state is an annihilation operator.
For a field specification `𝓕`, elements in `𝓕.CrAnFieldOp`, the type
of creation and annihilation field operators, corresponds to
- an incoming asymptotic field operator `.inAsymp` in `𝓕.FieldOp`.
- a position operator `.position` in `𝓕.FieldOp` and an element of
`CreateAnnihilate` specifying the creation or annihilation part of that position operator.
- an outgoing asymptotic field operator `.outAsymp` in `𝓕.FieldOp`.
Note that the incoming and outgoing asymptotic field operators are implicitly creation and
annihilation operators respectively.
-/
def CrAnFieldOp : Type := Σ (s : 𝓕.FieldOp), 𝓕.fieldOpToCrAnType s
@ -89,15 +90,23 @@ def crAnFieldOpToCreateAnnihilate : 𝓕.CrAnFieldOp → CreateAnnihilate
| ⟨FieldOp.position _, CreateAnnihilate.annihilate⟩ => CreateAnnihilate.annihilate
| ⟨FieldOp.outAsymp _, _⟩ => CreateAnnihilate.annihilate
/-- Takes a `CrAnFieldOp` state to its corresponding fields statistic (bosonic or fermionic). -/
def crAnStatistics : 𝓕.CrAnFieldOp → FieldStatistic :=
𝓕.statesStatistic ∘ 𝓕.crAnFieldOpToFieldOp
/-- For a field specification `𝓕`, and an element `φ` in `𝓕.CrAnFieldOp`, the field
statistic `crAnStatistics φ` is defined to be the statistic associated with the field `𝓕.Field`
(or equivalently `𝓕.FieldOp`) underlying `φ`.
/-- The field statistic of a `CrAnFieldOp`. -/
The following notation is used in relation to `crAnStatistics`:
- For `φ` an element of `𝓕.CrAnFieldOp`, `𝓕 |>ₛ φ` is `crAnStatistics φ`.
- For `φs` a list of `𝓕.CrAnFieldOp`, `𝓕 |>ₛ φs` is the product of `crAnStatistics φ` over
the list `φs`.
-/
def crAnStatistics : 𝓕.CrAnFieldOp → FieldStatistic :=
𝓕.fieldOpStatistic ∘ 𝓕.crAnFieldOpToFieldOp
@[inherit_doc crAnStatistics]
scoped[FieldSpecification] notation 𝓕 "|>ₛ" φ =>
(crAnStatistics 𝓕) φ
/-- The field statistic of a list of `CrAnFieldOp`s. -/
@[inherit_doc crAnStatistics]
scoped[FieldSpecification] notation 𝓕 "|>ₛ" φ => FieldStatistic.ofList
(crAnStatistics 𝓕) φ
@ -108,8 +117,14 @@ scoped[FieldSpecification] infixl:80 "|>ᶜ" =>
remark notation_remark := "When working with a field specification `𝓕` we will use
some notation within doc-strings and in code. The main notation used is:
- In docstrings when field statistics occur in exchange signs we may drop the `𝓕 |>ₛ`.
- In docstrings we will often write lists of `FieldOp` or `CrAnFieldOp` `φs` as e.g. `φ₀…φₙ`,
which should be interpreted within the context in which it appears."
- In doc-strings when field statistics occur in exchange signs we may drop the `𝓕 |>ₛ _`.
- In doc-strings we will often write lists of `FieldOp` or `CrAnFieldOp` `φs` as e.g. `φ₀…φₙ`,
which should be interpreted within the context in which it appears.
Some examples:
- `𝓢(φ, φs)` corresponds to `𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φs)`
- `φ₀…φᵢ₋₁φᵢ₊₁…φₙ` corresponds to a (given) list `φs = φ₀…φₙ` with the element at the `i`th position
removed.
"
end FieldSpecification

View file

@ -126,7 +126,7 @@ lemma timeOrder_maxTimeField (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
the state of greatest time to the left).
We pick up a minus sign for every fermion paired crossed. -/
def timeOrderSign (φs : List 𝓕.FieldOp) : :=
Wick.koszulSign 𝓕.statesStatistic 𝓕.timeOrderRel φs
Wick.koszulSign 𝓕.fieldOpStatistic 𝓕.timeOrderRel φs
@[simp]
lemma timeOrderSign_nil : timeOrderSign (𝓕 := 𝓕) [] = 1 := by
@ -354,7 +354,7 @@ lemma crAnTimeOrderList_swap_eq_time {φ ψ : 𝓕.CrAnFieldOp}
lemma koszulSignInsert_crAnTimeOrderRel_crAnSection {φ : 𝓕.FieldOp} {ψ : 𝓕.CrAnFieldOp}
(h : ψ.1 = φ) : {φs : List 𝓕.FieldOp} → (ψs : CrAnSection φs) →
Wick.koszulSignInsert 𝓕.crAnStatistics 𝓕.crAnTimeOrderRel ψ ψs.1 =
Wick.koszulSignInsert 𝓕.statesStatistic 𝓕.timeOrderRel φ φs
Wick.koszulSignInsert 𝓕.fieldOpStatistic 𝓕.timeOrderRel φ φs
| [], ⟨[], h⟩ => by
simp [Wick.koszulSignInsert]
| φ' :: φs, ⟨ψ' :: ψs, h1⟩ => by

View file

@ -26,8 +26,13 @@ namespace FieldStatistic
variable {𝓕 : Type}
/-- Field statistics form a commuative group isomorphic to `ℤ₂` in which `bosonic` is the identity
and `fermionic` is the non-trivial element. -/
/-- The type `FieldStatistic` carries an instance of a commuative group in which
- `bosonic * bosonic = bosonic`
- `bosonic * fermionic = fermionic`
- `fermionic * bosonic = fermionic`
- `fermionic * fermionic = bosonic`
This group is isomorphic to `ℤ₂`. -/
@[simp]
instance : CommGroup FieldStatistic where
one := bosonic

View file

@ -24,11 +24,12 @@ namespace FieldStatistic
variable {𝓕 : Type}
/-- The exchange sign is the group homomorphism `FieldStatistic →* FieldStatistic →* `,
which on two field statistics `a` and `b` is defined to be
`-1` if both `a` and `b` are `fermionic` and `1` otherwise.
It corresponds to the sign that one picks up when swapping fields `φ₁φ₂ → φ₂φ₁` with
`φ₁` and `φ₂` of statistics `a` and `b` respectively.
/-- The exchange sign, `exchangeSign`, is defined as the group homomorphism
`FieldStatistic →* FieldStatistic →* `,
for which `exchangeSign a b` is `-1` if both `a` and `b` are `fermionic` and `1` otherwise.
The exchange sign is sign one picks up on exchanging an operator or field `φ₁` of statistic `a`
with one `φ₂` of statistic `b`, i.e. `φ₁φ₂ → φ₂φ₁`.
The notation `𝓢(a, b)` is used for the exchange sign of `a` and `b`. -/
def exchangeSign : FieldStatistic →* FieldStatistic →* where
toFun a :=

View file

@ -162,7 +162,7 @@ lemma join_singleton_left_signFinset_eq_filter {φs : List 𝓕.FieldOp}
left
rw [join_singleton_signFinset_eq_filter]
rw [mul_comm]
exact (ofFinset_filter_mul_neg 𝓕.statesStatistic _ _ _).symm
exact (ofFinset_filter_mul_neg 𝓕.fieldOpStatistic _ _ _).symm
/-- The difference in sign between `φsucΛ.sign` and the direct contribution of `φsucΛ` to
`(join (singleton h) φsucΛ)`. -/

View file

@ -85,7 +85,7 @@ lemma staticContract_insert_some_of_lt
· simp
simp only [smul_smul]
congr 1
have h1 : ofList 𝓕.statesStatistic (List.take (↑(φsΛ.uncontractedIndexEquiv.symm k))
have h1 : ofList 𝓕.fieldOpStatistic (List.take (↑(φsΛ.uncontractedIndexEquiv.symm k))
(List.map φs.get φsΛ.uncontractedList))
= (𝓕 |>ₛ ⟨φs.get, (Finset.filter (fun x => x < k) φsΛ.uncontracted)⟩) := by
simp only [ofFinset]

View file

@ -97,7 +97,7 @@ lemma timeContract_insert_some_of_lt
· simp
simp only [smul_smul]
congr 1
have h1 : ofList 𝓕.statesStatistic (List.take (↑(φsΛ.uncontractedIndexEquiv.symm k))
have h1 : ofList 𝓕.fieldOpStatistic (List.take (↑(φsΛ.uncontractedIndexEquiv.symm k))
(List.map φs.get φsΛ.uncontractedList))
= (𝓕 |>ₛ ⟨φs.get, (Finset.filter (fun x => x < k) φsΛ.uncontracted)⟩) := by
simp only [ofFinset]
@ -128,7 +128,7 @@ lemma timeContract_insert_some_of_not_lt
rw [timeContract_of_not_timeOrderRel, timeContract_of_timeOrderRel]
simp only [instCommGroup.eq_1, Algebra.smul_mul_assoc, smul_smul]
congr
have h1 : ofList 𝓕.statesStatistic (List.take (↑(φsΛ.uncontractedIndexEquiv.symm k))
have h1 : ofList 𝓕.fieldOpStatistic (List.take (↑(φsΛ.uncontractedIndexEquiv.symm k))
(List.map φs.get φsΛ.uncontractedList))
= (𝓕 |>ₛ ⟨φs.get, (Finset.filter (fun x => x < k) φsΛ.uncontracted)⟩) := by
simp only [ofFinset]