refactor: Large, incomplete, refactor of index notation
This commit is contained in:
parent
85fc57750d
commit
a8474233ae
7 changed files with 1557 additions and 252 deletions
|
@ -5,6 +5,7 @@ Authors: Joseph Tooby-Smith
|
|||
-/
|
||||
import Mathlib.Data.Set.Finite
|
||||
import Mathlib.Data.Finset.Sort
|
||||
import Mathlib.Logic.Equiv.Fin
|
||||
/-!
|
||||
|
||||
# Index notation for a type
|
||||
|
@ -236,8 +237,6 @@ instance : Fintype l.toPosSet where
|
|||
def toPosFinset (l : IndexList X) : Finset (Fin l.numIndices × Index X) :=
|
||||
l.toPosSet.toFinset
|
||||
|
||||
instance : HAppend (IndexList X) (IndexList X) (IndexList X) :=
|
||||
@instHAppendOfAppend (List (Index X)) List.instAppend
|
||||
|
||||
/-- The construction of a list of indices from a map
|
||||
from `Fin n` to `Index X`. -/
|
||||
|
@ -249,233 +248,47 @@ lemma fromFinMap_numIndices {n : ℕ} (f : Fin n → Index X) :
|
|||
(fromFinMap f).numIndices = n := by
|
||||
simp [fromFinMap, numIndices]
|
||||
|
||||
/-!
|
||||
|
||||
## Contracted and non-contracting indices
|
||||
|
||||
-/
|
||||
|
||||
/-- The proposition on a element (or really index of element) of a index list
|
||||
`s` which is ture iff does not share an id with any other element.
|
||||
This tells us that it should not be contracted with any other element. -/
|
||||
def NoContr (i : Fin l.length) : Prop :=
|
||||
∀ j, i ≠ j → l.idMap i ≠ l.idMap j
|
||||
|
||||
instance (i : Fin l.length) : Decidable (l.NoContr i) :=
|
||||
Fintype.decidableForallFintype
|
||||
|
||||
/-- The finset of indices of an index list corresponding to elements which do not contract. -/
|
||||
def noContrFinset : Finset (Fin l.length) :=
|
||||
Finset.univ.filter l.NoContr
|
||||
|
||||
/-- An eqiuvalence between the subtype of indices of a index list `l` which do not contract and
|
||||
`Fin l.noContrFinset.card`. -/
|
||||
def noContrSubtypeEquiv : {i : Fin l.length // l.NoContr i} ≃ Fin l.noContrFinset.card :=
|
||||
(Equiv.subtypeEquivRight (fun x => by simp [noContrFinset])).trans
|
||||
(Finset.orderIsoOfFin l.noContrFinset rfl).toEquiv.symm
|
||||
|
||||
@[simp]
|
||||
lemma idMap_noContrSubtypeEquiv_neq (i j : Fin l.noContrFinset.card) (h : i ≠ j) :
|
||||
l.idMap (l.noContrSubtypeEquiv.symm i).1 ≠ l.idMap (l.noContrSubtypeEquiv.symm j).1 := by
|
||||
have hi := ((l.noContrSubtypeEquiv).symm i).2
|
||||
simp [NoContr] at hi
|
||||
have hj := hi ((l.noContrSubtypeEquiv).symm j)
|
||||
apply hj
|
||||
rw [@SetCoe.ext_iff]
|
||||
erw [Equiv.apply_eq_iff_eq]
|
||||
exact h
|
||||
|
||||
/-- The subtype of indices `l` which do contract. -/
|
||||
def contrSubtype : Type := {i : Fin l.length // ¬ l.NoContr i}
|
||||
|
||||
instance : Fintype l.contrSubtype :=
|
||||
Subtype.fintype fun x => ¬ l.NoContr x
|
||||
|
||||
instance : DecidableEq l.contrSubtype :=
|
||||
Subtype.instDecidableEq
|
||||
|
||||
/-!
|
||||
|
||||
## Getting the index which contracts with a given index
|
||||
## Appending index lists.
|
||||
|
||||
-/
|
||||
instance : HAppend (IndexList X) (IndexList X) (IndexList X) :=
|
||||
@instHAppendOfAppend (List (Index X)) List.instAppend
|
||||
|
||||
/-- Given a `i : l.contrSubtype` the proposition on a `j` in `Fin s.length` for
|
||||
it to be an index of `s` contracting with `i`. -/
|
||||
def getDualProp (i j : Fin l.length) : Prop :=
|
||||
i ≠ j ∧ l.idMap i = l.idMap j
|
||||
|
||||
instance (i j : Fin l.length) :
|
||||
Decidable (l.getDualProp i j) :=
|
||||
instDecidableAnd
|
||||
|
||||
/-- Given a `i : l.contrSubtype` the index of `s` contracting with `i`. -/
|
||||
def getDualFin (i : l.contrSubtype) : Fin l.length :=
|
||||
(Fin.find (l.getDualProp i.1)).get (by simpa [NoContr, Fin.isSome_find_iff] using i.prop)
|
||||
|
||||
lemma some_getDualFin_eq_find (i : l.contrSubtype) :
|
||||
Fin.find (l.getDualProp i.1) = some (l.getDualFin i) := by
|
||||
simp [getDualFin]
|
||||
|
||||
lemma getDualFin_not_NoContr (i : l.contrSubtype) : ¬ l.NoContr (l.getDualFin i) := by
|
||||
have h := l.some_getDualFin_eq_find i
|
||||
rw [Fin.find_eq_some_iff] at h
|
||||
simp [NoContr]
|
||||
exact ⟨i.1, And.intro (fun a => h.1.1 a.symm) h.1.2.symm⟩
|
||||
|
||||
/-- The dual index of an element of `𝓒.contrSubtype s`, that is the index
|
||||
contracting with it. -/
|
||||
def getDual (i : l.contrSubtype) : l.contrSubtype :=
|
||||
⟨l.getDualFin i, l.getDualFin_not_NoContr i⟩
|
||||
|
||||
lemma getDual_id (i : l.contrSubtype) : l.idMap i.1 = l.idMap (l.getDual i).1 := by
|
||||
simp [getDual]
|
||||
have h1 := l.some_getDualFin_eq_find i
|
||||
rw [Fin.find_eq_some_iff] at h1
|
||||
simp only [getDualProp, ne_eq, and_imp] at h1
|
||||
exact h1.1.2
|
||||
|
||||
lemma getDual_neq_self (i : l.contrSubtype) : i ≠ l.getDual i := by
|
||||
have h1 := l.some_getDualFin_eq_find i
|
||||
rw [Fin.find_eq_some_iff] at h1
|
||||
exact ne_of_apply_ne Subtype.val h1.1.1
|
||||
|
||||
lemma getDual_getDualProp (i : l.contrSubtype) : l.getDualProp i.1 (l.getDual i).1 := by
|
||||
simp [getDual]
|
||||
have h1 := l.some_getDualFin_eq_find i
|
||||
rw [Fin.find_eq_some_iff] at h1
|
||||
simp only [getDualProp, ne_eq, and_imp] at h1
|
||||
exact h1.1
|
||||
|
||||
/-!
|
||||
|
||||
## Index lists with no contracting indices
|
||||
|
||||
-/
|
||||
|
||||
/-- The proposition on a `IndexList` for it to have no contracting
|
||||
indices. -/
|
||||
def HasNoContr : Prop := ∀ i, l.NoContr i
|
||||
|
||||
lemma contrSubtype_is_empty_of_hasNoContr (h : l.HasNoContr) : IsEmpty l.contrSubtype := by
|
||||
rw [_root_.isEmpty_iff]
|
||||
intro a
|
||||
exact h a.1 a.1 (fun _ => a.2 (h a.1)) rfl
|
||||
|
||||
lemma hasNoContr_id_inj (h : l.HasNoContr) : Function.Injective l.idMap := fun i j => by
|
||||
simpa using (h i j).mt
|
||||
|
||||
lemma hasNoContr_color_eq_of_id_eq (h : l.HasNoContr) (i j : Fin l.length) :
|
||||
l.idMap i = l.idMap j → l.colorMap i = l.colorMap j := by
|
||||
intro h1
|
||||
apply l.hasNoContr_id_inj h at h1
|
||||
rw [h1]
|
||||
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
|
||||
|
||||
@[simp]
|
||||
lemma hasNoContr_noContrFinset_card (h : l.HasNoContr) :
|
||||
l.noContrFinset.card = List.length l := by
|
||||
simp only [noContrFinset]
|
||||
rw [Finset.filter_true_of_mem]
|
||||
simp only [Finset.card_univ, Fintype.card_fin]
|
||||
intro x _
|
||||
exact h x
|
||||
|
||||
/-!
|
||||
|
||||
## The contracted index list
|
||||
|
||||
-/
|
||||
|
||||
/-- The index list of those indices of `l` which do not contract. -/
|
||||
def contrIndexList : IndexList X :=
|
||||
IndexList.fromFinMap (fun i => l.get (l.noContrSubtypeEquiv.symm i))
|
||||
lemma idMap_append_inl {l l2 : IndexList X} (i : Fin l.length) :
|
||||
(l ++ l2).idMap (appendEquiv (Sum.inl i)) = l.idMap i := by
|
||||
simp [appendEquiv, idMap]
|
||||
rw [List.getElem_append_left]
|
||||
|
||||
@[simp]
|
||||
lemma contrIndexList_numIndices : l.contrIndexList.numIndices = l.noContrFinset.card := by
|
||||
simp [contrIndexList]
|
||||
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]
|
||||
rw [List.getElem_append_right]
|
||||
simp
|
||||
omega
|
||||
omega
|
||||
|
||||
@[simp]
|
||||
lemma contrIndexList_idMap_apply (i : Fin l.contrIndexList.numIndices) :
|
||||
l.contrIndexList.idMap i =
|
||||
l.idMap (l.noContrSubtypeEquiv.symm (Fin.cast (by simp) i)).1 := by
|
||||
simp [contrIndexList, IndexList.fromFinMap, IndexList.idMap]
|
||||
rfl
|
||||
lemma colorMap_append_inl {l l2 : IndexList X} (i : Fin l.length) :
|
||||
(l ++ l2).colorMap (appendEquiv (Sum.inl i)) = l.colorMap i := by
|
||||
simp [appendEquiv, colorMap]
|
||||
rw [List.getElem_append_left]
|
||||
|
||||
lemma contrIndexList_hasNoContr : HasNoContr l.contrIndexList := by
|
||||
intro i
|
||||
simp [NoContr]
|
||||
intro j h
|
||||
refine l.idMap_noContrSubtypeEquiv_neq _ _ ?_
|
||||
rw [@Fin.ne_iff_vne]
|
||||
simp only [Fin.coe_cast, ne_eq]
|
||||
exact Fin.val_ne_of_ne h
|
||||
|
||||
/-- Contracting indices on a index list that has no contractions does nothing. -/
|
||||
@[simp]
|
||||
lemma contrIndexList_of_hasNoContr (h : HasNoContr l) : l.contrIndexList = l := by
|
||||
simp only [contrIndexList, List.get_eq_getElem]
|
||||
have hn : (@Finset.univ (Fin (List.length l)) (Fin.fintype (List.length l))).card =
|
||||
(Finset.filter l.NoContr Finset.univ).card := by
|
||||
rw [Finset.filter_true_of_mem (fun a _ => h a)]
|
||||
have hx : (Finset.filter l.NoContr Finset.univ).card = (List.length l) := by
|
||||
rw [← hn]
|
||||
exact Finset.card_fin (List.length l)
|
||||
apply List.ext_get
|
||||
simpa [fromFinMap, noContrFinset] using hx
|
||||
intro n h1 h2
|
||||
simp only [noContrFinset, noContrSubtypeEquiv, OrderIso.toEquiv_symm, Equiv.symm_trans_apply,
|
||||
RelIso.coe_fn_toEquiv, Equiv.subtypeEquivRight_symm_apply_coe, fromFinMap, List.get_eq_getElem,
|
||||
OrderIso.symm_symm, Finset.coe_orderIsoOfFin_apply, List.getElem_map, Fin.getElem_list,
|
||||
Fin.cast_mk]
|
||||
simp only [Finset.filter_true_of_mem (fun a _ => h a)]
|
||||
congr
|
||||
rw [(Finset.orderEmbOfFin_unique' _
|
||||
(fun x => Finset.mem_univ ((Fin.castOrderIso hx).toOrderEmbedding x))).symm]
|
||||
rfl
|
||||
|
||||
/-- Applying contrIndexlist is equivalent to applying it once. -/
|
||||
@[simp]
|
||||
lemma contrIndexList_contrIndexList : l.contrIndexList.contrIndexList = l.contrIndexList :=
|
||||
l.contrIndexList.contrIndexList_of_hasNoContr (l.contrIndexList_hasNoContr)
|
||||
|
||||
/-!
|
||||
|
||||
## Pairs of contracting indices
|
||||
|
||||
-/
|
||||
|
||||
/-- The set of contracting ordered pairs of indices. -/
|
||||
def contrPairSet : Set (l.contrSubtype × l.contrSubtype) :=
|
||||
{p | p.1.1 < p.2.1 ∧ l.idMap p.1.1 = l.idMap p.2.1}
|
||||
|
||||
instance : DecidablePred fun x => x ∈ l.contrPairSet := fun _ =>
|
||||
And.decidable
|
||||
|
||||
instance : Fintype l.contrPairSet := setFintype _
|
||||
|
||||
lemma contrPairSet_isEmpty_of_hasNoContr (h : l.HasNoContr) :
|
||||
IsEmpty l.contrPairSet := by
|
||||
simp only [contrPairSet, Subtype.coe_lt_coe, Set.coe_setOf, isEmpty_subtype, not_and, Prod.forall]
|
||||
refine fun a b h' => h a.1 b.1 (Fin.ne_of_lt h')
|
||||
|
||||
lemma getDual_lt_self_mem_contrPairSet {i : l.contrSubtype}
|
||||
(h : (l.getDual i).1 < i.1) : (l.getDual i, i) ∈ l.contrPairSet :=
|
||||
And.intro h (l.getDual_id i).symm
|
||||
|
||||
lemma getDual_not_lt_self_mem_contrPairSet {i : l.contrSubtype}
|
||||
(h : ¬ (l.getDual i).1 < i.1) : (i, l.getDual i) ∈ l.contrPairSet := by
|
||||
apply And.intro
|
||||
have h1 := l.getDual_neq_self i
|
||||
simp only [Subtype.coe_lt_coe, gt_iff_lt]
|
||||
simp at h
|
||||
exact lt_of_le_of_ne h h1
|
||||
simp only
|
||||
exact l.getDual_id i
|
||||
|
||||
/-- The list of elements of `l` which contract together. -/
|
||||
def contrPairList : List (Fin l.length × Fin l.length) :=
|
||||
(List.product (Fin.list l.length) (Fin.list l.length)).filterMap fun (i, j) => if
|
||||
l.getDualProp i j then some (i, j) else none
|
||||
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]
|
||||
rw [List.getElem_append_right]
|
||||
simp
|
||||
omega
|
||||
omega
|
||||
|
||||
end IndexList
|
||||
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue