refactor: Lint

This commit is contained in:
jstoobysmith 2024-08-15 10:16:42 -04:00
parent e458300359
commit 0edce53795
15 changed files with 2406 additions and 2197 deletions

View file

@ -72,12 +72,17 @@ import HepLean.SpaceTime.LorentzGroup.Restricted
import HepLean.SpaceTime.LorentzGroup.Rotations
import HepLean.SpaceTime.LorentzTensor.Basic
import HepLean.SpaceTime.LorentzTensor.Contraction
import HepLean.SpaceTime.LorentzTensor.IndexNotation.AreDual
import HepLean.SpaceTime.LorentzTensor.IndexNotation.Basic
import HepLean.SpaceTime.LorentzTensor.IndexNotation.Color
import HepLean.SpaceTime.LorentzTensor.IndexNotation.Dual
import HepLean.SpaceTime.LorentzTensor.IndexNotation.Contraction
import HepLean.SpaceTime.LorentzTensor.IndexNotation.GetDual
import HepLean.SpaceTime.LorentzTensor.IndexNotation.IndexString
import HepLean.SpaceTime.LorentzTensor.IndexNotation.OnlyUniqueDuals
import HepLean.SpaceTime.LorentzTensor.IndexNotation.Relations
import HepLean.SpaceTime.LorentzTensor.IndexNotation.TensorIndex
import HepLean.SpaceTime.LorentzTensor.IndexNotation.WithDual
import HepLean.SpaceTime.LorentzTensor.IndexNotation.WithUniqueDual
import HepLean.SpaceTime.LorentzTensor.MulActionTensor
import HepLean.SpaceTime.LorentzTensor.Real.Basic
import HepLean.SpaceTime.LorentzTensor.Real.IndexNotation

View file

@ -0,0 +1,116 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.SpaceTime.LorentzTensor.IndexNotation.Basic
import Mathlib.Algebra.Order.Ring.Nat
import Mathlib.Data.Finset.Sort
/-!
# Indices which are dual in an index list
-/
namespace IndexNotation
namespace IndexList
variable {X : Type} [IndexNotation X] [Fintype X] [DecidableEq X]
variable (l l2 l3 : IndexList X)
/-!
## Are dual indices
-/
/-- Two indices are dual if they are not equivalent, but have the same id. -/
def AreDualInSelf (i j : Fin l.length) : Prop :=
i ≠ j ∧ l.idMap i = l.idMap j
/-- Two indices in different `IndexLists` are dual to one another if they have the same `id`. -/
def AreDualInOther (i : Fin l.length) (j : Fin l2.length) :
Prop := l.idMap i = l2.idMap j
namespace AreDualInSelf
variable {l l2 : IndexList X} {i j : Fin l.length}
instance (i j : Fin l.length) : Decidable (l.AreDualInSelf i j) :=
instDecidableAnd
@[symm]
lemma symm (h : l.AreDualInSelf i j) : l.AreDualInSelf j i := by
simp only [AreDualInSelf] at h ⊢
exact ⟨h.1.symm, h.2.symm⟩
@[simp]
lemma self_false (i : Fin l.length) : ¬ l.AreDualInSelf i i := by
simp [AreDualInSelf]
@[simp]
lemma append_inl_inl : (l ++ l2).AreDualInSelf (appendEquiv (Sum.inl i)) (appendEquiv (Sum.inl j))
↔ l.AreDualInSelf i j := by
simp [AreDualInSelf]
@[simp]
lemma append_inr_inr (l l2 : IndexList X) (i j : Fin l2.length) :
(l ++ l2).AreDualInSelf (appendEquiv (Sum.inr i)) (appendEquiv (Sum.inr j))
↔ l2.AreDualInSelf i j := by
simp [AreDualInSelf]
@[simp]
lemma append_inl_inr (l l2 : IndexList X) (i : Fin l.length) (j : Fin l2.length) :
(l ++ l2).AreDualInSelf (appendEquiv (Sum.inl i)) (appendEquiv (Sum.inr j)) =
l.AreDualInOther l2 i j := by
simp [AreDualInSelf, AreDualInOther]
@[simp]
lemma append_inr_inl (l l2 : IndexList X) (i : Fin l2.length) (j : Fin l.length) :
(l ++ l2).AreDualInSelf (appendEquiv (Sum.inr i)) (appendEquiv (Sum.inl j)) =
l2.AreDualInOther l i j := by
simp [AreDualInSelf, AreDualInOther]
end AreDualInSelf
namespace AreDualInOther
variable {l l2 l3 : IndexList X} {i : Fin l.length} {j : Fin l2.length}
instance {l : IndexList X} {l2 : IndexList X} (i : Fin l.length) (j : Fin l2.length) :
Decidable (l.AreDualInOther l2 i j) := (l.idMap i).decEq (l2.idMap j)
@[symm]
lemma symm (h : l.AreDualInOther l2 i j) : l2.AreDualInOther l j i := by
rw [AreDualInOther] at h ⊢
exact h.symm
@[simp]
lemma append_of_inl (i : Fin l.length) (j : Fin l3.length) :
(l ++ l2).AreDualInOther l3 (appendEquiv (Sum.inl i)) j ↔ l.AreDualInOther l3 i j := by
simp [AreDualInOther]
@[simp]
lemma append_of_inr (i : Fin l2.length) (j : Fin l3.length) :
(l ++ l2).AreDualInOther l3 (appendEquiv (Sum.inr i)) j ↔ l2.AreDualInOther l3 i j := by
simp [AreDualInOther]
@[simp]
lemma of_append_inl (i : Fin l.length) (j : Fin l2.length) :
l.AreDualInOther (l2 ++ l3) i (appendEquiv (Sum.inl j)) ↔ l.AreDualInOther l2 i j := by
simp [AreDualInOther]
@[simp]
lemma of_append_inr (i : Fin l.length) (j : Fin l3.length) :
l.AreDualInOther (l2 ++ l3) i (appendEquiv (Sum.inr j)) ↔ l.AreDualInOther l3 i j := by
simp [AreDualInOther]
end AreDualInOther
end IndexList
end IndexNotation

View file

@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import Mathlib.Data.Set.Finite
import Mathlib.Data.Finset.Sort
import Mathlib.Logic.Equiv.Fin
/-!
@ -176,6 +175,7 @@ end Index
/-- The type of lists of indices. -/
structure IndexList where
/-- The list of index values. For example `['ᵘ¹','ᵘ²','ᵤ₁']`. -/
val : List (Index X)
namespace IndexList
@ -278,9 +278,12 @@ lemma append_assoc : l ++ l2 ++ l3 = l ++ (l2 ++ l3) := by
change l.val ++ l2.val ++ l3.val = l.val ++ (l2.val ++ l3.val)
exact List.append_assoc l.val l2.val l3.val
/-- An equivalence between the sum of the types of indices of `l` an `l2` and the type
of indices of the joined index list `l ++ l2`. -/
def appendEquiv {l l2 : IndexList X} : Fin l.length ⊕ Fin l2.length ≃ Fin (l ++ l2).length :=
finSumFinEquiv.trans (Fin.castOrderIso (List.length_append _ _).symm).toEquiv
/-- The inclusion of the indices of `l` into the indices of `l ++ l2`. -/
def appendInl : Fin l.length ↪ Fin (l ++ l2).length where
toFun := appendEquiv ∘ Sum.inl
inj' := by
@ -288,6 +291,7 @@ def appendInl : Fin l.length ↪ Fin (l ++ l2).length where
simp [Function.comp] at h
exact h
/-- The inclusion of the indices of `l2` into the indices of `l ++ l2`. -/
def appendInr : Fin l2.length ↪ Fin (l ++ l2).length where
toFun := appendEquiv ∘ Sum.inr
inj' := by
@ -325,7 +329,7 @@ lemma idMap_append_inr {l l2 : IndexList X} (i : Fin l2.length) :
(l ++ l2).idMap (appendEquiv (Sum.inr i)) = l2.idMap i := by
simp [appendEquiv, idMap, IndexList.length]
rw [List.getElem_append_right]
simp
simp only [Nat.add_sub_cancel_left]
omega
omega
@ -346,7 +350,7 @@ lemma colorMap_append_inr {l l2 : IndexList X} (i : Fin l2.length) :
(l ++ l2).colorMap (appendEquiv (Sum.inr i)) = l2.colorMap i := by
simp [appendEquiv, colorMap, IndexList.length]
rw [List.getElem_append_right]
simp
simp only [Nat.add_sub_cancel_left]
omega
omega

View file

