refactor: Index notation
This commit is contained in:
parent
6b0ecc4405
commit
47d639bb1a
14 changed files with 1943 additions and 878 deletions
|
@ -101,27 +101,58 @@ instance : Decidable (listCharIndex X l) :=
|
|||
|
||||
/-!
|
||||
|
||||
## The definition of an index
|
||||
|
||||
-/
|
||||
|
||||
def Index : Type := X × ℕ
|
||||
|
||||
instance : DecidableEq (Index X) := instDecidableEqProd
|
||||
|
||||
namespace Index
|
||||
|
||||
variable {X : Type} [IndexNotation X] [Fintype X] [DecidableEq X]
|
||||
|
||||
/-- The color associated to an index. -/
|
||||
def toColor (I : Index X) : X := I.1
|
||||
|
||||
def id (I : Index X) : ℕ := I.2
|
||||
|
||||
lemma eq_iff_color_eq_and_id_eq (I J : Index X) : I = J ↔ I.toColor = J.toColor ∧ I.id = J.id := by
|
||||
refine Iff.intro ?_ ?_
|
||||
· intro h
|
||||
simp [h]
|
||||
· intro h
|
||||
cases I
|
||||
cases J
|
||||
simp [toColor, id] at h
|
||||
simp [h]
|
||||
|
||||
end Index
|
||||
|
||||
/-!
|
||||
|
||||
## The definition of an index and its properties
|
||||
|
||||
-/
|
||||
|
||||
/-- An index is a non-empty string satisfying the condtion `listCharIndex`,
|
||||
e.g. `ᵘ¹²` or `ᵤ₄₃` etc. -/
|
||||
def Index : Type := {s : String // listCharIndex X s.toList ∧ s.toList ≠ []}
|
||||
def IndexRep : Type := {s : String // listCharIndex X s.toList ∧ s.toList ≠ []}
|
||||
|
||||
instance : DecidableEq (Index X) := Subtype.instDecidableEq
|
||||
instance : DecidableEq (IndexRep X) := Subtype.instDecidableEq
|
||||
|
||||
namespace Index
|
||||
namespace IndexRep
|
||||
|
||||
variable {X : Type} [IndexNotation X] [Fintype X] [DecidableEq X]
|
||||
|
||||
/-- Creats an index from a non-empty list of characters satisfying `listCharIndex`. -/
|
||||
def ofCharList (l : List Char) (h : listCharIndex X l ∧ l ≠ []) : Index X := ⟨l.asString, h⟩
|
||||
def ofCharList (l : List Char) (h : listCharIndex X l ∧ l ≠ []) : IndexRep X := ⟨l.asString, h⟩
|
||||
|
||||
instance : ToString (Index X) := ⟨fun i => i.val⟩
|
||||
instance : ToString (IndexRep X) := ⟨fun i => i.val⟩
|
||||
|
||||
/-- Gets the first character in an index e.g. `ᵘ` as an element of `charList X`. -/
|
||||
def head (s : Index X) : charList X :=
|
||||
def head (s : IndexRep X) : charList X :=
|
||||
⟨s.val.toList.head (s.prop.2), by
|
||||
have h := s.prop.1
|
||||
have h2 := s.prop.2
|
||||
|
@ -130,7 +161,7 @@ def head (s : Index X) : charList X :=
|
|||
simpa [isNotationChar] using h.1⟩
|
||||
|
||||
/-- The color associated to an index. -/
|
||||
def toColor (s : Index X) : X := (IndexNotation.notaEquiv).invFun s.head
|
||||
def toColor (s : IndexRep X) : X := (IndexNotation.notaEquiv).invFun s.head
|
||||
|
||||
/-- A map from super and subscript numerical characters to the natural numbers,
|
||||
returning `0` on all other characters. -/
|
||||
|
@ -159,14 +190,13 @@ def charToNat (c : Char) : Nat :=
|
|||
| _ => 0
|
||||
|
||||
/-- The numerical characters associated with an index. -/
|
||||
def tail (s : Index X) : List Char := s.val.toList.tail
|
||||
def tail (s : IndexRep X) : List Char := s.val.toList.tail
|
||||
|
||||
/-- The natural numbers assocaited with an index. -/
|
||||
def tailNat (s : Index X) : List Nat := s.tail.map charToNat
|
||||
def tailNat (s : IndexRep X) : List ℕ := s.tail.map charToNat
|
||||
|
||||
/-- The id of an index, as a natural number. -/
|
||||
def id (s : Index X) : Nat := s.tailNat.foldl (fun a b => 10 * a + b) 0
|
||||
|
||||
end Index
|
||||
def id (s : IndexRep X) : ℕ := s.tailNat.foldl (fun a b => 10 * a + b) 0
|
||||
|
||||
end IndexRep
|
||||
end IndexNotation
|
||||
|
|
|
@ -1,589 +0,0 @@
|
|||
/-
|
||||
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.OnlyUniqueDuals
|
||||
import HepLean.SpaceTime.LorentzTensor.Basic
|
||||
import Init.Data.List.Lemmas
|
||||
/-!
|
||||
|
||||
# Index lists and color
|
||||
|
||||
In this file we look at the interaction of index lists and color.
|
||||
|
||||
The main definition of this file is `ColorCond`.
|
||||
-/
|
||||
|
||||
namespace IndexNotation
|
||||
|
||||
namespace IndexList
|
||||
|
||||
variable {𝓒 : TensorColor}
|
||||
variable [IndexNotation 𝓒.Color] [Fintype 𝓒.Color] [DecidableEq 𝓒.Color]
|
||||
variable (l l2 l3 : IndexList 𝓒.Color)
|
||||
|
||||
/-- The condition an index and its' dual, when it exists, have dual colors. -/
|
||||
def ColorCond : Prop := Option.map l.colorMap ∘
|
||||
l.getDual? = Option.map (𝓒.τ ∘ l.colorMap) ∘
|
||||
Option.guard fun i => (l.getDual? i).isSome
|
||||
|
||||
namespace ColorCond
|
||||
|
||||
variable {l l2 l3 : IndexList 𝓒.Color}
|
||||
|
||||
lemma iff_withDual :
|
||||
l.ColorCond ↔ ∀ (i : l.withDual), 𝓒.τ
|
||||
(l.colorMap ((l.getDual? i).get (l.withDual_isSome i))) = l.colorMap i := by
|
||||
refine Iff.intro (fun h i => ?_) (fun h => ?_)
|
||||
· have h' := congrFun h i
|
||||
simp at h'
|
||||
rw [show l.getDual? i = some ((l.getDual? i).get (l.withDual_isSome i)) by simp] at h'
|
||||
have h'' : (Option.guard (fun i => (l.getDual? i).isSome = true) ↑i) = i := by
|
||||
apply Option.guard_eq_some.mpr
|
||||
simp [l.withDual_isSome i]
|
||||
rw [h'', Option.map_some', Option.map_some'] at h'
|
||||
simp at h'
|
||||
rw [h']
|
||||
exact 𝓒.τ_involutive (l.colorMap i)
|
||||
· funext i
|
||||
by_cases hi : (l.getDual? i).isSome
|
||||
· have h'' : (Option.guard (fun i => (l.getDual? i).isSome = true) ↑i) = i := by
|
||||
apply Option.guard_eq_some.mpr
|
||||
simp only [true_and]
|
||||
exact hi
|
||||
simp only [Function.comp_apply, h'', Option.map_some']
|
||||
rw [show l.getDual? ↑i = some ((l.getDual? i).get hi) by simp]
|
||||
rw [Option.map_some']
|
||||
simp only [Option.some.injEq]
|
||||
have hii := h ⟨i, by simp [withDual, hi]⟩
|
||||
simp at hii
|
||||
rw [← hii]
|
||||
exact (𝓒.τ_involutive _).symm
|
||||
· simp [Option.guard, hi]
|
||||
exact Option.not_isSome_iff_eq_none.mp hi
|
||||
|
||||
lemma iff_on_isSome : l.ColorCond ↔ ∀ (i : Fin l.length) (h : (l.getDual? i).isSome), 𝓒.τ
|
||||
(l.colorMap ((l.getDual? i).get h)) = l.colorMap i := by
|
||||
rw [iff_withDual]
|
||||
simp only [Subtype.forall, mem_withDual_iff_isSome]
|
||||
|
||||
lemma color_quot_filter_of_countP_two (hl : l.withUniqueDual = l.withDual) (i : Fin l.length)
|
||||
(hi : (l.getDual? i).isSome) :
|
||||
(l.val.filter (fun J => (l.val.get i).id = J.id)).countP
|
||||
(fun J => (l.val.get i).toColor = J.toColor ∨ (l.val.get i).toColor = 𝓒.τ (J.toColor)) =
|
||||
(l.val.filter (fun J => (l.val.get i).id = J.id)).length ↔
|
||||
(l.colorMap i = l.colorMap ((l.getDual? i).get hi) ∨
|
||||
l.colorMap i = 𝓒.τ (l.colorMap ((l.getDual? i).get hi))) := by
|
||||
have hi1 := hi
|
||||
rw [← mem_withDual_iff_isSome, ← hl, mem_withUniqueDual_iff_countId_eq_two] at hi1
|
||||
rcases l.filter_id_of_countId_eq_two hi1 with hf | hf
|
||||
all_goals
|
||||
erw [hf]
|
||||
simp [List.countP, List.countP.go]
|
||||
refine Iff.intro (fun h => ?_) (fun h => ?_)
|
||||
· by_contra hn
|
||||
have hn' : (decide (l.val[↑i].toColor = l.val[↑((l.getDual? i).get hi)].toColor) ||
|
||||
decide (l.val[↑i].toColor = 𝓒.τ l.val[↑((l.getDual? i).get hi)].toColor)) = false := by
|
||||
simpa using hn
|
||||
erw [hn'] at h
|
||||
simp at h
|
||||
· have hn' : (decide (l.val[↑i].toColor = l.val[↑((l.getDual? i).get hi)].toColor) ||
|
||||
decide (l.val[↑i].toColor = 𝓒.τ l.val[↑((l.getDual? i).get hi)].toColor)) = true := by
|
||||
simpa using h
|
||||
erw [hn']
|
||||
simp
|
||||
|
||||
lemma color_dual_eq_self_filter_of_countP_two (hl : l.withUniqueDual = l.withDual)
|
||||
(i : Fin l.length) (hi : (l.getDual? i).isSome) :
|
||||
(l.val.filter (fun J => (l.val.get i).id = J.id)).countP
|
||||
(fun J => (l.val.get i).toColor = J.toColor) =
|
||||
(l.val.filter (fun J => (l.val.get i).id = J.id)).countP
|
||||
(fun J => (l.val.get i).toColor = 𝓒.τ (J.toColor))
|
||||
↔ l.colorMap i = 𝓒.τ (l.colorMap ((l.getDual? i).get hi))
|
||||
∨ (l.colorMap i) = 𝓒.τ (l.colorMap i) := by
|
||||
have hi1 := hi
|
||||
rw [← mem_withDual_iff_isSome, ← hl, mem_withUniqueDual_iff_countId_eq_two] at hi1
|
||||
rcases l.filter_id_of_countId_eq_two hi1 with hf | hf
|
||||
all_goals
|
||||
erw [hf]
|
||||
simp [List.countP, List.countP.go]
|
||||
by_cases hn : l.colorMap i = 𝓒.τ (l.colorMap ((l.getDual? i).get hi))
|
||||
· simp [hn]
|
||||
have hn' : decide (l.val[↑i].toColor = 𝓒.τ l.val[↑((l.getDual? i).get hi)].toColor)
|
||||
= true := by simpa [colorMap] using hn
|
||||
erw [hn']
|
||||
simp only [cond_true]
|
||||
have hτ : l.colorMap ((l.getDual? i).get hi) = 𝓒.τ (l.colorMap i) := by
|
||||
rw [hn]
|
||||
exact (𝓒.τ_involutive _).symm
|
||||
simp [colorMap] at hτ
|
||||
erw [hτ]
|
||||
· have hn' : decide (l.val[↑i].toColor = 𝓒.τ l.val[↑((l.getDual? i).get hi)].toColor) =
|
||||
false := by simpa [colorMap] using hn
|
||||
erw [hn']
|
||||
simp [hn]
|
||||
by_cases hm : l.colorMap i = 𝓒.τ (l.colorMap i)
|
||||
· trans True
|
||||
· simp
|
||||
have hm' : decide (l.val[↑i].toColor = 𝓒.τ l.val[↑i].toColor) = true := by simpa using hm
|
||||
erw [hm']
|
||||
simp only [cond_true]
|
||||
have hm'' : decide (l.val[↑i].toColor = l.val[↑((l.getDual? i).get hi)].toColor)
|
||||
= false := by
|
||||
simp only [Fin.getElem_fin, decide_eq_false_iff_not]
|
||||
simp [colorMap] at hm
|
||||
erw [hm]
|
||||
by_contra hn'
|
||||
have hn'' : l.colorMap i = 𝓒.τ (l.colorMap ((l.getDual? i).get hi)) := by
|
||||
simp [colorMap]
|
||||
rw [← hn']
|
||||
exact (𝓒.τ_involutive _).symm
|
||||
exact hn hn''
|
||||
erw [hm'']
|
||||
simp
|
||||
· exact true_iff_iff.mpr hm
|
||||
· simp [hm]
|
||||
simp [colorMap] at hm
|
||||
have hm' : decide (l.val[↑i].toColor = 𝓒.τ l.val[↑i].toColor) = false := by simpa using hm
|
||||
erw [hm']
|
||||
simp only [cond_false, ne_eq]
|
||||
by_cases hm'' : decide (l.val[↑i].toColor = l.val[↑((l.getDual? i).get hi)].toColor) = true
|
||||
· erw [hm'']
|
||||
simp
|
||||
· have hm''' : decide (l.val[↑i].toColor = l.val[↑((l.getDual? i).get hi)].toColor)
|
||||
= false := by
|
||||
simpa using hm''
|
||||
erw [hm''']
|
||||
simp
|
||||
|
||||
/-- A condition on an index list `l` and and index `I`. If the id of `I` appears
|
||||
twice in `l` (and `I` at least once) then this condition is equivalent to the dual of `I` having
|
||||
dual color to `I`, but written totally in terms of lists. -/
|
||||
abbrev countColorCond (l : IndexList 𝓒.Color) (I : Index 𝓒.Color) : Prop :=
|
||||
(l.val.filter (fun J => I.id = J.id)).countP
|
||||
(fun J => I.toColor = J.toColor ∨ I.toColor = 𝓒.τ (J.toColor)) =
|
||||
(l.val.filter (fun J => I.id = J.id)).length ∧
|
||||
(l.val.filter (fun J => I.id = J.id)).countP (fun J => I.toColor = J.toColor) =
|
||||
(l.val.filter (fun J => I.id = J.id)).countP (fun J => I.toColor = 𝓒.τ (J.toColor))
|
||||
|
||||
lemma countColorCond_cons_neg (l : IndexList 𝓒.Color) (I I' : Index 𝓒.Color) (hid : I.id ≠ I'.id) :
|
||||
countColorCond (l.cons I) I' ↔ countColorCond l I' := by
|
||||
have h1 : (l.cons I).val.filter (fun J => I'.id = J.id) =
|
||||
l.val.filter (fun J => I'.id = J.id) := by
|
||||
simp only [cons]
|
||||
rw [List.filter_cons_of_neg]
|
||||
simp only [decide_eq_true_eq]
|
||||
exact id (Ne.symm hid)
|
||||
rw [countColorCond, countColorCond, h1]
|
||||
|
||||
lemma countColorCond_of_filter_eq (l l2 : IndexList 𝓒.Color) {I : Index 𝓒.Color}
|
||||
(hf : l.val.filter (fun J => I.id = J.id) = l2.val.filter (fun J => I.id = J.id))
|
||||
(h1 : countColorCond l I) : countColorCond l2 I := by
|
||||
rw [countColorCond, ← hf]
|
||||
exact h1
|
||||
|
||||
lemma color_eq_of_countColorCond_cons_pos (l : IndexList 𝓒.Color) (I I' : Index 𝓒.Color)
|
||||
(hl : countColorCond (l.cons I) I') (hI : I.id = I'.id) : I.toColor = I'.toColor ∨
|
||||
I.toColor = 𝓒.τ I'.toColor := by
|
||||
have hl1 := hl.1
|
||||
rw [List.countP_eq_length] at hl1
|
||||
have h2 := hl1 I (by simp; exact hI.symm)
|
||||
simp at h2
|
||||
rcases h2 with h2 | h2
|
||||
· simp [h2]
|
||||
· rw [h2]
|
||||
apply Or.inr (𝓒.τ_involutive _).symm
|
||||
|
||||
lemma iff_countColorCond_isSome (hl : l.withUniqueDual = l.withDual) :
|
||||
l.ColorCond ↔
|
||||
∀ (i : Fin l.length) (_ : (l.getDual? i).isSome), countColorCond l (l.val.get i) := by
|
||||
rw [iff_on_isSome]
|
||||
simp only [countColorCond]
|
||||
refine Iff.intro (fun h i hi => ?_) (fun h i hi => ?_)
|
||||
· rw [color_quot_filter_of_countP_two hl i hi, color_dual_eq_self_filter_of_countP_two hl i hi]
|
||||
have hi' := h i hi
|
||||
exact And.intro (Or.inr hi'.symm) (Or.inl hi'.symm)
|
||||
· have hi' := h i hi
|
||||
rw [color_quot_filter_of_countP_two hl i hi,
|
||||
color_dual_eq_self_filter_of_countP_two hl i hi] at hi'
|
||||
rcases hi'.1 with hi1 | hi1
|
||||
<;> rcases hi'.2 with hi2 | hi2
|
||||
· exact hi2.symm
|
||||
· rw [← hi1]
|
||||
exact hi2.symm
|
||||
· exact hi1.symm
|
||||
· exact hi1.symm
|
||||
|
||||
lemma iff_countColorCond (hl : l.withUniqueDual = l.withDual) :
|
||||
l.ColorCond ↔ ∀ (i : Fin l.length), l.countId (l.val.get i) = 2
|
||||
→ countColorCond l (l.val.get i) := by
|
||||
rw [iff_countColorCond_isSome hl]
|
||||
refine Iff.intro (fun h i hi => ?_) (fun h i hi => ?_)
|
||||
· rw [← mem_withUniqueDual_iff_countId_eq_two] at hi
|
||||
exact h i (mem_withUniqueDual_isSome l i hi)
|
||||
· rw [← mem_withDual_iff_isSome, ← hl, mem_withUniqueDual_iff_countId_eq_two] at hi
|
||||
exact h i hi
|
||||
|
||||
lemma iff_countColorCond_mem (hl : l.withUniqueDual = l.withDual) :
|
||||
l.ColorCond ↔ ∀ (I : Index 𝓒.Color) (_ : I ∈ l.val),
|
||||
l.countId I = 2 → countColorCond l I := by
|
||||
rw [iff_countColorCond hl]
|
||||
refine Iff.intro (fun h I hI hi => ?_) (fun h i hi => ?_)
|
||||
· let i := l.val.indexOf I
|
||||
have hi' : i < l.length := List.indexOf_lt_length.mpr hI
|
||||
have hIi : I = l.val.get ⟨i, hi'⟩ := (List.indexOf_get hi').symm
|
||||
rw [hIi] at hi ⊢
|
||||
exact h ⟨i, hi'⟩ hi
|
||||
· exact h (l.val.get i) (List.getElem_mem l.val (↑i) i.isLt) hi
|
||||
|
||||
/-- The lemma `ColorCond` written totally in terms of lists. -/
|
||||
lemma iff_countColorCond_all (hl : l.withUniqueDual = l.withDual) :
|
||||
l.ColorCond ↔ l.val.all (fun I =>
|
||||
(l.countId I = 2 → countColorCond l I)) := by
|
||||
rw [iff_countColorCond_mem hl]
|
||||
simp only [List.all_eq_true, decide_eq_true_eq]
|
||||
|
||||
@[simp]
|
||||
lemma consDual_color {I : Index 𝓒.Color} (hI : l.countId I = 1)
|
||||
(hI2 : (l.countId I =
|
||||
l.val.countP (fun J => I.id = J.id ∧ I.toColor = 𝓒.τ (J.toColor)))) :
|
||||
(l.consDual hI).toColor = 𝓒.τ I.toColor := by
|
||||
have h1 : l.val.countP (fun J => I.id = J.id ∧ I.toColor = 𝓒.τ (J.toColor))
|
||||
= (l.val.filter (fun J => I.id = J.id)).countP (fun J => I.toColor = 𝓒.τ (J.toColor)) := by
|
||||
rw [List.countP_filter]
|
||||
apply congrFun
|
||||
apply congrArg
|
||||
funext J
|
||||
simp only [Bool.decide_and, decide_eq_true_eq]
|
||||
exact Bool.and_comm (decide (I.id = J.id)) (decide (I.toColor = 𝓒.τ J.toColor))
|
||||
rw [h1, countId, List.countP_eq_length_filter, l.consDual_filter hI] at hI2
|
||||
symm at hI2
|
||||
rw [List.countP_eq_length] at hI2
|
||||
simp only [List.mem_singleton, decide_eq_true_eq, forall_eq] at hI2
|
||||
rw [hI2, 𝓒.τ_involutive]
|
||||
|
||||
lemma of_cons (I : Index 𝓒.Color) (h : (l.cons I).ColorCond)
|
||||
(hl : (l.cons I).withUniqueDual = (l.cons I).withDual) : l.ColorCond := by
|
||||
rw [iff_countColorCond_mem hl] at h
|
||||
have hl' : l.withUniqueDual = l.withDual := withUniqueDual_eq_withDual_of_cons l hl
|
||||
rw [iff_countColorCond_mem hl']
|
||||
intro I' hI'mem hi
|
||||
have hI''mem : I' ∈ (l.cons I).val := by
|
||||
simp [hI'mem]
|
||||
have hI'' := h I' hI''mem
|
||||
by_cases hI'id : I'.id ≠ I.id
|
||||
· rw [countId_eq_length_filter, cons_val, List.filter_cons_of_neg,
|
||||
countColorCond_cons_neg] at hI''
|
||||
· rw [countId_eq_length_filter] at hi
|
||||
exact hI'' hi
|
||||
· exact id (Ne.symm hI'id)
|
||||
· simpa using hI'id
|
||||
· simp at hI'id
|
||||
rw [countId_eq_length_filter, hI'id] at hi
|
||||
rw [propext (withUniqueDual_eq_withDual_cons_iff l I hl'), countId_eq_length_filter, hi] at hl
|
||||
simp at hl
|
||||
|
||||
lemma countId_of_cons (I : Index 𝓒.Color) (h : (l.cons I).ColorCond)
|
||||
(hl : (l.cons I).withUniqueDual = (l.cons I).withDual) :
|
||||
l.countId I =
|
||||
l.val.countP (fun J => I.id = J.id ∧ I.toColor = 𝓒.τ (J.toColor)) := by
|
||||
have h1 := (l.withUniqueDual_eq_withDual_cons_iff I
|
||||
(l.withUniqueDual_eq_withDual_of_cons hl)).mp hl
|
||||
rw [List.countP_eq_length_filter]
|
||||
trans (l.val.filter (fun J => I.id = J.id)).countP (fun J => I.toColor = 𝓒.τ (J.toColor))
|
||||
· by_cases hc : l.countId I = 1
|
||||
· rw [l.consDual_filter hc]
|
||||
simp [List.countP, List.countP.go]
|
||||
rw [iff_withDual] at h
|
||||
have h' := h ⟨⟨0, by simp⟩, (by
|
||||
rw [mem_withDual_iff_countId_gt_one]
|
||||
simp_all [countId])⟩
|
||||
change 𝓒.τ (l.consDual hc).toColor = _ at h'
|
||||
rw [h']
|
||||
simpa [colorMap] using hc
|
||||
· have hc' : l.countId I = 0 := by
|
||||
omega
|
||||
rw [countId_eq_length_filter, List.length_eq_zero] at hc'
|
||||
simp [hc']
|
||||
omega
|
||||
· rw [List.countP_filter]
|
||||
simp only [decide_eq_true_eq, Bool.decide_and]
|
||||
rw [← List.countP_eq_length_filter]
|
||||
apply congrFun
|
||||
apply congrArg
|
||||
funext J
|
||||
exact Bool.and_comm (decide (I.toColor = 𝓒.τ J.toColor)) (decide (I.id = J.id))
|
||||
|
||||
lemma cons_of_countP (h : l.ColorCond) (I : Index 𝓒.Color) (hl : l.withUniqueDual = l.withDual)
|
||||
(hI1 : l.countId I ≤ 1)
|
||||
(hI2 : l.countId I =
|
||||
l.val.countP (fun J => I.id = J.id ∧ I.toColor = 𝓒.τ (J.toColor))) :
|
||||
(l.cons I).ColorCond := by
|
||||
rw [iff_countColorCond_mem]
|
||||
· intro I' hI'
|
||||
by_cases hI'' : I' ≠ I
|
||||
· have hI'mem : I' ∈ l.val := by
|
||||
simp only [cons, List.mem_cons] at hI'
|
||||
rcases hI' with hI' | hI'
|
||||
· exact False.elim (hI'' hI')
|
||||
· exact hI'
|
||||
by_cases hI'id : I'.id ≠ I.id
|
||||
· rw [countId_eq_length_filter, cons_val, List.filter_cons_of_neg]
|
||||
· rw [iff_countColorCond_mem] at h
|
||||
rw [countColorCond_cons_neg l I I' hI'id.symm]
|
||||
· rw [← countId_eq_length_filter]
|
||||
exact h I' hI'mem
|
||||
· exact hl
|
||||
· simpa using hI'id
|
||||
· simp at hI'id
|
||||
intro hI
|
||||
rw [countId_eq_length_filter, hI'id] at hI
|
||||
simp at hI
|
||||
rw [← List.countP_eq_length_filter] at hI
|
||||
have hI'dual : I' = l.consDual hI := by
|
||||
simp [consDual_iff, hI'id, hI'mem]
|
||||
subst hI'dual
|
||||
rw [countColorCond, l.filter_of_constDual hI]
|
||||
simp only [List.countP, List.countP.go, Bool.decide_or, true_or, decide_True, zero_add,
|
||||
Nat.reduceAdd, cond_true, List.length_cons, List.length_singleton]
|
||||
rw [consDual_color hI hI2, 𝓒.τ_involutive]
|
||||
simp
|
||||
· simp only [ne_eq, Decidable.not_not] at hI''
|
||||
symm at hI''
|
||||
subst hI''
|
||||
intro hIf
|
||||
rw [countId_cons_eq_two] at hIf
|
||||
rw [countColorCond]
|
||||
simp only [Bool.decide_or, cons_val, decide_True, List.filter_cons_of_pos, Bool.true_or,
|
||||
List.countP_cons_of_pos, List.length_cons, add_left_inj]
|
||||
rw [l.consDual_filter hIf]
|
||||
simp only [List.countP, List.countP.go, zero_add, List.length_singleton, Nat.reduceAdd]
|
||||
rw [consDual_color hIf hI2, 𝓒.τ_involutive]
|
||||
simp only [decide_True, Bool.or_true, cond_true, true_and]
|
||||
by_cases h1 : (I.toColor = 𝓒.τ I.toColor)
|
||||
· have h1' : decide (I.toColor = 𝓒.τ I.toColor) = true := by simpa using h1
|
||||
simp [h1']
|
||||
· have h1' : decide (I.toColor = 𝓒.τ I.toColor) = false := by simpa using h1
|
||||
simp [h1']
|
||||
· exact (withUniqueDual_eq_withDual_cons_iff l I hl).mpr hI1
|
||||
|
||||
lemma cons_iff (I : Index 𝓒.Color) :
|
||||
(l.cons I).withUniqueDual = (l.cons I).withDual ∧
|
||||
(l.cons I).ColorCond ↔
|
||||
l.withUniqueDual = l.withDual ∧ l.ColorCond ∧
|
||||
l.countId I ≤ 1 ∧
|
||||
(l.countId I =
|
||||
l.val.countP (fun J => I.id = J.id ∧ I.toColor = 𝓒.τ (J.toColor))) := by
|
||||
refine Iff.intro (fun h => ?_) (fun h => ?_)
|
||||
· apply And.intro
|
||||
· exact l.withUniqueDual_eq_withDual_of_cons h.1
|
||||
· apply And.intro
|
||||
· exact of_cons I h.2 h.1
|
||||
· apply And.intro
|
||||
· exact (l.withUniqueDual_eq_withDual_cons_iff I
|
||||
(l.withUniqueDual_eq_withDual_of_cons h.1)).mp h.1
|
||||
· exact countId_of_cons I h.2 h.1
|
||||
· apply And.intro
|
||||
· rw [withUniqueDual_eq_withDual_cons_iff]
|
||||
· exact h.2.2.1
|
||||
· exact h.1
|
||||
· exact cons_of_countP h.2.1 I h.1 h.2.2.1 h.2.2.2
|
||||
|
||||
lemma assoc (h : ColorCond (l ++ l2 ++ l3)) : ColorCond (l ++ (l2 ++ l3)) := by
|
||||
rw [← append_assoc]
|
||||
exact h
|
||||
|
||||
lemma inl (h : ColorCond (l ++ l2)) : ColorCond l := by
|
||||
rw [iff_withDual] at h ⊢
|
||||
intro i
|
||||
simpa only [withDual_isSome, getDual?_append_inl_of_getDual?_isSome, Option.get_some,
|
||||
colorMap_append_inl] using h ⟨appendEquiv (Sum.inl i), by simp only [mem_withDual_iff_isSome,
|
||||
withDual_isSome, getDual?_append_inl_of_getDual?_isSome, Option.isSome_some]⟩
|
||||
|
||||
lemma symm (hu : (l ++ l2).withUniqueDual = (l ++ l2).withDual) (h : ColorCond (l ++ l2)) :
|
||||
ColorCond (l2 ++ l) := by
|
||||
rw [iff_on_isSome] at h ⊢
|
||||
intro j hj
|
||||
obtain ⟨k, hk⟩ := appendEquiv.surjective j
|
||||
subst hk
|
||||
rw [append_withDual_eq_withUniqueDual_symm] at hu
|
||||
rw [← mem_withDual_iff_isSome, ← hu] at hj
|
||||
match k with
|
||||
| Sum.inl k =>
|
||||
have hn := l2.append_inl_not_mem_withDual_of_withDualInOther l k hj
|
||||
by_cases hk' : (l2.getDual? k).isSome
|
||||
· simp_all only [mem_withDual_iff_isSome, getDual?_append_inl_of_getDual?_isSome,
|
||||
Option.isSome_some, mem_withInDualOther_iff_isSome, Bool.not_eq_true, Option.not_isSome,
|
||||
Option.isNone_iff_eq_none, true_iff, Option.get_some, colorMap_append_inl]
|
||||
have hk'' := h (appendEquiv (Sum.inr k))
|
||||
simp only [getDual?_isSome_append_inr_iff, colorMap_append_inr] at hk''
|
||||
simp_all only [getDual?_append_inl_of_getDual?_isSome, Option.isSome_some, Option.isSome_none,
|
||||
Bool.false_eq_true, or_false, Option.isNone_none,
|
||||
getDual?_inr_getDualInOther?_isNone_getDual?_isSome, Option.get_some, colorMap_append_inr,
|
||||
true_implies]
|
||||
· simp_all only [mem_withDual_iff_isSome, Bool.false_eq_true, mem_withInDualOther_iff_isSome,
|
||||
Bool.not_eq_true, Option.not_isSome, Option.isNone_iff_eq_none, false_iff, Option.isNone_none,
|
||||
colorMap_append_inl]
|
||||
have hn' : (l2.getDualInOther? l k).isSome := by
|
||||
simp_all only [Option.isNone_none, getDual?_isSome_append_inl_iff, Option.isSome_none,
|
||||
Bool.false_eq_true, false_or]
|
||||
have hk'' := h (appendEquiv (Sum.inr k))
|
||||
simp only [getDual?_isSome_append_inr_iff, colorMap_append_inr] at hk''
|
||||
simp_all only [Option.isSome_none, Bool.false_eq_true, or_true,
|
||||
getDual?_append_inr_getDualInOther?_isSome, Option.get_some, colorMap_append_inl,
|
||||
true_implies, Option.isNone_none, getDual?_inl_of_getDual?_isNone_getDualInOther?_isSome,
|
||||
colorMap_append_inr]
|
||||
| Sum.inr k =>
|
||||
have hn := l2.append_inr_not_mem_withDual_of_withDualInOther l k hj
|
||||
by_cases hk' : (l.getDual? k).isSome
|
||||
· simp_all only [mem_withDual_iff_isSome, mem_withInDualOther_iff_isSome, Bool.not_eq_true,
|
||||
Option.not_isSome, Option.isNone_iff_eq_none, true_iff, Option.isNone_none,
|
||||
getDual?_inr_getDualInOther?_isNone_getDual?_isSome, Option.get_some, colorMap_append_inr]
|
||||
have hk'' := h (appendEquiv (Sum.inl k))
|
||||
simp only [getDual?_isSome_append_inl_iff, colorMap_append_inl] at hk''
|
||||
simp_all only [Option.isNone_none, getDual?_inr_getDualInOther?_isNone_getDual?_isSome,
|
||||
Option.isSome_some, Option.isSome_none, Bool.false_eq_true, or_false,
|
||||
getDual?_append_inl_of_getDual?_isSome, Option.get_some, colorMap_append_inl, true_implies]
|
||||
· simp_all only [mem_withDual_iff_isSome, Bool.false_eq_true, mem_withInDualOther_iff_isSome,
|
||||
Bool.not_eq_true, Option.not_isSome, Option.isNone_iff_eq_none, false_iff,
|
||||
colorMap_append_inr]
|
||||
have hn' : (l.getDualInOther? l2 k).isSome := by
|
||||
exact Option.ne_none_iff_isSome.mp hn
|
||||
have hk'' := h (appendEquiv (Sum.inl k))
|
||||
simp only [getDual?_isSome_append_inl_iff, colorMap_append_inl] at hk''
|
||||
simp_all only [Option.isSome_none, Bool.false_eq_true, or_true, Option.isNone_none,
|
||||
getDual?_inl_of_getDual?_isNone_getDualInOther?_isSome, Option.get_some,
|
||||
colorMap_append_inr, true_implies, getDual?_append_inr_getDualInOther?_isSome,
|
||||
colorMap_append_inl]
|
||||
|
||||
lemma inr (hu : (l ++ l2).withUniqueDual = (l ++ l2).withDual) (h : ColorCond (l ++ l2)) :
|
||||
ColorCond l2 := inl (symm hu h)
|
||||
|
||||
lemma triple_right (hu : (l ++ l2 ++ l3).withUniqueDual = (l ++ l2 ++ l3).withDual)
|
||||
(h : ColorCond (l ++ l2 ++ l3)) : ColorCond (l2 ++ l3) := by
|
||||
have h1 := assoc h
|
||||
rw [append_assoc] at hu
|
||||
exact h1.inr hu
|
||||
|
||||
lemma triple_drop_mid (hu : (l ++ l2 ++ l3).withUniqueDual = (l ++ l2 ++ l3).withDual)
|
||||
(h : ColorCond (l ++ l2 ++ l3)) : ColorCond (l ++ l3) := by
|
||||
rw [append_assoc] at hu
|
||||
refine (((assoc h).symm hu).assoc.inr ?_).symm ?_
|
||||
· rw [append_withDual_eq_withUniqueDual_symm, append_assoc] at hu
|
||||
exact hu
|
||||
· rw [append_withDual_eq_withUniqueDual_symm, append_assoc] at hu
|
||||
exact append_withDual_eq_withUniqueDual_inr _ _ hu
|
||||
|
||||
lemma swap (hu : (l ++ l2 ++ l3).withUniqueDual = (l ++ l2 ++ l3).withDual)
|
||||
(h : ColorCond (l ++ l2 ++ l3)) :
|
||||
ColorCond (l2 ++ l ++ l3) := by
|
||||
have hC := h
|
||||
have hu' := hu
|
||||
rw [iff_on_isSome] at h ⊢
|
||||
intro j hj
|
||||
obtain ⟨k, hk⟩ := appendEquiv.surjective j
|
||||
subst hk
|
||||
match k with
|
||||
| Sum.inl k =>
|
||||
have hj' := hj
|
||||
rw [append_withDual_eq_withUniqueDual_swap] at hu
|
||||
rw [← mem_withDual_iff_isSome, ← hu] at hj'
|
||||
have hn := (l2 ++ l).append_inl_not_mem_withDual_of_withDualInOther l3 k hj'
|
||||
simp only [mem_withDual_iff_isSome, mem_withInDualOther_iff_isSome, Bool.not_eq_true,
|
||||
Option.not_isSome, Option.isNone_iff_eq_none] at hn
|
||||
simp only [getDual?_isSome_append_inl_iff] at hj
|
||||
by_cases hk' : ((l2 ++ l).getDual? k).isSome
|
||||
· simp only [hk', getDual?_append_inl_of_getDual?_isSome, Option.get_some, colorMap_append_inl]
|
||||
have hu' := append_withDual_eq_withUniqueDual_inl (l2 ++ l) l3 hu
|
||||
have hC' := hC.inl.symm ((append_withDual_eq_withUniqueDual_symm l2 l).mp hu')
|
||||
rw [iff_on_isSome] at hC'
|
||||
exact hC' k hk'
|
||||
· simp only [hk', Bool.false_eq_true, false_iff] at hn
|
||||
rw [← @Option.not_isSome_iff_eq_none, not_not] at hn
|
||||
simp_all only [mem_withDual_iff_isSome, Bool.false_eq_true, or_true, Bool.not_eq_true,
|
||||
Option.not_isSome, Option.isNone_iff_eq_none, Option.isNone_none,
|
||||
getDual?_inl_of_getDual?_isNone_getDualInOther?_isSome, Option.get_some,
|
||||
colorMap_append_inr, colorMap_append_inl]
|
||||
obtain ⟨k', hk'⟩ := appendEquiv.surjective k
|
||||
subst hk'
|
||||
match k' with
|
||||
| Sum.inl k' =>
|
||||
simp only [getDualInOther?_append_of_inl] at hn
|
||||
simp only [getDualInOther?_append_of_inl, colorMap_append_inl]
|
||||
have hL := triple_right hu' hC
|
||||
rw [iff_on_isSome] at hL
|
||||
have hL' := hL (appendEquiv (Sum.inl k')) (by simp only [getDual?_isSome_append_inl_iff, hn,
|
||||
or_true])
|
||||
simp_all only [Option.isNone_none, getDualInOther?_append_of_inl,
|
||||
getDual?_inl_of_getDual?_isNone_getDualInOther?_isSome, Option.isSome_some,
|
||||
getDual?_eq_none_append_inl_iff, Option.get_some, colorMap_append_inr,
|
||||
colorMap_append_inl]
|
||||
| Sum.inr k' =>
|
||||
simp only [getDualInOther?_append_of_inr] at hn
|
||||
simp only [getDualInOther?_append_of_inr, colorMap_append_inr]
|
||||
have hR := triple_drop_mid hu' hC
|
||||
rw [iff_on_isSome] at hR
|
||||
have hR' := hR (appendEquiv (Sum.inl k')) (by simp only [getDual?_isSome_append_inl_iff, hn,
|
||||
or_true])
|
||||
simp_all only [Option.isNone_none, getDualInOther?_append_of_inr,
|
||||
getDual?_inl_of_getDual?_isNone_getDualInOther?_isSome, Option.isSome_some,
|
||||
getDual?_eq_none_append_inr_iff, Option.get_some, colorMap_append_inr,
|
||||
colorMap_append_inl]
|
||||
| Sum.inr k =>
|
||||
have hj' := hj
|
||||
rw [append_withDual_eq_withUniqueDual_swap] at hu
|
||||
rw [← mem_withDual_iff_isSome, ← hu] at hj'
|
||||
have hn := (l2 ++ l).append_inr_not_mem_withDual_of_withDualInOther l3 k hj'
|
||||
simp only [mem_withDual_iff_isSome, mem_withInDualOther_iff_isSome,
|
||||
getDualInOther?_isSome_of_append_iff, not_or, Bool.not_eq_true, Option.not_isSome,
|
||||
Option.isNone_iff_eq_none] at hn
|
||||
by_cases hk' : (l3.getDual? k).isSome
|
||||
· simp_all only [mem_withDual_iff_isSome, true_iff, Option.isNone_iff_eq_none,
|
||||
getDualInOther?_eq_none_of_append_iff, and_self,
|
||||
getDual?_inr_getDualInOther?_isNone_getDual?_isSome, Option.get_some, colorMap_append_inr]
|
||||
have hRR := hC.inr hu'
|
||||
rw [iff_on_isSome] at hRR
|
||||
exact hRR k hk'
|
||||
· simp_all only [mem_withDual_iff_isSome, Bool.false_eq_true, false_iff, not_and,
|
||||
Bool.not_eq_true, Option.not_isSome, Option.isNone_iff_eq_none, colorMap_append_inr]
|
||||
by_cases hk'' : (l3.getDualInOther? l2 k).isSome
|
||||
· simp_all only [getDualInOther?_of_append_of_isSome, Option.isSome_some,
|
||||
getDual?_append_inr_getDualInOther?_isSome, Option.get_some, colorMap_append_inl]
|
||||
have hL := triple_right hu' hC
|
||||
rw [iff_on_isSome] at hL
|
||||
have hL' := hL (appendEquiv (Sum.inr k)) (by simp [hk''])
|
||||
simp_all only [getDualInOther?_of_append_of_isSome, Option.isSome_some,
|
||||
getDual?_append_inr_getDualInOther?_isSome, Option.get_some, colorMap_append_inl,
|
||||
colorMap_append_inr]
|
||||
· simp_all only [Bool.not_eq_true, Option.not_isSome, Option.isNone_iff_eq_none,
|
||||
true_implies]
|
||||
rw [← Option.not_isSome_iff_eq_none, not_not] at hn
|
||||
simp_all only [getDualInOther?_of_append_of_isNone_isSome, Option.isSome_some,
|
||||
getDual?_append_inr_getDualInOther?_isSome, Option.get_some, colorMap_append_inl,
|
||||
colorMap_append_inr]
|
||||
have hR := triple_drop_mid hu' hC
|
||||
rw [iff_on_isSome] at hR
|
||||
have hR' := hR (appendEquiv (Sum.inr k)) (by simp [hn])
|
||||
simp_all only [getDualInOther?_of_append_of_isNone_isSome, Option.isSome_some,
|
||||
getDual?_append_inr_getDualInOther?_isSome, Option.get_some, colorMap_append_inl,
|
||||
colorMap_append_inr]
|
||||
|
||||
/-- A bool which is true for an index list if and only if every index and its' dual, when it exists,
|
||||
have dual colors. -/
|
||||
def bool (l : IndexList 𝓒.Color) : Bool :=
|
||||
if (∀ (i : l.withDual), 𝓒.τ
|
||||
(l.colorMap ((l.getDual? i).get (l.withDual_isSome i))) = l.colorMap i) then
|
||||
true
|
||||
else false
|
||||
|
||||
lemma iff_bool : l.ColorCond ↔ bool l := by
|
||||
rw [iff_withDual, bool]
|
||||
simp only [Subtype.forall, mem_withDual_iff_isSome, Bool.if_false_right, Bool.and_true,
|
||||
decide_eq_true_eq]
|
||||
|
||||
end ColorCond
|
||||
|
||||
end IndexList
|
||||
|
||||
end IndexNotation
|
|
@ -37,14 +37,13 @@ open IndexList TensorColor
|
|||
Note: `AppendCond` does not form an equivalence relation as it is not reflexive or
|
||||
transitive. -/
|
||||
def AppendCond : Prop :=
|
||||
(l.toIndexList ++ l2.toIndexList).withUniqueDual = (l.toIndexList ++ l2.toIndexList).withDual
|
||||
∧ ColorCond (l.toIndexList ++ l2.toIndexList)
|
||||
(l.toIndexList ++ l2.toIndexList).OnlyUniqueDuals ∧ ColorCond (l.toIndexList ++ l2.toIndexList)
|
||||
|
||||
/-- Given two `ColorIndexList`s satisfying `AppendCond`. The correponding combined
|
||||
`ColorIndexList`. -/
|
||||
def append (h : AppendCond l l2) : ColorIndexList 𝓒 where
|
||||
toIndexList := l.toIndexList ++ l2.toIndexList
|
||||
unique_duals := h.1.symm
|
||||
unique_duals := h.1
|
||||
dual_color := h.2
|
||||
|
||||
/-- The join of two `ColorIndexList` satisfying the condition `AppendCond` that they
|
||||
|
@ -53,8 +52,7 @@ scoped[IndexNotation.ColorIndexList] notation l " ++["h"] " l2 => append l l2 h
|
|||
|
||||
@[simp]
|
||||
lemma append_toIndexList (h : AppendCond l l2) :
|
||||
(l ++[h] l2).toIndexList = l.toIndexList ++ l2.toIndexList := by
|
||||
rfl
|
||||
(l ++[h] l2).toIndexList = l.toIndexList ++ l2.toIndexList := rfl
|
||||
|
||||
namespace AppendCond
|
||||
|
||||
|
@ -63,7 +61,7 @@ variable {l l2 l3 : ColorIndexList 𝓒}
|
|||
@[symm]
|
||||
lemma symm (h : AppendCond l l2) : AppendCond l2 l := by
|
||||
apply And.intro _ (h.2.symm h.1)
|
||||
rw [append_withDual_eq_withUniqueDual_symm]
|
||||
rw [OnlyUniqueDuals.symm]
|
||||
exact h.1
|
||||
|
||||
lemma inr (h : AppendCond l l2) (h' : AppendCond (l ++[h] l2) l3) :
|
||||
|
@ -72,7 +70,7 @@ lemma inr (h : AppendCond l l2) (h' : AppendCond (l ++[h] l2) l3) :
|
|||
· have h1 := h'.1
|
||||
simp at h1
|
||||
rw [append_assoc] at h1
|
||||
exact l.append_withDual_eq_withUniqueDual_inr (l2.toIndexList ++ l3.toIndexList) h1
|
||||
exact OnlyUniqueDuals.inr h1
|
||||
· have h1 := h'.2
|
||||
simp at h1
|
||||
rw [append_assoc] at h1
|
||||
|
@ -94,95 +92,19 @@ lemma swap (h : AppendCond l l2) (h' : AppendCond (l ++[h] l2) l3) :
|
|||
AppendCond (l2 ++[h.symm] l) l3:= by
|
||||
apply And.intro
|
||||
· simp only [append_toIndexList]
|
||||
rw [← append_withDual_eq_withUniqueDual_swap]
|
||||
apply OnlyUniqueDuals.swap
|
||||
simpa using h'.1
|
||||
· exact ColorCond.swap h'.1 h'.2
|
||||
|
||||
/-- If `AppendCond l l2` then `AppendCond l.contr l2`. Note that the inverse
|
||||
is generally not true. -/
|
||||
lemma contr_left (h : AppendCond l l2) : AppendCond l.contr l2 := by
|
||||
apply And.intro (append_withDual_eq_withUniqueDual_contr_left l.toIndexList l2.toIndexList h.1)
|
||||
· rw [ColorCond.iff_countColorCond_all]
|
||||
· simp only [append_val, countId_append, List.all_append, Bool.and_eq_true, List.all_eq_true,
|
||||
decide_eq_true_eq]
|
||||
apply And.intro
|
||||
· intro I hI hI2
|
||||
have hIC1 : l.contr.countId I = 1 :=
|
||||
mem_contrIndexList_countId_contrIndexList l.toIndexList hI
|
||||
have hIC2 : l2.countId I = 1 := by omega
|
||||
obtain ⟨I1, hI1, hI1'⟩ := countId_neq_zero_mem l.contr I (by omega)
|
||||
apply ColorCond.countColorCond_of_filter_eq (l.toIndexList ++ l2.toIndexList) _ _
|
||||
· have h2 := h.2
|
||||
rw [ColorCond.iff_countColorCond_all] at h2
|
||||
· simp only [append_val, countId_append, List.all_append, Bool.and_eq_true,
|
||||
List.all_eq_true, decide_eq_true_eq] at h2
|
||||
have hn := h2.1 I (List.mem_of_mem_filter hI)
|
||||
have hc : l.countId I + l2.countId I = 2 := by
|
||||
rw [contr, countId_contrIndexList_eq_one_iff_countId_eq_one] at hIC1
|
||||
omega
|
||||
exact hn hc
|
||||
· exact h.1
|
||||
· erw [List.filter_append, List.filter_append]
|
||||
apply congrFun
|
||||
apply congrArg
|
||||
rw [countId_congr _ hI1'] at hIC1
|
||||
rw [hI1', filter_id_of_countId_eq_one_mem l.contr hI1 hIC1]
|
||||
simp [contr, contrIndexList] at hI1
|
||||
exact filter_id_of_countId_eq_one_mem l hI1.1 hI1.2
|
||||
· intro I hI1 hI2
|
||||
by_cases hICz : l2.countId I = 0
|
||||
· rw [hICz] at hI2
|
||||
have hx : l.contr.countId I ≤ 1 := countId_contrIndexList_le_one l.toIndexList I
|
||||
omega
|
||||
· by_cases hICo : l2.countId I = 1
|
||||
· have hIC1 : l.contr.countId I = 1 := by omega
|
||||
obtain ⟨I1, hI', hI1'⟩ := countId_neq_zero_mem l.contr I (by omega)
|
||||
apply ColorCond.countColorCond_of_filter_eq (l.toIndexList ++ l2.toIndexList) _ _
|
||||
· have h2 := h.2
|
||||
rw [ColorCond.iff_countColorCond_all] at h2
|
||||
· simp only [append_val, countId_append, List.all_append, Bool.and_eq_true,
|
||||
List.all_eq_true, decide_eq_true_eq] at h2
|
||||
refine h2.2 I hI1 ?_
|
||||
rw [contr, countId_contrIndexList_eq_one_iff_countId_eq_one] at hIC1
|
||||
omega
|
||||
· exact h.1
|
||||
· erw [List.filter_append, List.filter_append]
|
||||
apply congrFun
|
||||
apply congrArg
|
||||
rw [hI1']
|
||||
rw [countId_congr _ hI1'] at hIC1
|
||||
rw [filter_id_of_countId_eq_one_mem l.contr hI' hIC1]
|
||||
simp [contr, contrIndexList] at hI'
|
||||
exact filter_id_of_countId_eq_one_mem l hI'.1 hI'.2
|
||||
· have hICt : l2.countId I = 2 := by
|
||||
omega
|
||||
apply ColorCond.countColorCond_of_filter_eq l2 _ _
|
||||
· have hl2C := l2.dual_color
|
||||
rw [ColorCond.iff_countColorCond_all] at hl2C
|
||||
· simp only [List.all_eq_true, decide_eq_true_eq] at hl2C
|
||||
exact hl2C I hI1 hICt
|
||||
· exact l2.unique_duals.symm
|
||||
· erw [List.filter_append]
|
||||
rw [filter_id_of_countId_eq_zero' l.contr.toIndexList]
|
||||
· rfl
|
||||
· omega
|
||||
· exact append_withDual_eq_withUniqueDual_contr_left l.toIndexList l2.toIndexList h.1
|
||||
lemma contr_left (h : AppendCond l l2) : AppendCond l.contr l2 :=
|
||||
And.intro (OnlyUniqueDuals.contrIndexList_left h.1) (ColorCond.contrIndexList_left h.1 h.2)
|
||||
|
||||
lemma contr_right (h : AppendCond l l2) : AppendCond l l2.contr := (contr_left h.symm).symm
|
||||
|
||||
lemma contr (h : AppendCond l l2) : AppendCond l.contr l2.contr := contr_left (contr_right h)
|
||||
|
||||
lemma of_eq (h1 : l.withUniqueDual = l.withDual)
|
||||
(h2 : l2.withUniqueDual = l2.withDual)
|
||||
(h3 : l.withUniqueDualInOther l2 = l.withDualInOther l2)
|
||||
(h4 : l2.withUniqueDualInOther l = l2.withDualInOther l)
|
||||
(h5 : ColorCond.bool (l.toIndexList ++ l2.toIndexList)) :
|
||||
AppendCond l l2 := by
|
||||
rw [AppendCond]
|
||||
rw [append_withDual_eq_withUniqueDual_iff']
|
||||
simp_all
|
||||
exact ColorCond.iff_bool.mpr h5
|
||||
|
||||
/-- A boolean which is true for two index lists `l` and `l2` if on appending
|
||||
they can form a `ColorIndexList`. -/
|
||||
def bool (l l2 : IndexList 𝓒.Color) : Bool :=
|
||||
|
@ -200,18 +122,17 @@ lemma iff_bool (l l2 : ColorIndexList 𝓒) : AppendCond l l2 ↔ bool l.toIndex
|
|||
simp [bool]
|
||||
rw [ColorCond.iff_bool]
|
||||
simp
|
||||
exact fun _ => Eq.to_iff rfl
|
||||
|
||||
lemma countId_contr_fst_eq_zero_mem_snd (h : AppendCond l l2) {I : Index 𝓒.Color}
|
||||
(hI : I ∈ l2.val) : l.contr.countId I = 0 ↔ l.countId I = 0 := by
|
||||
rw [countId_contr_eq_zero_iff]
|
||||
refine Iff.intro (fun h' => ?_) (fun h' => ?_)
|
||||
· have h1 := h.1
|
||||
rw [withUniqueDual_eq_withDual_iff_countId_leq_two'] at h1
|
||||
have h1I := h1 I
|
||||
simp at h1I
|
||||
have h2 : l2.countId I ≠ 0 := countId_mem l2.toIndexList I hI
|
||||
omega
|
||||
· simp [h']
|
||||
have h1 := countId_mem l2.toIndexList I hI
|
||||
have h1I := h.1
|
||||
rw [OnlyUniqueDuals.iff_countId_leq_two'] at h1I
|
||||
have h1I' := h1I I
|
||||
simp at h1I'
|
||||
omega
|
||||
|
||||
lemma countId_contr_snd_eq_zero_mem_fst (h : AppendCond l l2) {I : Index 𝓒.Color}
|
||||
(hI : I ∈ l.val) : l2.contr.countId I = 0 ↔ l2.countId I = 0 := by
|
||||
|
|
|
@ -3,8 +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.Color
|
||||
import HepLean.SpaceTime.LorentzTensor.IndexNotation.OnlyUniqueDuals
|
||||
import HepLean.SpaceTime.LorentzTensor.IndexNotation.IndexList.Color
|
||||
import HepLean.SpaceTime.LorentzTensor.Basic
|
||||
import Init.Data.List.Lemmas
|
||||
/-!
|
||||
|
@ -33,7 +32,7 @@ variable [IndexNotation 𝓒.Color] [Fintype 𝓒.Color] [DecidableEq 𝓒.Color
|
|||
structure ColorIndexList (𝓒 : TensorColor) [IndexNotation 𝓒.Color] extends IndexList 𝓒.Color where
|
||||
/-- The condition that for index with a dual, that dual is the unique other index with
|
||||
an identical `id`. -/
|
||||
unique_duals : toIndexList.withDual = toIndexList.withUniqueDual
|
||||
unique_duals : toIndexList.OnlyUniqueDuals
|
||||
/-- The condition that for an index with a dual, that dual has dual color to the index. -/
|
||||
dual_color : IndexList.ColorCond toIndexList
|
||||
|
||||
|
@ -90,122 +89,14 @@ def empty : ColorIndexList 𝓒 where
|
|||
unique_duals := rfl
|
||||
dual_color := rfl
|
||||
|
||||
/-- The `ColorIndexList` obtained by adding an index, subject to some conditions,
|
||||
to the start of `l`. -/
|
||||
def cons (I : Index 𝓒.Color) (hI1 : l.val.countP (fun J => I.id = J.id) ≤ 1)
|
||||
(hI2 : l.val.countP (fun J => I.id = J.id) =
|
||||
l.val.countP (fun J => I.id = J.id ∧ I.toColor = 𝓒.τ J.toColor)) : ColorIndexList 𝓒 where
|
||||
toIndexList := l.toIndexList.cons I
|
||||
unique_duals := by
|
||||
symm
|
||||
rw [withUniqueDual_eq_withDual_cons_iff]
|
||||
· exact hI1
|
||||
· exact l.unique_duals.symm
|
||||
dual_color := by
|
||||
have h1 : (l.toIndexList.cons I).withUniqueDual = (l.toIndexList.cons I).withDual ∧
|
||||
(l.toIndexList.cons I).ColorCond := by
|
||||
rw [ColorCond.cons_iff]
|
||||
exact ⟨l.unique_duals.symm, l.dual_color, hI1, hI2⟩
|
||||
exact h1.2
|
||||
|
||||
/-- The tail of a `ColorIndexList`. In other words, the `ColorIndexList` with the first index
|
||||
removed. -/
|
||||
def tail (l : ColorIndexList 𝓒) : ColorIndexList 𝓒 where
|
||||
toIndexList := l.toIndexList.tail
|
||||
unique_duals := by
|
||||
by_cases hl : l.toIndexList = {val := []}
|
||||
· rw [hl]
|
||||
simp [IndexList.tail]
|
||||
rfl
|
||||
· have hl' := l.head_cons_tail hl
|
||||
have h1 := l.unique_duals
|
||||
rw [hl'] at h1
|
||||
exact (withUniqueDual_eq_withDual_of_cons _ h1.symm).symm
|
||||
dual_color := by
|
||||
by_cases hl : l.toIndexList = {val := []}
|
||||
· rw [hl]
|
||||
simp [IndexList.tail]
|
||||
rfl
|
||||
· have hl' := l.head_cons_tail hl
|
||||
have h1 := l.unique_duals
|
||||
have h2 := l.dual_color
|
||||
rw [hl'] at h1 h2
|
||||
exact (ColorCond.of_cons _ h2 h1.symm)
|
||||
|
||||
@[simp]
|
||||
lemma tail_toIndexList : l.tail.toIndexList = l.toIndexList.tail := by
|
||||
rfl
|
||||
|
||||
/-- The first index in a `ColorIndexList`. -/
|
||||
def head (h : l ≠ empty) : Index 𝓒.Color :=
|
||||
l.toIndexList.head (by
|
||||
cases' l
|
||||
simpa [empty] using h)
|
||||
|
||||
lemma head_uniqueDual (h : l ≠ empty) :
|
||||
l.tail.val.countP (fun J => (l.head h).id = J.id) ≤ 1 := by
|
||||
have h1 := l.unique_duals.symm
|
||||
have hl : l.toIndexList = (l.tail.toIndexList.cons (l.head h)) := by
|
||||
simpa using l.head_cons_tail _
|
||||
rw [hl] at h1
|
||||
rw [withUniqueDual_eq_withDual_cons_iff'] at h1
|
||||
exact h1.2
|
||||
|
||||
lemma head_color_dual (h : l ≠ empty) :
|
||||
l.tail.val.countP (fun J => (l.head h).id = J.id) =
|
||||
l.tail.val.countP (fun J => (l.head h).id = J.id ∧ (l.head h).toColor = 𝓒.τ J.toColor) := by
|
||||
have h1 : l.withUniqueDual = l.withDual ∧ ColorCond l := ⟨l.unique_duals.symm, l.dual_color⟩
|
||||
have hl : l.toIndexList = (l.tail.toIndexList.cons (l.head h)) := by
|
||||
simpa using l.head_cons_tail _
|
||||
rw [hl, ColorCond.cons_iff] at h1
|
||||
exact h1.2.2.2
|
||||
|
||||
lemma head_cons_tail (h : l ≠ empty) :
|
||||
l = (l.tail).cons (l.head h) (l.head_uniqueDual h) (l.head_color_dual h) := by
|
||||
apply ext'
|
||||
exact IndexList.head_cons_tail _ _
|
||||
|
||||
/-!
|
||||
|
||||
## Induction for `ColorIndexList`
|
||||
|
||||
-/
|
||||
|
||||
lemma induction {P : ColorIndexList 𝓒 → Prop } (h_nil : P empty)
|
||||
(h_cons : ∀ (I : Index 𝓒.Color) (l : ColorIndexList 𝓒)
|
||||
(hI1 : l.val.countP (fun J => I.id = J.id) ≤ 1) (hI2 : l.val.countP (fun J => I.id = J.id) =
|
||||
l.val.countP (fun J => I.id = J.id ∧ I.toColor = 𝓒.τ J.toColor)), P l → P (l.cons I hI1 hI2))
|
||||
(l : ColorIndexList 𝓒) : P l := by
|
||||
by_cases h : l = empty
|
||||
· subst l
|
||||
exact h_nil
|
||||
· rw [l.head_cons_tail h]
|
||||
refine h_cons (l.head h) (l.tail) (l.head_uniqueDual h) (l.head_color_dual h) ?_
|
||||
exact induction h_nil h_cons l.tail
|
||||
termination_by l.length
|
||||
decreasing_by {
|
||||
by_cases hn : l = empty
|
||||
· subst hn
|
||||
simp
|
||||
exact False.elim (h rfl)
|
||||
· have h1 : l.tail.length < l.length := by
|
||||
simp [IndexList.tail, length]
|
||||
by_contra hn'
|
||||
simp at hn'
|
||||
have hv : l = empty := ext hn'
|
||||
exact False.elim (hn hv)
|
||||
exact h1
|
||||
}
|
||||
|
||||
/-!
|
||||
|
||||
## CountId for `ColorIndexList`
|
||||
|
||||
-/
|
||||
|
||||
lemma countId_le_two (l : ColorIndexList 𝓒) (I : Index 𝓒.Color) :
|
||||
l.countId I ≤ 2 :=
|
||||
(withUniqueDual_eq_withDual_iff_countId_leq_two' l.toIndexList).mp l.unique_duals.symm I
|
||||
lemma countId_le_two (l : ColorIndexList 𝓒) (I : Index 𝓒.Color) : l.countId I ≤ 2 :=
|
||||
(OnlyUniqueDuals.iff_countId_leq_two').mp l.unique_duals I
|
||||
|
||||
end ColorIndexList
|
||||
end IndexNotation
|
||||
|
|
|
@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
|
|||
Authors: Joseph Tooby-Smith
|
||||
-/
|
||||
import HepLean.SpaceTime.LorentzTensor.IndexNotation.ColorIndexList.Basic
|
||||
import HepLean.SpaceTime.LorentzTensor.IndexNotation.Contraction
|
||||
import HepLean.SpaceTime.LorentzTensor.IndexNotation.IndexList.Contraction
|
||||
import HepLean.SpaceTime.LorentzTensor.Contraction
|
||||
import HepLean.SpaceTime.LorentzTensor.Basic
|
||||
import Init.Data.List.Lemmas
|
||||
|
@ -36,7 +36,7 @@ open IndexList TensorColor
|
|||
def contr : ColorIndexList 𝓒 where
|
||||
toIndexList := l.toIndexList.contrIndexList
|
||||
unique_duals := by
|
||||
simp
|
||||
simp [OnlyUniqueDuals]
|
||||
dual_color := by
|
||||
funext i
|
||||
simp [Option.guard]
|
||||
|
@ -114,17 +114,9 @@ lemma contr_countId_eq_filter (I : Index 𝓒.Color) :
|
|||
exact Bool.and_comm (decide (I.id = J.id))
|
||||
(decide (List.countP (fun j => decide (J.id = j.id)) l.val = 1))
|
||||
|
||||
lemma countId_contr_le_one_of_countId (I : Index 𝓒.Color) (hI1 : l.countId I ≤ 1) :
|
||||
lemma countId_contr_le_one (I : Index 𝓒.Color) :
|
||||
l.contr.countId I ≤ 1 := by
|
||||
rw [contr_countId_eq_filter]
|
||||
by_cases hI1 : l.countId I = 0
|
||||
· rw [l.filter_id_of_countId_eq_zero' hI1]
|
||||
simp
|
||||
· have hI1 : l.countId I = 1 := by
|
||||
omega
|
||||
rw [l.consDual_filter hI1]
|
||||
apply (List.countP_le_length _).trans
|
||||
rfl
|
||||
exact l.countId_contrIndexList_le_one I
|
||||
|
||||
lemma countId_contr_eq_zero_iff (I : Index 𝓒.Color) :
|
||||
l.contr.countId I = 0 ↔ l.countId I = 0 ∨ l.countId I = 2 := by
|
||||
|
@ -155,7 +147,8 @@ lemma countId_contr_eq_zero_iff (I : Index 𝓒.Color) :
|
|||
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 only [l.unique_duals, implies_true]))
|
||||
rw [l.unique_duals]
|
||||
simp only [mem_withDual_iff_isSome, implies_true]))
|
||||
(Fin.castOrderIso l.contrIndexList_length).toEquiv).trans <|
|
||||
l.dualEquiv
|
||||
|
||||
|
|
|
@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
|
|||
Authors: Joseph Tooby-Smith
|
||||
-/
|
||||
import HepLean.SpaceTime.LorentzTensor.IndexNotation.ColorIndexList.Contraction
|
||||
import HepLean.SpaceTime.LorentzTensor.IndexNotation.IndexList.Subperm
|
||||
import HepLean.SpaceTime.LorentzTensor.Basic
|
||||
import Init.Data.List.Lemmas
|
||||
/-!
|
||||
|
@ -122,7 +123,7 @@ 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 = Finset.univ ∧
|
||||
IndexList.Subperm l.contr l'.contr.toIndexList ∧
|
||||
l'.contr.colorMap' ∘ Subtype.val ∘ (l.contr.getDualInOtherEquiv l'.contr)
|
||||
= l.contr.colorMap' ∘ Subtype.val
|
||||
|
||||
|
@ -130,47 +131,180 @@ namespace ContrPerm
|
|||
|
||||
variable {l l' l2 l3 : ColorIndexList 𝓒}
|
||||
|
||||
lemma getDualInOtherEquiv_eq (h : l.ContrPerm l2) (i : Fin l.contr.length) :
|
||||
l2.contr.val.get (l.contr.getDualInOtherEquiv l2.contr ⟨i, by
|
||||
rw [h.2.1]
|
||||
exact Finset.mem_univ i⟩).1 = l.contr.val.get i := by
|
||||
rw [Index.eq_iff_color_eq_and_id_eq]
|
||||
apply And.intro
|
||||
· trans (l2.contr.colorMap' ∘ Subtype.val ∘ (l.contr.getDualInOtherEquiv l2.contr)) ⟨i, by
|
||||
rw [h.2.1]
|
||||
exact Finset.mem_univ i⟩
|
||||
· simp
|
||||
rfl
|
||||
· simp only [h.2.2]
|
||||
rfl
|
||||
· trans l2.contr.idMap (l.contr.getDualInOtherEquiv l2.contr ⟨i, by
|
||||
rw [h.2.1]
|
||||
exact Finset.mem_univ i⟩).1
|
||||
· simp
|
||||
rfl
|
||||
· simp [getDualInOtherEquiv]
|
||||
rfl
|
||||
|
||||
lemma mem_snd_of_mem_snd (h : l.ContrPerm l2) {I : Index 𝓒.Color} (hI : I ∈ l.contr.val) :
|
||||
I ∈ l2.contr.val := by
|
||||
have h1 : l.contr.val.indexOf I < l.contr.val.length := List.indexOf_lt_length.mpr hI
|
||||
have h2 : I = l.contr.val.get ⟨l.contr.val.indexOf I, h1⟩ := (List.indexOf_get h1).symm
|
||||
rw [h2]
|
||||
rw [← getDualInOtherEquiv_eq h ⟨l.contr.val.indexOf I, h1⟩]
|
||||
simp only [List.get_eq_getElem]
|
||||
exact List.getElem_mem _ _ _
|
||||
|
||||
lemma countSelf_eq_one_snd_of_countSelf_eq_one_fst (h : l.ContrPerm l2) {I : Index 𝓒.Color}
|
||||
(h1 : l.contr.countSelf I = 1) : l2.contr.countSelf I = 1 := by
|
||||
refine countSelf_eq_one_of_countId_eq_one l2.contr I ?_ (mem_snd_of_mem_snd h ?_ )
|
||||
· have h2 := h.2.1
|
||||
rw [Subperm.iff_countId] at h2
|
||||
refine (h2 I).2 ?_
|
||||
have h1 := h2 I
|
||||
have h2' := countSelf_le_countId l.contr.toIndexList I
|
||||
omega
|
||||
· rw [← countSelf_neq_zero, h1]
|
||||
simp
|
||||
|
||||
lemma getDualInOtherEquiv_eq_of_countSelf
|
||||
(hn : IndexList.Subperm l.contr l2.contr.toIndexList) (i : Fin l.contr.length)
|
||||
(h1 : l2.contr.countId (l.contr.val.get i) = 1)
|
||||
(h2 : l2.contr.countSelf (l.contr.val.get i) = 1) :
|
||||
l2.contr.val.get (l.contr.getDualInOtherEquiv l2.contr ⟨i, by
|
||||
rw [hn]
|
||||
exact Finset.mem_univ i⟩).1 = l.contr.val.get i := by
|
||||
have h1' : (l.contr.val.get i) ∈ l2.contr.val := by
|
||||
rw [← countSelf_neq_zero, h2]
|
||||
simp
|
||||
rw [← List.mem_singleton, ← filter_id_of_countId_eq_one_mem l2.contr.toIndexList h1' h1 ]
|
||||
simp
|
||||
apply And.intro (List.getElem_mem _ _ _)
|
||||
simp [getDualInOtherEquiv]
|
||||
change _ = l2.contr.idMap (l.contr.getDualInOtherEquiv l2.contr ⟨i, by
|
||||
rw [hn]
|
||||
exact Finset.mem_univ i⟩).1
|
||||
simp [getDualInOtherEquiv]
|
||||
rfl
|
||||
|
||||
lemma colorMap_eq_of_countSelf (hn : IndexList.Subperm l.contr l2.contr.toIndexList)
|
||||
(h2 : ∀ i, l.contr.countSelf (l.contr.val.get i) = 1 → l2.contr.countSelf (l.contr.val.get i) = 1) :
|
||||
l2.contr.colorMap' ∘ Subtype.val ∘ (l.contr.getDualInOtherEquiv l2.contr)
|
||||
= l.contr.colorMap' ∘ Subtype.val := by
|
||||
funext a
|
||||
simp [colorMap', colorMap]
|
||||
change _ = (l.contr.val.get a.1).toColor
|
||||
rw [← getDualInOtherEquiv_eq_of_countSelf hn a.1]
|
||||
· rfl
|
||||
· rw [@Subperm.iff_countId_fin] at hn
|
||||
exact (hn a.1).2
|
||||
· refine h2 a.1
|
||||
(countSelf_eq_one_of_countId_eq_one l.contr.toIndexList (l.contr.val.get a.1) ?h1 ?hme)
|
||||
· rw [Subperm.iff_countId_fin] at hn
|
||||
exact (hn a.1).1
|
||||
· simp
|
||||
exact List.getElem_mem l.contr.val (↑↑a) a.1.isLt
|
||||
|
||||
lemma iff_count_fin : l.ContrPerm l2 ↔
|
||||
l.contr.length = l2.contr.length ∧ IndexList.Subperm l.contr l2.contr.toIndexList ∧
|
||||
∀ i, l.contr.countSelf (l.contr.val.get i) = 1 →
|
||||
l2.contr.countSelf (l.contr.val.get i) = 1 := by
|
||||
refine Iff.intro (fun h => ?_) (fun h => ?_)
|
||||
· refine And.intro h.1 (And.intro h.2.1 ?_)
|
||||
exact fun i a => countSelf_eq_one_snd_of_countSelf_eq_one_fst h a
|
||||
· refine And.intro h.1 (And.intro h.2.1 ?_)
|
||||
apply colorMap_eq_of_countSelf h.2.1 h.2.2
|
||||
|
||||
lemma iff_count' : l.ContrPerm l2 ↔
|
||||
l.contr.length = l2.contr.length ∧ IndexList.Subperm l.contr l2.contr.toIndexList ∧
|
||||
∀ I, l.contr.countSelf I = 1 → l2.contr.countSelf I = 1 := by
|
||||
rw [iff_count_fin]
|
||||
simp_all only [List.get_eq_getElem, and_congr_right_iff]
|
||||
intro _ _
|
||||
apply Iff.intro
|
||||
· intro ha I hI1
|
||||
have hI : I ∈ l.contr.val := by
|
||||
rw [← countSelf_neq_zero, hI1]
|
||||
simp
|
||||
have hId : l.contr.val.indexOf I < l.contr.val.length := List.indexOf_lt_length.mpr hI
|
||||
have hI' : I = l.contr.val.get ⟨l.contr.val.indexOf I, hId⟩ := (List.indexOf_get hId).symm
|
||||
rw [hI'] at hI1 ⊢
|
||||
exact ha ⟨l.contr.val.indexOf I, hId⟩ hI1
|
||||
· intro a_2 i a_3
|
||||
simp_all only
|
||||
|
||||
lemma iff_count : l.ContrPerm l2 ↔ l.contr.length = l2.contr.length ∧
|
||||
∀ I, l.contr.countSelf I = 1 → l2.contr.countSelf I = 1 := by
|
||||
rw [iff_count']
|
||||
refine Iff.intro (fun h => And.intro h.1 h.2.2) (fun h => And.intro h.1 (And.intro ?_ h.2))
|
||||
rw [Subperm.iff_countId]
|
||||
intro I
|
||||
apply And.intro (countId_contr_le_one l I)
|
||||
intro h'
|
||||
obtain ⟨I', hI1, hI2⟩ := countId_neq_zero_mem l.contr I (by omega)
|
||||
rw [countId_congr _ hI2] at h' ⊢
|
||||
have hi := h.2 I' (countSelf_eq_one_of_countId_eq_one l.contr.toIndexList I' h' hI1)
|
||||
have h1 := countSelf_le_countId l2.contr.toIndexList I'
|
||||
have h2 := countId_contr_le_one l2 I'
|
||||
omega
|
||||
|
||||
/-- The relation `ContrPerm` is symmetric. -/
|
||||
@[symm]
|
||||
lemma symm (h : ContrPerm l l') : ContrPerm l' l := by
|
||||
rw [ContrPerm] at h ⊢
|
||||
apply And.intro h.1.symm
|
||||
apply And.intro (l.contr.withUniqueDualInOther_eq_univ_symm l'.contr h.1 h.2.1)
|
||||
apply And.intro (Subperm.symm h.2.1 h.1)
|
||||
rw [← Function.comp.assoc, ← h.2.2, Function.comp.assoc, Function.comp.assoc]
|
||||
rw [show (l.contr.getDualInOtherEquiv l'.contr) =
|
||||
(l'.contr.getDualInOtherEquiv l.contr).symm from rfl]
|
||||
simp only [Equiv.symm_comp_self, CompTriple.comp_eq]
|
||||
|
||||
lemma iff_countSelf : l.ContrPerm l2 ↔ ∀ I, l.contr.countSelf I = l2.contr.countSelf I := by
|
||||
refine Iff.intro (fun h I => ?_) (fun h => ?_)
|
||||
· have hs := h.symm
|
||||
rw [iff_count] at hs h
|
||||
have ht := Iff.intro (fun h1 => h.2 I h1) (fun h2 => hs.2 I h2)
|
||||
have h1 : l.contr.countSelf I ≤ 1 := countSelf_contrIndexList_le_one l.toIndexList I
|
||||
have h2 : l2.contr.countSelf I ≤ 1 := countSelf_contrIndexList_le_one l2.toIndexList I
|
||||
omega
|
||||
· rw [iff_count]
|
||||
apply And.intro
|
||||
· have h1 : l.contr.val.Perm l2.contr.val := by
|
||||
rw [List.perm_iff_count]
|
||||
intro I
|
||||
rw [← countSelf_count, ← countSelf_count]
|
||||
exact h I
|
||||
exact List.Perm.length_eq h1
|
||||
· intro I
|
||||
rw [h I]
|
||||
simp
|
||||
|
||||
lemma contr_perm (h : l.ContrPerm l2) : l.contr.val.Perm l2.contr.val := by
|
||||
rw [List.perm_iff_count]
|
||||
intro I
|
||||
rw [← countSelf_count, ← countSelf_count]
|
||||
exact iff_countSelf.mp h I
|
||||
|
||||
/-- The relation `ContrPerm` is reflexive. -/
|
||||
@[simp]
|
||||
lemma refl (l : ColorIndexList 𝓒) : ContrPerm l l := by
|
||||
apply And.intro rfl
|
||||
apply And.intro l.withUniqueDualInOther_eq_univ_contr_refl
|
||||
simp only [getDualInOtherEquiv_self_refl, Equiv.coe_refl, CompTriple.comp_eq]
|
||||
rw [iff_countSelf]
|
||||
exact fun I => rfl
|
||||
|
||||
/-- The relation `ContrPerm` is transitive. -/
|
||||
@[trans]
|
||||
lemma trans (h1 : ContrPerm l l2) (h2 : ContrPerm l2 l3) : ContrPerm l l3 := by
|
||||
apply And.intro (h1.1.trans h2.1)
|
||||
apply And.intro (l.contr.withUniqueDualInOther_eq_univ_trans l2.contr l3.contr h1.2.1 h2.2.1)
|
||||
funext i
|
||||
simp only [Function.comp_apply]
|
||||
have h1' := congrFun h1.2.2 ⟨i, by simp [h1.2.1]⟩
|
||||
simp only [Function.comp_apply] at h1'
|
||||
rw [← h1']
|
||||
have h2' := congrFun h2.2.2 ⟨
|
||||
↑((l.contr.getDualInOtherEquiv l2.contr.toIndexList) ⟨↑i, by simp [h1.2.1]⟩), by simp [h2.2.1]⟩
|
||||
simp only [Function.comp_apply] at h2'
|
||||
rw [← h2']
|
||||
apply congrArg
|
||||
simp only [getDualInOtherEquiv, Equiv.coe_fn_mk]
|
||||
rw [← eq_getDualInOther?_get_of_withUniqueDualInOther_iff]
|
||||
· simp only [AreDualInOther, getDualInOther?_id]
|
||||
· rw [h2.2.1]
|
||||
simp
|
||||
rw [iff_countSelf] at h1 h2 ⊢
|
||||
exact fun I => (h1 I).trans (h2 I)
|
||||
|
||||
/-- `ContrPerm` forms an equivalence relation. -/
|
||||
lemma equivalence : Equivalence (@ContrPerm 𝓒 _) where
|
||||
lemma equivalence : Equivalence (@ContrPerm 𝓒 _ _) where
|
||||
refl := refl
|
||||
symm := symm
|
||||
trans := trans
|
||||
|
@ -181,12 +315,9 @@ lemma symm_trans (h1 : ContrPerm l l2) (h2 : ContrPerm l2 l3) :
|
|||
|
||||
@[simp]
|
||||
lemma contr_self : ContrPerm l l.contr := by
|
||||
rw [ContrPerm]
|
||||
simp only [contr_contr, true_and]
|
||||
have h1 := @refl _ _ l
|
||||
apply And.intro h1.2.1
|
||||
rw [show l.contr.contr = l.contr by simp]
|
||||
simp only [getDualInOtherEquiv_self_refl, Equiv.coe_refl, CompTriple.comp_eq]
|
||||
rw [iff_countSelf]
|
||||
intro I
|
||||
simp
|
||||
|
||||
@[simp]
|
||||
lemma self_contr : ContrPerm l.contr l := by
|
||||
|
@ -203,7 +334,9 @@ lemma mem_withUniqueDualInOther_of_no_contr (h : l.ContrPerm l') (h1 : l.withDua
|
|||
(h2 : l'.withDual = ∅) : ∀ (x : Fin l.length), x ∈ l.withUniqueDualInOther l'.toIndexList := by
|
||||
simp only [ContrPerm] at h
|
||||
rw [contr_of_withDual_empty l h1, contr_of_withDual_empty l' h2] at h
|
||||
simp [h.2.1]
|
||||
rw [h.2.1]
|
||||
intro x
|
||||
exact Finset.mem_univ x
|
||||
|
||||
end ContrPerm
|
||||
|
||||
|
@ -220,9 +353,9 @@ open ContrPerm
|
|||
permutation between the contracted indices. -/
|
||||
def contrPermEquiv {l l' : ColorIndexList 𝓒} (h : ContrPerm l l') :
|
||||
Fin l.contr.length ≃ Fin l'.contr.length :=
|
||||
(Equiv.subtypeUnivEquiv (by simp [h.2])).symm.trans <|
|
||||
(Equiv.subtypeUnivEquiv (by rw [h.2.1]; exact fun x => Finset.mem_univ x)).symm.trans <|
|
||||
(l.contr.getDualInOtherEquiv l'.contr.toIndexList).trans <|
|
||||
Equiv.subtypeUnivEquiv (by simp [h.symm.2])
|
||||
Equiv.subtypeUnivEquiv (by rw [h.symm.2.1]; exact fun x => Finset.mem_univ x)
|
||||
|
||||
lemma contrPermEquiv_colorMap_iso {l l' : ColorIndexList 𝓒} (h : ContrPerm l l') :
|
||||
ColorMap.MapIso (contrPermEquiv h).symm l'.contr.colorMap' l.contr.colorMap' := by
|
||||
|
@ -244,7 +377,7 @@ lemma contrPermEquiv_colorMap_iso' {l l' : ColorIndexList 𝓒} (h : ContrPerm l
|
|||
exact contrPermEquiv_colorMap_iso h
|
||||
|
||||
@[simp]
|
||||
lemma contrPermEquiv_refl : contrPermEquiv (@ContrPerm.refl 𝓒 _ l) = Equiv.refl _ := by
|
||||
lemma contrPermEquiv_refl : contrPermEquiv (ContrPerm.refl l) = Equiv.refl _ := by
|
||||
simp [contrPermEquiv, ContrPerm.refl]
|
||||
|
||||
@[simp]
|
||||
|
@ -262,14 +395,25 @@ lemma contrPermEquiv_trans {l l2 l3 : ColorIndexList 𝓒}
|
|||
simp only [getDualInOtherEquiv, Equiv.trans_apply, Equiv.subtypeUnivEquiv_symm_apply,
|
||||
Equiv.coe_fn_mk, Equiv.subtypeUnivEquiv_apply]
|
||||
apply congrArg
|
||||
rw [← eq_getDualInOther?_get_of_withUniqueDualInOther_iff]
|
||||
· simp [AreDualInOther]
|
||||
· rw [(h1.trans h2).2.1]
|
||||
simp
|
||||
have h1' : l.contr.countSelf (l.contr.val.get x) = 1 := by simp [contr]
|
||||
rw [iff_countSelf.mp h1, iff_countSelf.mp h2] at h1'
|
||||
have h3 : l3.contr.countId (l.contr.val.get x) = 1 := by
|
||||
have h' := countSelf_le_countId l3.contr.toIndexList (l.contr.val.get x)
|
||||
have h'' := countId_contr_le_one l3 (l.contr.val.get x)
|
||||
omega
|
||||
rw [countId_get_other, List.countP_eq_length_filter, List.length_eq_one] at h3
|
||||
obtain ⟨a, ha⟩ := h3
|
||||
trans a
|
||||
· rw [← List.mem_singleton, ← ha]
|
||||
simp [AreDualInOther]
|
||||
· symm
|
||||
rw [← List.mem_singleton, ← ha]
|
||||
simp [AreDualInOther]
|
||||
|
||||
|
||||
@[simp]
|
||||
lemma contrPermEquiv_self_contr {l : ColorIndexList 𝓒} :
|
||||
contrPermEquiv (by simp : ContrPerm l l.contr) =
|
||||
contrPermEquiv (contr_self : ContrPerm l l.contr) =
|
||||
(Fin.castOrderIso (by simp)).toEquiv := by
|
||||
simp [contrPermEquiv]
|
||||
ext1 x
|
||||
|
@ -277,11 +421,21 @@ lemma contrPermEquiv_self_contr {l : ColorIndexList 𝓒} :
|
|||
Equiv.coe_fn_mk, Equiv.subtypeUnivEquiv_apply, RelIso.coe_fn_toEquiv, Fin.castOrderIso_apply,
|
||||
Fin.coe_cast]
|
||||
symm
|
||||
rw [← eq_getDualInOther?_get_of_withUniqueDualInOther_iff]
|
||||
· simp only [AreDualInOther, contr_contr_idMap, Fin.cast_trans, Fin.cast_eq_self]
|
||||
· have h1 : ContrPerm l l.contr := by simp
|
||||
rw [h1.2.1]
|
||||
simp
|
||||
have h1' : l.contr.countSelf (l.contr.val.get x) = 1 := by simp [contr]
|
||||
rw [iff_countSelf.mp contr_self] at h1'
|
||||
have h3 : l.contr.contr.countId (l.contr.val.get x) = 1 := by
|
||||
have h' := countSelf_le_countId l.contr.contr.toIndexList (l.contr.val.get x)
|
||||
have h'' := countId_contr_le_one l.contr (l.contr.val.get x)
|
||||
omega
|
||||
rw [countId_get_other, List.countP_eq_length_filter, List.length_eq_one] at h3
|
||||
obtain ⟨a, ha⟩ := h3
|
||||
trans a
|
||||
· rw [← List.mem_singleton, ← ha]
|
||||
simp [AreDualInOther]
|
||||
· symm
|
||||
rw [← List.mem_singleton, ← ha]
|
||||
simp only [AreDualInOther, List.mem_filter, List.mem_finRange,
|
||||
decide_eq_true_eq, true_and, getDualInOther?_id]
|
||||
|
||||
@[simp]
|
||||
lemma contrPermEquiv_contr_self {l : ColorIndexList 𝓒} :
|
||||
|
|
|
@ -80,10 +80,6 @@ lemma mem_withoutDual_iff_countId_eq_one (i : Fin l.length) :
|
|||
rw [mem_withDual_iff_countId_gt_one] at h
|
||||
omega
|
||||
|
||||
/-- An equivalence from `Fin l.withoutDual.card` to `l.withoutDual` determined by
|
||||
the order on `l.withoutDual` inherted from `Fin`. -/
|
||||
def withoutDualEquiv : Fin l.withoutDual.card ≃ l.withoutDual :=
|
||||
(Finset.orderIsoOfFin l.withoutDual (by rfl)).toEquiv
|
||||
|
||||
lemma list_ofFn_withoutDualEquiv_eq_sort :
|
||||
List.ofFn (Subtype.val ∘ l.withoutDualEquiv) = l.withoutDual.sort (fun i j => i ≤ j) := by
|
||||
|
|
|
@ -0,0 +1,579 @@
|
|||
/-
|
||||
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.IndexList.Contraction
|
||||
import HepLean.SpaceTime.LorentzTensor.IndexNotation.IndexList.OnlyUniqueDuals
|
||||
import HepLean.SpaceTime.LorentzTensor.Basic
|
||||
import Init.Data.List.Lemmas
|
||||
/-!
|
||||
|
||||
# Index lists and color
|
||||
|
||||
In this file we look at the interaction of index lists and color.
|
||||
|
||||
The main definition of this file is `ColorCond`.
|
||||
-/
|
||||
|
||||
namespace IndexNotation
|
||||
|
||||
namespace Index
|
||||
|
||||
variable {𝓒 : TensorColor}
|
||||
variable [IndexNotation 𝓒.Color] [Fintype 𝓒.Color] [DecidableEq 𝓒.Color]
|
||||
variable (I : Index 𝓒.Color)
|
||||
|
||||
/-- The dual of an index is the index with the same id, but opposite color. -/
|
||||
def dual : Index 𝓒.Color := ⟨𝓒.τ I.toColor, I.id⟩
|
||||
|
||||
@[simp]
|
||||
lemma dual_dual : I.dual.dual = I := by
|
||||
simp [dual, toColor, id]
|
||||
rw [𝓒.τ_involutive]
|
||||
rfl
|
||||
|
||||
@[simp]
|
||||
lemma dual_id : I.dual.id = I.id := by
|
||||
simp [dual, id]
|
||||
|
||||
@[simp]
|
||||
lemma dual_color : I.dual.toColor = 𝓒.τ I.toColor := by
|
||||
simp [dual, toColor]
|
||||
|
||||
end Index
|
||||
|
||||
namespace IndexList
|
||||
|
||||
variable {𝓒 : TensorColor}
|
||||
variable [IndexNotation 𝓒.Color] [Fintype 𝓒.Color] [DecidableEq 𝓒.Color]
|
||||
variable (l l2 l3 : IndexList 𝓒.Color)
|
||||
|
||||
def countColorQuot (I : Index 𝓒.Color) : ℕ := l.val.countP (fun J => I = J ∨ I.dual = J)
|
||||
|
||||
lemma countColorQuot_eq_filter_id_countP (I : Index 𝓒.Color) :
|
||||
l.countColorQuot I =
|
||||
(l.val.filter (fun J => I.id = J.id)).countP
|
||||
(fun J => I.toColor = J.toColor ∨ I.toColor = 𝓒.τ (J.toColor)) := by
|
||||
simp [countColorQuot]
|
||||
rw [List.countP_filter]
|
||||
apply List.countP_congr
|
||||
intro I' _
|
||||
simp [Index.eq_iff_color_eq_and_id_eq]
|
||||
apply Iff.intro
|
||||
· intro a_1
|
||||
cases a_1 with
|
||||
| inl h => simp_all only [true_or, and_self]
|
||||
| inr h_1 =>
|
||||
simp_all only [and_true]
|
||||
obtain ⟨left, _⟩ := h_1
|
||||
rw [← left]
|
||||
rw [𝓒.τ_involutive]
|
||||
simp
|
||||
· intro a_1
|
||||
simp_all only [and_true]
|
||||
obtain ⟨left, _⟩ := a_1
|
||||
cases left with
|
||||
| inl h => simp_all only [true_or]
|
||||
| inr h_1 =>
|
||||
simp_all only
|
||||
rw [𝓒.τ_involutive]
|
||||
simp
|
||||
|
||||
lemma countColorQuot_eq_filter_color_countP (I : Index 𝓒.Color) :
|
||||
l.countColorQuot I = (l.val.filter (fun J => I.toColor = J.toColor ∨ I.toColor = 𝓒.τ (J.toColor))).countP
|
||||
(fun J => I.id = J.id) := by
|
||||
rw [countColorQuot_eq_filter_id_countP]
|
||||
rw [List.countP_filter, List.countP_filter]
|
||||
apply List.countP_congr
|
||||
intro I' _
|
||||
simpa using And.comm
|
||||
|
||||
@[simp]
|
||||
lemma countColorQuot_append (I : Index 𝓒.Color) :
|
||||
(l ++ l2).countColorQuot I = countColorQuot l I + countColorQuot l2 I := by
|
||||
simp [countColorQuot]
|
||||
|
||||
lemma countColorQuot_eq_countId_iff_of_isSome (hl : l.OnlyUniqueDuals) (i : Fin l.length)
|
||||
(hi : (l.getDual? i).isSome) :
|
||||
l.countColorQuot (l.val.get i) = l.countId (l.val.get i) ↔
|
||||
(l.colorMap i = l.colorMap ((l.getDual? i).get hi) ∨
|
||||
l.colorMap i = 𝓒.τ (l.colorMap ((l.getDual? i).get hi))) := by
|
||||
rw [countColorQuot_eq_filter_id_countP, countId_eq_length_filter]
|
||||
have hi1 := hi
|
||||
rw [← mem_withDual_iff_isSome, ← hl, mem_withUniqueDual_iff_countId_eq_two] at hi1
|
||||
rcases l.filter_id_of_countId_eq_two hi1 with hf | hf
|
||||
all_goals
|
||||
erw [hf]
|
||||
simp [List.countP, List.countP.go]
|
||||
refine Iff.intro (fun h => ?_) (fun h => ?_)
|
||||
· by_contra hn
|
||||
have hn' : (decide (l.val[↑i].toColor = l.val[↑((l.getDual? i).get hi)].toColor) ||
|
||||
decide (l.val[↑i].toColor = 𝓒.τ l.val[↑((l.getDual? i).get hi)].toColor)) = false := by
|
||||
simpa using hn
|
||||
erw [hn'] at h
|
||||
simp at h
|
||||
· have hn' : (decide (l.val[↑i].toColor = l.val[↑((l.getDual? i).get hi)].toColor) ||
|
||||
decide (l.val[↑i].toColor = 𝓒.τ l.val[↑((l.getDual? i).get hi)].toColor)) = true := by
|
||||
simpa using h
|
||||
erw [hn']
|
||||
simp
|
||||
|
||||
lemma countColorQuot_of_countId_zero {I : Index 𝓒.Color} (h : l.countId I = 0) :
|
||||
l.countColorQuot I = 0 := by
|
||||
rw [countColorQuot_eq_filter_id_countP]
|
||||
rw [countId_eq_length_filter, List.length_eq_zero] at h
|
||||
simp [h, countColorQuot]
|
||||
|
||||
lemma countColorQuot_le_countId (I : Index 𝓒.Color) :
|
||||
l.countColorQuot I ≤ l.countId I := by
|
||||
rw [countColorQuot_eq_filter_color_countP, countId]
|
||||
apply List.Sublist.countP_le
|
||||
exact List.filter_sublist l.val
|
||||
|
||||
lemma countColorQuot_contrIndexList_le_one (I : Index 𝓒.Color) :
|
||||
l.contrIndexList.countColorQuot I ≤ 1 :=
|
||||
(l.contrIndexList.countColorQuot_le_countId I).trans
|
||||
(countId_contrIndexList_le_one l I)
|
||||
|
||||
lemma countColorQuot_contrIndexList_eq_zero_or_one (I : Index 𝓒.Color) :
|
||||
l.contrIndexList.countColorQuot I = 0 ∨ l.contrIndexList.countColorQuot I = 1 := by
|
||||
have h1 := countColorQuot_contrIndexList_le_one l I
|
||||
omega
|
||||
|
||||
lemma countColorQuot_contrIndexList_le_countColorQuot (I : Index 𝓒.Color) :
|
||||
l.contrIndexList.countColorQuot I ≤ l.countColorQuot I := by
|
||||
rw [countColorQuot_eq_filter_color_countP, countColorQuot_eq_filter_color_countP]
|
||||
apply List.Sublist.countP_le
|
||||
exact List.Sublist.filter _ (List.filter_sublist l.val)
|
||||
|
||||
lemma countColorQuot_contrIndexList_eq_of_countId_eq
|
||||
(h1 : l.contrIndexList.countId I = l.countId I) :
|
||||
l.contrIndexList.countColorQuot I = l.countColorQuot I := by
|
||||
rw [countColorQuot_eq_filter_id_countP,
|
||||
l.filter_id_contrIndexList_eq_of_countId_contrIndexList I h1,
|
||||
countColorQuot_eq_filter_id_countP]
|
||||
|
||||
def countSelf (I : Index 𝓒.Color) : ℕ := l.val.countP (fun J => I = J)
|
||||
|
||||
lemma countSelf_eq_filter_id_countP : l.countSelf I =
|
||||
(l.val.filter (fun J => I.id = J.id)).countP (fun J => I.toColor = J.toColor) := by
|
||||
simp [countSelf]
|
||||
rw [List.countP_filter]
|
||||
apply List.countP_congr
|
||||
intro I' _
|
||||
simp [Index.eq_iff_color_eq_and_id_eq]
|
||||
|
||||
lemma countSelf_eq_filter_color_countP :
|
||||
l.countSelf I =
|
||||
(l.val.filter (fun J => I.toColor = J.toColor)).countP (fun J => I.id = J.id) := by
|
||||
simp [countSelf]
|
||||
rw [List.countP_filter]
|
||||
apply List.countP_congr
|
||||
intro I' _
|
||||
simpa [Index.eq_iff_color_eq_and_id_eq] using And.comm
|
||||
|
||||
lemma countSelf_of_countId_zero {I : Index 𝓒.Color} (h : l.countId I = 0) :
|
||||
l.countSelf I = 0 := by
|
||||
rw [countId_eq_length_filter, List.length_eq_zero] at h
|
||||
simp [h, countSelf_eq_filter_id_countP]
|
||||
|
||||
lemma countSelf_count (I : Index 𝓒.Color) : l.countSelf I = l.val.count I := by
|
||||
rw [countSelf, List.count]
|
||||
apply List.countP_congr
|
||||
intro I' _
|
||||
simp
|
||||
exact eq_comm
|
||||
|
||||
lemma countSelf_eq_zero (I : Index 𝓒.Color) : l.countSelf I = 0 ↔ I ∉ l.val := by
|
||||
rw [countSelf_count]
|
||||
exact List.count_eq_zero
|
||||
|
||||
lemma countSelf_neq_zero (I : Index 𝓒.Color) : l.countSelf I ≠ 0 ↔ I ∈ l.val := by
|
||||
refine Iff.intro (fun h => ?_) (fun h => ?_)
|
||||
· simpa using (l.countSelf_eq_zero I).mpr.mt h
|
||||
· exact (l.countSelf_eq_zero I).mp.mt (by simpa using h)
|
||||
|
||||
@[simp]
|
||||
lemma countSelf_append (I : Index 𝓒.Color) :
|
||||
(l ++ l2).countSelf I = countSelf l I + countSelf l2 I := by
|
||||
simp [countSelf]
|
||||
|
||||
lemma countSelf_le_countId (I : Index 𝓒.Color) :
|
||||
l.countSelf I ≤ l.countId I := by
|
||||
rw [countSelf_eq_filter_color_countP, countId]
|
||||
apply List.Sublist.countP_le
|
||||
exact List.filter_sublist l.val
|
||||
|
||||
lemma countSelf_eq_one_of_countId_eq_one (I : Index 𝓒.Color) (h1 : l.countId I = 1)
|
||||
(hme : I ∈ l.val) : l.countSelf I = 1 := by
|
||||
rw [countSelf_eq_filter_id_countP]
|
||||
rw [filter_id_of_countId_eq_one_mem l hme h1]
|
||||
simp
|
||||
|
||||
lemma countSelf_contrIndexList_le_one (I : Index 𝓒.Color) :
|
||||
l.contrIndexList.countSelf I ≤ 1 :=
|
||||
(l.contrIndexList.countSelf_le_countId I).trans (countId_contrIndexList_le_one l I)
|
||||
|
||||
lemma countSelf_contrIndexList_eq_zero_or_one (I : Index 𝓒.Color) :
|
||||
l.contrIndexList.countSelf I = 0 ∨ l.contrIndexList.countSelf I = 1 := by
|
||||
have h1 := countSelf_contrIndexList_le_one l I
|
||||
omega
|
||||
|
||||
lemma countSelf_contrIndexList_eq_zero_of_zero (I : Index 𝓒.Color) (h : l.countSelf I = 0) :
|
||||
l.contrIndexList.countSelf I = 0 := by
|
||||
rw [countSelf_eq_zero] at h ⊢
|
||||
simp_all [contrIndexList]
|
||||
|
||||
lemma countSelf_contrIndexList_le_countSelf (I : Index 𝓒.Color) :
|
||||
l.contrIndexList.countSelf I ≤ l.countSelf I := by
|
||||
rw [countSelf_eq_filter_color_countP, countSelf_eq_filter_color_countP]
|
||||
apply List.Sublist.countP_le
|
||||
exact List.Sublist.filter _ (List.filter_sublist l.val)
|
||||
|
||||
lemma countSelf_contrIndexList_eq_of_countId_eq
|
||||
(h1 : l.contrIndexList.countId I = l.countId I) :
|
||||
l.contrIndexList.countSelf I = l.countSelf I := by
|
||||
rw [countSelf_eq_filter_id_countP,
|
||||
l.filter_id_contrIndexList_eq_of_countId_contrIndexList I h1,
|
||||
countSelf_eq_filter_id_countP]
|
||||
|
||||
@[simp]
|
||||
lemma countSelf_contrIndexList_get (i : Fin l.contrIndexList.length) :
|
||||
l.contrIndexList.countSelf l.contrIndexList.val[Fin.val i] = 1 := by
|
||||
refine countSelf_eq_one_of_countId_eq_one _ _ ?h1 ?hme
|
||||
· refine mem_contrIndexList_countId_contrIndexList l ?_
|
||||
exact List.getElem_mem l.contrIndexList.val (↑i) _
|
||||
· exact List.getElem_mem l.contrIndexList.val (↑i) _
|
||||
|
||||
def countDual (I : Index 𝓒.Color) : ℕ := l.val.countP (fun J => I.dual = J)
|
||||
|
||||
lemma countDual_eq_countSelf_Dual (I : Index 𝓒.Color) : l.countDual I = l.countSelf I.dual := by
|
||||
rw [countDual, countSelf]
|
||||
|
||||
lemma countDual_eq_filter_id_countP : l.countDual I =
|
||||
(l.val.filter (fun J => I.id = J.id)).countP (fun J => I.toColor = 𝓒.τ (J.toColor)) := by
|
||||
simp [countDual]
|
||||
rw [List.countP_filter]
|
||||
apply List.countP_congr
|
||||
intro I' _
|
||||
simp [Index.eq_iff_color_eq_and_id_eq, Index.dual, Index.toColor, Index.id]
|
||||
intro _
|
||||
refine Iff.intro (fun h => ?_) (fun h => ?_)
|
||||
· rw [← h]
|
||||
exact (𝓒.τ_involutive _).symm
|
||||
· rw [h]
|
||||
exact (𝓒.τ_involutive _)
|
||||
|
||||
lemma countDual_of_countId_zero {I : Index 𝓒.Color} (h : l.countId I = 0) :
|
||||
l.countDual I = 0 := by
|
||||
rw [countId_eq_length_filter, List.length_eq_zero] at h
|
||||
simp [h, countDual_eq_filter_id_countP]
|
||||
|
||||
@[simp]
|
||||
lemma countDual_append (I : Index 𝓒.Color) :
|
||||
(l ++ l2).countDual I = countDual l I + countDual l2 I := by
|
||||
simp [countDual]
|
||||
|
||||
lemma countDual_contrIndexList_eq_of_countId_eq
|
||||
(h1 : l.contrIndexList.countId I = l.countId I) :
|
||||
l.contrIndexList.countDual I = l.countDual I := by
|
||||
rw [countDual_eq_countSelf_Dual, countDual_eq_countSelf_Dual]
|
||||
refine countSelf_contrIndexList_eq_of_countId_eq l h1
|
||||
|
||||
lemma countSelf_eq_countDual_iff_of_isSome (hl : l.OnlyUniqueDuals)
|
||||
(i : Fin l.length) (hi : (l.getDual? i).isSome) :
|
||||
l.countSelf (l.val.get i) = l.countDual (l.val.get i) ↔
|
||||
l.colorMap i = 𝓒.τ (l.colorMap ((l.getDual? i).get hi))
|
||||
∨ (l.colorMap i) = 𝓒.τ (l.colorMap i) := by
|
||||
rw [countSelf_eq_filter_id_countP, countDual_eq_filter_id_countP]
|
||||
have hi1 := hi
|
||||
rw [← mem_withDual_iff_isSome, ← hl, mem_withUniqueDual_iff_countId_eq_two] at hi1
|
||||
rcases l.filter_id_of_countId_eq_two hi1 with hf | hf
|
||||
all_goals
|
||||
erw [hf]
|
||||
simp [List.countP, List.countP.go]
|
||||
by_cases hn : l.colorMap i = 𝓒.τ (l.colorMap ((l.getDual? i).get hi))
|
||||
· simp [hn]
|
||||
have hn' : decide (l.val[↑i].toColor = 𝓒.τ l.val[↑((l.getDual? i).get hi)].toColor)
|
||||
= true := by simpa [colorMap] using hn
|
||||
erw [hn']
|
||||
simp only [cond_true]
|
||||
have hτ : l.colorMap ((l.getDual? i).get hi) = 𝓒.τ (l.colorMap i) := by
|
||||
rw [hn]
|
||||
exact (𝓒.τ_involutive _).symm
|
||||
simp [colorMap] at hτ
|
||||
erw [hτ]
|
||||
· have hn' : decide (l.val[↑i].toColor = 𝓒.τ l.val[↑((l.getDual? i).get hi)].toColor) =
|
||||
false := by simpa [colorMap] using hn
|
||||
erw [hn']
|
||||
simp [hn]
|
||||
by_cases hm : l.colorMap i = 𝓒.τ (l.colorMap i)
|
||||
· trans True
|
||||
· simp
|
||||
have hm' : decide (l.val[↑i].toColor = 𝓒.τ l.val[↑i].toColor) = true := by simpa using hm
|
||||
erw [hm']
|
||||
simp only [cond_true]
|
||||
have hm'' : decide (l.val[↑i].toColor = l.val[↑((l.getDual? i).get hi)].toColor)
|
||||
= false := by
|
||||
simp only [Fin.getElem_fin, decide_eq_false_iff_not]
|
||||
simp [colorMap] at hm
|
||||
erw [hm]
|
||||
by_contra hn'
|
||||
have hn'' : l.colorMap i = 𝓒.τ (l.colorMap ((l.getDual? i).get hi)) := by
|
||||
simp [colorMap]
|
||||
rw [← hn']
|
||||
exact (𝓒.τ_involutive _).symm
|
||||
exact hn hn''
|
||||
erw [hm'']
|
||||
simp
|
||||
· exact true_iff_iff.mpr hm
|
||||
· simp [hm]
|
||||
simp [colorMap] at hm
|
||||
have hm' : decide (l.val[↑i].toColor = 𝓒.τ l.val[↑i].toColor) = false := by simpa using hm
|
||||
erw [hm']
|
||||
simp only [cond_false, ne_eq]
|
||||
by_cases hm'' : decide (l.val[↑i].toColor = l.val[↑((l.getDual? i).get hi)].toColor) = true
|
||||
· erw [hm'']
|
||||
simp
|
||||
· have hm''' : decide (l.val[↑i].toColor = l.val[↑((l.getDual? i).get hi)].toColor)
|
||||
= false := by
|
||||
simpa using hm''
|
||||
erw [hm''']
|
||||
simp
|
||||
|
||||
/-- The condition an index and its' dual, when it exists, have dual colors. -/
|
||||
def ColorCond : Prop := Option.map l.colorMap ∘
|
||||
l.getDual? = Option.map (𝓒.τ ∘ l.colorMap) ∘
|
||||
Option.guard fun i => (l.getDual? i).isSome
|
||||
|
||||
namespace ColorCond
|
||||
|
||||
variable {l l2 l3 : IndexList 𝓒.Color}
|
||||
|
||||
lemma iff_withDual :
|
||||
l.ColorCond ↔ ∀ (i : l.withDual), 𝓒.τ
|
||||
(l.colorMap ((l.getDual? i).get (l.withDual_isSome i))) = l.colorMap i := by
|
||||
refine Iff.intro (fun h i => ?_) (fun h => ?_)
|
||||
· have h' := congrFun h i
|
||||
simp at h'
|
||||
rw [show l.getDual? i = some ((l.getDual? i).get (l.withDual_isSome i)) by simp] at h'
|
||||
have h'' : (Option.guard (fun i => (l.getDual? i).isSome = true) ↑i) = i := by
|
||||
apply Option.guard_eq_some.mpr
|
||||
simp [l.withDual_isSome i]
|
||||
rw [h'', Option.map_some', Option.map_some'] at h'
|
||||
simp at h'
|
||||
rw [h']
|
||||
exact 𝓒.τ_involutive (l.colorMap i)
|
||||
· funext i
|
||||
by_cases hi : (l.getDual? i).isSome
|
||||
· have h'' : (Option.guard (fun i => (l.getDual? i).isSome = true) ↑i) = i := by
|
||||
apply Option.guard_eq_some.mpr
|
||||
simp only [true_and]
|
||||
exact hi
|
||||
simp only [Function.comp_apply, h'', Option.map_some']
|
||||
rw [show l.getDual? ↑i = some ((l.getDual? i).get hi) by simp]
|
||||
rw [Option.map_some']
|
||||
simp only [Option.some.injEq]
|
||||
have hii := h ⟨i, by simp [withDual, hi]⟩
|
||||
simp at hii
|
||||
rw [← hii]
|
||||
exact (𝓒.τ_involutive _).symm
|
||||
· simp [Option.guard, hi]
|
||||
exact Option.not_isSome_iff_eq_none.mp hi
|
||||
|
||||
lemma iff_on_isSome : l.ColorCond ↔ ∀ (i : Fin l.length) (h : (l.getDual? i).isSome), 𝓒.τ
|
||||
(l.colorMap ((l.getDual? i).get h)) = l.colorMap i := by
|
||||
rw [iff_withDual]
|
||||
simp only [Subtype.forall, mem_withDual_iff_isSome]
|
||||
|
||||
/-- A condition on an index list `l` and and index `I`. If the id of `I` appears
|
||||
twice in `l` (and `I` at least once) then this condition is equivalent to the dual of `I` having
|
||||
dual color to `I`, but written totally in terms of lists. -/
|
||||
@[simp]
|
||||
abbrev countColorCond (l : IndexList 𝓒.Color) (I : Index 𝓒.Color) : Prop :=
|
||||
l.countColorQuot I = l.countId I ∧
|
||||
l.countSelf I = l.countDual I
|
||||
|
||||
lemma countColorCond_of_filter_eq (l l2 : IndexList 𝓒.Color) {I : Index 𝓒.Color}
|
||||
(hf : l.val.filter (fun J => I.id = J.id) = l2.val.filter (fun J => I.id = J.id))
|
||||
(h1 : countColorCond l I) : countColorCond l2 I := by
|
||||
rw [countColorCond, countColorQuot_eq_filter_id_countP, countId_eq_length_filter, countSelf_eq_filter_id_countP,
|
||||
countDual_eq_filter_id_countP, ← hf]
|
||||
rw [countColorCond, countColorQuot_eq_filter_id_countP, countId_eq_length_filter, countSelf_eq_filter_id_countP,
|
||||
countDual_eq_filter_id_countP] at h1
|
||||
exact h1
|
||||
|
||||
lemma iff_countColorCond_isSome (hl : l.OnlyUniqueDuals) : l.ColorCond ↔
|
||||
∀ (i : Fin l.length) (_ : (l.getDual? i).isSome), countColorCond l (l.val.get i) := by
|
||||
rw [iff_on_isSome]
|
||||
simp only [countColorCond]
|
||||
refine Iff.intro (fun h i hi => ?_) (fun h i hi => ?_)
|
||||
· rw [l.countColorQuot_eq_countId_iff_of_isSome hl i hi,
|
||||
l.countSelf_eq_countDual_iff_of_isSome hl i hi]
|
||||
have hi' := h i hi
|
||||
exact And.intro (Or.inr hi'.symm) (Or.inl hi'.symm)
|
||||
· have hi' := h i hi
|
||||
rw [l.countColorQuot_eq_countId_iff_of_isSome hl i hi,
|
||||
l.countSelf_eq_countDual_iff_of_isSome hl i hi] at hi'
|
||||
rcases hi'.1 with hi1 | hi1
|
||||
<;> rcases hi'.2 with hi2 | hi2
|
||||
· exact hi2.symm
|
||||
· rw [← hi1]
|
||||
exact hi2.symm
|
||||
· exact hi1.symm
|
||||
· exact hi1.symm
|
||||
|
||||
lemma iff_countColorCond_index (hl : l.OnlyUniqueDuals) :
|
||||
l.ColorCond ↔ ∀ (i : Fin l.length), l.countId (l.val.get i) = 2
|
||||
→ countColorCond l (l.val.get i) := by
|
||||
rw [iff_countColorCond_isSome hl]
|
||||
refine Iff.intro (fun h i hi => ?_) (fun h i hi => ?_)
|
||||
· rw [← mem_withUniqueDual_iff_countId_eq_two] at hi
|
||||
exact h i (mem_withUniqueDual_isSome l i hi)
|
||||
· rw [← mem_withDual_iff_isSome, ← hl, mem_withUniqueDual_iff_countId_eq_two] at hi
|
||||
exact h i hi
|
||||
|
||||
lemma iff_countColorCond_mem (hl : l.OnlyUniqueDuals) :
|
||||
l.ColorCond ↔ ∀ (I : Index 𝓒.Color) (_ : I ∈ l.val),
|
||||
l.countId I = 2 → countColorCond l I := by
|
||||
rw [iff_countColorCond_index hl]
|
||||
refine Iff.intro (fun h I hI hi => ?_) (fun h i hi => ?_)
|
||||
· let i := l.val.indexOf I
|
||||
have hi' : i < l.length := List.indexOf_lt_length.mpr hI
|
||||
have hIi : I = l.val.get ⟨i, hi'⟩ := (List.indexOf_get hi').symm
|
||||
rw [hIi] at hi ⊢
|
||||
exact h ⟨i, hi'⟩ hi
|
||||
· exact h (l.val.get i) (List.getElem_mem l.val (↑i) i.isLt) hi
|
||||
|
||||
|
||||
lemma mem_iff_dual_mem (hl : l.OnlyUniqueDuals) (hc : l.ColorCond) (I : Index 𝓒.Color)
|
||||
(hId : l.countId I = 2) : I ∈ l.val ↔ I.dual ∈ l.val := by
|
||||
rw [iff_countColorCond_mem hl] at hc
|
||||
refine Iff.intro (fun h => ?_) (fun h => ?_)
|
||||
· have hc' := hc I h hId
|
||||
simp only [countColorCond] at hc'
|
||||
rw [← countSelf_neq_zero] at h
|
||||
rw [← countSelf_neq_zero, ← countDual_eq_countSelf_Dual]
|
||||
omega
|
||||
· have hIdd : l.countId I.dual = 2 := by
|
||||
rw [← hId]
|
||||
apply countId_congr
|
||||
simp
|
||||
have hc' := (hc I.dual h hIdd).2
|
||||
simp at hc'
|
||||
rw [← countSelf_neq_zero]
|
||||
rw [← countSelf_neq_zero] at h
|
||||
rw [countDual_eq_countSelf_Dual] at hc'
|
||||
simp at hc'
|
||||
omega
|
||||
|
||||
lemma iff_countColorCond (hl : l.OnlyUniqueDuals) :
|
||||
l.ColorCond ↔ ∀ I, l.countSelf I ≠ 0 → l.countId I = 2 → countColorCond l I := by
|
||||
refine Iff.intro (fun h I hIs hi => ?_) (fun h => ?_)
|
||||
· rw [countSelf_neq_zero] at hIs
|
||||
rw [iff_countColorCond_mem hl] at h
|
||||
exact h I hIs hi
|
||||
· rw [iff_countColorCond_mem hl]
|
||||
intro I hmem hi
|
||||
refine h I ?_ hi
|
||||
rw [countSelf_neq_zero]
|
||||
exact hmem
|
||||
|
||||
lemma assoc (h : ColorCond (l ++ l2 ++ l3)) : ColorCond (l ++ (l2 ++ l3)) := by
|
||||
rw [← append_assoc]
|
||||
exact h
|
||||
|
||||
lemma symm (hl : (l ++ l2).OnlyUniqueDuals) (h : ColorCond (l ++ l2)) :
|
||||
ColorCond (l2 ++ l) := by
|
||||
rw [iff_countColorCond hl] at h
|
||||
rw [iff_countColorCond (OnlyUniqueDuals.symm' hl)]
|
||||
intro I hI1 hI2
|
||||
have hI' := h I (by simp_all; omega) (by simp_all; omega)
|
||||
simp_all
|
||||
omega
|
||||
|
||||
lemma inl (hl : (l ++ l2).OnlyUniqueDuals) (h : ColorCond (l ++ l2)) : ColorCond l := by
|
||||
rw [iff_countColorCond hl] at h
|
||||
rw [iff_countColorCond (OnlyUniqueDuals.inl hl)]
|
||||
intro I hI1 hI2
|
||||
have hI2' : l2.countId I = 0 := by
|
||||
rw [OnlyUniqueDuals.iff_countId_leq_two'] at hl
|
||||
have hlI := hl I
|
||||
simp at hlI
|
||||
omega
|
||||
have hI' := h I (by
|
||||
simp only [countSelf_append, ne_eq, add_eq_zero, not_and, hI1, false_implies])
|
||||
(by simp_all)
|
||||
simp at hI'
|
||||
rw [l2.countColorQuot_of_countId_zero hI2', l2.countSelf_of_countId_zero hI2',
|
||||
l2.countDual_of_countId_zero hI2', hI2'] at hI'
|
||||
simp at hI'
|
||||
omega
|
||||
|
||||
lemma inr (hl : (l ++ l2).OnlyUniqueDuals) (h : ColorCond (l ++ l2)) : ColorCond l2 := by
|
||||
have hs := symm hl h
|
||||
rw [OnlyUniqueDuals.symm] at hl
|
||||
exact inl hl hs
|
||||
|
||||
lemma swap (hl : (l ++ l2 ++ l3).OnlyUniqueDuals) (h : ColorCond (l ++ l2 ++ l3)) :
|
||||
ColorCond (l2 ++ l ++ l3) := by
|
||||
rw [iff_countColorCond hl] at h
|
||||
rw [iff_countColorCond (OnlyUniqueDuals.swap hl)]
|
||||
intro I hI1 hI2
|
||||
have hI' := h I (by simp_all) (by simp_all; omega)
|
||||
simp_all only [countSelf_append, ne_eq, add_eq_zero, not_and, and_imp, countId_append,
|
||||
countColorCond, countColorQuot_append, countDual_append, not_false_eq_true, implies_true]
|
||||
omega
|
||||
|
||||
/-!
|
||||
|
||||
## Contractions
|
||||
|
||||
-/
|
||||
|
||||
lemma contrIndexList : ColorCond l.contrIndexList := by
|
||||
funext i
|
||||
simp [Option.guard]
|
||||
|
||||
lemma contrIndexList_left (hl : (l ++ l2).OnlyUniqueDuals) (h1 : (l ++ l2).ColorCond) :
|
||||
ColorCond (l.contrIndexList ++ l2) := by
|
||||
rw [iff_countColorCond hl] at h1
|
||||
rw [iff_countColorCond (OnlyUniqueDuals.contrIndexList_left hl)]
|
||||
intro I hI1 hI2
|
||||
simp only [countSelf_append, ne_eq] at hI1
|
||||
have hc := countSelf_contrIndexList_le_countSelf l I
|
||||
have h2 := (countId_eq_two_ofcontrIndexList_left_of_OnlyUniqueDuals l l2 hl I hI2 )
|
||||
have hI1' := h1 I (by simp_all; omega) h2
|
||||
have hIdEq : l.contrIndexList.countId I = l.countId I := by
|
||||
simp at h2 hI2
|
||||
omega
|
||||
simp
|
||||
rw [l.countColorQuot_contrIndexList_eq_of_countId_eq hIdEq,
|
||||
l.countSelf_contrIndexList_eq_of_countId_eq hIdEq,
|
||||
l.countDual_contrIndexList_eq_of_countId_eq hIdEq, hIdEq]
|
||||
simpa using hI1'
|
||||
|
||||
/-!
|
||||
|
||||
## Bool
|
||||
|
||||
-/
|
||||
/-- A bool which is true for an index list if and only if every index and its' dual, when it exists,
|
||||
have dual colors. -/
|
||||
def bool (l : IndexList 𝓒.Color) : Bool :=
|
||||
if (∀ (i : l.withDual), 𝓒.τ
|
||||
(l.colorMap ((l.getDual? i).get (l.withDual_isSome i))) = l.colorMap i) then
|
||||
true
|
||||
else false
|
||||
|
||||
lemma iff_bool : l.ColorCond ↔ bool l := by
|
||||
rw [iff_withDual, bool]
|
||||
simp only [Subtype.forall, mem_withDual_iff_isSome, Bool.if_false_right, Bool.and_true,
|
||||
decide_eq_true_eq]
|
||||
|
||||
end ColorCond
|
||||
|
||||
end IndexList
|
||||
|
||||
end IndexNotation
|
|
@ -0,0 +1,341 @@
|
|||
/-
|
||||
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.IndexList.Equivs
|
||||
import Mathlib.Algebra.Order.Ring.Nat
|
||||
import Mathlib.Data.Finset.Sort
|
||||
import Mathlib.Tactic.FinCases
|
||||
/-!
|
||||
|
||||
# Contraction of an index list.
|
||||
|
||||
In this file we define the contraction of an index list `l` to be the index list formed by
|
||||
by the subset of indices of `l` which do not have a dual in `l`.
|
||||
|
||||
For example, the contraction of the index list `['ᵘ¹', 'ᵘ²', 'ᵤ₁', 'ᵘ¹']` is the index list
|
||||
`['ᵘ²']`.
|
||||
|
||||
We also define the following finite sets
|
||||
- `l.withoutDual` the finite set of indices of `l` which do not have a dual in `l`.
|
||||
- `l.withUniqueDualLT` the finite set of those indices of `l` which have a unique dual, and
|
||||
for which that dual is greater-then (determined by the ordering on `Fin`) then the index itself.
|
||||
- `l.withUniqueDualGT` the finite set of those indices of `l` which have a unique dual, and
|
||||
for which that dual is less-then (determined by the ordering on `Fin`) then the index itself.
|
||||
|
||||
We define an equivalence `l.withUniqueDualLT ⊕ l.withUniqueDualLT ≃ l.withUniqueDual`.
|
||||
|
||||
We prove properties related to when `l.withUniqueDualInOther l2 = Finset.univ` for two
|
||||
index lists `l` and `l2`.
|
||||
|
||||
-/
|
||||
|
||||
namespace IndexNotation
|
||||
|
||||
namespace IndexList
|
||||
|
||||
variable {X : Type} [IndexNotation X] [Fintype X] [DecidableEq X]
|
||||
variable (l l2 l3 : IndexList X)
|
||||
|
||||
/-- The index list formed from `l` by selecting only those indices in `l` which
|
||||
do not have a dual. -/
|
||||
def contrIndexList : IndexList X where
|
||||
val := l.val.filter (fun I => l.countId I = 1)
|
||||
|
||||
@[simp]
|
||||
lemma contrIndexList_empty : (⟨[]⟩ : IndexList X).contrIndexList = { val := [] } := by
|
||||
apply ext
|
||||
simp [contrIndexList]
|
||||
|
||||
lemma contrIndexList_val : l.contrIndexList.val =
|
||||
l.val.filter (fun I => l.countId I = 1) := by
|
||||
rfl
|
||||
|
||||
/-- An alternative form of the contracted index list. -/
|
||||
@[simp]
|
||||
def contrIndexList' : IndexList X where
|
||||
val := List.ofFn (l.val.get ∘ Subtype.val ∘ l.withoutDualEquiv)
|
||||
|
||||
lemma withoutDual_sort_eq_filter : l.withoutDual.sort (fun i j => i ≤ j) =
|
||||
(List.finRange l.length).filter (fun i => i ∈ l.withoutDual) := by
|
||||
rw [withoutDual]
|
||||
rw [← filter_sort_comm]
|
||||
simp only [Option.isNone_iff_eq_none, Finset.mem_filter, Finset.mem_univ, true_and]
|
||||
apply congrArg
|
||||
exact Fin.sort_univ l.length
|
||||
|
||||
lemma contrIndexList_eq_contrIndexList' : l.contrIndexList = l.contrIndexList' := by
|
||||
apply ext
|
||||
simp only [contrIndexList']
|
||||
trans List.map l.val.get (List.ofFn (Subtype.val ∘ ⇑l.withoutDualEquiv))
|
||||
· rw [list_ofFn_withoutDualEquiv_eq_sort, withoutDual_sort_eq_filter]
|
||||
simp only [contrIndexList]
|
||||
let f1 : Index X → Bool := fun (I : Index X) => l.countId I = 1
|
||||
let f2 : (Fin l.length) → Bool := fun i => i ∈ l.withoutDual
|
||||
change List.filter f1 l.val = List.map l.val.get (List.filter f2 (List.finRange l.length))
|
||||
have hf : f2 = f1 ∘ l.val.get := by
|
||||
funext i
|
||||
simp only [mem_withoutDual_iff_countId_eq_one l, List.get_eq_getElem, Function.comp_apply, f2,
|
||||
f1]
|
||||
rw [hf, ← List.filter_map]
|
||||
apply congrArg
|
||||
simp [length]
|
||||
· simp only [List.map_ofFn]
|
||||
|
||||
@[simp]
|
||||
lemma contrIndexList_length : l.contrIndexList.length = l.withoutDual.card := by
|
||||
simp [contrIndexList_eq_contrIndexList', withoutDual, length]
|
||||
|
||||
@[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_eq_contrIndexList', idMap]
|
||||
rfl
|
||||
|
||||
@[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_eq_contrIndexList', colorMap]
|
||||
rfl
|
||||
|
||||
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 [← Option.not_isSome_iff_eq_none, ← mem_withDual_iff_isSome, mem_withDual_iff_exists]
|
||||
simp only [contrIndexList_areDualInSelf, not_exists]
|
||||
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'
|
||||
exact fun x => forall_not_of_not_exists h1'
|
||||
↑(l.withoutDualEquiv (Fin.cast (contrIndexList_length l) x))
|
||||
|
||||
@[simp]
|
||||
lemma contrIndexList_withDual : l.contrIndexList.withDual = ∅ := by
|
||||
rw [Finset.eq_empty_iff_forall_not_mem]
|
||||
intro x
|
||||
simp [withDual]
|
||||
|
||||
@[simp]
|
||||
lemma contrIndexList_withUniqueDual : l.contrIndexList.withUniqueDual = ∅ := by
|
||||
rw [withUniqueDual]
|
||||
simp
|
||||
|
||||
@[simp]
|
||||
lemma contrIndexList_areDualInSelf_false (i j : Fin l.contrIndexList.length) :
|
||||
l.contrIndexList.AreDualInSelf i j ↔ False := by
|
||||
refine Iff.intro (fun h => ?_) (fun h => False.elim h)
|
||||
have h1 : i ∈ l.contrIndexList.withDual := by
|
||||
rw [@mem_withDual_iff_exists]
|
||||
exact Exists.intro j h
|
||||
simp_all
|
||||
|
||||
@[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 only [Finset.card_univ, Fintype.card_fin, List.get_eq_getElem, true_and]
|
||||
intro n h1' h2
|
||||
simp [contrIndexList_eq_contrIndexList']
|
||||
congr
|
||||
simp [withoutDualEquiv]
|
||||
simp [h1]
|
||||
rw [(Finset.orderEmbOfFin_unique' _
|
||||
(fun x => Finset.mem_univ ((Fin.castOrderIso _).toOrderEmbedding x))).symm]
|
||||
· exact Eq.symm (Nat.add_zero n)
|
||||
· rw [h1]
|
||||
exact Finset.card_fin l.length
|
||||
|
||||
lemma contrIndexList_contrIndexList : l.contrIndexList.contrIndexList = l.contrIndexList := by
|
||||
simp
|
||||
|
||||
@[simp]
|
||||
lemma contrIndexList_getDualInOther?_self (i : Fin l.contrIndexList.length) :
|
||||
l.contrIndexList.getDualInOther? l.contrIndexList i = some i := by
|
||||
simp [getDualInOther?]
|
||||
rw [@Fin.find_eq_some_iff]
|
||||
simp [AreDualInOther]
|
||||
intro j hj
|
||||
have h1 : i = j := by
|
||||
by_contra hn
|
||||
have h : l.contrIndexList.AreDualInSelf i j := by
|
||||
simp only [AreDualInSelf]
|
||||
simp [hj]
|
||||
exact hn
|
||||
exact (contrIndexList_areDualInSelf_false l i j).mp h
|
||||
exact Fin.ge_of_eq (id (Eq.symm h1))
|
||||
|
||||
lemma cons_contrIndexList_of_countId_eq_zero {I : Index X}
|
||||
(h : l.countId I = 0) :
|
||||
(l.cons I).contrIndexList = l.contrIndexList.cons I := by
|
||||
apply ext
|
||||
simp [contrIndexList, countId]
|
||||
rw [List.filter_cons_of_pos]
|
||||
· apply congrArg
|
||||
apply List.filter_congr
|
||||
intro J hJ
|
||||
simp only [decide_eq_decide]
|
||||
rw [countId, List.countP_eq_zero] at h
|
||||
simp only [decide_eq_true_eq] at h
|
||||
rw [List.countP_cons_of_neg]
|
||||
simp only [decide_eq_true_eq]
|
||||
exact fun a => h J hJ (id (Eq.symm a))
|
||||
· simp only [decide_True, List.countP_cons_of_pos, add_left_eq_self, decide_eq_true_eq]
|
||||
exact h
|
||||
|
||||
lemma cons_contrIndexList_of_countId_neq_zero {I : Index X} (h : l.countId I ≠ 0) :
|
||||
(l.cons I).contrIndexList.val = l.contrIndexList.val.filter (fun J => ¬ I.id = J.id) := by
|
||||
simp only [contrIndexList, countId, cons_val, decide_not, List.filter_filter, Bool.not_eq_true',
|
||||
decide_eq_false_iff_not, decide_eq_true_eq, Bool.decide_and]
|
||||
rw [List.filter_cons_of_neg]
|
||||
· apply List.filter_congr
|
||||
intro J hJ
|
||||
by_cases hI : I.id = J.id
|
||||
· simp only [hI, decide_True, List.countP_cons_of_pos, add_left_eq_self, Bool.not_true,
|
||||
Bool.false_and, decide_eq_false_iff_not, countId]
|
||||
rw [countId, hI] at h
|
||||
simp only [h, not_false_eq_true]
|
||||
· simp only [hI, decide_False, Bool.not_false, Bool.true_and, decide_eq_decide]
|
||||
rw [List.countP_cons_of_neg]
|
||||
simp only [decide_eq_true_eq]
|
||||
exact fun a => hI (id (Eq.symm a))
|
||||
· simp only [decide_True, List.countP_cons_of_pos, add_left_eq_self, decide_eq_true_eq]
|
||||
exact h
|
||||
|
||||
lemma mem_contrIndexList_countId {I : Index X} (h : I ∈ l.contrIndexList.val) :
|
||||
l.countId I = 1 := by
|
||||
simp only [contrIndexList, List.mem_filter, decide_eq_true_eq] at h
|
||||
exact h.2
|
||||
|
||||
lemma mem_contrIndexList_filter {I : Index X} (h : I ∈ l.contrIndexList.val) :
|
||||
l.val.filter (fun J => I.id = J.id) = [I] := by
|
||||
have h1 : (l.val.filter (fun J => I.id = J.id)).length = 1 := by
|
||||
rw [← List.countP_eq_length_filter]
|
||||
exact l.mem_contrIndexList_countId h
|
||||
have h2 : I ∈ l.val.filter (fun J => I.id = J.id) := by
|
||||
simp only [List.mem_filter, decide_True, and_true]
|
||||
simp only [contrIndexList, List.mem_filter, decide_eq_true_eq] at h
|
||||
exact h.1
|
||||
rw [List.length_eq_one] at h1
|
||||
obtain ⟨J, hJ⟩ := h1
|
||||
rw [hJ] at h2
|
||||
simp at h2
|
||||
subst h2
|
||||
exact hJ
|
||||
|
||||
lemma mem_contrIndexList_countId_contrIndexList {I : Index X} (h : I ∈ l.contrIndexList.val) :
|
||||
l.contrIndexList.countId I = 1 := by
|
||||
trans (l.val.filter (fun J => I.id = J.id)).countP
|
||||
(fun i => l.val.countP (fun j => i.id = j.id) = 1)
|
||||
· rw [contrIndexList]
|
||||
simp only [countId]
|
||||
rw [List.countP_filter, List.countP_filter]
|
||||
simp only [decide_eq_true_eq, Bool.decide_and]
|
||||
congr
|
||||
funext a
|
||||
rw [Bool.and_comm]
|
||||
· rw [l.mem_contrIndexList_filter h, List.countP_cons_of_pos]
|
||||
· rfl
|
||||
· simp only [decide_eq_true_eq]
|
||||
exact mem_contrIndexList_countId l h
|
||||
|
||||
lemma countId_contrIndexList_zero_of_countId (I : Index X)
|
||||
(h : l.countId I = 0) : l.contrIndexList.countId I = 0 := by
|
||||
trans (l.val.filter (fun J => I.id = J.id)).countP
|
||||
(fun i => l.val.countP (fun j => i.id = j.id) = 1)
|
||||
· rw [contrIndexList]
|
||||
simp only [countId]
|
||||
rw [List.countP_filter, List.countP_filter]
|
||||
simp only [decide_eq_true_eq, Bool.decide_and]
|
||||
congr
|
||||
funext a
|
||||
rw [Bool.and_comm]
|
||||
· rw [countId_eq_length_filter, List.length_eq_zero] at h
|
||||
rw [h]
|
||||
simp only [List.countP_nil]
|
||||
|
||||
lemma countId_contrIndexList_le_one (I : Index X) :
|
||||
l.contrIndexList.countId I ≤ 1 := by
|
||||
by_cases h : l.contrIndexList.countId I = 0
|
||||
· simp [h]
|
||||
· obtain ⟨I', hI1, hI2⟩ := countId_neq_zero_mem l.contrIndexList I h
|
||||
rw [countId_congr l.contrIndexList hI2, mem_contrIndexList_countId_contrIndexList l hI1]
|
||||
|
||||
lemma countId_contrIndexList_eq_one_iff_countId_eq_one (I : Index X) :
|
||||
l.contrIndexList.countId I = 1 ↔ l.countId I = 1 := by
|
||||
refine Iff.intro (fun h => ?_) (fun h => ?_)
|
||||
· obtain ⟨I', hI1, hI2⟩ := countId_neq_zero_mem l.contrIndexList I (by omega)
|
||||
simp [contrIndexList] at hI1
|
||||
rw [countId_congr l hI2]
|
||||
exact hI1.2
|
||||
· obtain ⟨I', hI1, hI2⟩ := countId_neq_zero_mem l I (by omega)
|
||||
rw [countId_congr l hI2] at h
|
||||
rw [countId_congr _ hI2]
|
||||
refine mem_contrIndexList_countId_contrIndexList l ?_
|
||||
simp [contrIndexList]
|
||||
exact ⟨hI1, h⟩
|
||||
|
||||
lemma countId_contrIndexList_le_countId (I : Index X) :
|
||||
l.contrIndexList.countId I ≤ l.countId I := by
|
||||
by_cases h : l.contrIndexList.countId I = 0
|
||||
· exact StrictMono.minimal_preimage_bot (fun ⦃a b⦄ a => a) h (l.countId I)
|
||||
· have ho : l.contrIndexList.countId I = 1 := by
|
||||
have h1 := l.countId_contrIndexList_le_one I
|
||||
omega
|
||||
rw [ho]
|
||||
rw [countId_contrIndexList_eq_one_iff_countId_eq_one] at ho
|
||||
rw [ho]
|
||||
|
||||
@[simp]
|
||||
lemma countId_contrIndexList_get (i : Fin l.contrIndexList.length) :
|
||||
l.contrIndexList.countId l.contrIndexList.val[Fin.val i] = 1 := by
|
||||
refine mem_contrIndexList_countId_contrIndexList l ?_
|
||||
exact List.getElem_mem l.contrIndexList.val (↑i) _
|
||||
|
||||
lemma filter_id_contrIndexList_eq_of_countId_contrIndexList (I : Index X)
|
||||
(h : l.contrIndexList.countId I = l.countId I) :
|
||||
l.contrIndexList.val.filter (fun J => I.id = J.id) =
|
||||
l.val.filter (fun J => I.id = J.id) := by
|
||||
apply List.Sublist.eq_of_length
|
||||
· rw [contrIndexList, List.filter_comm]
|
||||
exact List.filter_sublist (List.filter (fun J => decide (I.id = J.id)) l.val)
|
||||
· rw [← countId_eq_length_filter, h, countId_eq_length_filter]
|
||||
|
||||
lemma contrIndexList_append_eq_filter : (l ++ l2).contrIndexList.val =
|
||||
l.contrIndexList.val.filter (fun I => l2.countId I = 0) ++
|
||||
l2.contrIndexList.val.filter (fun I => l.countId I = 0) := by
|
||||
simp [contrIndexList]
|
||||
congr 1
|
||||
· apply List.filter_congr
|
||||
intro I hI
|
||||
have hIz : l.countId I ≠ 0 := countId_mem l I hI
|
||||
have hx : l.countId I + l2.countId I = 1 ↔ (l2.countId I = 0 ∧ l.countId I = 1) := by
|
||||
omega
|
||||
simp only [hx, Bool.decide_and]
|
||||
· apply List.filter_congr
|
||||
intro I hI
|
||||
have hIz : l2.countId I ≠ 0 := countId_mem l2 I hI
|
||||
have hx : l.countId I + l2.countId I = 1 ↔ (l2.countId I = 1 ∧ l.countId I = 0) := by
|
||||
omega
|
||||
simp only [hx, Bool.decide_and]
|
||||
exact Bool.and_comm (decide (l2.countId I = 1)) (decide (l.countId I = 0))
|
||||
|
||||
end IndexList
|
||||
|
||||
end IndexNotation
|
|
@ -168,10 +168,6 @@ lemma countId_get (i : Fin l.length) : l.countId (l.val.get i) =
|
|||
simp
|
||||
simp [List.countP, List.countP.go, AreDualInOther]
|
||||
|
||||
|
||||
|
||||
|
||||
|
||||
/-!
|
||||
|
||||
## Duals and countId
|
||||
|
@ -323,16 +319,220 @@ lemma mem_withUniqueDual_iff_countId_eq_two (i : Fin l.length) :
|
|||
Iff.intro (fun h => l.countId_eq_two_of_mem_withUniqueDual i h)
|
||||
(fun h => l.mem_withUniqueDual_of_countId_eq_two i h)
|
||||
|
||||
lemma countId_eq_one_of_mem_withUniqueDualInOther (i : Fin l.length)
|
||||
(h : i ∈ l.withUniqueDualInOther l2) :
|
||||
l.countId (l.val.get i) = 1 ∧ l2.countId (l.val.get i) = 1 := by
|
||||
let i' := (l.getDualInOther? l2 i).get (mem_withUniqueDualInOther_isSome l l2 i h)
|
||||
have h1 : [i'] = (List.finRange l2.length).filter (fun j => (l.AreDualInOther l2 i j)) := by
|
||||
trans List.filter (fun j => (l.AreDualInOther l2 i j)) [i']
|
||||
· simp [List.filter, i']
|
||||
trans List.filter (fun j => (l.AreDualInOther l2 i j))
|
||||
((List.finRange l2.length).filter (fun j => j = i'))
|
||||
· apply congrArg
|
||||
rw [← filter_finRange l2 i']
|
||||
apply List.filter_congr (fun j _ => ?_)
|
||||
simpa using eq_comm
|
||||
trans List.filter (fun j => j = i')
|
||||
((List.finRange l2.length).filter (fun j => (l.AreDualInOther l2 i j)))
|
||||
· simp
|
||||
apply List.filter_congr (fun j _ => ?_)
|
||||
exact Bool.and_comm (decide (l.AreDualInOther l2 i j)) (decide (j = i'))
|
||||
· simp
|
||||
refine List.filter_congr (fun j _ => ?_)
|
||||
simp [withUniqueDualInOther] at h
|
||||
simp
|
||||
intro hj
|
||||
have hj' := h.2.2 j hj
|
||||
apply Option.some_injective
|
||||
rw [hj']
|
||||
simp [i']
|
||||
apply And.intro
|
||||
· simp only [withUniqueDualInOther, Finset.mem_filter, Finset.mem_univ, true_and] at h
|
||||
rw [mem_withDual_iff_countId_gt_one] at h
|
||||
have h2 := countId_index_neq_zero l i
|
||||
omega
|
||||
· rw [countId_get_other, List.countP_eq_length_filter, ← h1]
|
||||
simp
|
||||
|
||||
lemma mem_withUniqueDualInOther_of_countId_eq_one (i : Fin l.length)
|
||||
(h : l.countId (l.val.get i) = 1 ∧ l2.countId (l.val.get i) = 1) :
|
||||
i ∈ l.withUniqueDualInOther l2 := by
|
||||
have hw : i ∈ l.withDualInOther l2 := by
|
||||
rw [mem_withDualInOther_iff_countId_neq_zero, h.2]
|
||||
exact Nat.one_ne_zero
|
||||
simp only [withUniqueDualInOther, Finset.mem_filter, Finset.mem_univ, true_and]
|
||||
apply And.intro
|
||||
· rw [mem_withDual_iff_countId_gt_one]
|
||||
omega
|
||||
· apply And.intro
|
||||
· rw [mem_withDualInOther_iff_countId_neq_zero]
|
||||
omega
|
||||
· intro j hj
|
||||
have h2 := h.2
|
||||
rw [countId_get_other l l2, List.countP_eq_length_filter, List.length_eq_one] at h2
|
||||
obtain ⟨a, ha⟩ := h2
|
||||
have hj : j ∈ List.filter (fun j => decide (l.AreDualInOther l2 i j))
|
||||
(List.finRange l2.length) := by
|
||||
simpa using hj
|
||||
rw [ha] at hj
|
||||
simp at hj
|
||||
subst hj
|
||||
have ht : (l.getDualInOther? l2 i).get ((mem_withInDualOther_iff_isSome l l2 i).mp hw) ∈
|
||||
(List.finRange l2.length).filter (fun j => decide (l.AreDualInOther l2 i j)) := by
|
||||
simp
|
||||
rw [ha] at ht
|
||||
simp at ht
|
||||
subst ht
|
||||
simp
|
||||
|
||||
lemma mem_withUniqueDualInOther_iff_countId_eq_one (i : Fin l.length) :
|
||||
i ∈ l.withUniqueDualInOther l2 ↔ l.countId (l.val.get i) = 1 ∧ l2.countId (l.val.get i) = 1 := by
|
||||
simp only [withUniqueDualInOther, Finset.mem_filter, Finset.mem_univ, true_and,
|
||||
List.get_eq_getElem]
|
||||
rw [mem_withDual_iff_countId_gt_one]
|
||||
rw [mem_withDualInOther_iff_countId_neq_zero]
|
||||
have hneq : l.countId (l.val.get i) ≠ 0 := by exact countId_index_neq_zero l i
|
||||
i ∈ l.withUniqueDualInOther l2 ↔ l.countId (l.val.get i) = 1 ∧ l2.countId (l.val.get i) = 1 :=
|
||||
Iff.intro (fun h => l.countId_eq_one_of_mem_withUniqueDualInOther l2 i h)
|
||||
(fun h => l.mem_withUniqueDualInOther_of_countId_eq_one l2 i h)
|
||||
|
||||
/-!
|
||||
|
||||
## getDual? and countId
|
||||
|
||||
-/
|
||||
|
||||
@[simp]
|
||||
lemma getDual?_countId (i : Fin l.length) (h : (l.getDual? i).isSome) :
|
||||
l.countId (l.val[Fin.val ((l.getDual? i).get h)]) = l.countId (l.val.get i) := by
|
||||
apply countId_congr
|
||||
change l.idMap ((l.getDual? i).get h) = _
|
||||
simp
|
||||
rfl
|
||||
|
||||
@[simp]
|
||||
lemma getDualInOther?_countId_right (i : Fin l.length) (h : (l.getDualInOther? l2 i).isSome) :
|
||||
l2.countId (l2.val[Fin.val ((l.getDualInOther? l2 i).get h)]) = l2.countId (l.val.get i) := by
|
||||
apply countId_congr
|
||||
change l2.idMap ((l.getDualInOther? l2 i).get h) = _
|
||||
simp
|
||||
rfl
|
||||
|
||||
@[simp]
|
||||
lemma getDualInOther?_countId_left (i : Fin l.length) (h : (l.getDualInOther? l2 i).isSome) :
|
||||
l.countId (l2.val[Fin.val ((l.getDualInOther? l2 i).get h)]) = l.countId (l.val.get i) := by
|
||||
apply countId_congr
|
||||
change l2.idMap ((l.getDualInOther? l2 i).get h) = _
|
||||
simp
|
||||
rfl
|
||||
|
||||
lemma getDual?_isSome_of_countId_eq_two {i : Fin l.length}
|
||||
(h : l.countId (l.val.get i) = 2) : (l.getDual? i).isSome := by
|
||||
rw [← l.mem_withUniqueDual_iff_countId_eq_two] at h
|
||||
exact mem_withUniqueDual_isSome l i h
|
||||
|
||||
/-!
|
||||
|
||||
## Filters
|
||||
|
||||
-/
|
||||
|
||||
|
||||
lemma filter_id_of_countId_eq_zero {i : Fin l.length} (h1 : l.countId (l.val.get i) = 0) :
|
||||
l.val.filter (fun J => (l.val.get i).id = J.id) = [] := by
|
||||
rw [countId_eq_length_filter, List.length_eq_zero] at h1
|
||||
exact h1
|
||||
|
||||
lemma filter_id_of_countId_eq_zero' {I : Index X} (h1 : l.countId I = 0) :
|
||||
l.val.filter (fun J => I.id = J.id) = [] := by
|
||||
rw [countId_eq_length_filter, List.length_eq_zero] at h1
|
||||
exact h1
|
||||
|
||||
lemma filter_id_of_countId_eq_one {i : Fin l.length} (h1 : l.countId (l.val.get i) = 1) :
|
||||
l.val.filter (fun J => (l.val.get i).id = J.id) = [l.val.get i] := by
|
||||
rw [countId_eq_length_filter, List.length_eq_one] at h1
|
||||
obtain ⟨J, hJ⟩ := h1
|
||||
have hme : (l.val.get i) ∈ List.filter (fun J => decide ((l.val.get i).id = J.id)) l.val := by
|
||||
simp only [List.get_eq_getElem, List.mem_filter, decide_True, and_true]
|
||||
exact List.getElem_mem l.val (↑i) i.isLt
|
||||
rw [hJ] at hme
|
||||
simp only [List.get_eq_getElem, List.mem_singleton] at hme
|
||||
erw [hJ]
|
||||
simp only [List.get_eq_getElem, List.cons.injEq, and_true]
|
||||
exact id (Eq.symm hme)
|
||||
|
||||
lemma filter_id_of_countId_eq_one_mem {I : Index X} (hI : I ∈ l.val) (h : l.countId I = 1) :
|
||||
l.val.filter (fun J => I.id = J.id) = [I] := by
|
||||
rw [countId_eq_length_filter, List.length_eq_one] at h
|
||||
obtain ⟨J, hJ⟩ := h
|
||||
have hme : I ∈ List.filter (fun J => decide (I.id = J.id)) l.val := by
|
||||
simp only [List.mem_filter, decide_True, and_true]
|
||||
exact hI
|
||||
rw [hJ] at hme
|
||||
simp only [List.mem_singleton] at hme
|
||||
erw [hJ]
|
||||
simp only [List.cons.injEq, and_true]
|
||||
exact id (Eq.symm hme)
|
||||
|
||||
lemma filter_id_of_countId_eq_two {i : Fin l.length}
|
||||
(h : l.countId (l.val.get i) = 2) :
|
||||
l.val.filter (fun J => (l.val.get i).id = J.id) =
|
||||
[l.val.get i, l.val.get ((l.getDual? i).get (l.getDual?_isSome_of_countId_eq_two h))] ∨
|
||||
l.val.filter (fun J => (l.val.get i).id = J.id) =
|
||||
[l.val.get ((l.getDual? i).get (l.getDual?_isSome_of_countId_eq_two h)), l.val.get i] := by
|
||||
have hc := h
|
||||
rw [countId_eq_length_filter] at hc
|
||||
by_cases hi : i < ((l.getDual? i).get (l.getDual?_isSome_of_countId_eq_two h))
|
||||
· refine Or.inl (List.Sublist.eq_of_length ?_ ?_).symm
|
||||
· have h1 : [l.val.get i, l.val.get
|
||||
((l.getDual? i).get (l.getDual?_isSome_of_countId_eq_two h))].filter
|
||||
(fun J => (l.val.get i).id = J.id) = [l.val.get i, l.val.get
|
||||
((l.getDual? i).get (l.getDual?_isSome_of_countId_eq_two h))] := by
|
||||
simp only [List.get_eq_getElem, decide_True, List.filter_cons_of_pos, List.cons.injEq,
|
||||
List.filter_eq_self, List.mem_singleton, decide_eq_true_eq, forall_eq, true_and]
|
||||
change l.idMap i = l.idMap ((l.getDual? i).get (l.getDual?_isSome_of_countId_eq_two h))
|
||||
simp
|
||||
rw [← h1]
|
||||
refine List.Sublist.filter (fun (J : Index X) => ((l.val.get i).id = J.id)) ?_
|
||||
rw [List.sublist_iff_exists_fin_orderEmbedding_get_eq]
|
||||
refine ⟨⟨⟨![i, (l.getDual? i).get (l.getDual?_isSome_of_countId_eq_two h)],
|
||||
List.nodup_ofFn.mp ?_⟩, ?_⟩, ?_⟩
|
||||
· simpa using Fin.ne_of_lt hi
|
||||
· intro a b
|
||||
fin_cases a, b
|
||||
<;> simp [hi]
|
||||
exact Fin.le_of_lt hi
|
||||
· intro a
|
||||
fin_cases a <;> rfl
|
||||
· rw [hc]
|
||||
simp
|
||||
· have hi' : ((l.getDual? i).get (l.getDual?_isSome_of_countId_eq_two h)) < i := by
|
||||
have h1 := l.getDual?_get_areDualInSelf i (getDual?_isSome_of_countId_eq_two l h)
|
||||
simp only [AreDualInSelf] at h1
|
||||
have h2 := h1.1
|
||||
omega
|
||||
refine Or.inr (List.Sublist.eq_of_length ?_ ?_).symm
|
||||
· have h1 : [l.val.get ((l.getDual? i).get (l.getDual?_isSome_of_countId_eq_two h)),
|
||||
l.val.get i].filter (fun J => (l.val.get i).id = J.id) = [l.val.get
|
||||
((l.getDual? i).get (l.getDual?_isSome_of_countId_eq_two h)), l.val.get i,] := by
|
||||
simp only [List.get_eq_getElem, List.filter_eq_self, List.mem_cons, List.mem_singleton,
|
||||
decide_eq_true_eq, forall_eq_or_imp, forall_eq, and_true]
|
||||
apply And.intro
|
||||
· change l.idMap i = l.idMap ((l.getDual? i).get (l.getDual?_isSome_of_countId_eq_two h))
|
||||
simp
|
||||
· simp only [List.not_mem_nil, false_implies, implies_true, and_self]
|
||||
rw [← h1]
|
||||
refine List.Sublist.filter (fun (J : Index X) => ((l.val.get i).id = J.id)) ?_
|
||||
rw [List.sublist_iff_exists_fin_orderEmbedding_get_eq]
|
||||
refine ⟨⟨⟨![(l.getDual? i).get (l.getDual?_isSome_of_countId_eq_two h), i], ?_⟩, ?_⟩, ?_⟩
|
||||
· refine List.nodup_ofFn.mp ?_
|
||||
simp only [List.get_eq_getElem, List.length_singleton, Nat.succ_eq_add_one, Nat.reduceAdd,
|
||||
List.length_nil, List.ofFn_succ, Fin.isValue, Matrix.cons_val_zero, Matrix.cons_val_succ,
|
||||
Matrix.cons_val_fin_one, List.ofFn_zero, List.nodup_cons, List.mem_singleton,
|
||||
List.not_mem_nil, not_false_eq_true, List.nodup_nil, and_self, and_true]
|
||||
exact Fin.ne_of_lt hi'
|
||||
· intro a b
|
||||
fin_cases a, b
|
||||
<;> simp [hi']
|
||||
exact Fin.le_of_lt hi'
|
||||
· intro a
|
||||
fin_cases a <;> rfl
|
||||
· rw [hc]
|
||||
simp
|
||||
|
||||
end IndexList
|
||||
|
||||
|
|
|
@ -74,6 +74,23 @@ def withUniqueDual : Finset (Fin l.length) :=
|
|||
Finset.filter (fun i => i ∈ l.withDual ∧
|
||||
∀ j, l.AreDualInSelf i j → j = l.getDual? i) Finset.univ
|
||||
|
||||
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
|
||||
|
||||
/-- The finite set of those indices of `l` which have a unique dual, and for which
|
||||
that dual is greater-then (determined by the ordering on `Fin`) then the index itself. -/
|
||||
def withUniqueDualLT : Finset (Fin l.length) :=
|
||||
Finset.filter (fun i => i < l.getDual? i) l.withUniqueDual
|
||||
|
||||
/-- The finite set of those indices of `l` which have a unique dual, and for which
|
||||
that dual is less-then (determined by the ordering on `Fin`) then the index itself. -/
|
||||
def withUniqueDualGT : Finset (Fin l.length) :=
|
||||
Finset.filter (fun i => ¬ i < l.getDual? i) l.withUniqueDual
|
||||
|
||||
/-- The finite set of indices of an index list which have a unique dual in another, provided, index
|
||||
list. -/
|
||||
def withUniqueDualInOther : Finset (Fin l.length) :=
|
||||
|
@ -160,6 +177,33 @@ lemma not_mem_withDual_of_mem_withoutDual (i : Fin l.length) (h : i ∈ l.withou
|
|||
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]
|
||||
|
||||
lemma mem_withUniqueDual_of_mem_withUniqueDualLt (i : Fin l.length) (h : i ∈ l.withUniqueDualLT) :
|
||||
i ∈ l.withUniqueDual := by
|
||||
simp only [withUniqueDualLT, Finset.mem_filter, Finset.mem_univ, true_and] at h
|
||||
exact h.1
|
||||
|
||||
lemma mem_withUniqueDual_of_mem_withUniqueDualGt (i : Fin l.length) (h : i ∈ l.withUniqueDualGT) :
|
||||
i ∈ l.withUniqueDual := by
|
||||
simp only [withUniqueDualGT, Finset.mem_filter, Finset.mem_univ, true_and] at h
|
||||
exact h.1
|
||||
|
||||
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 _ _
|
||||
/-!
|
||||
## Are dual properties
|
||||
-/
|
||||
|
@ -238,6 +282,7 @@ lemma of_append_inr (i : Fin l.length) (j : Fin l3.length) :
|
|||
simp [AreDualInOther]
|
||||
|
||||
end AreDualInOther
|
||||
|
||||
/-!
|
||||
|
||||
## Basic properties of getDual?
|
||||
|
|
|
@ -0,0 +1,229 @@
|
|||
/-
|
||||
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.IndexList.CountId
|
||||
import Mathlib.Algebra.Order.Ring.Nat
|
||||
import Mathlib.Data.Finset.Sort
|
||||
import Mathlib.Tactic.FinCases
|
||||
/-!
|
||||
|
||||
# Equivalences between finsets.
|
||||
|
||||
-/
|
||||
|
||||
namespace IndexNotation
|
||||
|
||||
namespace IndexList
|
||||
|
||||
variable {X : Type} [IndexNotation X] [Fintype X] [DecidableEq X]
|
||||
variable (l l2 l3 : IndexList X)
|
||||
|
||||
/-!
|
||||
|
||||
## withoutDualEquiv
|
||||
-/
|
||||
|
||||
/-- An equivalence from `Fin l.withoutDual.card` to `l.withoutDual` determined by
|
||||
the order on `l.withoutDual` inherted from `Fin`. -/
|
||||
def withoutDualEquiv : Fin l.withoutDual.card ≃ l.withoutDual :=
|
||||
(Finset.orderIsoOfFin l.withoutDual (by rfl)).toEquiv
|
||||
|
||||
lemma list_ofFn_withoutDualEquiv_eq_sort :
|
||||
List.ofFn (Subtype.val ∘ l.withoutDualEquiv) = l.withoutDual.sort (fun i j => i ≤ j) := by
|
||||
rw [@List.ext_get_iff]
|
||||
refine And.intro ?_ (fun n h1 h2 => ?_)
|
||||
· simp only [List.length_ofFn, Finset.length_sort]
|
||||
· simp only [List.get_eq_getElem, List.getElem_ofFn, Function.comp_apply]
|
||||
rfl
|
||||
|
||||
/-- An equivalence splitting the indices of an index list `l` into those indices
|
||||
which have a dual in `l` and those which do not have a dual in `l`. -/
|
||||
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])
|
||||
|
||||
|
||||
/-!
|
||||
|
||||
## getDualEquiv
|
||||
|
||||
-/
|
||||
|
||||
/-- An equivalence from `l.withUniqueDual` to itself obtained by taking the dual index. -/
|
||||
def getDualEquiv : l.withUniqueDual ≃ l.withUniqueDual where
|
||||
toFun x := ⟨(l.getDual? x).get (l.mem_withUniqueDual_isSome x.1 x.2), by
|
||||
rw [mem_withUniqueDual_iff_countId_eq_two]
|
||||
simpa using l.countId_eq_two_of_mem_withUniqueDual x.1 x.2⟩
|
||||
invFun x := ⟨(l.getDual? x).get (l.mem_withUniqueDual_isSome x.1 x.2), by
|
||||
rw [mem_withUniqueDual_iff_countId_eq_two]
|
||||
simpa using l.countId_eq_two_of_mem_withUniqueDual x.1 x.2⟩
|
||||
left_inv x := SetCoe.ext <| by
|
||||
let d := (l.getDual? x).get (l.mem_withUniqueDual_isSome x.1 x.2)
|
||||
have h1 : l.countId (l.val.get d) = 2 := by
|
||||
simpa [d] using l.countId_eq_two_of_mem_withUniqueDual x.1 x.2
|
||||
have h2 : ((List.finRange l.length).filter (fun j => l.AreDualInSelf d j)).length = 1 := by
|
||||
rw [countId_get, List.countP_eq_length_filter] at h1
|
||||
omega
|
||||
obtain ⟨a, ha⟩ := List.length_eq_one.mp h2
|
||||
trans a
|
||||
· rw [← List.mem_singleton, ← ha]
|
||||
simp [d]
|
||||
· symm
|
||||
rw [← List.mem_singleton, ← ha]
|
||||
simp [d]
|
||||
right_inv x := SetCoe.ext <| by
|
||||
let d := (l.getDual? x).get (l.mem_withUniqueDual_isSome x.1 x.2)
|
||||
have h1 : l.countId (l.val.get d) = 2 := by
|
||||
simpa [d] using l.countId_eq_two_of_mem_withUniqueDual x.1 x.2
|
||||
have h2 : ((List.finRange l.length).filter (fun j => l.AreDualInSelf d j)).length = 1 := by
|
||||
rw [countId_get, List.countP_eq_length_filter] at h1
|
||||
omega
|
||||
obtain ⟨a, ha⟩ := List.length_eq_one.mp h2
|
||||
trans a
|
||||
· rw [← List.mem_singleton, ← ha]
|
||||
simp [d]
|
||||
· symm
|
||||
rw [← List.mem_singleton, ← ha]
|
||||
simp [d]
|
||||
|
||||
@[simp]
|
||||
lemma getDual?_getDualEquiv (i : l.withUniqueDual) : l.getDual? (l.getDualEquiv i) = i := by
|
||||
have h1 := (Equiv.apply_symm_apply l.getDualEquiv i).symm
|
||||
nth_rewrite 2 [h1]
|
||||
nth_rewrite 2 [getDualEquiv]
|
||||
simp
|
||||
rfl
|
||||
|
||||
def getDualInOtherEquiv : l.withUniqueDualInOther l2 ≃ l2.withUniqueDualInOther l where
|
||||
toFun x := ⟨(l.getDualInOther? l2 x).get (l.mem_withUniqueDualInOther_isSome l2 x.1 x.2), by
|
||||
rw [mem_withUniqueDualInOther_iff_countId_eq_one]
|
||||
simpa using And.comm.mpr (l.countId_eq_one_of_mem_withUniqueDualInOther l2 x.1 x.2)⟩
|
||||
invFun x := ⟨(l2.getDualInOther? l x).get (l2.mem_withUniqueDualInOther_isSome l x.1 x.2), by
|
||||
rw [mem_withUniqueDualInOther_iff_countId_eq_one]
|
||||
simpa using And.comm.mpr (l2.countId_eq_one_of_mem_withUniqueDualInOther l x.1 x.2)⟩
|
||||
left_inv x := SetCoe.ext <| by
|
||||
let d := (l.getDualInOther? l2 x).get (l.mem_withUniqueDualInOther_isSome l2 x.1 x.2)
|
||||
have h1 : l.countId (l2.val.get d) = 1 := by
|
||||
simpa [d] using (l.countId_eq_one_of_mem_withUniqueDualInOther l2 x.1 x.2).1
|
||||
have h2 : ((List.finRange l.length).filter (fun j => l2.AreDualInOther l d j)).length = 1 := by
|
||||
rw [countId_get_other, List.countP_eq_length_filter] at h1
|
||||
omega
|
||||
obtain ⟨a, ha⟩ := List.length_eq_one.mp h2
|
||||
trans a
|
||||
· rw [← List.mem_singleton, ← ha]
|
||||
simp [d]
|
||||
· symm
|
||||
rw [← List.mem_singleton, ← ha]
|
||||
simp [d]
|
||||
right_inv x := SetCoe.ext <| by
|
||||
let d := (l2.getDualInOther? l x).get (l2.mem_withUniqueDualInOther_isSome l x.1 x.2)
|
||||
have h1 : l2.countId (l.val.get d) = 1 := by
|
||||
simpa [d] using (l2.countId_eq_one_of_mem_withUniqueDualInOther l x.1 x.2).1
|
||||
have h2 : ((List.finRange l2.length).filter (fun j => l.AreDualInOther l2 d j)).length = 1 := by
|
||||
rw [countId_get_other, List.countP_eq_length_filter] at h1
|
||||
omega
|
||||
obtain ⟨a, ha⟩ := List.length_eq_one.mp h2
|
||||
trans a
|
||||
· rw [← List.mem_singleton, ← ha]
|
||||
simp [d]
|
||||
· symm
|
||||
rw [← List.mem_singleton, ← ha]
|
||||
simp [d]
|
||||
|
||||
@[simp]
|
||||
lemma getDualInOtherEquiv_self_refl : l.getDualInOtherEquiv l = Equiv.refl _ := by
|
||||
apply Equiv.ext
|
||||
intro x
|
||||
simp [getDualInOtherEquiv]
|
||||
apply Subtype.eq
|
||||
simp
|
||||
have hx2 := x.2
|
||||
simp [withUniqueDualInOther] at hx2
|
||||
apply Option.some_injective
|
||||
rw [hx2.2 x.1 (by simp [ AreDualInOther]) ]
|
||||
simp
|
||||
|
||||
|
||||
@[simp]
|
||||
lemma getDualInOtherEquiv_symm : (l.getDualInOtherEquiv l2).symm = l2.getDualInOtherEquiv l := by
|
||||
rfl
|
||||
|
||||
|
||||
/-!
|
||||
|
||||
## Equivalences involving `withUniqueDualLT` and `withUniqueDualGT`.
|
||||
|
||||
-/
|
||||
|
||||
/-! 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 =>
|
||||
exact fun a _ _ => a
|
||||
| none, some k =>
|
||||
exact fun _ _ a => a
|
||||
| some i, none =>
|
||||
exact fun h _ _ => 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 =>
|
||||
exact fun a b => a (a (a (b rfl)))
|
||||
| none, some k =>
|
||||
exact fun _ _ => trivial
|
||||
| some i, none =>
|
||||
exact fun a _ => a trivial
|
||||
| some i, some j =>
|
||||
intro h h'
|
||||
simp_all
|
||||
change ¬ j < i at h
|
||||
change i < j
|
||||
omega
|
||||
|
||||
/-- The equivalence between `l.withUniqueDualLT` and `l.withUniqueDualGT` obtained by
|
||||
taking an index to its dual. -/
|
||||
def withUniqueDualLTEquivGT : l.withUniqueDualLT ≃ l.withUniqueDualGT where
|
||||
toFun i := ⟨l.getDualEquiv ⟨i, l.mem_withUniqueDual_of_mem_withUniqueDualLt i i.2⟩, by
|
||||
have hi := i.2
|
||||
simp only [withUniqueDualGT, Finset.mem_filter, Finset.coe_mem, true_and]
|
||||
simp only [withUniqueDualLT, Finset.mem_filter] at hi
|
||||
apply option_not_lt
|
||||
· rw [getDual?_getDualEquiv]
|
||||
simpa only [getDualEquiv, Equiv.coe_fn_mk, Option.some_get] using hi.2
|
||||
· simp only [ne_eq, getDual?_neq_self, not_false_eq_true]⟩
|
||||
invFun i := ⟨l.getDualEquiv.symm ⟨i, l.mem_withUniqueDual_of_mem_withUniqueDualGt i i.2⟩, by
|
||||
have hi := i.2
|
||||
simp only [withUniqueDualLT, Finset.mem_filter, Finset.coe_mem, true_and, gt_iff_lt]
|
||||
simp only [withUniqueDualGT, Finset.mem_filter] at hi
|
||||
apply lt_option_of_not
|
||||
· erw [getDual?_getDualEquiv]
|
||||
simpa [getDualEquiv] using hi.2
|
||||
· symm
|
||||
simp⟩
|
||||
left_inv x := SetCoe.ext <| by
|
||||
simp
|
||||
right_inv x := SetCoe.ext <| by
|
||||
simp
|
||||
|
||||
/-- An equivalence from `l.withUniqueDualLT ⊕ l.withUniqueDualLT` to `l.withUniqueDual`.
|
||||
The first `l.withUniqueDualLT` are taken to themselves, whilst the second copy are
|
||||
taken to their dual. -/
|
||||
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 only [l.withUniqueDualLT_union_withUniqueDualGT,
|
||||
implies_true]))
|
||||
|
||||
end IndexList
|
||||
|
||||
end IndexNotation
|
|
@ -0,0 +1,176 @@
|
|||
/-
|
||||
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.IndexList.CountId
|
||||
import HepLean.SpaceTime.LorentzTensor.IndexNotation.IndexList.Contraction
|
||||
import Mathlib.Algebra.Order.Ring.Nat
|
||||
import Mathlib.Data.Finset.Sort
|
||||
/-!
|
||||
|
||||
# withDuals equal to withUniqueDuals
|
||||
|
||||
In a permissible list of indices every index which has a dual has a unique dual.
|
||||
This corresponds to the condition that `l.withDual = l.withUniqueDual`.
|
||||
|
||||
We prove lemmas relating to this condition.
|
||||
|
||||
-/
|
||||
|
||||
namespace IndexNotation
|
||||
|
||||
namespace IndexList
|
||||
|
||||
variable {X : Type} [IndexNotation X] [Fintype X] [DecidableEq X]
|
||||
variable (l l2 l3 : IndexList X)
|
||||
|
||||
def OnlyUniqueDuals : Prop := l.withUniqueDual = l.withDual
|
||||
|
||||
namespace OnlyUniqueDuals
|
||||
|
||||
variable {l l2 l3 : IndexList X}
|
||||
|
||||
lemma iff_unique_forall :
|
||||
l.OnlyUniqueDuals ↔
|
||||
∀ (i : l.withDual) j, l.AreDualInSelf i j → j = l.getDual? i := by
|
||||
refine Iff.intro (fun h i j hj => ?_) (fun h => ?_)
|
||||
· rw [OnlyUniqueDuals, @Finset.ext_iff] at h
|
||||
simp only [withUniqueDual, mem_withDual_iff_isSome, Finset.mem_filter, Finset.mem_univ,
|
||||
true_and, and_iff_left_iff_imp] at h
|
||||
refine h i ?_ j hj
|
||||
exact withDual_isSome l i
|
||||
· refine Finset.ext (fun i => ?_)
|
||||
refine Iff.intro (fun hi => mem_withDual_of_mem_withUniqueDual l i hi) (fun hi => ?_)
|
||||
· simp only [withUniqueDual, mem_withDual_iff_isSome, Finset.mem_filter, Finset.mem_univ,
|
||||
true_and]
|
||||
exact And.intro ((mem_withDual_iff_isSome l i).mp hi) (fun j hj => h ⟨i, hi⟩ j hj)
|
||||
|
||||
lemma iff_countId_leq_two :
|
||||
l.OnlyUniqueDuals ↔ ∀ i, l.countId (l.val.get i) ≤ 2 := by
|
||||
refine Iff.intro (fun h i => ?_) (fun h => ?_)
|
||||
· by_cases hi : i ∈ l.withDual
|
||||
· rw [← h] at hi
|
||||
rw [mem_withUniqueDual_iff_countId_eq_two] at hi
|
||||
rw [hi]
|
||||
· rw [mem_withDual_iff_countId_gt_one] at hi
|
||||
simp at hi
|
||||
exact Nat.le_succ_of_le hi
|
||||
· refine Finset.ext (fun i => ?_)
|
||||
rw [mem_withUniqueDual_iff_countId_eq_two, mem_withDual_iff_countId_gt_one]
|
||||
have hi := h i
|
||||
omega
|
||||
|
||||
|
||||
lemma iff_countId_leq_two' : l.OnlyUniqueDuals ↔ ∀ I, l.countId I ≤ 2 := by
|
||||
rw [iff_countId_leq_two]
|
||||
refine Iff.intro (fun h I => ?_) (fun h i => h (l.val.get i))
|
||||
by_cases hI : l.countId I = 0
|
||||
· rw [hI]
|
||||
exact Nat.zero_le 2
|
||||
· obtain ⟨I', hI1', hI2'⟩ := countId_neq_zero_mem l I hI
|
||||
rw [countId_congr _ hI2']
|
||||
have hi : List.indexOf I' l.val < l.length := List.indexOf_lt_length.mpr hI1'
|
||||
have hI' : I' = l.val.get ⟨List.indexOf I' l.val, hi⟩ := (List.indexOf_get hi).symm
|
||||
rw [hI']
|
||||
exact h ⟨List.indexOf I' l.val, hi⟩
|
||||
|
||||
/-!
|
||||
|
||||
## On appended lists
|
||||
|
||||
-/
|
||||
|
||||
lemma inl (h : (l ++ l2).OnlyUniqueDuals) : l.OnlyUniqueDuals := by
|
||||
rw [iff_countId_leq_two'] at h ⊢
|
||||
intro I
|
||||
have hI := h I
|
||||
simp at hI
|
||||
omega
|
||||
|
||||
lemma inr (h : (l ++ l2).OnlyUniqueDuals) : l2.OnlyUniqueDuals := by
|
||||
rw [iff_countId_leq_two'] at h ⊢
|
||||
intro I
|
||||
have hI := h I
|
||||
simp at hI
|
||||
omega
|
||||
|
||||
lemma symm' (h : (l ++ l2).OnlyUniqueDuals) : (l2 ++ l).OnlyUniqueDuals := by
|
||||
rw [iff_countId_leq_two'] at h ⊢
|
||||
intro I
|
||||
have hI := h I
|
||||
simp at hI
|
||||
simp
|
||||
omega
|
||||
|
||||
lemma symm : (l ++ l2).OnlyUniqueDuals ↔ (l2 ++ l).OnlyUniqueDuals := by
|
||||
refine Iff.intro symm' symm'
|
||||
|
||||
lemma swap (h : (l ++ l2 ++ l3).OnlyUniqueDuals) : (l2 ++ l ++ l3).OnlyUniqueDuals := by
|
||||
rw [iff_countId_leq_two'] at h ⊢
|
||||
intro I
|
||||
have hI := h I
|
||||
simp at hI
|
||||
simp
|
||||
omega
|
||||
|
||||
/-!
|
||||
|
||||
## Contractions
|
||||
|
||||
-/
|
||||
|
||||
lemma contrIndexList (h : l.OnlyUniqueDuals) : l.contrIndexList.OnlyUniqueDuals := by
|
||||
rw [iff_countId_leq_two'] at h ⊢
|
||||
intro I
|
||||
have hI := h I
|
||||
have h1 := countId_contrIndexList_le_countId l I
|
||||
omega
|
||||
|
||||
lemma contrIndexList_left (h1 : (l ++ l2).OnlyUniqueDuals) :
|
||||
(l.contrIndexList ++ l2).OnlyUniqueDuals := by
|
||||
rw [iff_countId_leq_two'] at h1 ⊢
|
||||
intro I
|
||||
have hI := h1 I
|
||||
simp only [countId_append] at hI
|
||||
simp only [countId_append, ge_iff_le]
|
||||
have h1 := countId_contrIndexList_le_countId l I
|
||||
omega
|
||||
|
||||
lemma contrIndexList_right (h1 : (l ++ l2).OnlyUniqueDuals) :
|
||||
(l ++ l2.contrIndexList).OnlyUniqueDuals := by
|
||||
rw [symm] at h1 ⊢
|
||||
exact contrIndexList_left h1
|
||||
|
||||
lemma contrIndexList_append (h1 : (l ++ l2).OnlyUniqueDuals) :
|
||||
(l.contrIndexList ++ l2.contrIndexList).OnlyUniqueDuals := by
|
||||
rw [iff_countId_leq_two'] at h1 ⊢
|
||||
intro I
|
||||
have hI := h1 I
|
||||
simp only [countId_append] at hI
|
||||
simp only [countId_append, ge_iff_le]
|
||||
have h1 := countId_contrIndexList_le_countId l I
|
||||
have h2 := countId_contrIndexList_le_countId l2 I
|
||||
omega
|
||||
|
||||
end OnlyUniqueDuals
|
||||
|
||||
lemma countId_of_OnlyUniqueDuals (h : l.OnlyUniqueDuals) (I : Index X) :
|
||||
l.countId I ≤ 2 := by
|
||||
rw [OnlyUniqueDuals.iff_countId_leq_two'] at h
|
||||
exact h I
|
||||
|
||||
lemma countId_eq_two_ofcontrIndexList_left_of_OnlyUniqueDuals
|
||||
(h : (l ++ l2).OnlyUniqueDuals) (I : Index X) (h' : (l.contrIndexList ++ l2).countId I = 2 ) :
|
||||
(l ++ l2).countId I = 2 := by
|
||||
simp
|
||||
have h1 := countId_contrIndexList_le_countId l I
|
||||
have h3 := countId_of_OnlyUniqueDuals _ h I
|
||||
have h4 := countId_of_OnlyUniqueDuals _ h.contrIndexList_left I
|
||||
simp at h3 h4 h'
|
||||
omega
|
||||
|
||||
|
||||
end IndexList
|
||||
|
||||
end IndexNotation
|
|
@ -0,0 +1,99 @@
|
|||
/-
|
||||
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.IndexList.CountId
|
||||
import HepLean.SpaceTime.LorentzTensor.IndexNotation.IndexList.Contraction
|
||||
import HepLean.SpaceTime.LorentzTensor.IndexNotation.IndexList.Color
|
||||
import Mathlib.Algebra.Order.Ring.Nat
|
||||
import Mathlib.Data.Finset.Sort
|
||||
import Mathlib.Tactic.FinCases
|
||||
/-!
|
||||
|
||||
# Subperm for index lists.
|
||||
|
||||
-/
|
||||
|
||||
namespace IndexNotation
|
||||
|
||||
namespace IndexList
|
||||
|
||||
variable {X : Type} [IndexNotation X] [Fintype X] [DecidableEq X]
|
||||
variable (l l2 l3 : IndexList X)
|
||||
|
||||
def Subperm : Prop := l.withUniqueDualInOther l2 = Finset.univ
|
||||
|
||||
namespace Subperm
|
||||
|
||||
variable {l l2 l3 : IndexList X}
|
||||
|
||||
lemma iff_countId_fin : l.Subperm l2 ↔
|
||||
∀ i, l.countId (l.val.get i) = 1 ∧ l2.countId (l.val.get i) = 1 := by
|
||||
refine Iff.intro (fun h i => ?_) (fun h => ?_)
|
||||
· have hi : i ∈ l.withUniqueDualInOther l2 := by
|
||||
rw [h]
|
||||
exact Finset.mem_univ i
|
||||
exact countId_eq_one_of_mem_withUniqueDualInOther l l2 i hi
|
||||
· rw [Subperm, Finset.eq_univ_iff_forall]
|
||||
exact fun x => mem_withUniqueDualInOther_of_countId_eq_one l l2 x (h x)
|
||||
|
||||
lemma iff_countId : l.Subperm l2 ↔
|
||||
∀ I, l.countId I ≤ 1 ∧ (l.countId I = 1 → l2.countId I = 1) := by
|
||||
rw [iff_countId_fin]
|
||||
refine Iff.intro (fun h I => ?_) (fun h i => ?_)
|
||||
· by_cases h' : l.countId I = 0
|
||||
· simp_all
|
||||
· obtain ⟨I', hI'1, hI'2⟩ := l.countId_neq_zero_mem I h'
|
||||
rw [countId_congr _ hI'2, countId_congr _ hI'2]
|
||||
have hi' : l.val.indexOf I' < l.val.length := List.indexOf_lt_length.mpr hI'1
|
||||
have hIi' : I' = l.val.get ⟨l.val.indexOf I', hi'⟩ := (List.indexOf_get (hi')).symm
|
||||
have hi' := h ⟨l.val.indexOf I', hi'⟩
|
||||
rw [← hIi'] at hi'
|
||||
rw [hi'.1, hi'.2]
|
||||
simp
|
||||
· have hi := h (l.val.get i)
|
||||
have h1 : l.countId (l.val.get i) ≠ 0 := countId_index_neq_zero l i
|
||||
omega
|
||||
|
||||
lemma trans (h1 : l.Subperm l2) (h2 : l2.Subperm l3) : l.Subperm l3 := by
|
||||
rw [iff_countId] at *
|
||||
intro I
|
||||
have h1 := h1 I
|
||||
have h2 := h2 I
|
||||
omega
|
||||
|
||||
lemma fst_eq_contrIndexList (h : l.Subperm l2) : l.contrIndexList = l := by
|
||||
rw [iff_countId] at *
|
||||
apply ext
|
||||
simp [contrIndexList]
|
||||
intro I hI
|
||||
have h' : l.countId I ≠ 0 := by exact countId_mem l I hI
|
||||
have hI' := h I
|
||||
omega
|
||||
|
||||
lemma symm (h : l.Subperm l2) (hl : l.length = l2.length) : l2.Subperm l := by
|
||||
refine (Finset.card_eq_iff_eq_univ (l2.withUniqueDualInOther l)).mp ?_
|
||||
trans (Finset.map ⟨Subtype.val, Subtype.val_injective⟩ (Equiv.finsetCongr
|
||||
(l.getDualInOtherEquiv l2) Finset.univ)).card
|
||||
· apply congrArg
|
||||
simp [Finset.ext_iff]
|
||||
· trans l2.length
|
||||
· simp only [Finset.univ_eq_attach, Equiv.finsetCongr_apply, Finset.card_map,
|
||||
Finset.card_attach]
|
||||
rw [h, ← hl]
|
||||
exact Finset.card_fin l.length
|
||||
· exact Eq.symm (Finset.card_fin l2.length)
|
||||
|
||||
lemma contr_refl : l.contrIndexList.Subperm l.contrIndexList := by
|
||||
rw [iff_countId]
|
||||
intro I
|
||||
simp
|
||||
exact countId_contrIndexList_le_one l I
|
||||
|
||||
end Subperm
|
||||
|
||||
|
||||
end IndexList
|
||||
|
||||
end IndexNotation
|
Loading…
Add table
Add a link
Reference in a new issue