refactor: Index notation

This commit is contained in:
jstoobysmith 2024-08-12 14:14:45 -04:00
parent 33b83d850b
commit 26ed9a1831
10 changed files with 1133 additions and 434 deletions

View file

@ -33,6 +33,9 @@ def getDualInOther? (i : Fin l.length) : Option (Fin l2.length) :=
def withDual : Finset (Fin l.length) :=
Finset.filter (fun i => (l.getDual? i).isSome) Finset.univ
def getDual (i : l.withDual) : Fin l.length :=
(l.getDual? i).get (by simpa [withDual] using i.2)
def withoutDual : Finset (Fin l.length) :=
Finset.filter (fun i => (l.getDual? i).isNone) Finset.univ
@ -95,12 +98,16 @@ def InverDual : Prop :=
namespace InverDual
variable {l : IndexList X} (h : l.InverDual) (i : Fin l.length)
end InverDual
def InvertDualOther : Prop :=
def InverDualOther : Prop :=
∀ i, (l.getDualInOther? l2 i).bind (l2.getDualInOther? l) = some i
∧ ∀ i, (l2.getDualInOther? l i).bind (l.getDualInOther? l2) = some i
section Color
variable {𝓒 : TensorColor}

View file

@ -198,13 +198,17 @@ lemma contr_contr : l.contr.contr = l.contr :=
# Relations on IndexListColor
-/
/-
l.getDual? ∘ Option.guard l.HasDualInSelf =
l.getDual?
-/
def Relabel (l1 l2 : IndexListColor 𝓒) : Prop :=
l1.length = l2.length ∧
∀ (h : l1.length = l2.length),
Option.map (Fin.cast h) ∘ l1.getDual? = l2.getDual? ∘ Fin.cast h ∧
l1.getDual? = Option.map (Fin.cast h.symm) ∘ l2.getDual? ∘ Fin.cast h ∧
l1.1.colorMap = l2.1.colorMap ∘ Fin.cast h
/-! TODO: Rewrite in terms of HasSingDualInOther. -/
def PermAfterContr (l1 l2 : IndexListColor 𝓒) : Prop :=
List.Perm (l1.contr.1.map Index.id) (l2.contr.1.map Index.id)
∧ ∀ (i : Fin l1.contr.1.length) (j : Fin l2.contr.1.length),
@ -212,7 +216,7 @@ def PermAfterContr (l1 l2 : IndexListColor 𝓒) : Prop :=
l1.contr.1.colorMap i = l2.contr.1.colorMap j
def AppendHasSingColorDualsInSelf (l1 l2 : IndexListColor 𝓒) : Prop :=
(l1.contr.1 ++ l2.contr.1).HasSingColorDualsInSelf
(l1.contr.1 ++ l2.contr.1).HasSingColorDualsInSelf
end IndexListColor

View file