@ -3,7 +3,7 @@ Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.SpaceTime.LorentzTensor.IndexNotation.Dual
import HepLean.SpaceTime.LorentzTensor.IndexNotation.Contraction
import HepLean.SpaceTime.LorentzTensor.Basic
import Init.Data.List.Lemmas
import HepLean.SpaceTime.LorentzTensor.Contraction
@ -26,6 +26,7 @@ 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
@ -52,12 +53,12 @@ lemma iff_withDual :
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
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
simp only [Option.some.injEq]
have hii := h ⟨i, by simp [withDual, hi]⟩
simp at hii
rw [← hii]
@ -78,15 +79,7 @@ lemma assoc (h : ColorCond (l ++ l2 ++ l3)) :
lemma inl (h : ColorCond (l ++ l2)) : ColorCond l := by
rw [iff_withDual] at h ⊢
intro i
have hi' := h ⟨appendEquiv (Sum.inl i), by
simp_all⟩
have hn : (Option.map (appendEquiv ∘ Sum.inl) (l.getDual? ↑i) : Option (Fin (l ++ l2).length)) =
some (appendEquiv (Sum.inl ((l.getDual? i).get (l.withDual_isSome i)))) := by
trans Option.map (appendEquiv ∘ Sum.inl) (some ((l.getDual? i).get (l.withDual_isSome i)))
simp
rw [Option.map_some']
simp
simpa [hn] using hi'
simpa using h ⟨appendEquiv (Sum.inl i), by simp_all⟩
lemma symm (hu : (l ++ l2).withUniqueDual = (l ++ l2).withDual) (h : ColorCond (l ++ l2)) :
ColorCond (l2 ++ l) := by
@ -188,8 +181,8 @@ lemma swap (hu : (l ++ l2 ++ l3).withUniqueDual = (l ++ l2 ++ l3).withDual)
subst hk'
match k' with
| Sum.inl k' =>
simp at hn
simp
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 [hn])
@ -198,8 +191,8 @@ lemma swap (hu : (l ++ l2 ++ l3).withUniqueDual = (l ++ l2 ++ l3).withDual)
getDual?_eq_none_append_inl_iff, Option.get_some, colorMap_append_inr,
colorMap_append_inl]
| Sum.inr k' =>
simp at hn
simp
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 [hn])
@ -243,8 +236,8 @@ lemma swap (hu : (l ++ l2 ++ l3).withUniqueDual = (l ++ l2 ++ l3).withDual)
getDual?_append_inr_getDualInOther?_isSome, Option.get_some, colorMap_append_inl,
colorMap_append_inr]
/- l.getDual? = Option.map (𝓒.τ ∘ l.colorMap) ∘
Option.guard fun i => (l.getDual? i).isSome -/
/-- 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
@ -265,8 +258,17 @@ end IndexList
variable (𝓒 : TensorColor)
variable [IndexNotation 𝓒.Color] [Fintype 𝓒.Color] [DecidableEq 𝓒.Color]
/-- A list of indices with the additional constraint that if a index has a dual,
that dual is unique, and the dual of an index has dual color to that index.
This is the permissible type of indices which can be used for a tensor. For example,
the index list `['ᵘ¹', 'ᵤ₁']` can be extended to a `ColorIndexList` but the index list
`['ᵘ¹', 'ᵤ₁', 'ᵤ₁']` cannot. -/
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
/-- The condition that for an index with a dual, that dual has dual color to the index. -/
dual_color : IndexList.ColorCond toIndexList
namespace ColorIndexList
@ -277,13 +279,15 @@ variable (l l2 : ColorIndexList 𝓒)
open IndexList TensorColor
instance : Coe (ColorIndexList 𝓒) (IndexList 𝓒.Color) := ⟨fun l => l.toIndexList⟩
/-- The `ColorIndexList` whose underlying list of indices is empty. -/
def empty : ColorIndexList 𝓒 where
val := ∅
unique_duals := by
rfl
dual_color := by
rfl
unique_duals := rfl
dual_color := rfl
/-- The colorMap of a `ColorIndexList` as a `𝓒.ColorMap`.
This is to be compared with `colorMap` which is a map `Fin l.length → 𝓒.Color`. -/
def colorMap' : 𝓒.ColorMap (Fin l.length) :=
l.colorMap
@ -302,7 +306,8 @@ lemma orderEmbOfFin_univ (n m : ) (h : n = m):
(Fin.castOrderIso h.symm).toOrderEmbedding := by
symm
have h1 : (Fin.castOrderIso h.symm).toFun =
( Finset.orderEmbOfFin (Finset.univ : Finset (Fin n)) (by simp [h]: Finset.univ.card = m)).toFun := by
( Finset.orderEmbOfFin (Finset.univ : Finset (Fin n))
(by simp [h]: Finset.univ.card = m)).toFun := by
apply Finset.orderEmbOfFin_unique
intro x
exact Finset.mem_univ ((Fin.castOrderIso (Eq.symm h)).toFun x)
@ -316,6 +321,9 @@ lemma orderEmbOfFin_univ (n m : ) (h : n = m):
-/
/-- The contraction of a `ColorIndexList`, `l`.
That is, the `ColorIndexList` obtained by taking only those indices in `l` which do not
have a dual. This can be thought of as contracting all of those indices with a dual. -/
def contr : ColorIndexList 𝓒 where
toIndexList := l.toIndexList.contrIndexList
unique_duals := by
@ -369,6 +377,11 @@ lemma contr_areDualInSelf (i j : Fin l.contr.length) :
-/
/-- An equivalence splitting the indices of `l` into
a sum type of those indices and their duals (with choice determined by the ordering on `Fin`),
and those indices without duals.
This equivalence is used to contract the indices of tensors. -/
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
@ -404,9 +417,9 @@ lemma contrEquiv_colorMapIso :
rfl
lemma contrEquiv_contrCond : ColorMap.ContrCond l.contrEquiv l.colorMap := by
simp [ColorMap.ContrCond]
simp only [ColorMap.ContrCond]
funext i
simp
simp only [Function.comp_apply, contrEquiv_inl_inl_eq, contrEquiv_inl_inr_eq]
have h1 := l.dual_color
rw [ColorCond.iff_on_isSome] at h1
exact (h1 i.1 _).symm
@ -432,16 +445,21 @@ lemma contrEquiv_on_withDual_empty (i : Fin l.contr.length) (h : l.withDual =
-/
/-- The condition on the `ColorIndexList`s `l` and `l2` so that on appending they form
a `ColorIndexList`. -/
def AppendCond : Prop :=
(l.toIndexList ++ l2.toIndexList).withUniqueDual = (l.toIndexList ++ l2.toIndexList).withDual
∧ 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
dual_color := h.2
/-- The join of two `ColorIndexList` satisfying the condition `AppendCond` that they
can be appended to form a `ColorIndexList`. -/
scoped[IndexNotation.ColorIndexList] notation l " ++["h"] " l2 => append l l2 h
@[simp]
@ -489,11 +507,7 @@ lemma swap (h : AppendCond l l2) (h' : AppendCond (l ++[h] l2) l3) :
rw [← append_withDual_eq_withUniqueDual_swap]
simpa using h'.1
· exact ColorCond.swap h'.1 h'.2
/-
(l ++ l2).withUniqueDual = (l ++ l2).withDual ↔
l.withUniqueDual = l.withDual ∧ l2.withUniqueDual = l2.withDual
∧ l.withUniqueDualInOther l2 = l.withDualInOther l2 ∧
l2.withUniqueDualInOther l = l2.withDualInOther l -/
lemma appendCond_of_eq (h1 : l.withUniqueDual = l.withDual)
(h2 : l2.withUniqueDual = l2.withDual)
(h3 : l.withUniqueDualInOther l2 = l.withDualInOther l2)
@ -505,6 +519,8 @@ lemma appendCond_of_eq (h1 : l.withUniqueDual = l.withDual)
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 :=
if ¬ (l ++ l2).withUniqueDual = (l ++ l2).withDual then
false

View file

@ -0,0 +1,399 @@
/-
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 Mathlib.Algebra.Order.Ring.Nat
import Mathlib.Data.Finset.Sort
/-!
# Contraction of Dual indices
-/
namespace IndexNotation
namespace IndexList
variable {X : Type} [IndexNotation X] [Fintype X] [DecidableEq X]
variable (l l2 l3 : IndexList X)
/-!
## Indices which do not have duals.
-/
/-- The finite set of indices of an index list which do not have a dual. -/
def withoutDual : Finset (Fin l.length) :=
Finset.filter (fun i => (l.getDual? i).isNone) Finset.univ
lemma withDual_disjoint_withoutDual : Disjoint l.withDual l.withoutDual := by
rw [Finset.disjoint_iff_ne]
intro a ha b hb
by_contra hn
subst hn
simp_all only [withDual, Finset.mem_filter, Finset.mem_univ, true_and, withoutDual,
Option.isNone_iff_eq_none, Option.isSome_none, Bool.false_eq_true]
lemma not_mem_withDual_of_mem_withoutDual (i : Fin l.length) (h : i ∈ l.withoutDual) :
i ∉ l.withDual := by
have h1 := l.withDual_disjoint_withoutDual
exact Finset.disjoint_right.mp h1 h
lemma withDual_union_withoutDual : l.withDual l.withoutDual = Finset.univ := by
rw [Finset.eq_univ_iff_forall]
intro i
by_cases h : (l.getDual? i).isSome
· simp [withDual, Finset.mem_filter, Finset.mem_univ, h]
· simp at h
simp [withoutDual, Finset.mem_filter, Finset.mem_univ, h]
/-- 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
/-- 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])
/-!
## The contraction list
-/
/-- 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 := (Fin.list l.withoutDual.card).map (fun i => l.val.get (l.withoutDualEquiv i).1)
@[simp]
lemma contrIndexList_length : l.contrIndexList.length = l.withoutDual.card := by
simp [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, 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, 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 [contrIndexList_areDualInSelf]
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_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]
congr
simp [withoutDualEquiv]
simp [h1]
rw [(Finset.orderEmbOfFin_unique' _
(fun x => Finset.mem_univ ((Fin.castOrderIso _).toOrderEmbedding x))).symm]
refine 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))
/-!
## Splitting withUniqueDual
-/
instance (i j : Option (Fin l.length)) : Decidable (i < j) :=
match i, j with
| none, none => isFalse (fun a => a)
| none, some _ => isTrue (by trivial)
| some _, none => isFalse (fun a => a)
| some i, some j => Fin.decLt i j
/-- 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 casting of an element of `withUniqueDualLT` to an element of `withUniqueDual`. -/
def withUniqueDualLTToWithUniqueDual (i : l.withUniqueDualLT) : l.withUniqueDual :=
⟨i.1, Finset.mem_of_mem_filter i.1 i.2⟩
instance : Coe l.withUniqueDualLT l.withUniqueDual where
coe := l.withUniqueDualLTToWithUniqueDual
/-- 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 casting of an element in `withUniqueDualGT` to an element of `withUniqueDual`. -/
def withUniqueDualGTToWithUniqueDual (i : l.withUniqueDualGT) : l.withUniqueDual :=
⟨i.1, Finset.mem_of_mem_filter _ i.2⟩
instance : Coe l.withUniqueDualGT l.withUniqueDual where
coe := l.withUniqueDualGTToWithUniqueDual
lemma withUniqueDualLT_disjoint_withUniqueDualGT :
Disjoint l.withUniqueDualLT l.withUniqueDualGT := by
rw [Finset.disjoint_iff_inter_eq_empty]
exact @Finset.filter_inter_filter_neg_eq (Fin l.length) _ _ _ _ _
lemma withUniqueDualLT_union_withUniqueDualGT :
l.withUniqueDualLT l.withUniqueDualGT = l.withUniqueDual :=
Finset.filter_union_filter_neg_eq _ _
/-! TODO: Replace with a mathlib lemma. -/
lemma option_not_lt (i j : Option (Fin l.length)) : i < j → i ≠ j → ¬ j < i := by
match i, j with
| none, none =>
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, by
have hi := i.2
simp only [withUniqueDualGT, Finset.mem_filter, Finset.coe_mem, true_and]
simp only [getDualEquiv, Equiv.coe_fn_mk, Option.some_get, Finset.coe_mem,
getDual?_getDual?_get_of_withUniqueDual]
simp [withUniqueDualLT] at hi
apply option_not_lt
simpa [withUniqueDualLTToWithUniqueDual] using hi.2
exact Ne.symm (getDual?_neq_self l i)⟩
invFun i := ⟨l.getDualEquiv.symm i, by
have hi := i.2
simp only [withUniqueDualLT, Finset.mem_filter, Finset.coe_mem, true_and, gt_iff_lt]
simp only [getDualEquiv, Equiv.coe_fn_symm_mk, Option.some_get, Finset.coe_mem,
getDual?_getDual?_get_of_withUniqueDual]
simp [withUniqueDualGT] at hi
apply lt_option_of_not
simpa [withUniqueDualLTToWithUniqueDual] using hi.2
exact (getDual?_neq_self l i)⟩
left_inv x := SetCoe.ext (by simp [withUniqueDualGTToWithUniqueDual,
withUniqueDualLTToWithUniqueDual])
right_inv x := SetCoe.ext (by simp [withUniqueDualGTToWithUniqueDual,
withUniqueDualLTToWithUniqueDual])
/-- 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 [l.withUniqueDualLT_union_withUniqueDualGT]))
/-!
## withUniqueDualInOther equal univ
-/
/-! TODO: There is probably a better place for this section. -/
lemma withUniqueDualInOther_eq_univ_fst_withDual_empty
(h : l.withUniqueDualInOther l2 = Finset.univ) : l.withDual = ∅ := by
rw [@Finset.eq_empty_iff_forall_not_mem]
intro x
have hx : x ∈ l.withUniqueDualInOther l2 := by
rw [h]
exact Finset.mem_univ x
rw [withUniqueDualInOther] at hx
simp only [mem_withDual_iff_isSome, Bool.not_eq_true, Option.not_isSome,
Option.isNone_iff_eq_none, mem_withInDualOther_iff_isSome, Finset.mem_filter, Finset.mem_univ,
true_and] at hx
simpa using hx.1
lemma withUniqueDualInOther_eq_univ_fst_eq_contrIndexList
(h : l.withUniqueDualInOther l2 = Finset.univ) :
l = l.contrIndexList := by
refine Eq.symm (contrIndexList_of_withDual_empty l ?h)
exact withUniqueDualInOther_eq_univ_fst_withDual_empty l l2 h
lemma withUniqueDualInOther_eq_univ_symm (hl : l.length = l2.length)
(h : l.withUniqueDualInOther l2 = Finset.univ) :
l2.withUniqueDualInOther l = Finset.univ := by
let S' : Finset (Fin l2.length) :=
Finset.map ⟨Subtype.val, Subtype.val_injective⟩
(Equiv.finsetCongr
(l.getDualInOtherEquiv l2) Finset.univ )
have hSCard : S'.card = l2.length := by
rw [Finset.card_map]
simp only [Finset.univ_eq_attach, Equiv.finsetCongr_apply, Finset.card_map, Finset.card_attach]
rw [h, ← hl]
exact Finset.card_fin l.length
have hsCardUn : S'.card = (@Finset.univ (Fin l2.length)).card := by
rw [hSCard]
exact Eq.symm (Finset.card_fin l2.length)
have hS'Eq : S' = (l2.withUniqueDualInOther l) := by
rw [@Finset.ext_iff]
intro a
simp [S']
rw [hS'Eq] at hsCardUn
exact (Finset.card_eq_iff_eq_univ (l2.withUniqueDualInOther l)).mp hsCardUn
lemma withUniqueDualInOther_eq_univ_contr_refl :
l.contrIndexList.withUniqueDualInOther l.contrIndexList = Finset.univ := by
rw [@Finset.eq_univ_iff_forall]
intro x
simp only [withUniqueDualInOther, mem_withDual_iff_isSome,
Option.isSome_none, Bool.false_eq_true, not_false_eq_true, mem_withInDualOther_iff_isSome,
getDualInOther?_self_isSome, true_and, Finset.mem_filter, Finset.mem_univ]
simp only [contrIndexList_getDual?, Option.isSome_none, Bool.false_eq_true, not_false_eq_true,
contrIndexList_getDualInOther?_self, Option.some.injEq, true_and]
intro j hj
have h1 : j = x := by
by_contra hn
have hj : l.contrIndexList.AreDualInSelf x j := by
simp [AreDualInOther] at hj
simp only [AreDualInSelf, ne_eq, contrIndexList_idMap, hj, and_true]
exact fun a => hn (id (Eq.symm a))
exact (contrIndexList_areDualInSelf_false l x j).mp hj
exact h1
lemma withUniqueDualInOther_eq_univ_trans (h : l.withUniqueDualInOther l2 = Finset.univ)
(h' : l2.withUniqueDualInOther l3 = Finset.univ) :
l.withUniqueDualInOther l3 = Finset.univ := by
rw [Finset.eq_univ_iff_forall]
intro i
simp only [withUniqueDualInOther, mem_withDual_iff_isSome, Bool.not_eq_true, Option.not_isSome,
Option.isNone_iff_eq_none, mem_withInDualOther_iff_isSome, Finset.mem_filter, Finset.mem_univ,
true_and]
have hi : i ∈ l.withUniqueDualInOther l2 := by
rw [h]
exact Finset.mem_univ i
simp only [withUniqueDualInOther, mem_withDual_iff_isSome, Bool.not_eq_true, Option.not_isSome,
Option.isNone_iff_eq_none, mem_withInDualOther_iff_isSome, Finset.mem_filter, Finset.mem_univ,
true_and] at hi
have hi2 : ((l.getDualInOther? l2 i).get hi.2.1) ∈ l2.withUniqueDualInOther l3 := by
rw [h']
exact Finset.mem_univ ((l.getDualInOther? l2 i).get hi.right.left)
simp only [withUniqueDualInOther, mem_withDual_iff_isSome, Bool.not_eq_true, Option.not_isSome,
Option.isNone_iff_eq_none, mem_withInDualOther_iff_isSome, Finset.mem_filter, Finset.mem_univ,
true_and] at hi2
apply And.intro hi.1
apply And.intro
· rw [@getDualInOther?_isSome_iff_exists]
use (l2.getDualInOther? l3 ((l.getDualInOther? l2 i).get hi.2.1)).get hi2.2.1
simp [AreDualInOther]
intro j hj
have hj' : j = (l2.getDualInOther? l3 ((l.getDualInOther? l2 i).get hi.2.1)).get hi2.2.1 := by
rw [← eq_getDualInOther?_get_of_withUniqueDualInOther_iff]
simpa [AreDualInOther] using hj
rw [h']
exact Finset.mem_univ ((l.getDualInOther? l2 i).get hi.right.left)
have hs : (l.getDualInOther? l3 i).isSome := by
rw [@getDualInOther?_isSome_iff_exists]
exact Exists.intro j hj
have hs' : (l.getDualInOther? l3 i).get hs = (l2.getDualInOther? l3
((l.getDualInOther? l2 i).get hi.2.1)).get hi2.2.1 := by
rw [← eq_getDualInOther?_get_of_withUniqueDualInOther_iff]
simp [AreDualInOther]
rw [h']
exact Finset.mem_univ ((l.getDualInOther? l2 i).get hi.right.left)
rw [← hj'] at hs'
rw [← hs']
exact Eq.symm (Option.eq_some_of_isSome hs)
end IndexList
end IndexNotation

File diff suppressed because it is too large Load diff

View file

@ -0,0 +1,513 @@
/-
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.AreDual
import Mathlib.Algebra.Order.Ring.Nat
import Mathlib.Data.Finset.Sort
/-!
# Get dual index
We define the functions `getDual?` and `getDualInOther?` which return the
first dual index in an `IndexList`.
-/
namespace IndexNotation
namespace IndexList
variable {X : Type} [IndexNotation X] [Fintype X] [DecidableEq X]
variable (l l2 : IndexList X)
/-!
## The getDual? Function
-/
/-- Given an `i`, if a dual exists in the same list,
outputs the first such dual, otherwise outputs `none`. -/
def getDual? (i : Fin l.length) : Option (Fin l.length) :=
Fin.find (fun j => l.AreDualInSelf i j)
/-- Given an `i`, if a dual exists in the other list,
outputs the first such dual, otherwise outputs `none`. -/
def getDualInOther? (i : Fin l.length) : Option (Fin l2.length) :=
Fin.find (fun j => l.AreDualInOther l2 i j)
/-!
## Basic properties of getDual?
-/
lemma getDual?_isSome_iff_exists (i : Fin l.length) :
(l.getDual? i).isSome ↔ ∃ j, l.AreDualInSelf i j := by
rw [getDual?, Fin.isSome_find_iff]
lemma getDual?_of_areDualInSelf (h : l.AreDualInSelf i j) :
l.getDual? j = i l.getDual? i = j l.getDual? j = l.getDual? i := by
have h3 : (l.getDual? i).isSome := by
simpa [getDual?, Fin.isSome_find_iff] using ⟨j, h⟩
obtain ⟨k, hk⟩ := Option.isSome_iff_exists.mp h3
rw [hk]
rw [getDual?, Fin.find_eq_some_iff] at hk
by_cases hik : i < k
· apply Or.inl
rw [getDual?, Fin.find_eq_some_iff]
apply And.intro h.symm
intro k' hk'
by_cases hik' : i = k'
· exact Fin.ge_of_eq (id (Eq.symm hik'))
· have hik'' : l.AreDualInSelf i k' := by
simp [AreDualInSelf, hik']
simp_all [AreDualInSelf]
have hk'' := hk.2 k' hik''
exact (lt_of_lt_of_le hik hk'').le
· by_cases hjk : j ≤ k
· apply Or.inr
apply Or.inl
have hj := hk.2 j h
simp only [Option.some.injEq]
exact Fin.le_antisymm hj hjk
· apply Or.inr
apply Or.inr
rw [getDual?, Fin.find_eq_some_iff]
apply And.intro
· simp_all [AreDualInSelf]
exact Fin.ne_of_gt hjk
intro k' hk'
by_cases hik' : i = k'
subst hik'
exact Lean.Omega.Fin.le_of_not_lt hik
have hik'' : l.AreDualInSelf i k' := by
simp [AreDualInSelf, hik']
simp_all [AreDualInSelf]
exact hk.2 k' hik''
@[simp]
lemma getDual?_neq_self (i : Fin l.length) : ¬ l.getDual? i = some i := by
intro h
simp [getDual?] at h
rw [Fin.find_eq_some_iff] at h
simp [AreDualInSelf] at h
@[simp]
lemma getDual?_get_neq_self (i : Fin l.length) (h : (l.getDual? i).isSome) :
¬ (l.getDual? i).get h = i := by
have h1 := l.getDual?_neq_self i
by_contra hn
have h' : l.getDual? i = some i := by
nth_rewrite 2 [← hn]
exact Option.eq_some_of_isSome h
exact h1 h'
@[simp]
lemma getDual?_get_id (i : Fin l.length) (h : (l.getDual? i).isSome) :
l.idMap ((l.getDual? i).get h) = l.idMap i := by
have h1 : l.getDual? i = some ((l.getDual? i).get h) := Option.eq_some_of_isSome h
nth_rewrite 1 [getDual?] at h1
rw [Fin.find_eq_some_iff] at h1
simp [AreDualInSelf] at h1
exact h1.1.2.symm
@[simp]
lemma getDual?_get_areDualInSelf (i : Fin l.length) (h : (l.getDual? i).isSome) :
l.AreDualInSelf ((l.getDual? i).get h) i := by
simp [AreDualInSelf]
@[simp]
lemma getDual?_areDualInSelf_get (i : Fin l.length) (h : (l.getDual? i).isSome) :
l.AreDualInSelf i ((l.getDual? i).get h):=
AreDualInSelf.symm (getDual?_get_areDualInSelf l i h)
@[simp]
lemma getDual?_getDual?_get_isSome (i : Fin l.length) (h : (l.getDual? i).isSome) :
(l.getDual? ((l.getDual? i).get h)).isSome := by
rw [getDual?, Fin.isSome_find_iff]
exact ⟨i, getDual?_get_areDualInSelf l i h⟩
/-!
## Basic properties of getDualInOther?
-/
lemma getDualInOther?_isSome_iff_exists (i : Fin l.length) :
(l.getDualInOther? l2 i).isSome ↔ ∃ j, l.AreDualInOther l2 i j := by
rw [getDualInOther?, Fin.isSome_find_iff]
@[simp]
lemma getDualInOther?_id (i : Fin l.length) (h : (l.getDualInOther? l2 i).isSome) :
l2.idMap ((l.getDualInOther? l2 i).get h) = l.idMap i := by
have h1 : l.getDualInOther? l2 i = some ((l.getDualInOther? l2 i).get h) :=
Option.eq_some_of_isSome h
nth_rewrite 1 [getDualInOther?] at h1
rw [Fin.find_eq_some_iff] at h1
simp [AreDualInOther] at h1
exact h1.1.symm
@[simp]
lemma getDualInOther?_get_areDualInOther (i : Fin l.length) (h : (l.getDualInOther? l2 i).isSome) :
l2.AreDualInOther l ((l.getDualInOther? l2 i).get h) i := by
simp [AreDualInOther]
@[simp]
lemma getDualInOther?_areDualInOther_get (i : Fin l.length) (h : (l.getDualInOther? l2 i).isSome) :
l.AreDualInOther l2 i ((l.getDualInOther? l2 i).get h) :=
AreDualInOther.symm (getDualInOther?_get_areDualInOther l l2 i h)
@[simp]
lemma getDualInOther?_getDualInOther?_get_isSome (i : Fin l.length)
(h : (l.getDualInOther? l2 i).isSome) :
(l2.getDualInOther? l ((l.getDualInOther? l2 i).get h)).isSome := by
rw [getDualInOther?, Fin.isSome_find_iff]
exact ⟨i, getDualInOther?_get_areDualInOther l l2 i h⟩
@[simp]
lemma getDualInOther?_self_isSome (i : Fin l.length) :
(l.getDualInOther? l i).isSome := by
simp [getDualInOther?]
rw [@Fin.isSome_find_iff]
exact ⟨i, rfl⟩
/-!
## Append properties of getDual?
-/
@[simp]
lemma getDual?_isSome_append_inl_iff (i : Fin l.length) :
((l ++ l2).getDual? (appendEquiv (Sum.inl i))).isSome ↔
(l.getDual? i).isSome (l.getDualInOther? l2 i).isSome := by
rw [getDual?_isSome_iff_exists, getDual?_isSome_iff_exists, getDualInOther?_isSome_iff_exists]
refine Iff.intro (fun h => ?_) (fun h => ?_)
· obtain ⟨j, hj⟩ := h
obtain ⟨k, hk⟩ := appendEquiv.surjective j
subst hk
match k with
| Sum.inl k =>
exact Or.inl ⟨k, AreDualInSelf.append_inl_inl.mp hj⟩
| Sum.inr k =>
exact Or.inr ⟨k, by simpa only [AreDualInSelf.append_inl_inr] using hj⟩
· cases' h with h h <;>
obtain ⟨j, hj⟩ := h
· exact ⟨appendEquiv (Sum.inl j), AreDualInSelf.append_inl_inl.mpr hj⟩
· use appendEquiv (Sum.inr j)
simpa only [AreDualInSelf.append_inl_inr] using hj
@[simp]
lemma getDual?_isSome_append_inr_iff (i : Fin l2.length) :
((l ++ l2).getDual? (appendEquiv (Sum.inr i))).isSome ↔
(l2.getDual? i).isSome (l2.getDualInOther? l i).isSome := by
rw [getDual?_isSome_iff_exists, getDual?_isSome_iff_exists, getDualInOther?_isSome_iff_exists]
refine Iff.intro (fun h => ?_) (fun h => ?_)
· obtain ⟨j, hj⟩ := h
obtain ⟨k, hk⟩ := appendEquiv.surjective j
subst hk
match k with
| Sum.inl k =>
exact Or.inr ⟨k, by simpa using hj⟩
| Sum.inr k =>
exact Or.inl ⟨k, (AreDualInSelf.append_inr_inr l l2 i k).mp hj⟩
· cases' h with h h <;>
obtain ⟨j, hj⟩ := h
· exact ⟨appendEquiv (Sum.inr j), (AreDualInSelf.append_inr_inr l l2 i j).mpr hj⟩
· use appendEquiv (Sum.inl j)
simpa only [AreDualInSelf.append_inr_inl] using hj
lemma getDual?_isSome_append_symm (i : Fin l.length) :
((l ++ l2).getDual? (appendEquiv (Sum.inl i))).isSome ↔
((l2 ++ l).getDual? (appendEquiv (Sum.inr i))).isSome := by
simp
@[simp]
lemma getDual?_eq_none_append_inl_iff (i : Fin l.length) :
(l ++ l2).getDual? (appendEquiv (Sum.inl i)) = none ↔
l.getDual? i = none ∧ l.getDualInOther? l2 i = none := by
apply Iff.intro
· intro h
have h1 := (l.getDual?_isSome_append_inl_iff l2 i).mpr.mt
simp only [not_or, Bool.not_eq_true, Option.not_isSome,
Option.isNone_iff_eq_none, imp_self] at h1
exact h1 h
· intro h
have h1 := (l.getDual?_isSome_append_inl_iff l2 i).mp.mt
simp only [not_or, Bool.not_eq_true, Option.not_isSome,
Option.isNone_iff_eq_none, imp_self] at h1
exact h1 h
@[simp]
lemma getDual?_eq_none_append_inr_iff (i : Fin l2.length) :
(l ++ l2).getDual? (appendEquiv (Sum.inr i)) = none ↔
(l2.getDual? i = none ∧ l2.getDualInOther? l i = none) := by
apply Iff.intro
· intro h
have h1 := (l.getDual?_isSome_append_inr_iff l2 i).mpr.mt
simp only [not_or, Bool.not_eq_true, Option.not_isSome,
Option.isNone_iff_eq_none, imp_self] at h1
exact h1 h
· intro h
have h1 := (l.getDual?_isSome_append_inr_iff l2 i).mp.mt
simp only [not_or, Bool.not_eq_true, Option.not_isSome,
Option.isNone_iff_eq_none, imp_self] at h1
exact h1 h
@[simp]
lemma getDual?_append_inl_of_getDual?_isSome (i : Fin l.length) (h : (l.getDual? i).isSome) :
(l ++ l2).getDual? (appendEquiv (Sum.inl i)) =
some (appendEquiv (Sum.inl ((l.getDual? i).get h))) := by
rw [getDual?, Fin.find_eq_some_iff, AreDualInSelf.append_inl_inl]
apply And.intro (getDual?_get_areDualInSelf l i h).symm
intro j hj
obtain ⟨k, hk⟩ := appendEquiv.surjective j
subst hk
match k with
| Sum.inl k =>
simp only [appendEquiv, Equiv.trans_apply, finSumFinEquiv_apply_left, RelIso.coe_fn_toEquiv,
Fin.castOrderIso_apply, ge_iff_le]
rw [Fin.le_def]
have h2 : l.getDual? i = some (((l.getDual? i).get h)) := Option.eq_some_of_isSome h
nth_rewrite 1 [getDual?] at h2
rw [Fin.find_eq_some_iff] at h2
exact h2.2 k (AreDualInSelf.append_inl_inl.mp hj)
| Sum.inr k =>
simp only [appendEquiv, Equiv.trans_apply, finSumFinEquiv_apply_left, RelIso.coe_fn_toEquiv,
Fin.castOrderIso_apply, finSumFinEquiv_apply_right, ge_iff_le]
rw [Fin.le_def]
simp only [length, append_val, RelIso.coe_fn_toEquiv, Fin.castOrderIso_apply, Fin.coe_cast,
Fin.coe_castAdd, Fin.coe_natAdd]
omega
@[simp]
lemma getDual?_inl_of_getDual?_isNone_getDualInOther?_isSome (i : Fin l.length)
(hi : (l.getDual? i).isNone) (h : (l.getDualInOther? l2 i).isSome) :
(l ++ l2).getDual? (appendEquiv (Sum.inl i)) = some
(appendEquiv (Sum.inr ((l.getDualInOther? l2 i).get h))) := by
rw [getDual?, Fin.find_eq_some_iff, AreDualInSelf.append_inl_inr]
apply And.intro (getDualInOther?_areDualInOther_get l l2 i h)
intro j hj
obtain ⟨k, hk⟩ := appendEquiv.surjective j
subst hk
match k with
| Sum.inl k =>
rw [AreDualInSelf.append_inl_inl] at hj
simp only [getDual?, Finset.mem_filter, Finset.mem_univ, true_and, Bool.not_eq_true,
Option.not_isSome, Option.isNone_iff_eq_none, Fin.find_eq_none_iff] at hi
exact False.elim (hi k hj)
| Sum.inr k =>
simp [appendEquiv]
rw [Fin.le_def]
have h1 : l.getDualInOther? l2 i = some (((l.getDualInOther? l2 i).get h)) :=
Option.eq_some_of_isSome h
nth_rewrite 1 [getDualInOther?] at h1
rw [Fin.find_eq_some_iff] at h1
simp only [length, append_val, RelIso.coe_fn_toEquiv, Fin.castOrderIso_apply, Fin.coe_cast,
Fin.coe_natAdd, add_le_add_iff_left, Fin.val_fin_le, ge_iff_le]
refine h1.2 k (by simpa using hj)
lemma getDual?_append_inl (i : Fin l.length) : (l ++ l2).getDual? (appendEquiv (Sum.inl i)) =
Option.or (Option.map (appendEquiv ∘ Sum.inl) (l.getDual? i))
(Option.map (appendEquiv ∘ Sum.inr) (l.getDualInOther? l2 i)) := by
by_cases h : (l.getDual? i).isSome
· simp_all
rw [congrArg (Option.map (appendEquiv ∘ Sum.inl)) (Option.eq_some_of_isSome h)]
rfl
by_cases ho : (l.getDualInOther? l2 i).isSome
· simp_all
rw [congrArg (Option.map (appendEquiv ∘ Sum.inr)) (Option.eq_some_of_isSome ho)]
rfl
simp_all
@[simp]
lemma getDual?_append_inr_getDualInOther?_isSome (i : Fin l2.length)
(h : (l2.getDualInOther? l i).isSome) :
(l ++ l2).getDual? (appendEquiv (Sum.inr i)) =
some (appendEquiv (Sum.inl ((l2.getDualInOther? l i).get h))) := by
rw [getDual?, Fin.find_eq_some_iff, AreDualInSelf.append_inr_inl]
apply And.intro
exact getDualInOther?_areDualInOther_get l2 l i h
intro j hj
obtain ⟨k, hk⟩ := appendEquiv.surjective j
subst hk
match k with
| Sum.inl k =>
simp only [appendEquiv, Equiv.trans_apply, finSumFinEquiv_apply_left, RelIso.coe_fn_toEquiv,
Fin.castOrderIso_apply, ge_iff_le]
rw [Fin.le_def]
have h1 : l2.getDualInOther? l i = some (((l2.getDualInOther? l i).get h)) :=
Option.eq_some_of_isSome h
nth_rewrite 1 [getDualInOther?] at h1
rw [Fin.find_eq_some_iff] at h1
simp only [Fin.coe_cast, Fin.coe_natAdd, add_le_add_iff_left, Fin.val_fin_le, ge_iff_le]
refine h1.2 k (by simpa using hj)
| Sum.inr k =>
simp only [appendEquiv, Equiv.trans_apply, finSumFinEquiv_apply_left, RelIso.coe_fn_toEquiv,
Fin.castOrderIso_apply, finSumFinEquiv_apply_right, ge_iff_le]
rw [Fin.le_def]
simp only [length, append_val, RelIso.coe_fn_toEquiv, Fin.castOrderIso_apply, Fin.coe_cast,
Fin.coe_castAdd, Fin.coe_natAdd]
omega
@[simp]
lemma getDual?_inr_getDualInOther?_isNone_getDual?_isSome (i : Fin l2.length)
(h : (l2.getDualInOther? l i).isNone) (hi : (l2.getDual? i).isSome) :
(l ++ l2).getDual? (appendEquiv (Sum.inr i)) = some
(appendEquiv (Sum.inr ((l2.getDual? i).get hi))) := by
rw [getDual?, Fin.find_eq_some_iff, AreDualInSelf.append_inr_inr]
apply And.intro
exact getDual?_areDualInSelf_get l2 i hi
intro j hj
obtain ⟨k, hk⟩ := appendEquiv.surjective j
subst hk
match k with
| Sum.inl k =>
simp at hj
simp only [getDualInOther?, Option.isNone_iff_eq_none, Fin.find_eq_none_iff] at h
exact False.elim (h k hj)
| Sum.inr k =>
simp [appendEquiv, IndexList.length]
rw [Fin.le_def]
simp only [Fin.coe_cast, Fin.coe_natAdd, add_le_add_iff_left, Fin.val_fin_le]
have h2 : l2.getDual? i = some ((l2.getDual? i).get hi) := Option.eq_some_of_isSome hi
nth_rewrite 1 [getDual?] at h2
rw [Fin.find_eq_some_iff] at h2
simp only [AreDualInSelf.append_inr_inr] at hj
exact h2.2 k hj
lemma getDual?_append_inr (i : Fin l2.length) :
(l ++ l2).getDual? (appendEquiv (Sum.inr i)) =
Option.or (Option.map (appendEquiv ∘ Sum.inl) (l2.getDualInOther? l i))
(Option.map (appendEquiv ∘ Sum.inr) (l2.getDual? i)) := by
by_cases h : (l2.getDualInOther? l i).isSome
· simp_all
rw [congrArg (Option.map (appendEquiv ∘ Sum.inl)) (Option.eq_some_of_isSome h)]
rfl
by_cases ho : (l2.getDual? i).isSome
· simp_all
rw [congrArg (Option.map (appendEquiv ∘ Sum.inr)) (Option.eq_some_of_isSome ho)]
rfl
simp_all
/-!
## Properties of getDualInOther? and append
-/
variable (l3 : IndexList X)
@[simp]
lemma getDualInOther?_append_of_inl (i : Fin l.length) :
(l ++ l2).getDualInOther? l3 (appendEquiv (Sum.inl i)) = l.getDualInOther? l3 i := by
simp [getDualInOther?]
@[simp]
lemma getDualInOther?_append_of_inr (i : Fin l2.length) :
(l ++ l2).getDualInOther? l3 (appendEquiv (Sum.inr i)) = l2.getDualInOther? l3 i := by
simp [getDualInOther?]
@[simp]
lemma getDualInOther?_isSome_of_append_iff (i : Fin l.length) :
(l.getDualInOther? (l2 ++ l3) i).isSome ↔
(l.getDualInOther? l2 i).isSome (l.getDualInOther? l3 i).isSome := by
rw [getDualInOther?_isSome_iff_exists, getDualInOther?_isSome_iff_exists,
getDualInOther?_isSome_iff_exists]
refine Iff.intro (fun h => ?_) (fun h => ?_)
· obtain ⟨j, hj⟩ := h
obtain ⟨k, hk⟩ := appendEquiv.surjective j
subst hk
match k with
| Sum.inl k =>
exact Or.inl ⟨k, by simpa using hj⟩
| Sum.inr k =>
exact Or.inr ⟨k, (AreDualInOther.of_append_inr i k).mp hj⟩
· cases' h with h h <;>
obtain ⟨j, hj⟩ := h
· use appendEquiv (Sum.inl j)
exact (AreDualInOther.of_append_inl i j).mpr hj
· use appendEquiv (Sum.inr j)
exact (AreDualInOther.of_append_inr i j).mpr hj
@[simp]
lemma getDualInOther?_eq_none_of_append_iff (i : Fin l.length) :
(l.getDualInOther? (l2 ++ l3) i) = none ↔
(l.getDualInOther? l2 i) = none ∧ (l.getDualInOther? l3 i) = none := by
apply Iff.intro
· intro h
have h1 := (l.getDualInOther?_isSome_of_append_iff l2 l3 i).mpr.mt
simp only [not_or, Bool.not_eq_true, Option.not_isSome,
Option.isNone_iff_eq_none, imp_self] at h1
exact h1 h
· intro h
have h1 := (l.getDualInOther?_isSome_of_append_iff l2 l3 i).mp.mt
simp only [not_or, Bool.not_eq_true, Option.not_isSome,
Option.isNone_iff_eq_none, imp_self] at h1
exact h1 h
@[simp]
lemma getDualInOther?_of_append_of_isSome (i : Fin l.length)
(hi : (l.getDualInOther? l2 i).isSome) : l.getDualInOther? (l2 ++ l3) i =
some (appendEquiv (Sum.inl ((l.getDualInOther? l2 i).get hi))) := by
rw [getDualInOther?, Fin.find_eq_some_iff, AreDualInOther.of_append_inl]
apply And.intro
exact getDualInOther?_areDualInOther_get l l2 i hi
intro j hj
obtain ⟨k, hk⟩ := appendEquiv.surjective j
subst hk
match k with
| Sum.inl k =>
simp only [appendEquiv, Equiv.trans_apply, finSumFinEquiv_apply_left, RelIso.coe_fn_toEquiv,
Fin.castOrderIso_apply, ge_iff_le]
rw [Fin.le_def]
have h1 : l.getDualInOther? l2 i = some (((l.getDualInOther? l2 i).get hi)) :=
Option.eq_some_of_isSome hi
nth_rewrite 1 [getDualInOther?] at h1
rw [Fin.find_eq_some_iff] at h1
simp only [Fin.coe_cast, Fin.coe_natAdd, add_le_add_iff_left, Fin.val_fin_le, ge_iff_le]
refine h1.2 k ((AreDualInOther.of_append_inl i k).mp hj)
| Sum.inr k =>
simp only [appendEquiv, Equiv.trans_apply, finSumFinEquiv_apply_left, RelIso.coe_fn_toEquiv,
Fin.castOrderIso_apply, finSumFinEquiv_apply_right, ge_iff_le]
rw [Fin.le_def]
simp only [length, append_val, RelIso.coe_fn_toEquiv, Fin.castOrderIso_apply, Fin.coe_cast,
Fin.coe_castAdd, Fin.coe_natAdd]
omega
@[simp]
lemma getDualInOther?_of_append_of_isNone_isSome (i : Fin l.length)
(hi : (l.getDualInOther? l2 i) = none) (h2 : (l.getDualInOther? l3 i).isSome) :
l.getDualInOther? (l2 ++ l3) i =
some (appendEquiv (Sum.inr ((l.getDualInOther? l3 i).get h2))) := by
rw [getDualInOther?, Fin.find_eq_some_iff, AreDualInOther.of_append_inr]
apply And.intro
exact getDualInOther?_areDualInOther_get l l3 i h2
intro j hj
obtain ⟨k, hk⟩ := appendEquiv.surjective j
subst hk
match k with
| Sum.inl k =>
simp at hj
simp only [getDualInOther?, Option.isNone_iff_eq_none, Fin.find_eq_none_iff] at hi
exact False.elim (hi k hj)
| Sum.inr k =>
simp [appendEquiv, IndexList.length]
rw [Fin.le_def]
simp only [Fin.coe_cast, Fin.coe_natAdd, add_le_add_iff_left, Fin.val_fin_le]
have h1 : l.getDualInOther? l3 i = some ((l.getDualInOther? l3 i).get h2) :=
Option.eq_some_of_isSome h2
nth_rewrite 1 [getDualInOther?] at h1
rw [Fin.find_eq_some_iff] at h1
simp only [AreDualInOther.of_append_inr] at hj
exact h1.2 k hj
end IndexList
end IndexNotation

View file

@ -239,6 +239,7 @@ def toIndexList (s : IndexString X) : IndexList X :=
{val := (listCharIndexStringTolistCharIndex X s.toCharList (listCharIndexString s)).map
fun x => Index.ofCharList x.1 x.2}
/-- The formation of an index list from a string `s` statisfying `listCharIndexStringBool`. -/
def toIndexList' (s : String) (hs : listCharIndexStringBool X s.toList = true) : IndexList X :=
toIndexList ⟨s, hs⟩

View file

@ -0,0 +1,293 @@
/-
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.WithUniqueDual
import Mathlib.Algebra.Order.Ring.Nat
import Mathlib.Data.Finset.Sort
/-!
# withDuals equal to withUniqueDuals
-/
namespace IndexNotation
namespace IndexList
variable {X : Type} [IndexNotation X] [Fintype X] [DecidableEq X]
variable (l l2 l3 : IndexList X)
/-!
## withDual equal to withUniqueDual
-/
lemma withUnqiueDual_eq_withDual_iff_unique_forall :
l.withUniqueDual = l.withDual ↔
∀ (i : l.withDual) j, l.AreDualInSelf i j → j = l.getDual? i := by
apply Iff.intro
· intro h i j hj
rw [@Finset.ext_iff] at h
simp [withUniqueDual] at h
refine h i ?_ j hj
exact withDual_isSome l i
· intro h
apply Finset.ext
intro i
apply Iff.intro
· exact fun hi => mem_withDual_of_mem_withUniqueDual l i hi
· intro hi
simp [withUniqueDual]
apply And.intro
exact (mem_withDual_iff_isSome l i).mp hi
intro j hj
exact h ⟨i, hi⟩ j hj
lemma withUnqiueDual_eq_withDual_iff :
l.withUniqueDual = l.withDual ↔
∀ i, (l.getDual? i).bind l.getDual? = Option.guard (fun i => i ∈ l.withDual) i := by
apply Iff.intro
· intro h i
by_cases hi : i ∈ l.withDual
· have hii : i ∈ l.withUniqueDual := by
simp_all only
change (l.getDual? i).bind l.getDual? = _
simp [hii]
symm
rw [Option.guard_eq_some]
exact ⟨rfl, mem_withUniqueDual_isSome l i hii⟩
· simp at hi
simp [Option.guard, hi]
· intro h
rw [withUnqiueDual_eq_withDual_iff_unique_forall]
intro i j hj
rcases l.getDual?_of_areDualInSelf hj with hi | hi | hi
· have hj' := h j
rw [hi] at hj'
simp at hj'
rw [hj']
symm
rw [Option.guard_eq_some, hi]
exact ⟨rfl, rfl⟩
· exact hi.symm
· have hj' := h j
rw [hi] at hj'
rw [h i] at hj'
have hi : Option.guard (fun i => i ∈ l.withDual) ↑i = some i := by
apply Option.guard_eq_some.mpr
simp
rw [hi] at hj'
simp at hj'
have hj'' := Option.guard_eq_some.mp hj'.symm
have hj''' := hj''.1
rw [hj'''] at hj
simp at hj
lemma withUnqiueDual_eq_withDual_iff_list_apply :
l.withUniqueDual = l.withDual ↔
(Fin.list l.length).map (fun i => (l.getDual? i).bind l.getDual?) =
(Fin.list l.length).map (fun i => Option.guard (fun i => i ∈ l.withDual) i) := by
rw [withUnqiueDual_eq_withDual_iff]
apply Iff.intro
exact fun h => List.map_eq_map_iff.mpr fun a _ => h a
intro h
intro i
simp only [List.map_inj_left] at h
have h1 {n : } (m : Fin n) : m ∈ Fin.list n := by
have h1' : (Fin.list n)[m] = m := Fin.getElem_list _ _
exact h1' ▸ List.getElem_mem _ _ _
exact h i (h1 i)
/-- A boolean which is true for an index list `l` if for every index in `l` with a dual,
that dual is unique. -/
def withUnqiueDualEqWithDualBool : Bool :=
if (Fin.list l.length).map (fun i => (l.getDual? i).bind l.getDual?) =
(Fin.list l.length).map (fun i => Option.guard (fun i => i ∈ l.withDual) i) then
true
else
false
lemma withUnqiueDual_eq_withDual_iff_list_apply_bool :
l.withUniqueDual = l.withDual ↔ l.withUnqiueDualEqWithDualBool := by
rw [withUnqiueDual_eq_withDual_iff_list_apply]
apply Iff.intro
intro h
simp [withUnqiueDualEqWithDualBool, h]
intro h
simpa [withUnqiueDualEqWithDualBool] using h
@[simp]
lemma withUnqiueDual_eq_withDual_of_empty (h : l.withDual = ∅) :
l.withUniqueDual = l.withDual := by
rw [h, Finset.eq_empty_iff_forall_not_mem]
intro x
by_contra hx
have x' : l.withDual := ⟨x, l.mem_withDual_of_withUniqueDual ⟨x, hx⟩⟩
have hx' := x'.2
simp [h] at hx'
/-!
## withUniqueDualInOther equal to withDualInOther append conditions
-/
lemma withUniqueDualInOther_eq_withDualInOther_append_of_symm'
(h : (l ++ l2).withUniqueDualInOther l3 = (l ++ l2).withDualInOther l3) :
(l2 ++ l).withUniqueDualInOther l3 = (l2 ++ l).withDualInOther l3 := by
rw [Finset.ext_iff] at h ⊢
intro j
obtain ⟨k, hk⟩ := appendEquiv.surjective j
subst hk
match k with
| Sum.inl k =>
rw [mem_append_withUniqueDualInOther_symm]
simpa using h (appendEquiv (Sum.inr k))
| Sum.inr k =>
rw [← mem_append_withUniqueDualInOther_symm]
simpa using h (appendEquiv (Sum.inl k))
lemma withUniqueDualInOther_eq_withDualInOther_append_of_symm :
(l ++ l2).withUniqueDualInOther l3 = (l ++ l2).withDualInOther l3 ↔
(l2 ++ l).withUniqueDualInOther l3 = (l2 ++ l).withDualInOther l3 := by
apply Iff.intro
exact l.withUniqueDualInOther_eq_withDualInOther_append_of_symm' l2 l3
exact l2.withUniqueDualInOther_eq_withDualInOther_append_of_symm' l l3
lemma withUniqueDualInOther_eq_withDualInOther_of_append_symm'
(h : l.withUniqueDualInOther (l2 ++ l3) = l.withDualInOther (l2 ++ l3)) :
l.withUniqueDualInOther (l3 ++ l2) = l.withDualInOther (l3 ++ l2) := by
rw [Finset.ext_iff] at h ⊢
intro j
rw [mem_withUniqueDualInOther_symm]
rw [h j]
simp only [mem_withInDualOther_iff_isSome, getDualInOther?_isSome_of_append_iff]
exact Or.comm
lemma withUniqueDualInOther_eq_withDualInOther_of_append_symm :
l.withUniqueDualInOther (l2 ++ l3) = l.withDualInOther (l2 ++ l3) ↔
l.withUniqueDualInOther (l3 ++ l2) = l.withDualInOther (l3 ++ l2) := by
apply Iff.intro
exact l.withUniqueDualInOther_eq_withDualInOther_of_append_symm' l2 l3
exact l.withUniqueDualInOther_eq_withDualInOther_of_append_symm' l3 l2
/-!
## withDual equal to withUniqueDual append conditions
-/
lemma append_withDual_eq_withUniqueDual_iff :
(l ++ l2).withUniqueDual = (l ++ l2).withDual ↔
((l.withUniqueDual ∩ (l.withDualInOther l2)ᶜ) l.withUniqueDualInOther l2)
= l.withDual l.withDualInOther l2
∧ (l2.withUniqueDual ∩ (l2.withDualInOther l)ᶜ) l2.withUniqueDualInOther l
= l2.withDual l2.withDualInOther l := by
rw [append_withUniqueDual_disjSum, withDual_append_eq_disjSum]
simp only [Equiv.finsetCongr_apply, Finset.map_inj]
have h (s s' : Finset (Fin l.length)) (t t' : Finset (Fin l2.length)) :
s.disjSum t = s'.disjSum t' ↔ s = s' ∧ t = t' := by
simp [Finset.ext_iff]
exact h _ _ _ _
lemma append_withDual_eq_withUniqueDual_symm :
(l ++ l2).withUniqueDual = (l ++ l2).withDual ↔
(l2 ++ l).withUniqueDual = (l2 ++ l).withDual := by
rw [append_withDual_eq_withUniqueDual_iff, append_withDual_eq_withUniqueDual_iff]
exact And.comm
@[simp]
lemma append_withDual_eq_withUniqueDual_inl (h : (l ++ l2).withUniqueDual = (l ++ l2).withDual) :
l.withUniqueDual = l.withDual := by
rw [Finset.ext_iff]
intro i
refine Iff.intro (fun h' => ?_) (fun h' => ?_)
· exact mem_withDual_of_mem_withUniqueDual l i h'
· have hn : appendEquiv (Sum.inl i) ∈ (l ++ l2).withUniqueDual := by
rw [h]
simp_all
refine l.mem_withUniqueDual_of_inl l2 i hn ?_
exact (mem_withDual_iff_isSome l i).mp h'
@[simp]
lemma append_withDual_eq_withUniqueDual_inr (h : (l ++ l2).withUniqueDual = (l ++ l2).withDual) :
l2.withUniqueDual = l2.withDual := by
rw [append_withDual_eq_withUniqueDual_symm] at h
exact append_withDual_eq_withUniqueDual_inl l2 l h
@[simp]
lemma append_withDual_eq_withUniqueDual_withUniqueDualInOther_inl
(h : (l ++ l2).withUniqueDual = (l ++ l2).withDual) :
l.withUniqueDualInOther l2 = l.withDualInOther l2 := by
rw [Finset.ext_iff]
intro i
refine Iff.intro (fun h' => ?_) (fun h' => ?_)
· simp [h']
· have hn : appendEquiv (Sum.inl i) ∈ (l ++ l2).withUniqueDual := by
rw [h]
simp_all
refine l.mem_withUniqueDualInOther_of_inl_withDualInOther l2 i hn ?_
exact (mem_withInDualOther_iff_isSome l l2 i).mp h'
@[simp]
lemma append_withDual_eq_withUniqueDual_withUniqueDualInOther_inr
(h : (l ++ l2).withUniqueDual = (l ++ l2).withDual) :
l2.withUniqueDualInOther l = l2.withDualInOther l := by
rw [append_withDual_eq_withUniqueDual_symm] at h
exact append_withDual_eq_withUniqueDual_withUniqueDualInOther_inl l2 l h
lemma append_withDual_eq_withUniqueDual_iff' :
(l ++ l2).withUniqueDual = (l ++ l2).withDual ↔
l.withUniqueDual = l.withDual ∧ l2.withUniqueDual = l2.withDual
∧ l.withUniqueDualInOther l2 = l.withDualInOther l2 ∧
l2.withUniqueDualInOther l = l2.withDualInOther l := by
apply Iff.intro
intro h
exact ⟨append_withDual_eq_withUniqueDual_inl l l2 h, append_withDual_eq_withUniqueDual_inr l l2 h,
append_withDual_eq_withUniqueDual_withUniqueDualInOther_inl l l2 h,
append_withDual_eq_withUniqueDual_withUniqueDualInOther_inr l l2 h⟩
intro h
rw [append_withDual_eq_withUniqueDual_iff]
rw [h.1, h.2.1, h.2.2.1, h.2.2.2]
have h1 : l.withDual ∩ (l.withDualInOther l2)ᶜ = l.withDual := by
rw [Finset.inter_eq_left]
rw [Finset.subset_iff]
rw [← h.1, ← h.2.2.1]
intro i hi
simp only [withUniqueDualInOther, mem_withDual_iff_isSome, Bool.not_eq_true, Option.not_isSome,
Option.isNone_iff_eq_none, mem_withInDualOther_iff_isSome, Finset.compl_filter, not_and,
not_forall, Classical.not_imp, Finset.mem_filter, Finset.mem_univ, true_and]
intro hn
simp_all
have h2 : l2.withDual ∩ (l2.withDualInOther l)ᶜ = l2.withDual := by
rw [Finset.inter_eq_left]
rw [Finset.subset_iff]
rw [← h.2.1, ← h.2.2.2]
intro i hi
simp only [withUniqueDualInOther, mem_withDual_iff_isSome, Bool.not_eq_true, Option.not_isSome,
Option.isNone_iff_eq_none, mem_withInDualOther_iff_isSome, Finset.compl_filter, not_and,
not_forall, Classical.not_imp, Finset.mem_filter, Finset.mem_univ, true_and]
intro hn
simp_all
exact ⟨congrFun (congrArg Union.union h1) (l.withDualInOther l2),
congrFun (congrArg Union.union h2) (l2.withDualInOther l)⟩
lemma append_withDual_eq_withUniqueDual_swap :
(l ++ l2 ++ l3).withUniqueDual = (l ++ l2 ++ l3).withDual
↔ (l2 ++ l ++ l3).withUniqueDual = (l2 ++ l ++ l3).withDual := by
rw [append_withDual_eq_withUniqueDual_iff']
rw [append_withDual_eq_withUniqueDual_iff' (l2 ++ l) l3]
rw [append_withDual_eq_withUniqueDual_symm]
rw [withUniqueDualInOther_eq_withDualInOther_of_append_symm]
rw [withUniqueDualInOther_eq_withDualInOther_append_of_symm]
end IndexList
end IndexNotation

View file

@ -31,6 +31,11 @@ open IndexList TensorColor
-/
/-- Two `ColorIndexList` are said to be reindexes of one another if they:
(1) have the same length and (2) every corresponding index has the same color,
and duals which correspond.
Note: This does not allow for reordrings of indices. -/
def Reindexing : Prop := l.length = l'.length ∧
∀ (h : l.length = l'.length), l.colorMap = l'.colorMap ∘ Fin.cast h ∧
Option.map (Fin.cast h) ∘ l.getDual? = l'.getDual? ∘ Fin.cast h
@ -44,6 +49,10 @@ To prevent choice problems, this has to be done after contraction.
-/
/-- Two `ColorIndexList` are said to be related by contracted permutations, `ContrPerm`,
if and only if: 1) Their contractions are the same length.
2) Every index in the contracted list of one has a unqiue dual in the contracted
list of the other and that dual has a the same color. -/
def ContrPerm : Prop :=
l.contr.length = l'.contr.length ∧
l.contr.withUniqueDualInOther l'.contr = Finset.univ ∧
@ -76,7 +85,7 @@ 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
simp only [Function.comp_apply]
have h1' := congrFun h1.2.2 ⟨i, by simp [h1.2.1]⟩
simp at h1'
rw [← h1']
@ -91,15 +100,14 @@ lemma trans (h1 : ContrPerm l l2) (h2 : ContrPerm l2 l3) : ContrPerm l l3 := by
rw [h2.2.1]
simp
@[simp]
lemma symm_trans (h1 : ContrPerm l l2) (h2 : ContrPerm l2 l3) :
(h1.trans h2).symm = h2.symm.trans h1.symm := by
rfl
simp only
@[simp]
lemma contr_self : ContrPerm l l.contr := by
rw [ContrPerm]
simp
simp only [contr_contr, true_and]
have h1 := @refl _ _ l
apply And.intro h1.2.1
erw [contr_contr]
@ -112,6 +120,9 @@ lemma self_contr : ContrPerm l.contr l := by
end ContrPerm
/-- Given two `ColorIndexList` related by contracted permutations, the equivalence between
the indices of the corresponding contracted index lists. This equivalence is the
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 <|
@ -144,7 +155,7 @@ lemma contrPermEquiv_refl : contrPermEquiv (@ContrPerm.refl 𝓒 _ l) = Equiv.re
@[simp]
lemma contrPermEquiv_symm {l l' : ColorIndexList 𝓒} (h : ContrPerm l l') :
(contrPermEquiv h).symm = contrPermEquiv h.symm := by
(contrPermEquiv h).symm = contrPermEquiv h.symm := by
simp only [contrPermEquiv]
rfl

View file

@ -131,8 +131,8 @@ lemma contr_of_withDual_empty (T : 𝓣.TensorIndex) (h : T.withDual = ∅) :
intro a b hx hy
simp [map_add, add_mul, hx, hy])
intro r f
simp only [PiTensorProduct.tprodCoeff_eq_smul_tprod, LinearMapClass.map_smul, mapIso_tprod, id_eq,
eq_mpr_eq_cast, OrderIso.toEquiv_symm, RelIso.coe_fn_toEquiv]
simp only [PiTensorProduct.tprodCoeff_eq_smul_tprod, LinearMapClass.map_smul, mapIso_tprod,
id_eq, eq_mpr_eq_cast, OrderIso.toEquiv_symm, RelIso.coe_fn_toEquiv]
apply congrArg
have hEm : IsEmpty { x // x ∈ i.withUniqueDualLT } := by
rw [Finset.isEmpty_coe_sort]
@ -161,10 +161,12 @@ lemma contr_contr (T : 𝓣.TensorIndex) : T.contr.contr = T.contr :=
T.contr.contr_of_withDual_empty (by simp [contr, ColorIndexList.contr])
@[simp]
lemma contr_toColorIndexList (T : 𝓣.TensorIndex) : T.contr.toColorIndexList = T.toColorIndexList.contr := rfl
lemma contr_toColorIndexList (T : 𝓣.TensorIndex) :
T.contr.toColorIndexList = T.toColorIndexList.contr := rfl
@[simp]
lemma contr_toIndexList (T : 𝓣.TensorIndex) : T.contr.toIndexList = T.toIndexList.contrIndexList := rfl
lemma contr_toIndexList (T : 𝓣.TensorIndex) :
T.contr.toIndexList = T.toIndexList.contrIndexList := by
rfl
/-!
## Scalar multiplication of
@ -210,7 +212,7 @@ namespace Rel
/-- Rel is reflexive. -/
lemma refl (T : 𝓣.TensorIndex) : Rel T T := by
apply And.intro
simp
simp only [ContrPerm.refl]
simp
/-- Rel is symmetric. -/
@ -231,7 +233,7 @@ lemma trans {T₁ T₂ T₃ : 𝓣.TensorIndex} (h1 : Rel T₁ T₂) (h2 : Rel T
intro h
change _ = (𝓣.mapIso (contrPermEquiv (h1.1.trans h2.1)).symm _) T₃.contr.tensor
trans (𝓣.mapIso ((contrPermEquiv h1.1).trans (contrPermEquiv h2.1)).symm (by
simp
simp only [contrPermEquiv_trans, contrPermEquiv_symm, contr_toColorIndexList]
have h1 := contrPermEquiv_colorMap_iso (ContrPerm.symm (ContrPerm.trans h1.left h2.left))
rw [← ColorMap.MapIso.symm'] at h1
exact h1)) T₃.contr.tensor
@ -252,7 +254,8 @@ lemma isEquivalence : Equivalence (@Rel _ _ 𝓣 _) where
/-- The equality of tensors corresponding to related tensor indices. -/
lemma to_eq {T₁ T₂ : 𝓣.TensorIndex} (h : Rel T₁ T₂) :
T₁.contr.tensor = 𝓣.mapIso (contrPermEquiv h.1).symm (contrPermEquiv_colorMap_iso h.1) T₂.contr.tensor := h.2 h.1
T₁.contr.tensor = 𝓣.mapIso (contrPermEquiv h.1).symm
(contrPermEquiv_colorMap_iso h.1) T₂.contr.tensor := h.2 h.1
end Rel
@ -262,7 +265,7 @@ instance asSetoid : Setoid 𝓣.TensorIndex := ⟨Rel, Rel.isEquivalence⟩
/-- A tensor index is equivalent to its contraction. -/
lemma rel_contr (T : 𝓣.TensorIndex) : T ≈ T.contr := by
apply And.intro
simp
simp only [contr_toColorIndexList, ContrPerm.contr_self]
intro h
rw [tensor_eq_of_eq T.contr_contr]
simp only [contr_toColorIndexList, colorMap', contrPermEquiv_self_contr, OrderIso.toEquiv_symm,
@ -390,7 +393,7 @@ lemma add_toColorIndexList (T₁ T₂ : 𝓣.TensorIndex) (h : AddCond T₁ T₂
@[simp]
lemma add_tensor (T₁ T₂ : 𝓣.TensorIndex) (h : AddCond T₁ T₂) :
(add T₁ T₂ h).tensor =
(𝓣.mapIso (addCondEquiv h) (addCondEquiv_colorMap h) T₁.contr.tensor) + T₂.contr.tensor := by rfl
(𝓣.mapIso (addCondEquiv h) (addCondEquiv_colorMap h) T₁.contr.tensor) + T₂.contr.tensor := by rfl
/-- Scalar multiplication commutes with addition. -/
lemma smul_add (r : R) (T₁ T₂ : 𝓣.TensorIndex) (h : AddCond T₁ T₂) :
@ -483,6 +486,9 @@ lemma add_assoc {T₁ T₂ T₃ : 𝓣.TensorIndex} {h' : AddCond T₁ T₂} (h
-/
/-- The condition on two `TensorIndex` which is true if and only if their `ColorIndexList`s
are related by the condition `AppendCond`. That is, they can be appended to form a
`ColorIndexList`. -/
def ProdCond (T₁ T₂ : 𝓣.TensorIndex) : Prop :=
AppendCond T₁.toColorIndexList T₂.toColorIndexList

View file

@ -0,0 +1,126 @@
/-
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.GetDual
import Mathlib.Algebra.Order.Ring.Nat
import Mathlib.Data.Finset.Sort
/-!
# With dual
We define the finite sets of indices in an index list which have a dual
-/
namespace IndexNotation
namespace IndexList
variable {X : Type} [IndexNotation X] [Fintype X] [DecidableEq X]
variable (l l2 l3 : IndexList X)
/-!
## Finsets on which getDual? and getDualInOther? are some.
-/
/-- The finite set of indices of an index list which have a dual in that index list. -/
def withDual : Finset (Fin l.length) :=
Finset.filter (fun i => (l.getDual? i).isSome) Finset.univ
/-- The finite set of indices of an index list which have a dual in another provided index list. -/
def withDualInOther : Finset (Fin l.length) :=
Finset.filter (fun i => (l.getDualInOther? l2 i).isSome) Finset.univ
/-!
## Basic properties of withDual
-/
@[simp]
lemma withDual_isSome (i : l.withDual) : (l.getDual? i).isSome := by
simpa [withDual] using i.2
@[simp]
lemma mem_withDual_iff_isSome (i : Fin l.length) : i ∈ l.withDual ↔ (l.getDual? i).isSome := by
simp [withDual]
lemma not_mem_withDual_iff_isNone (i : Fin l.length) :
i ∉ l.withDual ↔ (l.getDual? i).isNone := by
simp only [mem_withDual_iff_isSome, Bool.not_eq_true, Option.not_isSome,
Option.isNone_iff_eq_none]
lemma mem_withDual_iff_exists : i ∈ l.withDual ↔ ∃ j, l.AreDualInSelf i j := by
simp [withDual, Finset.mem_filter, Finset.mem_univ, getDual?]
exact Fin.isSome_find_iff
/-!
## Basic properties of withDualInOther
-/
@[simp]
lemma mem_withInDualOther_iff_isSome (i : Fin l.length) :
i ∈ l.withDualInOther l2 ↔ (l.getDualInOther? l2 i).isSome := by
simp only [withDualInOther, getDualInOther?, Finset.mem_filter, Finset.mem_univ, true_and]
lemma mem_withInDualOther_iff_exists :
i ∈ l.withDualInOther l2 ↔ ∃ (j : Fin l2.length), l.AreDualInOther l2 i j := by
simp [withDualInOther, Finset.mem_filter, Finset.mem_univ, getDualInOther?]
exact Fin.isSome_find_iff
/-!
## Append properties of withDual
-/
lemma withDual_append_eq_disjSum : (l ++ l2).withDual =
Equiv.finsetCongr appendEquiv
((l.withDual l.withDualInOther l2).disjSum
(l2.withDual l2.withDualInOther l)) := by
ext i
obtain ⟨k, hk⟩ := appendEquiv.surjective i
subst hk
match k with
| Sum.inl k =>
simp
| Sum.inr k =>
simp
/-!
## Append properties of withDualInOther
-/
lemma append_withDualInOther_eq :
(l ++ l2).withDualInOther l3 =
Equiv.finsetCongr appendEquiv ((l.withDualInOther l3).disjSum (l2.withDualInOther l3)) := by
rw [Finset.ext_iff]
intro i
obtain ⟨k, hk⟩ := appendEquiv.surjective i
subst hk
match k with
| Sum.inl k =>
simp
| Sum.inr k =>
simp
lemma withDualInOther_append_eq : l.withDualInOther (l2 ++ l3) =
l.withDualInOther l2 l.withDualInOther l3 := by
ext i
simp
end IndexList
end IndexNotation

View file

@ -0,0 +1,854 @@
/-
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.WithDual
import Mathlib.Algebra.Order.Ring.Nat
import Mathlib.Data.Finset.Sort
/-!
# With Unique Dual
Finite sets of indices with a unique dual.
-/
namespace IndexNotation
namespace IndexList
variable {X : Type} [IndexNotation X] [Fintype X] [DecidableEq X]
variable (l l2 l3 : IndexList X)
/-!
## Unique duals
-/
/-- The finite set of indices of an index list which have a unique dual in that index list. -/
def withUniqueDual : Finset (Fin l.length) :=
Finset.filter (fun i => i ∈ l.withDual ∧
∀ j, l.AreDualInSelf i j → j = l.getDual? i) Finset.univ
/-- 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) :=
Finset.filter (fun i => i ∉ l.withDual ∧ i ∈ l.withDualInOther l2
∧ (∀ j, l.AreDualInOther l2 i j → j = l.getDualInOther? l2 i)) Finset.univ
/-!
## Basic properties of withUniqueDual
-/
@[simp, nolint simpNF]
lemma mem_withDual_of_mem_withUniqueDual (i : Fin l.length) (h : i ∈ l.withUniqueDual) :
i ∈ l.withDual := by
simp only [withUniqueDual, mem_withDual_iff_isSome, Finset.mem_filter, Finset.mem_univ,
true_and] at h
simpa using h.1
@[simp]
lemma mem_withUniqueDual_isSome (i : Fin l.length) (h : i ∈ l.withUniqueDual) :
(l.getDual? i).isSome := by
simpa using mem_withDual_of_mem_withUniqueDual l i h
lemma mem_withDual_of_withUniqueDual (i : l.withUniqueDual) :
i.1 ∈ l.withDual :=
mem_withDual_of_mem_withUniqueDual l (↑i) i.2
/-!
## Basic properties of withUniqueDualInOther
-/
lemma not_mem_withDual_of_withUniqueDualInOther (i : l.withUniqueDualInOther l2) :
i.1 ∉ l.withDual := by
have hi := i.2
simp only [withUniqueDualInOther, Finset.univ_eq_attach, Finset.mem_filter, Finset.mem_attach,
true_and] at hi
exact hi.2.1
lemma mem_withDualInOther_of_withUniqueDualInOther (i : l.withUniqueDualInOther l2) :
i.1 ∈ l.withDualInOther l2 := by
have hi := i.2
simp only [withUniqueDualInOther, Finset.univ_eq_attach, Finset.mem_filter, Finset.mem_attach,
true_and] at hi
exact hi.2.2.1
@[simp]
lemma mem_withUniqueDualInOther_isSome (i : Fin l.length) (hi : i ∈ l.withUniqueDualInOther l2) :
(l.getDualInOther? l2 i).isSome := by
simp only [withUniqueDualInOther, Finset.mem_filter, Finset.mem_univ, true_and] at hi
exact (mem_withInDualOther_iff_isSome l l2 i).mp hi.2.1
/-!
## Properties of getDual? and withUniqueDual
-/
lemma all_dual_eq_getDual?_of_mem_withUniqueDual (i : Fin l.length) (h : i ∈ l.withUniqueDual) :
∀ j, l.AreDualInSelf i j → j = l.getDual? i := by
simp [withUniqueDual] at h
exact fun j hj => h.2 j hj
lemma some_eq_getDual?_of_withUniqueDual_iff (i j : Fin l.length) (h : i ∈ l.withUniqueDual) :
l.AreDualInSelf i j ↔ some j = l.getDual? i := by
apply Iff.intro
exact fun h' => all_dual_eq_getDual?_of_mem_withUniqueDual l i h j h'
intro h'
have hj : j = (l.getDual? i).get (mem_withUniqueDual_isSome l i h) :=
Eq.symm (Option.get_of_mem (mem_withUniqueDual_isSome l i h) (id (Eq.symm h')))
subst hj
exact (getDual?_get_areDualInSelf l i (mem_withUniqueDual_isSome l i h)).symm
@[simp]
lemma eq_getDual?_get_of_withUniqueDual_iff (i j : Fin l.length) (h : i ∈ l.withUniqueDual) :
l.AreDualInSelf i j ↔ j = (l.getDual? i).get (mem_withUniqueDual_isSome l i h) := by
rw [l.some_eq_getDual?_of_withUniqueDual_iff i j h]
refine Iff.intro (fun h' => ?_) (fun h' => ?_)
exact Eq.symm (Option.get_of_mem (mem_withUniqueDual_isSome l i h) (id (Eq.symm h')))
simp [h']
lemma eq_of_areDualInSelf_withUniqueDual {j k : Fin l.length} (i : l.withUniqueDual)
(hj : l.AreDualInSelf i j) (hk : l.AreDualInSelf i k) : j = k := by
simp at hj hk
exact hj.trans hk.symm
@[simp, nolint simpNF]
lemma getDual?_get_getDual?_get_of_withUniqueDual (i : Fin l.length) (h : i ∈ l.withUniqueDual) :
(l.getDual? ((l.getDual? i).get (mem_withUniqueDual_isSome l i h))).get
(l.getDual?_getDual?_get_isSome i (mem_withUniqueDual_isSome l i h)) = i := by
by_contra hn
have h' : l.AreDualInSelf i ((l.getDual? ((l.getDual? i).get
(mem_withUniqueDual_isSome l i h))).get (
getDual?_getDual?_get_isSome l i (mem_withUniqueDual_isSome l i h))) := by
simp [AreDualInSelf, hn]
exact fun a => hn (id (Eq.symm a))
simp [h] at h'
@[simp, nolint simpNF]
lemma getDual?_getDual?_get_of_withUniqueDual (i : Fin l.length) (h : i ∈ l.withUniqueDual) :
l.getDual? ((l.getDual? i).get (mem_withUniqueDual_isSome l i h)) = some i := by
nth_rewrite 3 [← l.getDual?_get_getDual?_get_of_withUniqueDual i h]
exact Option.eq_some_of_isSome
(getDual?_getDual?_get_isSome l i (mem_withUniqueDual_isSome l i h))
@[simp]
lemma getDual?_getDual?_of_withUniqueDual (i : Fin l.length) (h : i ∈ l.withUniqueDual) :
(l.getDual? i).bind l.getDual? = some i := by
have h1 : (l.getDual? i) = some ((l.getDual? i).get (mem_withUniqueDual_isSome l i h)) :=
Option.eq_some_of_isSome (mem_withUniqueDual_isSome l i h)
nth_rewrite 1 [h1]
rw [Option.some_bind']
exact getDual?_getDual?_get_of_withUniqueDual l i h
@[simp, nolint simpNF]
lemma getDual?_get_of_mem_withUnique_mem (i : Fin l.length) (h : i ∈ l.withUniqueDual) :
(l.getDual? i).get (l.mem_withUniqueDual_isSome i h) ∈ l.withUniqueDual := by
simp only [withUniqueDual, mem_withDual_iff_isSome, Finset.mem_filter, Finset.mem_univ,
getDual?_getDual?_get_isSome, true_and]
intro j hj
have h1 : i = j := by
by_contra hn
have h' : l.AreDualInSelf i j := by
simp only [AreDualInSelf, ne_eq, hn, not_false_eq_true, true_and]
simp_all only [AreDualInSelf, ne_eq, getDual?_get_id]
simp [h] at h'
subst h'
simp_all
subst h1
exact Eq.symm (getDual?_getDual?_get_of_withUniqueDual l i h)
/-!
## Properties of getDualInOther? and withUniqueDualInOther
-/
lemma all_dual_eq_of_withUniqueDualInOther (i : Fin l.length)
(h : i ∈ l.withUniqueDualInOther l2) :
∀ j, l.AreDualInOther l2 i j → j = l.getDualInOther? l2 i := by
simp [withUniqueDualInOther] at h
exact fun j hj => h.2.2 j hj
lemma some_eq_getDualInOther?_of_withUniqueDualInOther_iff (i : Fin l.length) (j : Fin l2.length)
(h : i ∈ l.withUniqueDualInOther l2) :
l.AreDualInOther l2 i j ↔ some j = l.getDualInOther? l2 i := by
apply Iff.intro
exact fun h' => l.all_dual_eq_of_withUniqueDualInOther l2 i h j h'
intro h'
have hj : j = (l.getDualInOther? l2 i).get (mem_withUniqueDualInOther_isSome l l2 i h) :=
Eq.symm (Option.get_of_mem (mem_withUniqueDualInOther_isSome l l2 i h) (id (Eq.symm h')))
subst hj
exact getDualInOther?_areDualInOther_get l l2 i (mem_withUniqueDualInOther_isSome l l2 i h)
@[simp]
lemma eq_getDualInOther?_get_of_withUniqueDualInOther_iff (i : Fin l.length) (j : Fin l2.length)
(h : i ∈ l.withUniqueDualInOther l2) :
l.AreDualInOther l2 i j ↔ j = (l.getDualInOther? l2 i).get
(mem_withUniqueDualInOther_isSome l l2 i h) := by
rw [l.some_eq_getDualInOther?_of_withUniqueDualInOther_iff l2 i j h]
refine Iff.intro (fun h' => ?_) (fun h' => ?_)
exact Eq.symm (Option.get_of_mem (mem_withUniqueDualInOther_isSome l l2 i h) (id (Eq.symm h')))
simp [h']
@[simp, nolint simpNF]
lemma getDualInOther?_get_getDualInOther?_get_of_withUniqueDualInOther
(i : Fin l.length) (h : i ∈ l.withUniqueDualInOther l2) :
(l2.getDualInOther? l ((l.getDualInOther? l2 i).get
(mem_withUniqueDualInOther_isSome l l2 i h))).get
(l.getDualInOther?_getDualInOther?_get_isSome l2 i
(mem_withUniqueDualInOther_isSome l l2 i h)) = i := by
by_contra hn
have h' : l.AreDualInSelf i ((l2.getDualInOther? l ((l.getDualInOther? l2 i).get
(mem_withUniqueDualInOther_isSome l l2 i h))).get
(l.getDualInOther?_getDualInOther?_get_isSome l2 i
(mem_withUniqueDualInOther_isSome l l2 i h))):= by
simp [AreDualInSelf, hn]
exact fun a => hn (id (Eq.symm a))
have h1 := l.not_mem_withDual_of_withUniqueDualInOther l2 ⟨i, h⟩
simp [getDual?] at h1
rw [Fin.find_eq_none_iff] at h1
exact h1 _ h'
@[simp, nolint simpNF]
lemma getDualInOther?_getDualInOther?_get_of_withUniqueDualInOther (i : Fin l.length)
(h : i ∈ l.withUniqueDualInOther l2) :
l2.getDualInOther? l ((l.getDualInOther? l2 i).get (mem_withUniqueDualInOther_isSome l l2 i h))
= some i := by
nth_rewrite 3 [← l.getDualInOther?_get_getDualInOther?_get_of_withUniqueDualInOther l2 i h]
exact Option.eq_some_of_isSome
(getDualInOther?_getDualInOther?_get_isSome l l2 i (mem_withUniqueDualInOther_isSome l l2 i h))
@[simp]
lemma getDualInOther?_getDualInOther?_of_withUniqueDualInOther
(i : Fin l.length) (h : i ∈ l.withUniqueDualInOther l2) :
(l.getDualInOther? l2 i).bind (l2.getDualInOther? l) = some i := by
have h1 : (l.getDualInOther? l2 i) = some ((l.getDualInOther? l2 i).get
(mem_withUniqueDualInOther_isSome l l2 i h)) :=
Option.eq_some_of_isSome (mem_withUniqueDualInOther_isSome l l2 i h)
nth_rewrite 1 [h1]
rw [Option.some_bind']
exact getDualInOther?_getDualInOther?_get_of_withUniqueDualInOther l l2 i h
lemma eq_of_areDualInOther_withUniqueDualInOther {j k : Fin l2.length}
(i : l.withUniqueDualInOther l2)
(hj : l.AreDualInOther l2 i j) (hk : l.AreDualInOther l2 i k) : j = k := by
simp at hj hk
exact hj.trans hk.symm
lemma getDual?_of_getDualInOther?_of_mem_withUniqueInOther_eq_none (i : Fin l.length)
(h : i ∈ l.withUniqueDualInOther l2) : l2.getDual?
((l.getDualInOther? l2 i).get (l.mem_withUniqueDualInOther_isSome l2 i h)) = none := by
by_contra hn
rw [← @Option.not_isSome_iff_eq_none, not_not] at hn
rw [@getDual?_isSome_iff_exists] at hn
obtain ⟨j, hj⟩ := hn
have hx : l.AreDualInOther l2 i j := by
simp [AreDualInOther, hj]
simp [AreDualInSelf] at hj
exact hj.2
have hn := l.eq_of_areDualInOther_withUniqueDualInOther l2 ⟨i, h⟩ hx
(getDualInOther?_areDualInOther_get l l2 i (mem_withUniqueDualInOther_isSome l l2 i h))
subst hn
simp_all
@[simp, nolint simpNF]
lemma getDualInOther?_get_of_mem_withUniqueInOther_mem (i : Fin l.length)
(h : i ∈ l.withUniqueDualInOther l2) :
(l.getDualInOther? l2 i).get
(l.mem_withUniqueDualInOther_isSome l2 i h) ∈ l2.withUniqueDualInOther l := by
simp only [withUniqueDualInOther, mem_withDual_iff_isSome, Bool.not_eq_true, Option.not_isSome,
Option.isNone_iff_eq_none, mem_withInDualOther_iff_isSome, Finset.mem_filter, Finset.mem_univ,
getDualInOther?_getDualInOther?_get_isSome, true_and]
apply And.intro
exact getDual?_of_getDualInOther?_of_mem_withUniqueInOther_eq_none l l2 i h
intro j hj
simp only [h, getDualInOther?_getDualInOther?_get_of_withUniqueDualInOther, Option.some.injEq]
by_contra hn
have hx : l.AreDualInSelf i j := by
simp [AreDualInSelf, hn]
simp [AreDualInOther] at hj
exact ⟨fun a => hn (id (Eq.symm a)), hj⟩
simp only [withUniqueDualInOther, mem_withDual_iff_isSome, getDual?, Bool.not_eq_true,
Option.not_isSome, Option.isNone_iff_eq_none, mem_withInDualOther_iff_isSome, Finset.mem_filter,
Finset.mem_univ, true_and] at h
rw [Fin.find_eq_none_iff] at h
simp_all only
@[simp]
lemma getDualInOther?_self_of_mem_withUniqueInOther (i : Fin l.length)
(h : i ∈ l.withUniqueDualInOther l) :
l.getDualInOther? l i = some i := by
rw [all_dual_eq_of_withUniqueDualInOther l l i h i rfl]
/-!
## Properties of getDual?, withUniqueDual and append
-/
lemma append_inl_not_mem_withDual_of_withDualInOther (i : Fin l.length)
(h : appendEquiv (Sum.inl i) ∈ (l ++ l2).withUniqueDual) :
i ∈ l.withDual ↔ ¬ i ∈ l.withDualInOther l2 := by
refine Iff.intro (fun hs => ?_) (fun ho => ?_)
· by_contra ho
obtain ⟨j, hj⟩ := (l.mem_withDual_iff_exists).mp hs
obtain ⟨k, hk⟩ := (l.mem_withInDualOther_iff_exists l2).mp ho
have h1 : appendEquiv (Sum.inl j) = appendEquiv (Sum.inr k) := by
refine (l ++ l2).eq_of_areDualInSelf_withUniqueDual ⟨appendEquiv (Sum.inl i), h⟩ ?_ ?_
exact AreDualInSelf.append_inl_inl.mpr hj
simpa using hk
simp at h1
· have ht : ((l ++ l2).getDual? (appendEquiv (Sum.inl i))).isSome :=
mem_withUniqueDual_isSome (l ++ l2) (appendEquiv (Sum.inl i)) h
simp only [getDual?_isSome_append_inl_iff] at ht
simp_all
lemma append_inr_not_mem_withDual_of_withDualInOther (i : Fin l2.length)
(h : appendEquiv (Sum.inr i) ∈ (l ++ l2).withUniqueDual) :
i ∈ l2.withDual ↔ ¬ i ∈ l2.withDualInOther l := by
refine Iff.intro (fun hs => ?_) (fun ho => ?_)
· by_contra ho
obtain ⟨j, hj⟩ := (l2.mem_withDual_iff_exists).mp hs
obtain ⟨k, hk⟩ := (l2.mem_withInDualOther_iff_exists l).mp ho
have h1 : appendEquiv (Sum.inr j) = appendEquiv (Sum.inl k) := by
refine (l ++ l2).eq_of_areDualInSelf_withUniqueDual ⟨appendEquiv (Sum.inr i), h⟩ ?_ ?_
exact (AreDualInSelf.append_inr_inr l l2 i j).mpr hj
simpa using hk
simp at h1
· have ht : ((l ++ l2).getDual? (appendEquiv (Sum.inr i))).isSome :=
mem_withUniqueDual_isSome (l ++ l2) (appendEquiv (Sum.inr i)) h
simp only [getDual?_isSome_append_inr_iff] at ht
simp_all
lemma getDual?_append_symm_of_withUniqueDual_of_inl (i : Fin l.length) (k : Fin l2.length)
(h : appendEquiv (Sum.inl i) ∈ (l ++ l2).withUniqueDual) :
(l ++ l2).getDual? (appendEquiv (Sum.inl i)) = some (appendEquiv (Sum.inr k))
↔ (l2 ++ l).getDual? (appendEquiv (Sum.inr i)) = some (appendEquiv (Sum.inl k)) := by
have h := l.append_inl_not_mem_withDual_of_withDualInOther l2 i h
by_cases hs : (l.getDual? i).isSome
<;> by_cases ho : (l.getDualInOther? l2 i).isSome
<;> simp_all [hs, ho]
lemma getDual?_append_symm_of_withUniqueDual_of_inl' (i k : Fin l.length)
(h : appendEquiv (Sum.inl i) ∈ (l ++ l2).withUniqueDual) :
(l ++ l2).getDual? (appendEquiv (Sum.inl i)) = some (appendEquiv (Sum.inl k))
↔ (l2 ++ l).getDual? (appendEquiv (Sum.inr i)) = some (appendEquiv (Sum.inr k)) := by
have h := l.append_inl_not_mem_withDual_of_withDualInOther l2 i h
by_cases hs : (l.getDual? i).isSome
<;> by_cases ho : (l.getDualInOther? l2 i).isSome
<;> simp_all [hs, ho]
lemma getDual?_append_symm_of_withUniqueDual_of_inr (i : Fin l2.length) (k : Fin l.length)
(h : appendEquiv (Sum.inr i) ∈ (l ++ l2).withUniqueDual) :
(l ++ l2).getDual? (appendEquiv (Sum.inr i)) = some (appendEquiv (Sum.inl k))
↔ (l2 ++ l).getDual? (appendEquiv (Sum.inl i)) = some (appendEquiv (Sum.inr k)) := by
have h := l.append_inr_not_mem_withDual_of_withDualInOther l2 i h
by_cases hs : (l2.getDual? i).isSome
<;> by_cases ho : (l2.getDualInOther? l i).isSome
<;> simp_all [hs, ho]
lemma getDual?_append_symm_of_withUniqueDual_of_inr' (i k : Fin l2.length)
(h : appendEquiv (Sum.inr i) ∈ (l ++ l2).withUniqueDual) :
(l ++ l2).getDual? (appendEquiv (Sum.inr i)) = some (appendEquiv (Sum.inr k))
↔ (l2 ++ l).getDual? (appendEquiv (Sum.inl i)) = some (appendEquiv (Sum.inl k)) := by
have h := l.append_inr_not_mem_withDual_of_withDualInOther l2 i h
by_cases hs : (l2.getDual? i).isSome
<;> by_cases ho : (l2.getDualInOther? l i).isSome
<;> simp_all [hs, ho]
lemma mem_withUniqueDual_append_symm (i : Fin l.length) :
appendEquiv (Sum.inl i) ∈ (l ++ l2).withUniqueDual ↔
appendEquiv (Sum.inr i) ∈ (l2 ++ l).withUniqueDual := by
simp only [withUniqueDual, mem_withDual_iff_isSome, Finset.mem_filter, Finset.mem_univ,
getDual?_isSome_append_inl_iff, true_and, getDual?_isSome_append_inr_iff, and_congr_right_iff]
intro h
refine Iff.intro (fun h' j hj => ?_) (fun h' j hj => ?_)
· obtain ⟨k, hk⟩ := appendEquiv.surjective j
subst hk
match k with
| Sum.inl k =>
have hk' := h' (appendEquiv (Sum.inr k))
simp only [AreDualInSelf.append_inl_inr] at hk'
simp only [AreDualInSelf.append_inr_inl] at hj
refine ((l.getDual?_append_symm_of_withUniqueDual_of_inl l2 _ _ ?_).mp (hk' hj).symm).symm
simp_all only [AreDualInSelf.append_inl_inr, imp_self, withUniqueDual,
mem_withDual_iff_isSome, Finset.mem_filter, Finset.mem_univ, getDual?_isSome_append_inl_iff,
implies_true, and_self, mem_withUniqueDual_isSome, eq_getDual?_get_of_withUniqueDual_iff,
getDual?_areDualInSelf_get]
| Sum.inr k =>
have hk' := h' (appendEquiv (Sum.inl k))
simp only [AreDualInSelf.append_inl_inl] at hk'
simp only [AreDualInSelf.append_inr_inr] at hj
refine ((l.getDual?_append_symm_of_withUniqueDual_of_inl' l2 _ _ ?_).mp (hk' hj).symm).symm
simp_all only [AreDualInSelf.append_inl_inl, imp_self, withUniqueDual,
mem_withDual_iff_isSome, Finset.mem_filter, Finset.mem_univ, getDual?_isSome_append_inl_iff,
implies_true, and_self, mem_withUniqueDual_isSome, eq_getDual?_get_of_withUniqueDual_iff,
getDual?_areDualInSelf_get]
· obtain ⟨k, hk⟩ := appendEquiv.surjective j
subst hk
match k with
| Sum.inl k =>
have hk' := h' (appendEquiv (Sum.inr k))
simp only [AreDualInSelf.append_inr_inr] at hk'
simp only [AreDualInSelf.append_inl_inl] at hj
refine ((l2.getDual?_append_symm_of_withUniqueDual_of_inr' l _ _ ?_).mp (hk' hj).symm).symm
simp_all only [AreDualInSelf.append_inr_inr, imp_self, withUniqueDual,
mem_withDual_iff_isSome, Finset.mem_filter, Finset.mem_univ, getDual?_isSome_append_inr_iff,
implies_true, and_self, mem_withUniqueDual_isSome, eq_getDual?_get_of_withUniqueDual_iff,
getDual?_areDualInSelf_get]
| Sum.inr k =>
have hk' := h' (appendEquiv (Sum.inl k))
simp only [AreDualInSelf.append_inr_inl] at hk'
simp only [AreDualInSelf.append_inl_inr] at hj
refine ((l2.getDual?_append_symm_of_withUniqueDual_of_inr l _ _ ?_).mp (hk' hj).symm).symm
simp_all only [AreDualInSelf.append_inr_inl, imp_self, withUniqueDual,
mem_withDual_iff_isSome, Finset.mem_filter, Finset.mem_univ, getDual?_isSome_append_inr_iff,
implies_true, and_self, mem_withUniqueDual_isSome, eq_getDual?_get_of_withUniqueDual_iff,
getDual?_areDualInSelf_get]
@[simp]
lemma not_mem_withDualInOther_of_inl_mem_withUniqueDual (i : Fin l.length)
(h : appendEquiv (Sum.inl i) ∈ (l ++ l2).withUniqueDual) (hs : i ∈ l.withUniqueDual) :
¬ i ∈ l.withUniqueDualInOther l2 := by
have hn := l.append_inl_not_mem_withDual_of_withDualInOther l2 i h
simp_all
by_contra ho
have ho' : (l.getDualInOther? l2 i).isSome := by
exact mem_withUniqueDualInOther_isSome l l2 i ho
simp_all [Option.isSome_none, Bool.false_eq_true]
@[simp]
lemma not_mem_withUniqueDual_of_inl_mem_withUnqieuDual (i : Fin l.length)
(h : appendEquiv (Sum.inl i) ∈ (l ++ l2).withUniqueDual) (hi : i ∈ l.withUniqueDualInOther l2) :
¬ i ∈ l.withUniqueDual := by
by_contra hs
simp_all only [not_mem_withDualInOther_of_inl_mem_withUniqueDual]
@[simp]
lemma mem_withUniqueDual_of_inl (i : Fin l.length)
(h : appendEquiv (Sum.inl i) ∈ (l ++ l2).withUniqueDual) (hi : (l.getDual? i).isSome) :
i ∈ l.withUniqueDual := by
simp only [withUniqueDual, mem_withDual_iff_isSome, Finset.mem_filter, Finset.mem_univ,
getDual?_isSome_append_inl_iff, true_and] at h ⊢
apply And.intro hi
intro j hj
have hj' := h.2 (appendEquiv (Sum.inl j))
simp at hj'
have hj'' := hj' hj
simp [hi] at hj''
simp_all
@[simp]
lemma mem_withUniqueDualInOther_of_inl_withDualInOther (i : Fin l.length)
(h : appendEquiv (Sum.inl i) ∈ (l ++ l2).withUniqueDual)
(hi : (l.getDualInOther? l2 i).isSome) : i ∈ l.withUniqueDualInOther l2 := by
simp only [withUniqueDualInOther, mem_withDual_iff_isSome, Bool.not_eq_true, Option.not_isSome,
Option.isNone_iff_eq_none, mem_withInDualOther_iff_isSome, Finset.mem_filter, Finset.mem_univ,
true_and]
have hn := l.append_inl_not_mem_withDual_of_withDualInOther l2 i h
simp_all only [mem_withDual_iff_isSome, mem_withInDualOther_iff_isSome, not_true_eq_false,
iff_false, Bool.not_eq_true, Option.not_isSome, Option.isNone_iff_eq_none, true_and]
intro j hj
simp only [withUniqueDual, mem_withDual_iff_isSome, Finset.mem_filter, Finset.mem_univ,
getDual?_isSome_append_inl_iff, true_and] at h
have hj' := h.2 (appendEquiv (Sum.inr j))
simp only [AreDualInSelf.append_inl_inr] at hj'
have hj'' := hj' hj
simp only [hn, Option.isNone_none, hi, getDual?_inl_of_getDual?_isNone_getDualInOther?_isSome,
Option.some.injEq, EmbeddingLike.apply_eq_iff_eq, Sum.inr.injEq] at hj''
simp_all only [getDualInOther?_areDualInOther_get, Option.isSome_none, Bool.false_eq_true,
or_true, Option.isNone_none, getDual?_inl_of_getDual?_isNone_getDualInOther?_isSome,
Option.some.injEq, true_and, AreDualInSelf.append_inl_inr, imp_self, Option.some_get]
lemma withUniqueDual_iff_not_withUniqueDualInOther_of_inl_withUniqueDualInOther
(i : Fin l.length) (h : appendEquiv (Sum.inl i) ∈ (l ++ l2).withUniqueDual) :
i ∈ l.withUniqueDual ↔ ¬ i ∈ l.withUniqueDualInOther l2 := by
by_cases h' : (l.getDual? i).isSome
have hn : i ∈ l.withUniqueDual := mem_withUniqueDual_of_inl l l2 i h h'
simp_all
have hn := l.append_inl_not_mem_withDual_of_withDualInOther l2 i h
simp_all
simp [withUniqueDual]
simp_all
exact mem_withUniqueDualInOther_of_inl_withDualInOther l l2 i h (Option.ne_none_iff_isSome.mp hn)
lemma append_inl_mem_withUniqueDual_of_withUniqueDual (i : Fin l.length)
(h : i ∈ l.withUniqueDual) (hn : i ∉ l.withDualInOther l2) :
appendEquiv (Sum.inl i) ∈ (l ++ l2).withUniqueDual := by
simp only [withUniqueDual, mem_withDual_iff_isSome, Finset.mem_filter, Finset.mem_univ,
getDual?_isSome_append_inl_iff, true_and]
apply And.intro
simp_all
intro j hj
obtain ⟨k, hk⟩ := appendEquiv.surjective j
subst hk
match k with
| Sum.inl k => simp_all
| Sum.inr k =>
simp at hj
refine False.elim (hn ?_)
exact (l.mem_withInDualOther_iff_exists _).mpr ⟨k, hj⟩
lemma append_inl_mem_of_withUniqueDualInOther (i : Fin l.length)
(ho : i ∈ l.withUniqueDualInOther l2) :
appendEquiv (Sum.inl i) ∈ (l ++ l2).withUniqueDual := by
simp only [withUniqueDual, mem_withDual_iff_isSome, Finset.mem_filter, Finset.mem_univ,
getDual?_isSome_append_inl_iff, true_and]
apply And.intro
simp_all only [mem_withUniqueDualInOther_isSome, or_true]
intro j hj
obtain ⟨k, hk⟩ := appendEquiv.surjective j
subst hk
have hs := l.not_mem_withDual_of_withUniqueDualInOther l2 ⟨i, ho⟩
match k with
| Sum.inl k =>
refine False.elim (hs ?_)
simp at hj
exact (l.mem_withDual_iff_exists).mpr ⟨k, hj⟩
| Sum.inr k =>
simp_all
@[simp]
lemma append_inl_mem_withUniqueDual_iff (i : Fin l.length) :
appendEquiv (Sum.inl i) ∈ (l ++ l2).withUniqueDual ↔
((i ∈ l.withUniqueDual ∧ i ∉ l.withDualInOther l2) ↔ ¬ i ∈ l.withUniqueDualInOther l2) := by
apply Iff.intro
· intro h
apply Iff.intro
· intro hs
exact (l.withUniqueDual_iff_not_withUniqueDualInOther_of_inl_withUniqueDualInOther l2 i
h).mp hs.1
· intro ho
have hs := ((l.withUniqueDual_iff_not_withUniqueDualInOther_of_inl_withUniqueDualInOther l2 i
h).mpr ho)
apply And.intro hs
refine (l.append_inl_not_mem_withDual_of_withDualInOther l2 i h).mp ?_
exact (l.mem_withDual_of_withUniqueDual ⟨i, hs⟩)
· intro h
by_cases ho : i ∈ l.withUniqueDualInOther l2
· exact append_inl_mem_of_withUniqueDualInOther l l2 i ho
· exact append_inl_mem_withUniqueDual_of_withUniqueDual l l2 i (h.mpr ho).1 (h.mpr ho).2
@[simp]
lemma append_inr_mem_withUniqueDual_iff (i : Fin l2.length) :
appendEquiv (Sum.inr i) ∈ (l ++ l2).withUniqueDual ↔
((i ∈ l2.withUniqueDual ∧ i ∉ l2.withDualInOther l) ↔ ¬ i ∈ l2.withUniqueDualInOther l) := by
rw [← mem_withUniqueDual_append_symm]
exact append_inl_mem_withUniqueDual_iff l2 l i
lemma append_withUniqueDual : (l ++ l2).withUniqueDual =
Finset.map (l.appendInl l2)
((l.withUniqueDual ∩ (l.withDualInOther l2)ᶜ) l.withUniqueDualInOther l2)
Finset.map (l.appendInr l2)
((l2.withUniqueDual ∩ (l2.withDualInOther l)ᶜ) l2.withUniqueDualInOther l) := by
rw [Finset.ext_iff]
intro j
obtain ⟨k, hk⟩ := appendEquiv.surjective j
subst hk
match k with
| Sum.inl k =>
simp only [append_inl_mem_withUniqueDual_iff, Finset.mem_union]
apply Iff.intro
· intro h
apply Or.inl
simp only [Finset.mem_map, Finset.mem_union, Finset.mem_inter, Finset.mem_compl,
mem_withInDualOther_iff_isSome, Bool.not_eq_true, Option.not_isSome,
Option.isNone_iff_eq_none]
use k
simp only [appendInl, Function.Embedding.coeFn_mk, Function.comp_apply, and_true]
by_cases hk : k ∈ l.withUniqueDualInOther l2
exact Or.inr hk
have hk' := h.mpr hk
simp_all only [not_false_eq_true, and_self, mem_withInDualOther_iff_isSome, Bool.not_eq_true,
Option.not_isSome, Option.isNone_iff_eq_none, or_false]
· intro h
simp only [Finset.mem_map, Finset.mem_union, Finset.mem_inter, Finset.mem_compl,
mem_withInDualOther_iff_isSome, Bool.not_eq_true, Option.not_isSome,
Option.isNone_iff_eq_none] at h
cases' h with h h
· obtain ⟨j, hj⟩ := h
have hjk : j = k := by
simp only [appendInl, Function.Embedding.coeFn_mk, Function.comp_apply,
EmbeddingLike.apply_eq_iff_eq, Sum.inl.injEq] at hj
exact hj.2
subst hjk
have hj1 := hj.1
cases' hj1 with hj1 hj1
· simp_all only [and_self, true_or, true_and, mem_withInDualOther_iff_isSome,
Option.isSome_none, Bool.false_eq_true, not_false_eq_true, true_iff]
by_contra hn
have h' := l.mem_withDualInOther_of_withUniqueDualInOther l2 ⟨j, hn⟩
simp_all only [mem_withInDualOther_iff_isSome, Option.isSome_none, Bool.false_eq_true]
· simp_all only [or_true, true_and, mem_withInDualOther_iff_isSome,
mem_withUniqueDualInOther_isSome, not_true_eq_false, and_false]
· obtain ⟨j, hj⟩ := h
simp only [appendInr, Function.Embedding.coeFn_mk, Function.comp_apply,
EmbeddingLike.apply_eq_iff_eq, and_false] at hj
| Sum.inr k =>
simp only [append_inr_mem_withUniqueDual_iff, Finset.mem_union]
apply Iff.intro
· intro h
apply Or.inr
simp only [Finset.mem_map, Finset.mem_union, Finset.mem_inter, Finset.mem_compl,
mem_withInDualOther_iff_isSome, Bool.not_eq_true, Option.not_isSome,
Option.isNone_iff_eq_none]
use k
simp [appendInr]
by_cases hk : k ∈ l2.withUniqueDualInOther l
exact Or.inr hk
have hk' := h.mpr hk
simp_all only [not_false_eq_true, and_self, mem_withInDualOther_iff_isSome, Bool.not_eq_true,
Option.not_isSome, Option.isNone_iff_eq_none, or_false]
· intro h
simp only [Finset.mem_map, Finset.mem_union, Finset.mem_inter, Finset.mem_compl,
mem_withInDualOther_iff_isSome, Bool.not_eq_true, Option.not_isSome,
Option.isNone_iff_eq_none] at h
cases' h with h h
· obtain ⟨j, hj⟩ := h
simp only [appendInl, Function.Embedding.coeFn_mk, Function.comp_apply,
EmbeddingLike.apply_eq_iff_eq, and_false] at hj
· obtain ⟨j, hj⟩ := h
have hjk : j = k := by
simp only [appendInr, Function.Embedding.coeFn_mk, Function.comp_apply,
EmbeddingLike.apply_eq_iff_eq, Sum.inr.injEq] at hj
exact hj.2
subst hjk
have hj1 := hj.1
cases' hj1 with hj1 hj1
· simp_all only [and_self, true_or, true_and, mem_withInDualOther_iff_isSome,
Option.isSome_none, Bool.false_eq_true, not_false_eq_true, true_iff]
by_contra hn
have h' := l2.mem_withDualInOther_of_withUniqueDualInOther l ⟨j, hn⟩
simp_all only [mem_withInDualOther_iff_isSome, Option.isSome_none, Bool.false_eq_true]
· simp_all only [or_true, true_and, mem_withInDualOther_iff_isSome,
mem_withUniqueDualInOther_isSome, not_true_eq_false, and_false]
lemma append_withUniqueDual_disjSum : (l ++ l2).withUniqueDual =
Equiv.finsetCongr appendEquiv
(((l.withUniqueDual ∩ (l.withDualInOther l2)ᶜ) l.withUniqueDualInOther l2).disjSum
((l2.withUniqueDual ∩ (l2.withDualInOther l)ᶜ) l2.withUniqueDualInOther l)) := by
rw [← Equiv.symm_apply_eq]
simp only [Equiv.finsetCongr_symm, append_withUniqueDual, Equiv.finsetCongr_apply]
rw [Finset.map_union]
rw [Finset.map_map, Finset.map_map]
ext1 a
cases a with
| inl val => simp
| inr val_1 => simp
/-!
## Properties of getDualInOther?, withUniqueDualInOther and appendInOther
-/
lemma mem_append_withUniqueDualInOther_symm (i : Fin l.length) :
appendEquiv (Sum.inl i) ∈ (l ++ l2).withUniqueDualInOther l3 ↔
appendEquiv (Sum.inr i) ∈ (l2 ++ l).withUniqueDualInOther l3 := by
simp only [withUniqueDualInOther, mem_withDual_iff_isSome, Bool.not_eq_true, Option.not_isSome,
Option.isNone_iff_eq_none, mem_withInDualOther_iff_isSome, Finset.mem_filter, Finset.mem_univ,
getDual?_eq_none_append_inl_iff, getDualInOther?_append_of_inl, AreDualInOther.append_of_inl,
true_and, getDual?_eq_none_append_inr_iff, getDualInOther?_append_of_inr,
AreDualInOther.append_of_inr]
@[simp]
lemma withUniqueDualInOther_append_not_isSome_snd_of_isSome_fst (i : Fin l.length)
(h1 : i ∈ l.withUniqueDualInOther (l2 ++ l3)) (h2 : (l.getDualInOther? l2 i).isSome) :
(l.getDualInOther? l3 i) = none := by
by_contra hn
simp only [getDualInOther?] at h2 hn
rw [← @Option.not_isSome_iff_eq_none, not_not] at hn
rw [Fin.isSome_find_iff] at h2 hn
obtain ⟨j2, hj2⟩ := h2
obtain ⟨j3, hj3⟩ := hn
have h1' : l.AreDualInOther (l2 ++ l3) i (appendEquiv (Sum.inl j2)) :=
(AreDualInOther.of_append_inl i j2).mpr hj2
have h2 : l.AreDualInOther (l2 ++ l3) i (appendEquiv (Sum.inr j3)) :=
(AreDualInOther.of_append_inr i j3).mpr hj3
have h3 := l.eq_of_areDualInOther_withUniqueDualInOther (l2 ++ l3) ⟨i, h1⟩ h1' h2
simp only [EmbeddingLike.apply_eq_iff_eq] at h3
@[simp]
lemma withUniqueDualInOther_append_not_isSome_fst_of_isSome_snd (i : Fin l.length)
(h1 : i ∈ l.withUniqueDualInOther (l2 ++ l3)) (h2 : (l.getDualInOther? l3 i).isSome) :
(l.getDualInOther? l2 i) = none := by
by_contra hn
simp only [getDualInOther?] at h2 hn
rw [← @Option.not_isSome_iff_eq_none, not_not] at hn
rw [Fin.isSome_find_iff] at h2 hn
obtain ⟨j2, hj2⟩ := h2
obtain ⟨j3, hj3⟩ := hn
have h1' : l.AreDualInOther (l2 ++ l3) i (appendEquiv (Sum.inr j2)) :=
(AreDualInOther.of_append_inr i j2).mpr hj2
have h2 : l.AreDualInOther (l2 ++ l3) i (appendEquiv (Sum.inl j3)) :=
(AreDualInOther.of_append_inl i j3).mpr hj3
have h3 := l.eq_of_areDualInOther_withUniqueDualInOther (l2 ++ l3) ⟨i, h1⟩ h1' h2
simp only [EmbeddingLike.apply_eq_iff_eq] at h3
@[simp]
lemma withUniqueDualInOther_append_isSome_snd_of_not_isSome_fst (i : Fin l.length)
(h1 : i ∈ l.withUniqueDualInOther (l2 ++ l3)) (h2 : ¬ (l.getDualInOther? l2 i).isSome) :
(l.getDualInOther? l3 i).isSome := by
by_contra hn
simp only [withUniqueDualInOther, mem_withDual_iff_isSome, Bool.not_eq_true, Option.not_isSome,
Option.isNone_iff_eq_none, mem_withInDualOther_iff_isSome, getDualInOther?_isSome_of_append_iff,
Finset.mem_filter, Finset.mem_univ, true_and] at h1
simp_all only [Bool.not_eq_true, Option.not_isSome, Option.isNone_iff_eq_none, Option.isSome_none,
Bool.false_eq_true, or_self, false_and, and_false]
lemma withUniqueDualInOther_append_isSome_fst_iff_not_isSome_snd (i : Fin l.length)
(h1 : i ∈ l.withUniqueDualInOther (l2 ++ l3)) :
(l.getDualInOther? l2 i).isSome ↔ (l.getDualInOther? l3 i) = none := by
by_cases hs : (l.getDualInOther? l2 i).isSome
simp [hs, h1]
exact l.withUniqueDualInOther_append_not_isSome_snd_of_isSome_fst l2 l3 i h1 hs
simp [hs]
rw [← @Option.not_isSome_iff_eq_none, not_not]
exact withUniqueDualInOther_append_isSome_snd_of_not_isSome_fst l l2 l3 i h1 hs
lemma getDualInOther?_append_symm_of_withUniqueDual_of_inl (i : Fin l.length)
(k : Fin l2.length) (h : i ∈ l.withUniqueDualInOther (l2 ++ l3)) :
l.getDualInOther? (l2 ++ l3) i = some (appendEquiv (Sum.inl k))
↔ l.getDualInOther? (l3 ++ l2) i = some (appendEquiv (Sum.inr k)) := by
have h := l.withUniqueDualInOther_append_isSome_fst_iff_not_isSome_snd l2 l3 i h
by_cases hs : (l.getDualInOther? l2 i).isSome
<;> by_cases ho : (l.getDualInOther? l3 i).isSome
<;> simp_all [hs]
lemma getDualInOther?_append_symm_of_withUniqueDual_of_inr (i : Fin l.length)
(k : Fin l3.length) (h : i ∈ l.withUniqueDualInOther (l2 ++ l3)) :
l.getDualInOther? (l2 ++ l3) i = some (appendEquiv (Sum.inr k))
↔ l.getDualInOther? (l3 ++ l2) i = some (appendEquiv (Sum.inl k)) := by
have h := l.withUniqueDualInOther_append_isSome_fst_iff_not_isSome_snd l2 l3 i h
by_cases hs : (l.getDualInOther? l2 i).isSome
<;> by_cases ho : (l.getDualInOther? l3 i).isSome
<;> simp_all [hs]
lemma mem_withUniqueDualInOther_symm' (i : Fin l.length)
(h : i ∈ l.withUniqueDualInOther (l2 ++ l3)):
i ∈ l.withUniqueDualInOther (l3 ++ l2) := by
have h' := h
simp only [withUniqueDualInOther, mem_withDual_iff_isSome, Bool.not_eq_true,
Option.not_isSome, Option.isNone_iff_eq_none, mem_withInDualOther_iff_isSome,
getDualInOther?_isSome_of_append_iff, Finset.mem_filter, Finset.mem_univ, true_and,
implies_true, and_self, eq_getDualInOther?_get_of_withUniqueDualInOther_iff,
getDualInOther?_areDualInOther_get] at h ⊢
apply And.intro h.1
have hc := l.withUniqueDualInOther_append_isSome_fst_iff_not_isSome_snd l2 l3 i h'
by_cases h1 : (l.getDualInOther? l2 i).isSome <;>
by_cases h2 : (l.getDualInOther? l3 i).isSome
· simp only [h1, h2, not_true_eq_false, imp_false] at hc
rw [← @Option.not_isSome_iff_eq_none] at hc
simp [h2] at hc
· simp only [h1, or_true, true_and]
intro j
obtain ⟨k, hk⟩ := appendEquiv.surjective j
subst hk
match k with
| Sum.inl k =>
simp only [AreDualInOther.of_append_inl]
have hk'' := h.2.2 (appendEquiv (Sum.inr k))
simp at hk''
exact fun h'' => ((getDualInOther?_append_symm_of_withUniqueDual_of_inr l l2 l3 i k h').mp
(hk'' h'').symm).symm
| Sum.inr k =>
simp only [AreDualInOther.of_append_inr]
have hk'' := h.2.2 (appendEquiv (Sum.inl k))
simp at hk''
exact fun h'' => ((getDualInOther?_append_symm_of_withUniqueDual_of_inl l l2 l3 i k h').mp
(hk'' h'').symm).symm
· simp only [h2, true_or, Option.some.injEq, true_and]
intro j
obtain ⟨k, hk⟩ := appendEquiv.surjective j
subst hk
match k with
| Sum.inl k =>
simp only [AreDualInOther.of_append_inl]
have hk'' := h.2.2 (appendEquiv (Sum.inr k))
simp at hk''
exact fun h'' => ((getDualInOther?_append_symm_of_withUniqueDual_of_inr l l2 l3 i k h').mp
(hk'' h'').symm).symm
| Sum.inr k =>
simp only [AreDualInOther.of_append_inr]
have hk'' := h.2.2 (appendEquiv (Sum.inl k))
simp only [AreDualInOther.of_append_inl] at hk''
exact fun h'' => ((getDualInOther?_append_symm_of_withUniqueDual_of_inl l l2 l3 i k h').mp
(hk'' h'').symm).symm
· simp only [h1, Bool.false_eq_true, false_iff] at hc
simp_all only [Bool.false_eq_true, or_self, eq_getDualInOther?_get_of_withUniqueDualInOther_iff,
Option.some_get, implies_true, and_true, and_false]
lemma mem_withUniqueDualInOther_symm (i : Fin l.length) :
i ∈ l.withUniqueDualInOther (l2 ++ l3) ↔ i ∈ l.withUniqueDualInOther (l3 ++ l2) :=
Iff.intro (l.mem_withUniqueDualInOther_symm' l2 l3 i) (l.mem_withUniqueDualInOther_symm' l3 l2 i)
/-!
## The equivalence defined by getDual?
-/
/-- 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
simp only [Finset.coe_mem, getDual?_get_of_mem_withUnique_mem]⟩
invFun x := ⟨(l.getDual? x).get (l.mem_withUniqueDual_isSome x.1 x.2), by simp⟩
left_inv x := SetCoe.ext (by
simp only [Finset.coe_mem, getDual?_getDual?_get_of_withUniqueDual, Option.get_some])
right_inv x := SetCoe.ext (by
simp only [Finset.coe_mem, getDual?_getDual?_get_of_withUniqueDual, Option.get_some])
@[simp]
lemma getDualEquiv_involutive : Function.Involutive l.getDualEquiv := by
intro x
exact (Equiv.apply_eq_iff_eq_symm_apply l.getDualEquiv).mpr rfl
/-!
## Equivalence for withUniqueDualInOther
-/
/-- An equivalence from `l.withUniqueDualInOther l2` to `l2.withUniqueDualInOther l`
obtained by taking the dual index. -/
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 simp⟩
invFun x := ⟨(l2.getDualInOther? l x).get (l2.mem_withUniqueDualInOther_isSome l x.1 x.2),
by simp⟩
left_inv x := SetCoe.ext (by simp)
right_inv x := SetCoe.ext (by simp)
@[simp]
lemma getDualInOtherEquiv_self_refl : l.getDualInOtherEquiv l = Equiv.refl _ := by
apply Equiv.ext
intro x
simp [getDualInOtherEquiv]
end IndexList
end IndexNotation

View file

@ -82,7 +82,8 @@ noncomputable def fromIndexStringColor {cn : Fin n → realTensorColor.Color}
(hn : n = (toIndexList' s hs).length)
(hD : (toIndexList' s hs).withDual = (toIndexList' s hs).withUniqueDual)
(hC : IndexList.ColorCond.bool (toIndexList' s hs))
(hd : TensorColor.ColorMap.DualMap.boolFin (toIndexList' s hs).colorMap (cn ∘ Fin.cast hn.symm)) :
(hd : TensorColor.ColorMap.DualMap.boolFin
(toIndexList' s hs).colorMap (cn ∘ Fin.cast hn.symm)) :
(realLorentzTensor d).TensorIndex :=
TensorStructure.TensorIndex.mkDualMap T ⟨(toIndexList' s hs), hD,
IndexList.ColorCond.iff_bool.mpr hC⟩ hn
@ -97,16 +98,17 @@ macro "dualMapTactic" : tactic =>
/-- Notation for the construction of a tensor index from a tensor and a string.
Conditions are checked automatically. -/
notation:20 T "|" S:21 => fromIndexStringColor T S (by rfl) (by rfl) (by rfl) (by rfl) (by dualMapTactic)
notation:20 T "|" S:21 => fromIndexStringColor T S
(by rfl) (by rfl) (by rfl) (by rfl) (by dualMapTactic)
/-- A tactics used to prove `colorPropBool` for real Lorentz tensors. -/
macro "prodTactic" : tactic =>
`(tactic| {
apply (ColorIndexList.AppendCond.iff_bool _ _).mpr
change @ColorIndexList.AppendCond.bool realTensorColor instIndexNotationColorRealTensorColor instDecidableEqColorRealTensorColor
_ _
simp only [prod_toIndexList, indexNotation_eq_color, fromIndexStringColor, mkDualMap, toTensorColor_eq,
decidableEq_eq_color]
change @ColorIndexList.AppendCond.bool realTensorColor
instIndexNotationColorRealTensorColor instDecidableEqColorRealTensorColor _ _
simp only [prod_toIndexList, indexNotation_eq_color, fromIndexStringColor, mkDualMap,
toTensorColor_eq, decidableEq_eq_color]
rfl})
/-- The product of Real Lorentz tensors. Conditions on indices are checked automatically. -/

View file

@ -1,7 +1,7 @@
#!/usr/bin/env python3
"""
Taken from Mathlib4.
Taken from Mathlib4.
Lint a file or files from mathlib for style.
@ -30,7 +30,7 @@ exceptions by redirecting the output to ``style-exceptions.txt``. Use:
$ ./scripts/update-style-exceptions.sh
to perform this update.
to perform this update.
"""
# TODO: This is adapted from the linter for mathlib3. It should be rewritten in Lean.
@ -354,7 +354,8 @@ def output_message(path, line_nr, code, msg):
if len(exceptions) == 0:
# we are generating a new exceptions file
# filename first, then line so that we can call "sort" on the output
print(f"{path} : line {line_nr} : {code} : {msg}")
#print(f"{path} : line {line_nr} : {code} : {msg}")
print(f"{path}:{line_nr} : {code} : {msg}")
else:
if code.startswith("ERR"):
msg_type = "error"