refactor: Spelling

This commit is contained in:
jstoobysmith 2025-02-07 15:43:59 +00:00
parent b82791d671
commit f8f1e1757f
25 changed files with 80 additions and 80 deletions

View file

@ -6,7 +6,7 @@ Authors: Joseph Tooby-Smith
import Mathlib.Algebra.BigOperators.Group.Finset
/-!
# Creation and annihlation parts of fields
# Creation and annihilation parts of fields
-/
@ -33,8 +33,8 @@ instance : Fintype CreateAnnihilate where
lemma eq_create_or_annihilate (φ : CreateAnnihilate) : φ = create φ = annihilate := by
cases φ <;> simp
/-- The normal ordering on creation and annihlation operators.
Under this relation, `normalOrder a b` is false only if `a` is annihlate and `b` is create. -/
/-- The normal ordering on creation and annihilation operators.
Under this relation, `normalOrder a b` is false only if `a` is annihilate and `b` is create. -/
def normalOrder : CreateAnnihilate → CreateAnnihilate → Prop
| create, _ => True
| annihilate, annihilate => True

View file

@ -219,7 +219,7 @@ lemma ι_normalOrderF_eq_of_equiv (a b : 𝓕.FieldOpFreeAlgebra) (h : a ≈ b)
simp only [LinearMap.mem_ker, ← map_sub]
exact ι_normalOrderF_zero_of_mem_ideal (a - b) h
/-- For a field specification `𝓕`, `normalOrder` is the linera map
/-- For a field specification `𝓕`, `normalOrder` is the linear map
`FieldOpAlgebra 𝓕 →ₗ[] FieldOpAlgebra 𝓕`

View file

@ -118,7 +118,7 @@ lemma crPart_mul_normalOrder (φ : 𝓕.FieldOp) (a : 𝓕.FieldOpAlgebra) :
-/
/-- For a field specfication `𝓕`, and `a` and `b` in `𝓕.FieldOpAlgebra` the normal ordering
/-- For a field specification `𝓕`, and `a` and `b` in `𝓕.FieldOpAlgebra` the normal ordering
of the super commutator of `a` and `b` vanishes. I.e. `𝓝([a,b]ₛ) = 0`. -/
@[simp]
lemma normalOrder_superCommute_eq_zero (a b : 𝓕.FieldOpAlgebra) :

View file

@ -7,7 +7,7 @@ import HepLean.PerturbationTheory.FieldSpecification.CrAnFieldOp
import HepLean.PerturbationTheory.FieldSpecification.CrAnSection
/-!
# Creation and annihlation free-algebra
# Creation and annihilation free-algebra
This module defines the creation and annihilation algebra for a field structure.
@ -142,7 +142,7 @@ lemma ofFieldOpListF_sum (φs : List 𝓕.FieldOp) :
-/
/-- The algebra map taking an element of the free-state algbra to
the part of it in the creation and annihlation free algebra
the part of it in the creation and annihilation free algebra
spanned by creation operators. -/
def crPartF : 𝓕.FieldOp → 𝓕.FieldOpFreeAlgebra := fun φ =>
match φ with
@ -201,7 +201,7 @@ lemma ofFieldOpF_eq_crPartF_add_anPartF (φ : 𝓕.FieldOp) :
/-!
## The basis of the creation and annihlation free-algebra.
## The basis of the creation and annihilation free-algebra.
-/

View file

@ -27,7 +27,7 @@ namespace FieldOpFreeAlgebra
noncomputable section
/-- For a field specification `𝓕`, `normalOrderF` is the linera map
/-- For a field specification `𝓕`, `normalOrderF` is the linear map
`FieldOpFreeAlgebra 𝓕 →ₗ[] FieldOpFreeAlgebra 𝓕`
defined by its action on the basis `ofCrAnListF φs`, taking `ofCrAnListF φs` to
`normalOrderSign φs • ofCrAnListF (normalOrderList φs)`.
@ -143,7 +143,7 @@ lemma normalOrderF_create_mul (φ : 𝓕.CrAnFieldOp)
lemma normalOrderF_ofCrAnListF_append_annihilate (φ : 𝓕.CrAnFieldOp)
(hφ : 𝓕 |>ᶜ φ = CreateAnnihilate.annihilate) (φs : List 𝓕.CrAnFieldOp) :
𝓝ᶠ(ofCrAnListF (φs ++ [φ])) = 𝓝ᶠ(ofCrAnListF φs) * ofCrAnOpF φ := by
rw [normalOrderF_ofCrAnListF, normalOrderSign_append_annihlate φ hφ φs,
rw [normalOrderF_ofCrAnListF, normalOrderSign_append_annihilate φ hφ φs,
normalOrderList_append_annihilate φ hφ φs, ofCrAnListF_append, ofCrAnListF_singleton,
normalOrderF_ofCrAnListF, smul_mul_assoc]
@ -189,18 +189,18 @@ lemma normalOrderF_mul_anPartF (φ : 𝓕.FieldOp) (a : FieldOpFreeAlgebra 𝓕)
The main result of this section is `normalOrderF_superCommuteF_annihilate_create`.
-/
lemma normalOrderF_swap_create_annihlate_ofCrAnListF_ofCrAnListF (φc φa : 𝓕.CrAnFieldOp)
lemma normalOrderF_swap_create_annihilate_ofCrAnListF_ofCrAnListF (φc φa : 𝓕.CrAnFieldOp)
(hφc : 𝓕 |>ᶜ φc = CreateAnnihilate.create) (hφa : 𝓕 |>ᶜ φa = CreateAnnihilate.annihilate)
(φs φs' : List 𝓕.CrAnFieldOp) :
𝓝ᶠ(ofCrAnListF φs' * ofCrAnOpF φc * ofCrAnOpF φa * ofCrAnListF φs) = 𝓢(𝓕 |>ₛ φc, 𝓕 |>ₛ φa) •
𝓝ᶠ(ofCrAnListF φs' * ofCrAnOpF φa * ofCrAnOpF φc * ofCrAnListF φs) := by
rw [mul_assoc, mul_assoc, ← ofCrAnListF_cons, ← ofCrAnListF_cons, ← ofCrAnListF_append]
rw [normalOrderF_ofCrAnListF, normalOrderSign_swap_create_annihlate φc φa hφc hφa]
rw [normalOrderList_swap_create_annihlate φc φa hφc hφa, ← smul_smul, ← normalOrderF_ofCrAnListF]
rw [normalOrderF_ofCrAnListF, normalOrderSign_swap_create_annihilate φc φa hφc hφa]
rw [normalOrderList_swap_create_annihilate φc φa hφc hφa, ← smul_smul, ← normalOrderF_ofCrAnListF]
rw [ofCrAnListF_append, ofCrAnListF_cons, ofCrAnListF_cons]
noncomm_ring
lemma normalOrderF_swap_create_annihlate_ofCrAnListF (φc φa : 𝓕.CrAnFieldOp)
lemma normalOrderF_swap_create_annihilate_ofCrAnListF (φc φa : 𝓕.CrAnFieldOp)
(hφc : 𝓕 |>ᶜ φc = CreateAnnihilate.create) (hφa : 𝓕 |>ᶜ φa = CreateAnnihilate.annihilate)
(φs : List 𝓕.CrAnFieldOp) (a : 𝓕.FieldOpFreeAlgebra) :
𝓝ᶠ(ofCrAnListF φs * ofCrAnOpF φc * ofCrAnOpF φa * a) = 𝓢(𝓕 |>ₛ φc, 𝓕 |>ₛ φa) •
@ -211,10 +211,10 @@ lemma normalOrderF_swap_create_annihlate_ofCrAnListF (φc φa : 𝓕.CrAnFieldOp
refine LinearMap.congr_fun (ofCrAnListFBasis.ext fun l ↦ ?_) a
simp only [mulLinearMap, LinearMap.coe_mk, AddHom.coe_mk, ofListBasis_eq_ofList,
LinearMap.coe_comp, Function.comp_apply, instCommGroup.eq_1]
rw [normalOrderF_swap_create_annihlate_ofCrAnListF_ofCrAnListF φc φa hφc hφa]
rw [normalOrderF_swap_create_annihilate_ofCrAnListF_ofCrAnListF φc φa hφc hφa]
rfl
lemma normalOrderF_swap_create_annihlate (φc φa : 𝓕.CrAnFieldOp)
lemma normalOrderF_swap_create_annihilate (φc φa : 𝓕.CrAnFieldOp)
(hφc : 𝓕 |>ᶜ φc = CreateAnnihilate.create) (hφa : 𝓕 |>ᶜ φa = CreateAnnihilate.annihilate)
(a b : 𝓕.FieldOpFreeAlgebra) :
𝓝ᶠ(a * ofCrAnOpF φc * ofCrAnOpF φa * b) = 𝓢(𝓕 |>ₛ φc, 𝓕 |>ₛ φa) •
@ -226,7 +226,7 @@ lemma normalOrderF_swap_create_annihlate (φc φa : 𝓕.CrAnFieldOp)
refine LinearMap.congr_fun (ofCrAnListFBasis.ext fun l ↦ ?_) _
simp only [mulLinearMap, ofListBasis_eq_ofList, LinearMap.coe_comp, Function.comp_apply,
LinearMap.flip_apply, LinearMap.coe_mk, AddHom.coe_mk, instCommGroup.eq_1, ← mul_assoc,
normalOrderF_swap_create_annihlate_ofCrAnListF φc φa hφc hφa]
normalOrderF_swap_create_annihilate_ofCrAnListF φc φa hφc hφa]
rfl
lemma normalOrderF_superCommuteF_create_annihilate (φc φa : 𝓕.CrAnFieldOp)
@ -235,7 +235,7 @@ lemma normalOrderF_superCommuteF_create_annihilate (φc φa : 𝓕.CrAnFieldOp)
𝓝ᶠ(a * [ofCrAnOpF φc, ofCrAnOpF φa]ₛca * b) = 0 := by
simp only [superCommuteF_ofCrAnOpF_ofCrAnOpF, instCommGroup.eq_1, Algebra.smul_mul_assoc]
rw [mul_sub, sub_mul, map_sub, ← smul_mul_assoc, ← mul_assoc, ← mul_assoc,
normalOrderF_swap_create_annihlate φc φa hφc hφa]
normalOrderF_swap_create_annihilate φc φa hφc hφa]
simp
lemma normalOrderF_superCommuteF_annihilate_create (φc φa : 𝓕.CrAnFieldOp)
@ -256,22 +256,22 @@ lemma normalOrderF_swap_crPartF_anPartF (φ φ' : 𝓕.FieldOp) (a b : FieldOpFr
| .outAsymp φ, _ => simp
| .position φ, .position φ' =>
simp only [crPartF_position, anPartF_position, instCommGroup.eq_1]
rw [normalOrderF_swap_create_annihlate]
rw [normalOrderF_swap_create_annihilate]
simp only [instCommGroup.eq_1, crAnStatistics, Function.comp_apply, crAnFieldOpToFieldOp_prod]
rfl; rfl
| .inAsymp φ, .outAsymp φ' =>
simp only [crPartF_negAsymp, anPartF_posAsymp, instCommGroup.eq_1]
rw [normalOrderF_swap_create_annihlate]
rw [normalOrderF_swap_create_annihilate]
simp only [instCommGroup.eq_1, crAnStatistics, Function.comp_apply, crAnFieldOpToFieldOp_prod]
rfl; rfl
| .inAsymp φ, .position φ' =>
simp only [crPartF_negAsymp, anPartF_position, instCommGroup.eq_1]
rw [normalOrderF_swap_create_annihlate]
rw [normalOrderF_swap_create_annihilate]
simp only [instCommGroup.eq_1, crAnStatistics, Function.comp_apply, crAnFieldOpToFieldOp_prod]
rfl; rfl
| .position φ, .outAsymp φ' =>
simp only [crPartF_position, anPartF_posAsymp, instCommGroup.eq_1]
rw [normalOrderF_swap_create_annihlate]
rw [normalOrderF_swap_create_annihilate]
simp only [instCommGroup.eq_1, crAnStatistics, Function.comp_apply, crAnFieldOpToFieldOp_prod]
rfl; rfl

View file

@ -34,8 +34,8 @@ In this module in addition to defining `CrAnFieldOp` we also define some maps:
namespace FieldSpecification
variable (𝓕 : FieldSpecification)
/-- To each field operator the specificaition of the type of creation and annihlation parts.
For asymptotic staes there is only one allowed part, whilst for position
/-- To each field operator the specificaition of the type of creation and annihilation parts.
For asymptotic states there is only one allowed part, whilst for position
field operator there is two. -/
def fieldOpToCrAnType : 𝓕.FieldOp → Type
| FieldOp.inAsymp _ => Unit
@ -75,14 +75,14 @@ annihilation operators respectively.
-/
def CrAnFieldOp : Type := Σ (s : 𝓕.FieldOp), 𝓕.fieldOpToCrAnType s
/-- The map from creation and annihlation field operator to their underlying states. -/
/-- The map from creation and annihilation field operator to their underlying states. -/
def crAnFieldOpToFieldOp : 𝓕.CrAnFieldOp → 𝓕.FieldOp := Sigma.fst
@[simp]
lemma crAnFieldOpToFieldOp_prod (s : 𝓕.FieldOp) (t : 𝓕.fieldOpToCrAnType s) :
𝓕.crAnFieldOpToFieldOp ⟨s, t⟩ = s := rfl
/-- The map from creation and annihlation states to the type `CreateAnnihilate`
/-- The map from creation and annihilation states to the type `CreateAnnihilate`
specifying if a state is a creation or an annihilation state. -/
def crAnFieldOpToCreateAnnihilate : 𝓕.CrAnFieldOp → CreateAnnihilate
| ⟨FieldOp.inAsymp _, _⟩ => CreateAnnihilate.create

View file

@ -6,7 +6,7 @@ Authors: Joseph Tooby-Smith
import HepLean.PerturbationTheory.FieldSpecification.CrAnFieldOp
/-!
# Creation and annihlation sections
# Creation and annihilation sections
In the module
`HepLean.PerturbationTheory.FieldSpecification.Basic`
@ -34,8 +34,8 @@ variable {𝓕 : FieldSpecification}
/-- The sections in `𝓕.CrAnFieldOp` over a list `φs : List 𝓕.FieldOp`.
In terms of physics, given some fields `φ₁...φₙ`, the different ways one can associate
each field as a `creation` or an `annilation` operator. E.g. the number of terms
`φ₁⁰φ₂¹...φₙ⁰` `φ₁¹φ₂¹...φₙ⁰` etc. If some fields are exclusively creation or annhilation
operators at this point (e.g. ansymptotic states) this is accounted for. -/
`φ₁⁰φ₂¹...φₙ⁰` `φ₁¹φ₂¹...φₙ⁰` etc. If some fields are exclusively creation or annihilation
operators at this point (e.g. asymptotic states) this is accounted for. -/
def CrAnSection (φs : List 𝓕.FieldOp) : Type :=
{ψs : List 𝓕.CrAnFieldOp // ψs.map 𝓕.crAnFieldOpToFieldOp = φs}
-- Π i, 𝓕.fieldOpToCreateAnnihilateType (φs.get i)
@ -107,7 +107,7 @@ def nilEquiv : CrAnSection (𝓕 := 𝓕) [] ≃ Unit where
right_inv _ := by
simp
/-- The creation and annihlation sections for a singleton list is given by
/-- The creation and annihilation sections for a singleton list is given by
a choice of `𝓕.fieldOpToCreateAnnihilateType φ`. If `φ` is a asymptotic state
there is no choice here, else there are two choices. -/
def singletonEquiv {φ : 𝓕.FieldOp} : CrAnSection [φ] ≃
@ -126,7 +126,7 @@ def singletonEquiv {φ : 𝓕.FieldOp} : CrAnSection [φ] ≃
simp only [head]
rfl
/-- An equivalence seperating the head of a creation and annhilation section
/-- An equivalence seperating the head of a creation and annihilation section
from the tail. -/
def consEquiv {φ : 𝓕.FieldOp} {φs : List 𝓕.FieldOp} : CrAnSection (φ :: φs) ≃
𝓕.fieldOpToCrAnType φ × CrAnSection φs where

View file

@ -43,7 +43,7 @@ instance (φ φ' : 𝓕.CrAnFieldOp) : Decidable (normalOrderRel φ φ') :=
-/
/-- For a field speciication `𝓕`, and a list `φs` of `𝓕.CrAnFieldOp`, `𝓕.normalOrderSign φs` is the
/-- For a field specification `𝓕`, and a list `φs` of `𝓕.CrAnFieldOp`, `𝓕.normalOrderSign φs` is the
sign corresponding to the number of `fermionic`-`fermionic` exchanges undertaken to normal-order
`φs` using the insertion sort algorithm. -/
def normalOrderSign (φs : List 𝓕.CrAnFieldOp) : :=
@ -91,14 +91,14 @@ lemma koszulSignInsert_append_annihilate (φ' φ : 𝓕.CrAnFieldOp)
dsimp only [List.cons_append, Wick.koszulSignInsert]
rw [koszulSignInsert_append_annihilate φ' φ hφ φs]
lemma normalOrderSign_append_annihlate (φ : 𝓕.CrAnFieldOp)
lemma normalOrderSign_append_annihilate (φ : 𝓕.CrAnFieldOp)
(hφ : 𝓕 |>ᶜ φ = CreateAnnihilate.annihilate) :
(φs : List 𝓕.CrAnFieldOp) →
normalOrderSign (φs ++ [φ]) = normalOrderSign φs
| [] => by simp
| φ' :: φs => by
dsimp only [List.cons_append, normalOrderSign, Wick.koszulSign]
have hi := normalOrderSign_append_annihlate φ hφ φs
have hi := normalOrderSign_append_annihilate φ hφ φs
dsimp only [normalOrderSign] at hi
rw [hi, koszulSignInsert_append_annihilate φ' φ hφ φs]
@ -117,7 +117,7 @@ lemma koszulSignInsert_annihilate_cons_create (φc φa : 𝓕.CrAnFieldOp)
rw [normalOrderRel, hφa, hφc]
simp [CreateAnnihilate.normalOrder]
lemma normalOrderSign_swap_create_annihlate_fst (φc φa : 𝓕.CrAnFieldOp)
lemma normalOrderSign_swap_create_annihilate_fst (φc φa : 𝓕.CrAnFieldOp)
(hφc : 𝓕 |>ᶜ φc = CreateAnnihilate.create)
(hφa : 𝓕 |>ᶜ φa = CreateAnnihilate.annihilate)
(φs : List 𝓕.CrAnFieldOp) :
@ -137,16 +137,16 @@ lemma koszulSignInsert_swap (φ φc φa : 𝓕.CrAnFieldOp) (φs φs' : List
= Wick.koszulSignInsert 𝓕.crAnStatistics normalOrderRel φ (φs ++ φc :: φa :: φs') :=
Wick.koszulSignInsert_eq_perm _ _ _ _ _ (List.Perm.append_left φs (List.Perm.swap φc φa φs'))
lemma normalOrderSign_swap_create_annihlate (φc φa : 𝓕.CrAnFieldOp)
lemma normalOrderSign_swap_create_annihilate (φc φa : 𝓕.CrAnFieldOp)
(hφc : 𝓕 |>ᶜ φc = CreateAnnihilate.create) (hφa : 𝓕 |>ᶜ φa = CreateAnnihilate.annihilate) :
(φs φs' : List 𝓕.CrAnFieldOp) → normalOrderSign (φs ++ φc :: φa :: φs') =
FieldStatistic.exchangeSign (𝓕.crAnStatistics φc) (𝓕.crAnStatistics φa) *
normalOrderSign (φs ++ φa :: φc :: φs')
| [], φs' => normalOrderSign_swap_create_annihlate_fst φc φa hφc hφa φs'
| [], φs' => normalOrderSign_swap_create_annihilate_fst φc φa hφc hφa φs'
| φ :: φs, φs' => by
rw [normalOrderSign]
dsimp only [List.cons_append, Wick.koszulSign, FieldStatistic.instCommGroup.eq_1]
rw [← normalOrderSign, normalOrderSign_swap_create_annihlate φc φa hφc hφa φs φs']
rw [← normalOrderSign, normalOrderSign_swap_create_annihilate φc φa hφc hφa φs φs']
rw [← mul_assoc, mul_comm _ (FieldStatistic.exchangeSign _ _), mul_assoc]
simp only [FieldStatistic.instCommGroup.eq_1, mul_eq_mul_left_iff]
apply Or.inl
@ -283,7 +283,7 @@ lemma normalOrderList_append_annihilate (φ : 𝓕.CrAnFieldOp)
dsimp only [normalOrderList] at hi
rw [hi, orderedInsert_append_annihilate φ' φ hφ]
lemma normalOrder_swap_create_annihlate_fst (φc φa : 𝓕.CrAnFieldOp)
lemma normalOrder_swap_create_annihilate_fst (φc φa : 𝓕.CrAnFieldOp)
(hφc : 𝓕 |>ᶜ φc = CreateAnnihilate.create)
(hφa : 𝓕 |>ᶜ φa = CreateAnnihilate.annihilate)
(φs : List 𝓕.CrAnFieldOp) :
@ -301,19 +301,19 @@ lemma normalOrder_swap_create_annihlate_fst (φc φa : 𝓕.CrAnFieldOp)
dsimp [CreateAnnihilate.normalOrder] at h
· rfl
lemma normalOrderList_swap_create_annihlate (φc φa : 𝓕.CrAnFieldOp)
lemma normalOrderList_swap_create_annihilate (φc φa : 𝓕.CrAnFieldOp)
(hφc : 𝓕 |>ᶜ φc = CreateAnnihilate.create)
(hφa : 𝓕 |>ᶜ φa = CreateAnnihilate.annihilate) :
(φs φs' : List 𝓕.CrAnFieldOp) →
normalOrderList (φs ++ φc :: φa :: φs') = normalOrderList (φs ++ φa :: φc :: φs')
| [], φs' => normalOrder_swap_create_annihlate_fst φc φa hφc hφa φs'
| [], φs' => normalOrder_swap_create_annihilate_fst φc φa hφc hφa φs'
| φ :: φs, φs' => by
dsimp only [List.cons_append, normalOrderList, List.insertionSort]
have hi := normalOrderList_swap_create_annihlate φc φa hφc hφa φs φs'
have hi := normalOrderList_swap_create_annihilate φc φa hφc hφa φs φs'
dsimp only [normalOrderList] at hi
rw [hi]
/-- For a list of creation and annihlation states, the equivalence between
/-- For a list of creation and annihilation states, the equivalence between
`Fin φs.length` and `Fin (normalOrderList φs).length` taking each position in `φs`
to it's corresponding position in the normal ordered list. This assumes that
we are using the insertion sort method.

View file

@ -103,7 +103,7 @@ lemma maxTimeFieldPos_lt_eraseMaxTimeField_length_succ (φ : 𝓕.FieldOp) (φs
exact maxTimeFieldPos_lt_length φ φs
/-- Given a list `φ :: φs` of states, the position of the left-most state of maximum
time as an eement of `Fin (eraseMaxTimeField φ φs).length.succ`.
time as an element of `Fin (eraseMaxTimeField φ φs).length.succ`.
As an example:
- for the list `[φ1(t = 4), φ2(t = 5), φ3(t = 3), φ4(t = 5)]` this would return `⟨1,...⟩`.
-/

View file

@ -143,7 +143,7 @@ lemma mul_eq_iff_eq_mul' (a b c : FieldStatistic) : a * b = c ↔ b = a * c := b
reduceCtorEq, fermionic_mul_bosonic, true_iff, iff_true]
all_goals rfl
/-- The field statistics of a list of fields is fermionic if ther is an odd number of fermions,
/-- The field statistics of a list of fields is fermionic if there is an odd number of fermions,
otherwise it is bosonic. -/
def ofList (s : 𝓕 → FieldStatistic) : (φs : List 𝓕) → FieldStatistic
| [] => bosonic

View file

@ -146,7 +146,7 @@ lemma fromInvolution_toInvolution (f : {f : Fin n → Fin n // Function.Involuti
simp only [fromInvolution_getDual?_isSome, ne_eq, Decidable.not_not] at h
exact id (Eq.symm h)
/-- The equivalence btween Wick contractions for `n` and involutions of `Fin n`.
/-- The equivalence between Wick contractions for `n` and involutions of `Fin n`.
The involution of `Fin n` associated with a Wick contraction `c : WickContraction n` as follows.
If `i : Fin n` is contracted in `c` then it is taken to its dual, otherwise
it is taken to itself. -/

View file

@ -188,7 +188,7 @@ lemma orderedInsert_eq_insertIdx_of_fin_list_sorted (l : List (Fin n)) (hl : l.S
-/
/-- Given a Wick contraction `c`, the ordered list of elements of `Fin n` which are not contracted,
i.e. do not appera anywhere in `c.1`. -/
i.e. do not appear anywhere in `c.1`. -/
def uncontractedList : List (Fin n) := List.filter (fun x => x ∈ c.uncontracted) (List.finRange n)
lemma uncontractedList_mem_iff (i : Fin n) :