@ -3,7 +3,8 @@ 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
-/
import HepLean.SpaceTime.LorentzTensor.IndexNotation.DualIndex
import HepLean.SpaceTime.LorentzTensor.IndexNotation.Indices.UniqueDualInOther
import HepLean.SpaceTime.LorentzTensor.IndexNotation.Indices.UniqueDual
import HepLean.SpaceTime.LorentzTensor.Basic
/-!
@ -25,22 +26,57 @@ variable (l l2 : IndexList X)
-/
instance : HAppend (IndexList X) (IndexList X) (IndexList X) :=
@instHAppendOfAppend (List (Index X)) List.instAppend
instance : HAppend (IndexList X) (IndexList X) (IndexList X) where
hAppend := fun l l2 => {val := l.val ++ l2.val}
def appendEquiv {l l2 : IndexList X} : Fin l.length ⊕ Fin l2.length ≃ Fin (l ++ l2).length :=
finSumFinEquiv.trans (Fin.castOrderIso (List.length_append _ _).symm).toEquiv
def appendInl : Fin l.length ↪ Fin (l ++ l2).length where
toFun := appendEquiv ∘ Sum.inl
inj' := by
intro i j h
simp [Function.comp] at h
exact h
def appendInr : Fin l2.length ↪ Fin (l ++ l2).length where
toFun := appendEquiv ∘ Sum.inr
inj' := by
intro i j h
simp [Function.comp] at h
exact h
@[simp]
lemma appendInl_appendEquiv :
(l.appendInl l2).trans appendEquiv.symm.toEmbedding =
{toFun := Sum.inl, inj' := Sum.inl_injective} := by
ext i
simp [appendInl]
@[simp]
lemma appendInr_appendEquiv :
(l.appendInr l2).trans appendEquiv.symm.toEmbedding =
{toFun := Sum.inr, inj' := Sum.inr_injective} := by
ext i
simp [appendInr]
@[simp]
lemma append_val {l l2 : IndexList X} : (l ++ l2).val = l.val ++ l2.val := by
rfl
@[simp]
lemma idMap_append_inl {l l2 : IndexList X} (i : Fin l.length) :
(l ++ l2).idMap (appendEquiv (Sum.inl i)) = l.idMap i := by
simp [appendEquiv, idMap]
rw [List.getElem_append_left]
rfl
@[simp]
lemma idMap_append_inr {l l2 : IndexList X} (i : Fin l2.length) :
(l ++ l2).idMap (appendEquiv (Sum.inr i)) = l2.idMap i := by
simp [appendEquiv, idMap]
simp [appendEquiv, idMap, IndexList.length]
rw [List.getElem_append_right]
simp
omega
@ -49,13 +85,13 @@ lemma idMap_append_inr {l l2 : IndexList X} (i : Fin l2.length) :
@[simp]
lemma colorMap_append_inl {l l2 : IndexList X} (i : Fin l.length) :
(l ++ l2).colorMap (appendEquiv (Sum.inl i)) = l.colorMap i := by
simp [appendEquiv, colorMap]
simp [appendEquiv, colorMap, IndexList.length]
rw [List.getElem_append_left]
@[simp]
lemma colorMap_append_inr {l l2 : IndexList X} (i : Fin l2.length) :
(l ++ l2).colorMap (appendEquiv (Sum.inr i)) = l2.colorMap i := by
simp [appendEquiv, colorMap]
simp [appendEquiv, colorMap, IndexList.length]
rw [List.getElem_append_right]
simp
omega
@ -146,6 +182,60 @@ lemma inr_mem_withDual_append_iff (i : Fin l2.length) :
def appendInlWithDual (i : l.withDual) : (l ++ l2).withDual :=
⟨appendEquiv (Sum.inl i), by simp⟩
lemma append_withDual : (l ++ l2).withDual =
(Finset.map (l.appendInl l2) (l.withDual l.withDualInOther l2))
(Finset.map (l.appendInr l2) (l2.withDual l2.withDualInOther l)) := by
ext i
obtain ⟨k, hk⟩ := appendEquiv.surjective i
subst hk
match k with
| Sum.inl k =>
simp
apply Iff.intro
intro h
· apply Or.inl
use k
simp_all [appendInl]
· intro h
cases' h with h h
· obtain ⟨j, hj⟩ := h
simp [appendInl] at hj
have hj2 := hj.2
subst hj2
exact hj.1
· obtain ⟨j, hj⟩ := h
simp [appendInr] at hj
| Sum.inr k =>
simp
apply Iff.intro
intro h
· apply Or.inr
use k
simp_all [appendInr]
· intro h
cases' h with h h
· obtain ⟨j, hj⟩ := h
simp [appendInl] at hj
· obtain ⟨j, hj⟩ := h
simp [appendInr] at hj
have hj2 := hj.2
subst hj2
exact hj.1
lemma append_withDual_disjSum : (l ++ l2).withDual =
Equiv.finsetCongr appendEquiv
((l.withDual l.withDualInOther l2).disjSum
(l2.withDual l2.withDualInOther l)) := by
rw [← Equiv.symm_apply_eq]
simp [append_withDual]
rw [Finset.map_union]
rw [Finset.map_map, Finset.map_map]
ext1 a
cases a with
| inl val => simp
| inr val_1 => simp
/-!
@ -173,7 +263,8 @@ lemma getDual?_append_inl_withDual (i : l.withDual) :
simp only [appendEquiv, Equiv.trans_apply, finSumFinEquiv_apply_left, RelIso.coe_fn_toEquiv,
Fin.castOrderIso_apply, finSumFinEquiv_apply_right, ge_iff_le]
rw [Fin.le_def]
simp only [Fin.coe_cast, Fin.coe_castAdd, Fin.coe_natAdd]
simp only [length, append_val, RelIso.coe_fn_toEquiv, Fin.castOrderIso_apply, Fin.coe_cast,
Fin.coe_castAdd, Fin.coe_natAdd]
omega
@[simp]
@ -203,7 +294,8 @@ lemma getDual?_append_inl_not_withDual (i : Fin l.length) (hi : i ∉ l.withDual
rw [Fin.le_def]
have h1 := l.getDualInOther?_eq_some_getDualInOther! l2 ⟨i, by simp_all⟩
rw [getDualInOther?, Fin.find_eq_some_iff] at h1
simp only [Fin.coe_cast, Fin.coe_natAdd, add_le_add_iff_left, Fin.val_fin_le, ge_iff_le]
simp only [length, append_val, RelIso.coe_fn_toEquiv, Fin.castOrderIso_apply, Fin.coe_cast,
Fin.coe_natAdd, add_le_add_iff_left, Fin.val_fin_le, ge_iff_le]
refine h1.2 k (by simpa using hj)
@[simp]
@ -236,7 +328,8 @@ lemma getDual?_append_inr_withDualInOther (i : l2.withDualInOther l) :
simp only [appendEquiv, Equiv.trans_apply, finSumFinEquiv_apply_left, RelIso.coe_fn_toEquiv,
Fin.castOrderIso_apply, finSumFinEquiv_apply_right, ge_iff_le]
rw [Fin.le_def]
simp only [Fin.coe_cast, Fin.coe_castAdd, Fin.coe_natAdd]
simp only [length, append_val, RelIso.coe_fn_toEquiv, Fin.castOrderIso_apply, Fin.coe_cast,
Fin.coe_castAdd, Fin.coe_natAdd]
omega
@[simp]
@ -263,7 +356,7 @@ lemma getDual?_append_inr_not_withDualInOther (i : Fin l2.length) (hi : i ∉ l2
simp [withDualInOther, getDualInOther?, Fin.isSome_find_iff]
exact ⟨k, hj⟩
| Sum.inr k =>
simp [appendEquiv]
simp [appendEquiv, IndexList.length]
rw [Fin.le_def]
simp
have h2 := l2.getDual?_eq_some_getDual! ⟨i, by simp_all⟩
@ -429,7 +522,7 @@ lemma append_inr_mem_withUniqueDual_of_append_inl (i : Fin l.length)
rw [l.getDualInOther?_eq_none_on_not_mem _ _ ho]
simp only [Option.map_none', Option.none_or]
rw [← Option.map_some', ← Option.map_some', Option.map_map, h2]
simp [some_dual!_eq_gual?]
simp [some_dual!_eq_getDual?]
· have ho := (l.append_inl_not_mem_withDual_of_withDualInOther l2 i h').mpr.mt hs
simp at ho
rw [l.getDualInOther?_eq_some_getDualInOther! l2 ⟨i, ho⟩] at h2 ⊢
@ -573,14 +666,14 @@ lemma append_inl_mem_withUniqueDual_of_append_inr (i : Fin l.length)
refine (l2.append_inr_not_mem_withDual_of_withDualInOther l i h).mp ?_
exact (l.mem_withDual_of_withUniqueDual ⟨i, hs⟩)
lemma append_mem_withUniqueDual_symm (i : Fin l.length) :
appendEquiv (Sum.inl i) ∈ (l ++ l2).withUniqueDual ↔
appendEquiv (Sum.inr i) ∈ (l2 ++ l).withUniqueDual := by
apply Iff.intro
· intro h'
exact append_inr_mem_withUniqueDual_of_append_inl l l2 i h'
· intro h'
exact append_inl_mem_withUniqueDual_of_append_inr l l2 i h'
lemma append_mem_withUniqueDual_symm (i : Fin l.length) :
appendEquiv (Sum.inl i) ∈ (l ++ l2).withUniqueDual ↔
appendEquiv (Sum.inr i) ∈ (l2 ++ l).withUniqueDual := by
apply Iff.intro
· intro h'
exact append_inr_mem_withUniqueDual_of_append_inl l l2 i h'
· intro h'
exact append_inl_mem_withUniqueDual_of_append_inr l l2 i h'
@[simp]
lemma append_inr_mem_withUniqueDual_iff (i : Fin l2.length) :
@ -590,6 +683,116 @@ lemma append_inr_mem_withUniqueDual_iff (i : Fin l2.length) :
simp
lemma append_withUniqueDual : (l ++ l2).withUniqueDual =
Finset.map (l.appendInl l2) ((l.withUniqueDual ∩ (l.withDualInOther l2)ᶜ) l.withUniqueDualInOther l2)
Finset.map (l.appendInr l2) ((l2.withUniqueDual ∩ (l2.withDualInOther l)ᶜ) l2.withUniqueDualInOther l) := by
rw [Finset.ext_iff]
intro j
obtain ⟨k, hk⟩ := appendEquiv.surjective j
subst hk
match k with
| Sum.inl k =>
simp only [append_inl_mem_withUniqueDual_iff, Finset.mem_union]
apply Iff.intro
· intro h
apply Or.inl
simp
use k
simp [appendInl]
by_cases hk : k ∈ l.withUniqueDualInOther l2
simp_all
have hk' := h.mpr hk
exact Or.symm (Or.inr hk')
· intro h
simp at h
cases' h with h h
· obtain ⟨j, hj⟩ := h
have hjk : j = k := by
simp [appendInl] at hj
exact hj.2
subst hjk
have hj1 := hj.1
cases' hj1 with hj1 hj1
· simp_all
by_contra hn
have h' := l.mem_withDualInOther_of_withUniqueDualInOther l2 ⟨j, hn⟩
simp_all
· simp_all
intro _
exact l.mem_withDualInOther_of_withUniqueDualInOther l2 ⟨j, hj1⟩
· obtain ⟨j, hj⟩ := h
simp [appendInr] at hj
| Sum.inr k =>
simp only [append_inr_mem_withUniqueDual_iff, Finset.mem_union]
apply Iff.intro
· intro h
apply Or.inr
simp
use k
simp [appendInr]
by_cases hk : k ∈ l2.withUniqueDualInOther l
simp_all
have hk' := h.mpr hk
exact Or.symm (Or.inr hk')
· intro h
simp at h
cases' h with h h
· obtain ⟨j, hj⟩ := h
simp [appendInl] at hj
· obtain ⟨j, hj⟩ := h
have hjk : j = k := by
simp [appendInr] at hj
exact hj.2
subst hjk
have hj1 := hj.1
cases' hj1 with hj1 hj1
· simp_all
by_contra hn
have h' := l2.mem_withDualInOther_of_withUniqueDualInOther l ⟨j, hn⟩
simp_all
· simp_all
intro _
exact l2.mem_withDualInOther_of_withUniqueDualInOther l ⟨j, hj1⟩
lemma append_withUniqueDual_disjSum : (l ++ l2).withUniqueDual =
Equiv.finsetCongr appendEquiv
(((l.withUniqueDual ∩ (l.withDualInOther l2)ᶜ) l.withUniqueDualInOther l2).disjSum
((l2.withUniqueDual ∩ (l2.withDualInOther l)ᶜ) l2.withUniqueDualInOther l)) := by
rw [← Equiv.symm_apply_eq]
simp [append_withUniqueDual]
rw [Finset.map_union]
rw [Finset.map_map, Finset.map_map]
ext1 a
cases a with
| inl val => simp
| inr val_1 => simp
lemma append_withDual_eq_withUniqueDual_iff :
(l ++ l2).withUniqueDual = (l ++ l2).withDual ↔
((l.withUniqueDual ∩ (l.withDualInOther l2)ᶜ) l.withUniqueDualInOther l2)
= l.withDual l.withDualInOther l2
∧ (l2.withUniqueDual ∩ (l2.withDualInOther l)ᶜ) l2.withUniqueDualInOther l
= l2.withDual l2.withDualInOther l := by
rw [append_withUniqueDual_disjSum, append_withDual_disjSum]
simp
have h (s s' : Finset (Fin l.length)) (t t' : Finset (Fin l2.length)) :
s.disjSum t = s'.disjSum t' ↔ s = s' ∧ t = t' := by
simp [Finset.ext_iff]
rw [h]
lemma append_withDual_eq_withUniqueDual_symm :
(l ++ l2).withUniqueDual = (l ++ l2).withDual ↔
(l2 ++ l).withUniqueDual = (l2 ++ l).withDual := by
rw [append_withDual_eq_withUniqueDual_iff, append_withDual_eq_withUniqueDual_iff]
apply Iff.intro
· intro a
simp_all only [and_self]
· intro a
simp_all only [and_self]
end IndexList
end IndexNotation

View file

@ -175,7 +175,8 @@ end Index
-/
/-- The type of lists of indices. -/
def IndexList : Type := List (Index X)
structure IndexList where
val : List (Index X)
namespace IndexList
@ -184,15 +185,20 @@ variable {X : Type} [IndexNotation X] [Fintype X] [DecidableEq X]
variable (l : IndexList X)
/-- The number of indices in an index list. -/
def numIndices : Nat := l.length
def length : := l.val.length
lemma ext (h : l.val = l2.val) : l = l2 := by
cases l
cases l2
simp_all
/-- The map of from `Fin s.numIndices` into colors associated to an index list. -/
def colorMap : Fin l.numIndices → X :=
fun i => (l.get i).toColor
def colorMap : Fin l.length → X :=
fun i => (l.val.get i).toColor
/-- The map of from `Fin s.numIndices` into the natural numbers associated to an index list. -/
def idMap : Fin l.length → Nat :=
fun i => (l.get i).id
fun i => (l.val.get i).id
lemma idMap_cast {l1 l2 : IndexList X} (h : l1 = l2) (i : Fin l1.length) :
l1.idMap i = l2.idMap (Fin.cast (by rw [h]) i) := by
@ -201,13 +207,13 @@ lemma idMap_cast {l1 l2 : IndexList X} (h : l1 = l2) (i : Fin l1.length) :
/-- Given a list of indices a subset of `Fin l.numIndices × Index X`
of pairs of positions in `l` and the corresponding item in `l`. -/
def toPosSet (l : IndexList X) : Set (Fin l.numIndices × Index X) :=
{(i, l.get i) | i : Fin l.numIndices}
def toPosSet (l : IndexList X) : Set (Fin l.length × Index X) :=
{(i, l.val.get i) | i : Fin l.length}
/-- Equivalence between `toPosSet` and `Fin l.numIndices`. -/
def toPosSetEquiv (l : IndexList X) : l.toPosSet ≃ Fin l.numIndices where
def toPosSetEquiv (l : IndexList X) : l.toPosSet ≃ Fin l.length where
toFun := fun x => x.1.1
invFun := fun x => ⟨(x, l.get x), by simp [toPosSet]⟩
invFun := fun x => ⟨(x, l.val.get x), by simp [toPosSet]⟩
left_inv x := by
have hx := x.prop
simp [toPosSet] at hx
@ -232,21 +238,21 @@ instance : Fintype l.toPosSet where
intro x
simp_all only [Finset.mem_map_equiv, Equiv.symm_symm, Finset.mem_univ]
/-- Given a list of indices a finite set of `Fin l.numIndices × Index X`
/-- Given a list of indices a finite set of `Fin l.length × Index X`
of pairs of positions in `l` and the corresponding item in `l`. -/
def toPosFinset (l : IndexList X) : Finset (Fin l.numIndices × Index X) :=
def toPosFinset (l : IndexList X) : Finset (Fin l.length × Index X) :=
l.toPosSet.toFinset
/-- The construction of a list of indices from a map
from `Fin n` to `Index X`. -/
def fromFinMap {n : } (f : Fin n → Index X) : IndexList X :=
(Fin.list n).map f
def fromFinMap {n : } (f : Fin n → Index X) : IndexList X where
val := (Fin.list n).map f
@[simp]
lemma fromFinMap_numIndices {n : } (f : Fin n → Index X) :
(fromFinMap f).numIndices = n := by
simp [fromFinMap, numIndices]
(fromFinMap f).length = n := by
simp [fromFinMap, length]
end IndexList

View file

@ -0,0 +1,89 @@
/-
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
-/
import HepLean.SpaceTime.LorentzTensor.IndexNotation.Indices.UniqueDual
import HepLean.SpaceTime.LorentzTensor.IndexNotation.Indices.Append
import HepLean.SpaceTime.LorentzTensor.Basic
import Init.Data.List.Lemmas
/-!
# Index lists with color conditions
Here we consider `IndexListColor` which is a subtype of all lists of indices
on those where the elements to be contracted have consistent colors with respect to
a Tensor Color structure.
-/
namespace IndexNotation
variable (𝓒 : TensorColor)
variable [IndexNotation 𝓒.Color] [Fintype 𝓒.Color] [DecidableEq 𝓒.Color]
structure ColorIndexList (𝓒 : TensorColor) [IndexNotation 𝓒.Color] extends IndexList 𝓒.Color where
unique_duals : toIndexList.withDual = toIndexList.withUniqueDual
dual_color : (Option.map toIndexList.colorMap) ∘ toIndexList.getDual?
= (Option.map (𝓒.τ ∘ toIndexList.colorMap)) ∘
Option.guard (fun i => (toIndexList.getDual? i).isSome)
namespace ColorIndexList
variable {𝓒 : TensorColor} [IndexNotation 𝓒.Color] [Fintype 𝓒.Color] [DecidableEq 𝓒.Color]
variable (l l2 : ColorIndexList 𝓒)
@[ext]
lemma ext {l l' : ColorIndexList 𝓒} (h : l.val = l'.val) : l = l' := by
cases l
cases l'
simp_all
apply IndexList.ext
exact h
/-!
## Contracting an `ColorIndexList`
-/
def contr : ColorIndexList 𝓒 where
toIndexList := l.toIndexList.contrIndexList
unique_duals := by
simp
dual_color := by
funext i
simp [Option.guard]
@[simp]
lemma contr_contr : l.contr.contr = l.contr := by
apply ext
simp [contr]
/-!
## Contract equiv
-/
def contrEquiv :
(l.withUniqueDualLT ⊕ l.withUniqueDualLT) ⊕ Fin l.contr.length ≃ Fin l.length :=
(Equiv.sumCongr (l.withUniqueLTGTEquiv) (Equiv.refl _)).trans <|
(Equiv.sumCongr (Equiv.subtypeEquivRight (by
simp [l.unique_duals])) (Fin.castOrderIso l.contrIndexList_length).toEquiv).trans <|
l.dualEquiv
/-!
## Append
-/
def append (h : (l.toIndexList ++ l2.toIndexList).withUniqueDual =
(l.toIndexList ++ l2.toIndexList).withDual ) : ColorIndexList 𝓒 := by
end ColorIndexList
end IndexNotation

View file

@ -3,7 +3,7 @@ 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
-/
import HepLean.SpaceTime.LorentzTensor.IndexNotation.Basic
import HepLean.SpaceTime.LorentzTensor.IndexNotation.Indices.Basic
import HepLean.SpaceTime.LorentzTensor.Basic
/-!
@ -32,13 +32,6 @@ def AreDualInSelf (i j : Fin l.length) : Prop :=
instance (i j : Fin l.length) : Decidable (l.AreDualInSelf i j) :=
instDecidableAnd
def AreDualInOther (i : Fin l.length) (j : Fin l2.length) :
Prop := l.idMap i = l2.idMap j
instance {l : IndexList X} {l2 : IndexList X} (i : Fin l.length) (j : Fin l2.length) :
Decidable (l.AreDualInOther l2 i j) := (l.idMap i).decEq (l2.idMap j)
namespace AreDualInSelf
variable {l l2 : IndexList X} {i j : Fin l.length}
@ -54,16 +47,6 @@ lemma self_false (i : Fin l.length) : ¬ l.AreDualInSelf i i := by
end AreDualInSelf
namespace AreDualInOther
variable {l l2 : IndexList X} {i : Fin l.length} {j : Fin l2.length}
@[symm]
lemma symm (h : l.AreDualInOther l2 i j) : l2.AreDualInOther l j i := by
rw [AreDualInOther] at h ⊢
exact h.symm
end AreDualInOther
/-!
@ -76,24 +59,6 @@ end AreDualInOther
def getDual? (i : Fin l.length) : Option (Fin l.length) :=
Fin.find (fun j => l.AreDualInSelf i j)
/-- Given an `i`, if a dual exists in the other list,
outputs the first such dual, otherwise outputs `none`. -/
def getDualInOther? (i : Fin l.length) : Option (Fin l2.length) :=
Fin.find (fun j => l.AreDualInOther l2 i j)
/-!
## With dual self.
-/
def withDual : Finset (Fin l.length) :=
Finset.filter (fun i => (l.getDual? i).isSome) Finset.univ
lemma mem_withDual_iff_exists : i ∈ l.withDual ↔ ∃ j, l.AreDualInSelf i j := by
simp [withDual, Finset.mem_filter, Finset.mem_univ, getDual?]
rw [Fin.isSome_find_iff]
lemma getDual?_of_areDualInSelf (h : l.AreDualInSelf i j) :
l.getDual? j = i l.getDual? i = j l.getDual? j = l.getDual? i := by
have h3 : (l.getDual? i).isSome := by
@ -135,6 +100,26 @@ lemma getDual?_of_areDualInSelf (h : l.AreDualInSelf i j) :
simp_all [AreDualInSelf]
exact hk.2 k' hik''
@[simp]
lemma getDual?_neq_self (i : Fin l.length) : ¬ l.getDual? i = some i := by
intro h
simp [getDual?] at h
rw [Fin.find_eq_some_iff] at h
simp [AreDualInSelf] at h
/-!
## Indices which have duals.
-/
def withDual : Finset (Fin l.length) :=
Finset.filter (fun i => (l.getDual? i).isSome) Finset.univ
lemma mem_withDual_iff_exists : i ∈ l.withDual ↔ ∃ j, l.AreDualInSelf i j := by
simp [withDual, Finset.mem_filter, Finset.mem_univ, getDual?]
rw [Fin.isSome_find_iff]
def getDual! (i : l.withDual) : Fin l.length :=
(l.getDual? i).get (by simpa [withDual] using i.2)
@ -146,7 +131,7 @@ lemma getDual?_eq_none_on_not_mem (i : Fin l.length) (h : i ∉ l.withDual) :
l.getDual? i = none := by
simpa [withDual, getDual?, Fin.find_eq_none_iff] using h
lemma some_dual!_eq_gual? (i : l.withDual) : some (l.getDual! i) = l.getDual? i := by
lemma some_dual!_eq_getDual? (i : l.withDual) : some (l.getDual! i) = l.getDual? i := by
rw [getDual?_eq_some_getDual!]
lemma areDualInSelf_getDual! (i : l.withDual) : l.AreDualInSelf i (l.getDual! i) := by
@ -170,387 +155,133 @@ lemma getDual_id (i : l.withDual) :
l.idMap (l.getDual i) = l.idMap i := by
simp [getDual]
/-!
## With dual other.
-/
def withDualInOther : Finset (Fin l.length) :=
Finset.filter (fun i => (l.getDualInOther? l2 i).isSome) Finset.univ
lemma mem_withInDualOther_iff_exists :
i ∈ l.withDualInOther l2 ↔ ∃ (j : Fin l2.length), l.AreDualInOther l2 i j := by
simp [withDualInOther, Finset.mem_filter, Finset.mem_univ, getDualInOther?]
rw [Fin.isSome_find_iff]
def getDualInOther! (i : l.withDualInOther l2) : Fin l2.length :=
(l.getDualInOther? l2 i).get (by simpa [withDualInOther] using i.2)
lemma getDualInOther?_eq_some_getDualInOther! (i : l.withDualInOther l2) :
l.getDualInOther? l2 i = some (l.getDualInOther! l2 i) := by
simp [getDualInOther!]
lemma getDualInOther?_eq_none_on_not_mem (i : Fin l.length) (h : i ∉ l.withDualInOther l2) :
l.getDualInOther? l2 i = none := by
simpa [getDualInOther?, Fin.find_eq_none_iff, mem_withInDualOther_iff_exists] using h
lemma areDualInOther_getDualInOther! (i : l.withDualInOther l2) :
l.AreDualInOther l2 i (l.getDualInOther! l2 i) := by
have h := l.getDualInOther?_eq_some_getDualInOther! l2 i
rw [getDualInOther?, Fin.find_eq_some_iff] at h
exact h.1
@[simp]
lemma getDualInOther!_id (i : l.withDualInOther l2) :
l2.idMap (l.getDualInOther! l2 i) = l.idMap i := by
simpa using (l.areDualInOther_getDualInOther! l2 i).symm
lemma getDualInOther_mem_withDualInOther (i : l.withDualInOther l2) :
l.getDualInOther! l2 i ∈ l2.withDualInOther l :=
(l2.mem_withInDualOther_iff_exists l).mpr ⟨i, (areDualInOther_getDualInOther! l l2 i).symm⟩
def getDualInOther (i : l.withDualInOther l2) : l2.withDualInOther l :=
⟨l.getDualInOther! l2 i, l.getDualInOther_mem_withDualInOther l2 i⟩
@[simp]
lemma getDualInOther_id (i : l.withDualInOther l2) :
l2.idMap (l.getDualInOther l2 i) = l.idMap i := by
simp [getDualInOther]
lemma getDualInOther_coe (i : l.withDualInOther l2) :
(l.getDualInOther l2 i).1 = l.getDualInOther! l2 i := by
rfl
/-!
## Has single dual in self.
-/
def withUniqueDual : Finset (Fin l.length) :=
Finset.filter (fun i => i ∈ l.withDual ∧
∀ j, l.AreDualInSelf i j → j = l.getDual? i) Finset.univ
lemma mem_withDual_of_withUniqueDual (i : l.withUniqueDual) :
i.1 ∈ l.withDual := by
have hi := i.2
simp [withUniqueDual, Finset.mem_filter, Finset.mem_univ] at hi
exact hi.1
def fromWithUnique (i : l.withUniqueDual) : l.withDual :=
⟨i.1, l.mem_withDual_of_withUniqueDual i⟩
instance : Coe l.withUniqueDual l.withDual where
coe i := ⟨i.1, l.mem_withDual_of_withUniqueDual i⟩
@[simp]
lemma fromWithUnique_coe (i : l.withUniqueDual) : (l.fromWithUnique i).1 = i.1 := by
rfl
lemma all_dual_eq_getDual_of_withUniqueDual (i : l.withUniqueDual) :
∀ j, l.AreDualInSelf i j → j = l.getDual! i := by
have hi := i.2
simp [withUniqueDual] at hi
intro j hj
simpa [getDual!, fromWithUnique] using (Option.get_of_mem _ (hi.2 j hj ).symm).symm
lemma eq_of_areDualInSelf_withUniqueDual {j k : Fin l.length} (i : l.withUniqueDual)
(hj : l.AreDualInSelf i j) (hk : l.AreDualInSelf i k) :
j = k := by
have hj' := all_dual_eq_getDual_of_withUniqueDual l i j hj
have hk' := all_dual_eq_getDual_of_withUniqueDual l i k hk
rw [hj', hk']
lemma eq_getDual_of_withUniqueDual_iff (i : l.withUniqueDual) (j : Fin l.length) :
l.AreDualInSelf i j ↔ j = l.getDual! i := by
apply Iff.intro
intro h
exact (l.all_dual_eq_getDual_of_withUniqueDual i) j h
intro h
subst h
exact areDualInSelf_getDual! l i
@[simp]
lemma getDual!_getDual_of_withUniqueDual (i : l.withUniqueDual) :
l.getDual! (l.getDual i) = i := by
by_contra hn
have h' : l.AreDualInSelf i (l.getDual! (l.getDual i)) := by
simp [AreDualInSelf, hn]
simp_all [AreDualInSelf, getDual, fromWithUnique]
exact fun a => hn (id (Eq.symm a))
rw [eq_getDual_of_withUniqueDual_iff] at h'
have hx := l.areDualInSelf_getDual! (l.getDual i)
simp_all [getDual]
@[simp]
lemma getDual_getDual_of_withUniqueDual (i : l.withUniqueDual) :
l.getDual (l.getDual i) = l.fromWithUnique i :=
SetCoe.ext (getDual!_getDual_of_withUniqueDual l i)
@[simp]
lemma getDual?_getDual!_of_withUniqueDual (i : l.withUniqueDual) :
l.getDual? (l.getDual! i) = some i := by
change l.getDual? (l.getDual i) = some i
lemma some_getDual_eq_getDual? (i : l.withDual) : some (l.getDual i).1 = l.getDual? i := by
rw [getDual?_eq_some_getDual!]
simp
@[simp]
lemma getDual?_getDual?_of_withUniqueDualInOther (i : l.withUniqueDual) :
(l.getDual? i).bind l.getDual? = some i := by
rw [getDual?_eq_some_getDual! l i]
simp
lemma getDual_of_withUniqueDual_mem (i : l.withUniqueDual) :
l.getDual! i ∈ l.withUniqueDual := by
simp [withUniqueDual]
apply And.intro (getDual_mem_withDual l ⟨↑i, mem_withDual_of_withUniqueDual l i⟩)
intro j hj
have h1 : i = j := by
by_contra hn
have h' : l.AreDualInSelf i j := by
simp [AreDualInSelf, hn]
simp_all [AreDualInSelf, getDual, fromWithUnique]
rw [eq_getDual_of_withUniqueDual_iff] at h'
subst h'
simp_all [getDual]
subst h1
rfl
def getDualEquiv : l.withUniqueDual ≃ l.withUniqueDual where
toFun x := ⟨l.getDual x, l.getDual_of_withUniqueDual_mem x⟩
invFun x := ⟨l.getDual x, l.getDual_of_withUniqueDual_mem x⟩
left_inv x := SetCoe.ext (by
simp only [Subtype.coe_eta, getDual_getDual_of_withUniqueDual]
rfl)
right_inv x := SetCoe.ext (by
simp only [Subtype.coe_eta, getDual_getDual_of_withUniqueDual]
rfl)
@[simp]
lemma getDualEquiv_involutive : Function.Involutive l.getDualEquiv := by
intro x
simp [getDualEquiv, fromWithUnique]
/-!
## Has single in other.
## Indices which do not have duals.
-/
def withUniqueDualInOther : Finset (Fin l.length) :=
Finset.filter (fun i => i ∉ l.withDual ∧ i ∈ l.withDualInOther l2
∧ (∀ j, l.AreDualInOther l2 i j → j = l.getDualInOther? l2 i)) Finset.univ
def withoutDual : Finset (Fin l.length) :=
Finset.filter (fun i => (l.getDual? i).isNone) Finset.univ
lemma not_mem_withDual_of_withUniqueDualInOther (i : l.withUniqueDualInOther l2) :
i.1 ∉ l.withDual := by
have hi := i.2
simp only [withUniqueDualInOther, Finset.univ_eq_attach, Finset.mem_filter, Finset.mem_attach,
true_and] at hi
exact hi.2.1
lemma mem_withDualInOther_of_withUniqueDualInOther (i : l.withUniqueDualInOther l2) :
i.1 ∈ l.withDualInOther l2 := by
have hi := i.2
simp only [withUniqueDualInOther, Finset.univ_eq_attach, Finset.mem_filter, Finset.mem_attach,
true_and] at hi
exact hi.2.2.1
def fromWithUniqueDualInOther (i : l.withUniqueDualInOther l2) : l.withDualInOther l2 :=
⟨i.1, l.mem_withDualInOther_of_withUniqueDualInOther l2 i⟩
instance : Coe (l.withUniqueDualInOther l2) (l.withDualInOther l2) where
coe i := ⟨i.1, l.mem_withDualInOther_of_withUniqueDualInOther l2 i⟩
lemma all_dual_eq_of_withUniqueDualInOther (i : l.withUniqueDualInOther l2) :
∀ j, l.AreDualInOther l2 i j → j = l.getDualInOther! l2 i := by
have hi := i.2
simp only [withUniqueDualInOther, Finset.univ_eq_attach, Finset.mem_filter, Finset.mem_attach,
true_and] at hi
intro j hj
refine (Option.get_of_mem _ (hi.2.2.2 j hj).symm).symm
lemma eq_getDualInOther_of_withUniqueDual_iff (i : l.withUniqueDualInOther l2) (j : Fin l2.length) :
l.AreDualInOther l2 i j ↔ j = l.getDualInOther! l2 i := by
apply Iff.intro
intro h
exact l.all_dual_eq_of_withUniqueDualInOther l2 i j h
intro h
subst h
exact areDualInOther_getDualInOther! l l2 i
@[simp]
lemma getDualInOther!_getDualInOther_of_withUniqueDualInOther (i : l.withUniqueDualInOther l2) :
l2.getDualInOther! l (l.getDualInOther l2 i) = i := by
lemma withDual_disjoint_withoutDual : Disjoint l.withDual l.withoutDual := by
rw [Finset.disjoint_iff_ne]
intro a ha b hb
by_contra hn
refine (l.not_mem_withDual_of_withUniqueDualInOther l2 i)
(l.mem_withDual_iff_exists.mpr ⟨(l2.getDualInOther! l (l.getDualInOther l2 i)), ?_⟩ )
simp [AreDualInSelf, hn]
exact (fun a => hn (id (Eq.symm a)))
subst hn
simp_all only [withDual, Finset.mem_filter, Finset.mem_univ, true_and, withoutDual,
Option.isNone_iff_eq_none, Option.isSome_none, Bool.false_eq_true]
lemma not_mem_withDual_of_mem_withoutDual (i : Fin l.length) (h : i ∈ l.withoutDual) :
i ∉ l.withDual := by
have h1 := l.withDual_disjoint_withoutDual
exact Finset.disjoint_right.mp h1 h
lemma withDual_union_withoutDual : l.withDual l.withoutDual = Finset.univ := by
rw [Finset.eq_univ_iff_forall]
intro i
by_cases h : (l.getDual? i).isSome
· simp [withDual, Finset.mem_filter, Finset.mem_univ, h]
· simp at h
simp [withoutDual, Finset.mem_filter, Finset.mem_univ, h]
def withoutDualEquiv : Fin l.withoutDual.card ≃ l.withoutDual :=
(Finset.orderIsoOfFin l.withoutDual (by rfl)).toEquiv
def dualEquiv : l.withDual ⊕ Fin l.withoutDual.card ≃ Fin l.length :=
(Equiv.sumCongr (Equiv.refl _) l.withoutDualEquiv).trans <|
(Equiv.Finset.union _ _ l.withDual_disjoint_withoutDual).trans <|
Equiv.subtypeUnivEquiv (by simp [withDual_union_withoutDual])
/-!
## The contraction list
-/
def contrIndexList : IndexList X where
val := (Fin.list l.withoutDual.card).map (fun i => l.val.get (l.withoutDualEquiv i).1)
@[simp]
lemma getDualInOther_getDualInOther_of_withUniqueDualInOther (i : l.withUniqueDualInOther l2) :
l2.getDualInOther l (l.getDualInOther l2 i) = i :=
SetCoe.ext (getDualInOther!_getDualInOther_of_withUniqueDualInOther l l2 i)
lemma contrIndexList_length : l.contrIndexList.length = l.withoutDual.card := by
simp [contrIndexList, withoutDual, length]
@[simp]
lemma getDualInOther?_getDualInOther!_of_withUniqueDualInOther (i : l.withUniqueDualInOther l2) :
l2.getDualInOther? l (l.getDualInOther! l2 i) = some i := by
change l2.getDualInOther? l (l.getDualInOther l2 i) = some i
rw [getDualInOther?_eq_some_getDualInOther!]
simp
lemma contrIndexList_idMap (i : Fin l.contrIndexList.length) : l.contrIndexList.idMap i
= l.idMap (l.withoutDualEquiv (Fin.cast l.contrIndexList_length i)).1 := by
simp [contrIndexList, idMap]
rfl
lemma getDualInOther_of_withUniqueDualInOther_not_mem_of_withDual (i : l.withUniqueDualInOther l2) :
(l.getDualInOther l2 i).1 ∉ l2.withDual := by
@[simp]
lemma contrIndexList_colorMap (i : Fin l.contrIndexList.length) : l.contrIndexList.colorMap i
= l.colorMap (l.withoutDualEquiv (Fin.cast l.contrIndexList_length i)).1 := by
simp [contrIndexList, colorMap]
rfl
@[simp]
lemma contrIndexList_areDualInSelf (i j : Fin l.contrIndexList.length) :
l.contrIndexList.AreDualInSelf i j ↔
l.AreDualInSelf (l.withoutDualEquiv (Fin.cast l.contrIndexList_length i)).1
(l.withoutDualEquiv (Fin.cast l.contrIndexList_length j)).1 := by
simp [AreDualInSelf]
intro _
trans ¬ (l.withoutDualEquiv (Fin.cast l.contrIndexList_length i)) =
(l.withoutDualEquiv (Fin.cast l.contrIndexList_length j))
rw [l.withoutDualEquiv.apply_eq_iff_eq]
simp [Fin.ext_iff]
exact Iff.symm Subtype.coe_ne_coe
@[simp]
lemma contrIndexList_getDual? (i : Fin l.contrIndexList.length) :
l.contrIndexList.getDual? i = none := by
rw [getDual?_eq_none_on_not_mem]
rw [mem_withDual_iff_exists]
simp
intro j
simp [AreDualInSelf]
intro hj
by_contra hn
have hn' : l.AreDualInOther l2 i j := by
simp [AreDualInOther, hn, hj]
rw [eq_getDualInOther_of_withUniqueDual_iff] at hn'
simp_all
simp only [getDualInOther_coe, not_true_eq_false] at hj
have h1 := (l.withoutDualEquiv (Fin.cast l.contrIndexList_length i)).2
have h1' := Finset.disjoint_right.mp l.withDual_disjoint_withoutDual h1
rw [mem_withDual_iff_exists] at h1'
simp at h1'
exact fun x => h1' ↑(l.withoutDualEquiv (Fin.cast (contrIndexList_length l) x))
lemma getDualInOther_of_withUniqueDualInOther_mem (i : l.withUniqueDualInOther l2) :
(l.getDualInOther l2 i).1 ∈ l2.withUniqueDualInOther l := by
simp only [withUniqueDualInOther, Finset.mem_filter, Finset.mem_univ, true_and]
apply And.intro (l.getDualInOther_of_withUniqueDualInOther_not_mem_of_withDual l2 i)
apply And.intro
exact Finset.coe_mem
(l.getDualInOther l2 ⟨↑i, mem_withDualInOther_of_withUniqueDualInOther l l2 i⟩)
intro j hj
by_contra hn
have h' : l.AreDualInSelf i j := by
simp [AreDualInSelf, hn]
simp_all [AreDualInOther, getDual]
simp [getDualInOther_coe] at hn
exact fun a => hn (id (Eq.symm a))
have hi := l.not_mem_withDual_of_withUniqueDualInOther l2 i
rw [mem_withDual_iff_exists] at hi
simp_all
@[simp]
lemma contrIndexList_withDual : l.contrIndexList.withDual = ∅ := by
rw [Finset.eq_empty_iff_forall_not_mem]
intro x
simp [withDual]
def getDualInOtherEquiv : l.withUniqueDualInOther l2 ≃ l2.withUniqueDualInOther l where
toFun x := ⟨l.getDualInOther l2 x, l.getDualInOther_of_withUniqueDualInOther_mem l2 x⟩
invFun x := ⟨l2.getDualInOther l x, l2.getDualInOther_of_withUniqueDualInOther_mem l x⟩
left_inv x := SetCoe.ext (by simp)
right_inv x := SetCoe.ext (by simp)
@[simp]
lemma contrIndexList_of_withDual_empty (h : l.withDual = ∅) : l.contrIndexList = l := by
have h1 := l.withDual_union_withoutDual
rw [h , Finset.empty_union] at h1
apply ext
rw [@List.ext_get_iff]
change l.contrIndexList.length = l.length ∧ _
rw [contrIndexList_length, h1]
simp
intro n h1' h2
simp [contrIndexList]
congr
simp [withoutDualEquiv]
simp [h1]
rw [(Finset.orderEmbOfFin_unique' _
(fun x => Finset.mem_univ ((Fin.castOrderIso _).toOrderEmbedding x))).symm]
simp
rw [h1]
exact Finset.card_fin l.length
/-!
## Equality of withUnqiueDual and withDual
-/
lemma withUnqiueDual_eq_withDual_iff_unique_forall :
l.withUniqueDual = l.withDual ↔
∀ (i : l.withDual) j, l.AreDualInSelf i j → j = l.getDual? i := by
apply Iff.intro
· intro h i j hj
rw [@Finset.ext_iff] at h
simp [withUniqueDual] at h
exact h i i.2 j hj
· intro h
apply Finset.ext
intro i
apply Iff.intro
· intro hi
simp [withUniqueDual] at hi
exact hi.1
· intro hi
simp [withUniqueDual]
apply And.intro hi
intro j hj
exact h ⟨i, hi⟩ j hj
lemma withUnqiueDual_eq_withDual_iff :
l.withUniqueDual = l.withDual ↔
∀ i, (l.getDual? i).bind l.getDual? = Option.guard (fun i => i ∈ l.withDual) i := by
apply Iff.intro
· intro h i
by_cases hi : i ∈ l.withDual
· have hii : i ∈ l.withUniqueDual := by
simp_all only
change (l.getDual? i).bind l.getDual? = _
rw [getDual?_getDual?_of_withUniqueDualInOther l ⟨i, hii⟩]
simp
symm
rw [Option.guard_eq_some]
exact ⟨rfl, hi⟩
· rw [getDual?_eq_none_on_not_mem l i hi]
simp
symm
simpa only [Option.guard, ite_eq_right_iff, imp_false] using hi
· intro h
rw [withUnqiueDual_eq_withDual_iff_unique_forall]
intro i j hj
rcases l.getDual?_of_areDualInSelf hj with hi | hi | hi
· have hj' := h j
rw [hi] at hj'
simp at hj'
rw [hj']
symm
rw [Option.guard_eq_some]
exact ⟨rfl, l.mem_withDual_iff_exists.mpr ⟨i, hj.symm⟩⟩
· exact hi.symm
· have hj' := h j
rw [hi] at hj'
rw [h i] at hj'
have hi : Option.guard (fun i => i ∈ l.withDual) ↑i = some i := by
apply Option.guard_eq_some.mpr
simp
rw [hi] at hj'
simp at hj'
have hj'' := Option.guard_eq_some.mp hj'.symm
have hj''' := hj''.1
rw [hj'''] at hj
simp at hj
lemma withUnqiueDual_eq_withDual_iff_list_apply :
l.withUniqueDual = l.withDual ↔
(Fin.list l.length).map (fun i => (l.getDual? i).bind l.getDual?) =
(Fin.list l.length).map (fun i => Option.guard (fun i => i ∈ l.withDual) i) := by
rw [withUnqiueDual_eq_withDual_iff]
apply Iff.intro
intro h
apply congrFun
apply congrArg
exact (Set.eqOn_univ (fun i => (l.getDual? i).bind l.getDual?) fun i =>
Option.guard (fun i => i ∈ l.withDual) i).mp fun ⦃x⦄ _ => h x
intro h
intro i
simp only [List.map_inj_left] at h
have h1 {n : } (m : Fin n) : m ∈ Fin.list n := by
have h1' : (Fin.list n)[m] = m := Fin.getElem_list _ _
exact h1' ▸ List.getElem_mem _ _ _
exact h i (h1 i)
def withUnqiueDualEqWithDualBool : Bool :=
if (Fin.list l.length).map (fun i => (l.getDual? i).bind l.getDual?) =
(Fin.list l.length).map (fun i => Option.guard (fun i => i ∈ l.withDual) i) then
true
else
false
lemma withUnqiueDual_eq_withDual_iff_list_apply_bool :
l.withUniqueDual = l.withDual ↔ l.withUnqiueDualEqWithDualBool := by
rw [withUnqiueDual_eq_withDual_iff_list_apply]
apply Iff.intro
intro h
simp [withUnqiueDualEqWithDualBool, h]
intro h
simpa [withUnqiueDualEqWithDualBool] using h
lemma contrIndexList_contrIndexList : l.contrIndexList.contrIndexList = l.contrIndexList := by
simp
end IndexList
end IndexNotation
/-
@ -1386,6 +1117,3 @@ end HasSingColorDualsInSelf
end Color
-/
end IndexList
end IndexNotation

View file

@ -0,0 +1,101 @@
/-
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
-/
import HepLean.SpaceTime.LorentzTensor.IndexNotation.Indices.Basic
import HepLean.SpaceTime.LorentzTensor.Basic
/-!
# Dual in other list
-/
namespace IndexNotation
namespace IndexList
variable {X : Type} [IndexNotation X] [Fintype X] [DecidableEq X]
variable (l l2 : IndexList X)
def AreDualInOther (i : Fin l.length) (j : Fin l2.length) :
Prop := l.idMap i = l2.idMap j
instance {l : IndexList X} {l2 : IndexList X} (i : Fin l.length) (j : Fin l2.length) :
Decidable (l.AreDualInOther l2 i j) := (l.idMap i).decEq (l2.idMap j)
namespace AreDualInOther
variable {l l2 : IndexList X} {i : Fin l.length} {j : Fin l2.length}
@[symm]
lemma symm (h : l.AreDualInOther l2 i j) : l2.AreDualInOther l j i := by
rw [AreDualInOther] at h ⊢
exact h.symm
end AreDualInOther
/-- Given an `i`, if a dual exists in the other list,
outputs the first such dual, otherwise outputs `none`. -/
def getDualInOther? (i : Fin l.length) : Option (Fin l2.length) :=
Fin.find (fun j => l.AreDualInOther l2 i j)
/-!
## With dual other.
-/
def withDualInOther : Finset (Fin l.length) :=
Finset.filter (fun i => (l.getDualInOther? l2 i).isSome) Finset.univ
lemma mem_withInDualOther_iff_exists :
i ∈ l.withDualInOther l2 ↔ ∃ (j : Fin l2.length), l.AreDualInOther l2 i j := by
simp [withDualInOther, Finset.mem_filter, Finset.mem_univ, getDualInOther?]
rw [Fin.isSome_find_iff]
def getDualInOther! (i : l.withDualInOther l2) : Fin l2.length :=
(l.getDualInOther? l2 i).get (by simpa [withDualInOther] using i.2)
lemma getDualInOther?_eq_some_getDualInOther! (i : l.withDualInOther l2) :
l.getDualInOther? l2 i = some (l.getDualInOther! l2 i) := by
simp [getDualInOther!]
lemma getDualInOther?_eq_none_on_not_mem (i : Fin l.length) (h : i ∉ l.withDualInOther l2) :
l.getDualInOther? l2 i = none := by
simpa [getDualInOther?, Fin.find_eq_none_iff, mem_withInDualOther_iff_exists] using h
lemma areDualInOther_getDualInOther! (i : l.withDualInOther l2) :
l.AreDualInOther l2 i (l.getDualInOther! l2 i) := by
have h := l.getDualInOther?_eq_some_getDualInOther! l2 i
rw [getDualInOther?, Fin.find_eq_some_iff] at h
exact h.1
@[simp]
lemma getDualInOther!_id (i : l.withDualInOther l2) :
l2.idMap (l.getDualInOther! l2 i) = l.idMap i := by
simpa using (l.areDualInOther_getDualInOther! l2 i).symm
lemma getDualInOther_mem_withDualInOther (i : l.withDualInOther l2) :
l.getDualInOther! l2 i ∈ l2.withDualInOther l :=
(l2.mem_withInDualOther_iff_exists l).mpr ⟨i, (areDualInOther_getDualInOther! l l2 i).symm⟩
def getDualInOther (i : l.withDualInOther l2) : l2.withDualInOther l :=
⟨l.getDualInOther! l2 i, l.getDualInOther_mem_withDualInOther l2 i⟩
@[simp]
lemma getDualInOther_id (i : l.withDualInOther l2) :
l2.idMap (l.getDualInOther l2 i) = l.idMap i := by
simp [getDualInOther]
lemma getDualInOther_coe (i : l.withDualInOther l2) :
(l.getDualInOther l2 i).1 = l.getDualInOther! l2 i := by
rfl
end IndexList
end IndexNotation

View file

@ -0,0 +1,57 @@
/-
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
-/
import HepLean.SpaceTime.LorentzTensor.IndexNotation.Indices.Color
import HepLean.SpaceTime.LorentzTensor.IndexNotation.Indices.UniqueDualInOther
import HepLean.SpaceTime.LorentzTensor.Basic
import Init.Data.List.Lemmas
/-!
# Index lists with color conditions
Here we consider `IndexListColor` which is a subtype of all lists of indices
on those where the elements to be contracted have consistent colors with respect to
a Tensor Color structure.
-/
namespace IndexNotation
namespace ColorIndexList
variable {𝓒 : TensorColor} [IndexNotation 𝓒.Color] [Fintype 𝓒.Color] [DecidableEq 𝓒.Color]
variable (l l' : ColorIndexList 𝓒)
/-!
## Reindexing
-/
def Reindexing : Prop := l.length = l'.length ∧
∀ (h : l.length = l'.length), l.colorMap = l'.colorMap ∘ Fin.cast h ∧
Option.map (Fin.cast h) ∘ l.getDual? = l'.getDual? ∘ Fin.cast h
/-!
## Permutation
Test whether two `ColorIndexList`s are permutations of each other.
To prevent choice problems, this has to be done after contraction.
-/
def ContrPerm : Prop := l.contr.length = l'.contr.length ∧
l.contr.withUniqueDualInOther l'.contr.toIndexList = Finset.univ ∧
l'.contr.colorMap ∘ Subtype.val ∘ (l.contr.getDualInOtherEquiv l'.contr.toIndexList)
= l.contr.colorMap ∘ Subtype.val
end ColorIndexList
end IndexNotation

View file

@ -0,0 +1,371 @@
/-
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
-/
import HepLean.SpaceTime.LorentzTensor.IndexNotation.Indices.Dual
import HepLean.SpaceTime.LorentzTensor.Basic
/-!
# Unique Dual indices
-/
namespace IndexNotation
namespace IndexList
variable {X : Type} [IndexNotation X] [Fintype X] [DecidableEq X]
variable (l l2 : IndexList X)
/-!
## Has single dual in self.
-/
def withUniqueDual : Finset (Fin l.length) :=
Finset.filter (fun i => i ∈ l.withDual ∧
∀ j, l.AreDualInSelf i j → j = l.getDual? i) Finset.univ
lemma mem_withDual_of_withUniqueDual (i : l.withUniqueDual) :
i.1 ∈ l.withDual := by
have hi := i.2
simp [withUniqueDual, Finset.mem_filter, Finset.mem_univ] at hi
exact hi.1
def fromWithUnique (i : l.withUniqueDual) : l.withDual :=
⟨i.1, l.mem_withDual_of_withUniqueDual i⟩
instance : Coe l.withUniqueDual l.withDual where
coe i := ⟨i.1, l.mem_withDual_of_withUniqueDual i⟩
@[simp]
lemma fromWithUnique_coe (i : l.withUniqueDual) : (l.fromWithUnique i).1 = i.1 := by
rfl
lemma all_dual_eq_getDual_of_withUniqueDual (i : l.withUniqueDual) :
∀ j, l.AreDualInSelf i j → j = l.getDual! i := by
have hi := i.2
simp [withUniqueDual] at hi
intro j hj
simpa [getDual!, fromWithUnique] using (Option.get_of_mem _ (hi.2 j hj ).symm).symm
lemma eq_of_areDualInSelf_withUniqueDual {j k : Fin l.length} (i : l.withUniqueDual)
(hj : l.AreDualInSelf i j) (hk : l.AreDualInSelf i k) :
j = k := by
have hj' := all_dual_eq_getDual_of_withUniqueDual l i j hj
have hk' := all_dual_eq_getDual_of_withUniqueDual l i k hk
rw [hj', hk']
lemma eq_getDual_of_withUniqueDual_iff (i : l.withUniqueDual) (j : Fin l.length) :
l.AreDualInSelf i j ↔ j = l.getDual! i := by
apply Iff.intro
intro h
exact (l.all_dual_eq_getDual_of_withUniqueDual i) j h
intro h
subst h
exact areDualInSelf_getDual! l i
@[simp]
lemma getDual!_getDual_of_withUniqueDual (i : l.withUniqueDual) :
l.getDual! (l.getDual i) = i := by
by_contra hn
have h' : l.AreDualInSelf i (l.getDual! (l.getDual i)) := by
simp [AreDualInSelf, hn]
simp_all [AreDualInSelf, getDual, fromWithUnique]
exact fun a => hn (id (Eq.symm a))
rw [eq_getDual_of_withUniqueDual_iff] at h'
have hx := l.areDualInSelf_getDual! (l.getDual i)
simp_all [getDual]
@[simp]
lemma getDual_getDual_of_withUniqueDual (i : l.withUniqueDual) :
l.getDual (l.getDual i) = l.fromWithUnique i :=
SetCoe.ext (getDual!_getDual_of_withUniqueDual l i)
@[simp]
lemma getDual?_getDual!_of_withUniqueDual (i : l.withUniqueDual) :
l.getDual? (l.getDual! i) = some i := by
change l.getDual? (l.getDual i) = some i
rw [getDual?_eq_some_getDual!]
simp
@[simp]
lemma getDual?_getDual_of_withUniqueDual (i : l.withUniqueDual) :
l.getDual? (l.getDual i) = some i := by
change l.getDual? (l.getDual i) = some i
rw [getDual?_eq_some_getDual!]
simp
@[simp]
lemma getDual?_getDual?_of_withUniqueDualInOther (i : l.withUniqueDual) :
(l.getDual? i).bind l.getDual? = some i := by
rw [getDual?_eq_some_getDual! l i]
simp
lemma getDual_of_withUniqueDual_mem (i : l.withUniqueDual) :
l.getDual! i ∈ l.withUniqueDual := by
simp [withUniqueDual]
apply And.intro (getDual_mem_withDual l ⟨↑i, mem_withDual_of_withUniqueDual l i⟩)
intro j hj
have h1 : i = j := by
by_contra hn
have h' : l.AreDualInSelf i j := by
simp [AreDualInSelf, hn]
simp_all [AreDualInSelf, getDual, fromWithUnique]
rw [eq_getDual_of_withUniqueDual_iff] at h'
subst h'
simp_all [getDual]
subst h1
rfl
def getDualEquiv : l.withUniqueDual ≃ l.withUniqueDual where
toFun x := ⟨l.getDual x, l.getDual_of_withUniqueDual_mem x⟩
invFun x := ⟨l.getDual x, l.getDual_of_withUniqueDual_mem x⟩
left_inv x := SetCoe.ext (by
simp only [Subtype.coe_eta, getDual_getDual_of_withUniqueDual]
rfl)
right_inv x := SetCoe.ext (by
simp only [Subtype.coe_eta, getDual_getDual_of_withUniqueDual]
rfl)
@[simp]
lemma getDualEquiv_involutive : Function.Involutive l.getDualEquiv := by
intro x
simp [getDualEquiv, fromWithUnique]
/-!
## Equality of withUnqiueDual and withDual
-/
lemma withUnqiueDual_eq_withDual_iff_unique_forall :
l.withUniqueDual = l.withDual ↔
∀ (i : l.withDual) j, l.AreDualInSelf i j → j = l.getDual? i := by
apply Iff.intro
· intro h i j hj
rw [@Finset.ext_iff] at h
simp [withUniqueDual] at h
exact h i i.2 j hj
· intro h
apply Finset.ext
intro i
apply Iff.intro
· intro hi
simp [withUniqueDual] at hi
exact hi.1
· intro hi
simp [withUniqueDual]
apply And.intro hi
intro j hj
exact h ⟨i, hi⟩ j hj
lemma withUnqiueDual_eq_withDual_iff :
l.withUniqueDual = l.withDual ↔
∀ i, (l.getDual? i).bind l.getDual? = Option.guard (fun i => i ∈ l.withDual) i := by
apply Iff.intro
· intro h i
by_cases hi : i ∈ l.withDual
· have hii : i ∈ l.withUniqueDual := by
simp_all only
change (l.getDual? i).bind l.getDual? = _
rw [getDual?_getDual?_of_withUniqueDualInOther l ⟨i, hii⟩]
simp
symm
rw [Option.guard_eq_some]
exact ⟨rfl, hi⟩
· rw [getDual?_eq_none_on_not_mem l i hi]
simp
symm
simpa only [Option.guard, ite_eq_right_iff, imp_false] using hi
· intro h
rw [withUnqiueDual_eq_withDual_iff_unique_forall]
intro i j hj
rcases l.getDual?_of_areDualInSelf hj with hi | hi | hi
· have hj' := h j
rw [hi] at hj'
simp at hj'
rw [hj']
symm
rw [Option.guard_eq_some]
exact ⟨rfl, l.mem_withDual_iff_exists.mpr ⟨i, hj.symm⟩⟩
· exact hi.symm
· have hj' := h j
rw [hi] at hj'
rw [h i] at hj'
have hi : Option.guard (fun i => i ∈ l.withDual) ↑i = some i := by
apply Option.guard_eq_some.mpr
simp
rw [hi] at hj'
simp at hj'
have hj'' := Option.guard_eq_some.mp hj'.symm
have hj''' := hj''.1
rw [hj'''] at hj
simp at hj
lemma withUnqiueDual_eq_withDual_iff_list_apply :
l.withUniqueDual = l.withDual ↔
(Fin.list l.length).map (fun i => (l.getDual? i).bind l.getDual?) =
(Fin.list l.length).map (fun i => Option.guard (fun i => i ∈ l.withDual) i) := by
rw [withUnqiueDual_eq_withDual_iff]
apply Iff.intro
intro h
apply congrFun
apply congrArg
exact (Set.eqOn_univ (fun i => (l.getDual? i).bind l.getDual?) fun i =>
Option.guard (fun i => i ∈ l.withDual) i).mp fun ⦃x⦄ _ => h x
intro h
intro i
simp only [List.map_inj_left] at h
have h1 {n : } (m : Fin n) : m ∈ Fin.list n := by
have h1' : (Fin.list n)[m] = m := Fin.getElem_list _ _
exact h1' ▸ List.getElem_mem _ _ _
exact h i (h1 i)
def withUnqiueDualEqWithDualBool : Bool :=
if (Fin.list l.length).map (fun i => (l.getDual? i).bind l.getDual?) =
(Fin.list l.length).map (fun i => Option.guard (fun i => i ∈ l.withDual) i) then
true
else
false
lemma withUnqiueDual_eq_withDual_iff_list_apply_bool :
l.withUniqueDual = l.withDual ↔ l.withUnqiueDualEqWithDualBool := by
rw [withUnqiueDual_eq_withDual_iff_list_apply]
apply Iff.intro
intro h
simp [withUnqiueDualEqWithDualBool, h]
intro h
simpa [withUnqiueDualEqWithDualBool] using h
@[simp]
lemma withUnqiueDual_eq_withDual_of_empty (h : l.withDual = ∅) :
l.withUniqueDual = l.withDual := by
rw [h, Finset.eq_empty_iff_forall_not_mem]
intro x
by_contra hx
have x' : l.withDual := ⟨x, l.mem_withDual_of_withUniqueDual ⟨x, hx⟩⟩
have hx' := x'.2
simp [h] at hx'
/-!
## Splitting withUniqueDual
-/
instance (i j : Option (Fin l.length)) : Decidable (i < j) :=
match i, j with
| none, none => isFalse (fun a => a)
| none, some _ => isTrue (by trivial)
| some _, none => isFalse (fun a => a)
| some i, some j => Fin.decLt i j
def withUniqueDualLT : Finset (Fin l.length) :=
Finset.filter (fun i => i < l.getDual? i) l.withUniqueDual
def withUniqueDualLTToWithUniqueDual (i : l.withUniqueDualLT) : l.withUniqueDual :=
⟨i.1, by
have hi := i.2
simp only [withUniqueDualLT, Finset.mem_filter] at hi
exact hi.1⟩
instance : Coe l.withUniqueDualLT l.withUniqueDual where
coe := l.withUniqueDualLTToWithUniqueDual
def withUniqueDualGT : Finset (Fin l.length) :=
Finset.filter (fun i => ¬ i < l.getDual? i) l.withUniqueDual
def withUniqueDualGTToWithUniqueDual (i : l.withUniqueDualGT) : l.withUniqueDual :=
⟨i.1, by
have hi := i.2
simp only [withUniqueDualGT, Finset.mem_filter] at hi
exact hi.1⟩
instance : Coe l.withUniqueDualGT l.withUniqueDual where
coe := l.withUniqueDualGTToWithUniqueDual
lemma withUniqueDualLT_disjoint_withUniqueDualGT :
Disjoint l.withUniqueDualLT l.withUniqueDualGT := by
rw [Finset.disjoint_iff_inter_eq_empty]
exact @Finset.filter_inter_filter_neg_eq (Fin l.length) _ _ _ _ _
lemma withUniqueDualLT_union_withUniqueDualGT :
l.withUniqueDualLT l.withUniqueDualGT = l.withUniqueDual :=
Finset.filter_union_filter_neg_eq _ _
/-! TODO: Replace with a mathlib lemma. -/
lemma option_not_lt (i j : Option (Fin l.length)) : i < j → i ≠ j → ¬ j < i := by
match i, j with
| none, none =>
intro _
simp
| none, some k =>
intro _
exact fun _ a => a
| some i, none =>
intro h
exact fun _ _ => h
| some i, some j =>
intro h h'
simp_all
change i < j at h
change ¬ j < i
exact Fin.lt_asymm h
/-! TODO: Replace with a mathlib lemma. -/
lemma lt_option_of_not (i j : Option (Fin l.length)) : ¬ j < i → i ≠ j → i < j := by
match i, j with
| none, none =>
intro _
simp
| none, some k =>
intro _
exact fun _ => trivial
| some i, none =>
intro h
exact fun _ => h trivial
| some i, some j =>
intro h h'
simp_all
change ¬ j < i at h
change i < j
omega
def withUniqueDualLTEquivGT : l.withUniqueDualLT ≃ l.withUniqueDualGT where
toFun i := ⟨l.getDualEquiv i, by
have hi := i.2
simp [withUniqueDualGT]
simp [getDualEquiv]
simp [withUniqueDualLT] at hi
apply option_not_lt
simpa [withUniqueDualLTToWithUniqueDual] using hi.2
exact Ne.symm (getDual?_neq_self l i)⟩
invFun i := ⟨l.getDualEquiv.symm i, by
have hi := i.2
simp [withUniqueDualLT]
simp [getDualEquiv]
simp [withUniqueDualGT] at hi
apply lt_option_of_not
simpa [withUniqueDualLTToWithUniqueDual] using hi.2
exact (getDual?_neq_self l i)⟩
left_inv x := SetCoe.ext (by simp [withUniqueDualGTToWithUniqueDual,
withUniqueDualLTToWithUniqueDual])
right_inv x := SetCoe.ext (by simp [withUniqueDualGTToWithUniqueDual,
withUniqueDualLTToWithUniqueDual])
def withUniqueLTGTEquiv : l.withUniqueDualLT ⊕ l.withUniqueDualLT ≃ l.withUniqueDual := by
apply (Equiv.sumCongr (Equiv.refl _ ) l.withUniqueDualLTEquivGT).trans
apply (Equiv.Finset.union _ _ l.withUniqueDualLT_disjoint_withUniqueDualGT).trans
apply (Equiv.subtypeEquivRight (by simp [l.withUniqueDualLT_union_withUniqueDualGT]))
end IndexList
end IndexNotation

View file

@ -0,0 +1,133 @@
/-
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
-/
import HepLean.SpaceTime.LorentzTensor.IndexNotation.Indices.DualInOther
import HepLean.SpaceTime.LorentzTensor.IndexNotation.Indices.Dual
import HepLean.SpaceTime.LorentzTensor.Basic
/-!
# Unique Dual indices
-/
namespace IndexNotation
namespace IndexList
variable {X : Type} [IndexNotation X] [Fintype X] [DecidableEq X]
variable (l l2 : IndexList X)
/-!
## Has single in other.
-/
def withUniqueDualInOther : Finset (Fin l.length) :=
Finset.filter (fun i => i ∉ l.withDual ∧ i ∈ l.withDualInOther l2
∧ (∀ j, l.AreDualInOther l2 i j → j = l.getDualInOther? l2 i)) Finset.univ
lemma not_mem_withDual_of_withUniqueDualInOther (i : l.withUniqueDualInOther l2) :
i.1 ∉ l.withDual := by
have hi := i.2
simp only [withUniqueDualInOther, Finset.univ_eq_attach, Finset.mem_filter, Finset.mem_attach,
true_and] at hi
exact hi.2.1
lemma mem_withDualInOther_of_withUniqueDualInOther (i : l.withUniqueDualInOther l2) :
i.1 ∈ l.withDualInOther l2 := by
have hi := i.2
simp only [withUniqueDualInOther, Finset.univ_eq_attach, Finset.mem_filter, Finset.mem_attach,
true_and] at hi
exact hi.2.2.1
def fromWithUniqueDualInOther (i : l.withUniqueDualInOther l2) : l.withDualInOther l2 :=
⟨i.1, l.mem_withDualInOther_of_withUniqueDualInOther l2 i⟩
instance : Coe (l.withUniqueDualInOther l2) (l.withDualInOther l2) where
coe i := ⟨i.1, l.mem_withDualInOther_of_withUniqueDualInOther l2 i⟩
lemma all_dual_eq_of_withUniqueDualInOther (i : l.withUniqueDualInOther l2) :
∀ j, l.AreDualInOther l2 i j → j = l.getDualInOther! l2 i := by
have hi := i.2
simp only [withUniqueDualInOther, Finset.univ_eq_attach, Finset.mem_filter, Finset.mem_attach,
true_and] at hi
intro j hj
refine (Option.get_of_mem _ (hi.2.2.2 j hj).symm).symm
lemma eq_getDualInOther_of_withUniqueDual_iff (i : l.withUniqueDualInOther l2) (j : Fin l2.length) :
l.AreDualInOther l2 i j ↔ j = l.getDualInOther! l2 i := by
apply Iff.intro
intro h
exact l.all_dual_eq_of_withUniqueDualInOther l2 i j h
intro h
subst h
exact areDualInOther_getDualInOther! l l2 i
@[simp]
lemma getDualInOther!_getDualInOther_of_withUniqueDualInOther (i : l.withUniqueDualInOther l2) :
l2.getDualInOther! l (l.getDualInOther l2 i) = i := by
by_contra hn
refine (l.not_mem_withDual_of_withUniqueDualInOther l2 i)
(l.mem_withDual_iff_exists.mpr ⟨(l2.getDualInOther! l (l.getDualInOther l2 i)), ?_⟩ )
simp [AreDualInSelf, hn]
exact (fun a => hn (id (Eq.symm a)))
@[simp]
lemma getDualInOther_getDualInOther_of_withUniqueDualInOther (i : l.withUniqueDualInOther l2) :
l2.getDualInOther l (l.getDualInOther l2 i) = i :=
SetCoe.ext (getDualInOther!_getDualInOther_of_withUniqueDualInOther l l2 i)
@[simp]
lemma getDualInOther?_getDualInOther!_of_withUniqueDualInOther (i : l.withUniqueDualInOther l2) :
l2.getDualInOther? l (l.getDualInOther! l2 i) = some i := by
change l2.getDualInOther? l (l.getDualInOther l2 i) = some i
rw [getDualInOther?_eq_some_getDualInOther!]
simp
lemma getDualInOther_of_withUniqueDualInOther_not_mem_of_withDual (i : l.withUniqueDualInOther l2) :
(l.getDualInOther l2 i).1 ∉ l2.withDual := by
rw [mem_withDual_iff_exists]
simp
intro j
simp [AreDualInSelf]
intro hj
by_contra hn
have hn' : l.AreDualInOther l2 i j := by
simp [AreDualInOther, hn, hj]
rw [eq_getDualInOther_of_withUniqueDual_iff] at hn'
simp_all
simp only [getDualInOther_coe, not_true_eq_false] at hj
lemma getDualInOther_of_withUniqueDualInOther_mem (i : l.withUniqueDualInOther l2) :
(l.getDualInOther l2 i).1 ∈ l2.withUniqueDualInOther l := by
simp only [withUniqueDualInOther, Finset.mem_filter, Finset.mem_univ, true_and]
apply And.intro (l.getDualInOther_of_withUniqueDualInOther_not_mem_of_withDual l2 i)
apply And.intro
exact Finset.coe_mem
(l.getDualInOther l2 ⟨↑i, mem_withDualInOther_of_withUniqueDualInOther l l2 i⟩)
intro j hj
by_contra hn
have h' : l.AreDualInSelf i j := by
simp [AreDualInSelf, hn]
simp_all [AreDualInOther, getDual]
simp [getDualInOther_coe] at hn
exact fun a => hn (id (Eq.symm a))
have hi := l.not_mem_withDual_of_withUniqueDualInOther l2 i
rw [mem_withDual_iff_exists] at hi
simp_all
def getDualInOtherEquiv : l.withUniqueDualInOther l2 ≃ l2.withUniqueDualInOther l where
toFun x := ⟨l.getDualInOther l2 x, l.getDualInOther_of_withUniqueDualInOther_mem l2 x⟩
invFun x := ⟨l2.getDualInOther l x, l2.getDualInOther_of_withUniqueDualInOther_mem l x⟩
left_inv x := SetCoe.ext (by simp)
right_inv x := SetCoe.ext (by simp)
end IndexList
end IndexNotation