commit
3b00f963bf
36 changed files with 204 additions and 148 deletions
|
@ -181,6 +181,7 @@ lemma consAddContract_surjective_on_zero_contract (i : Fin n.succ)
|
|||
rcases h with h | h
|
||||
· obtain ⟨b, hb, rfl⟩ := h
|
||||
rw [Finset.mapEmbedding_apply, Finset.mapEmbedding_apply]
|
||||
simp only [succ_eq_add_one, Finset.mem_filter, Finset.mem_univ, true_and, c'] at hb
|
||||
exact hb
|
||||
· subst h
|
||||
rw [← h2]
|
||||
|
@ -203,9 +204,9 @@ lemma consAddContract_surjective_on_zero_contract (i : Fin n.succ)
|
|||
obtain ⟨y, rfl⟩ := (Fin.exists_succAbove_eq (x := y) (y := i)) (by omega)
|
||||
use {x, y}
|
||||
simp only [Finset.map_insert, Fin.succAboveEmb_apply, Finset.map_singleton, Fin.val_succEmb,
|
||||
h, true_and]
|
||||
h, true_and, c']
|
||||
rw [Finset.mapEmbedding_apply, Finset.mapEmbedding_apply]
|
||||
simp
|
||||
simpa using h
|
||||
|
||||
lemma consAddContract_bijection (i : Fin n.succ) :
|
||||
Function.Bijective (fun c => (⟨(consAddContract i c), by simp⟩ :
|
||||
|
|
|
@ -220,7 +220,7 @@ lemma insertAndContract_none_prod_contractions (φ : 𝓕.FieldOp) (φs : List
|
|||
let e2 := Equiv.ofBijective (congrLift (insertIdx_length_fin φ φs i).symm)
|
||||
((φsΛ.insertAndContractNat i none).congrLift_bijective ((insertIdx_length_fin φ φs i).symm))
|
||||
erw [← e2.prod_comp]
|
||||
erw [← e1.prod_comp]
|
||||
rw [← e1.prod_comp]
|
||||
rfl
|
||||
|
||||
lemma insertAndContract_some_prod_contractions (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
|
||||
|
@ -233,7 +233,7 @@ lemma insertAndContract_some_prod_contractions (φ : 𝓕.FieldOp) (φs : List
|
|||
((φsΛ.insertAndContractNat i (some j)).congrLift_bijective ((insertIdx_length_fin φ φs i).symm))
|
||||
erw [← e2.prod_comp]
|
||||
let e1 := Equiv.ofBijective (φsΛ.insertLiftSome i j) (insertLiftSome_bijective i j)
|
||||
erw [← e1.prod_comp]
|
||||
rw [← e1.prod_comp]
|
||||
rw [Fintype.prod_sum_type]
|
||||
simp only [Finset.univ_unique, PUnit.default_eq_unit, Nat.succ_eq_add_one, Finset.prod_singleton,
|
||||
Finset.univ_eq_attach]
|
||||
|
|
|
@ -276,7 +276,7 @@ def uncontractedIndexEquiv (c : WickContraction n) :
|
|||
Fin (c.uncontractedList).length ≃ c.uncontracted where
|
||||
toFun i := ⟨c.uncontractedList.get i, c.uncontractedList_get_mem_uncontracted i⟩
|
||||
invFun i := ⟨List.indexOf i.1 c.uncontractedList,
|
||||
List.indexOf_lt_length.mpr ((c.uncontractedList_mem_iff i.1).mpr i.2)⟩
|
||||
List.indexOf_lt_length_iff.mpr ((c.uncontractedList_mem_iff i.1).mpr i.2)⟩
|
||||
left_inv i := by
|
||||
ext
|
||||
exact List.get_indexOf (uncontractedList_nodup c) _
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue