chore: bump toolchain to v4.15.0

#281 adapt code to v4.15.0 and fix long heartbeats, e.g., toDualRep_apply_eq_contrOneTwoLeft.

---------

Co-authored-by: jstoobysmith <72603918+jstoobysmith@users.noreply.github.com>
This commit is contained in:
KUO-TSAN HSU (Gordon) 2025-01-20 15:42:53 +08:00 committed by GitHub
parent 6e31281a5b
commit 656a3e422f
No known key found for this signature in database
GPG key ID: B5690EEEBB952194
49 changed files with 484 additions and 472 deletions

View file

@ -49,13 +49,13 @@ def uncontractedFromInvolution : {φs : List 𝓕} →
if hn : n' = none then
have hn' := involutionAddEquiv_none_image_zero (n := φs.length) (f := f) hn
⟨optionEraseZ luc φ none, by
simp only [optionEraseZ, Nat.succ_eq_add_one, List.length_cons, Mathlib.Vector.length_val]
simp only [optionEraseZ, Nat.succ_eq_add_one, List.length_cons, List.Vector.length_val]
rw [← luc.2]
conv_rhs => rw [Finset.card_filter]
rw [Fin.sum_univ_succ]
conv_rhs => erw [if_pos hn']
ring_nf
simp only [Nat.succ_eq_add_one, Mathlib.Vector.length_val, Nat.cast_id,
simp only [Nat.succ_eq_add_one, List.Vector.length_val, Nat.cast_id,
add_right_inj]
rw [Finset.card_filter]
apply congrArg
@ -85,7 +85,7 @@ def uncontractedFromInvolution : {φs : List 𝓕} →
conv_rhs => rw [hksucc]
exact fun hn => Fin.succ_ne_zero k (Function.Involutive.injective f.2 hn)
have hluc : 1 ≤ luc.1.length := by
simp only [Nat.succ_eq_add_one, Mathlib.Vector.length_val, Finset.one_le_card]
simp only [Nat.succ_eq_add_one, List.Vector.length_val, Finset.one_le_card]
use k
simp only [involutionCons, Nat.succ_eq_add_one, Fin.cons_update, Equiv.coe_fn_mk,
dite_eq_left_iff, Finset.mem_filter, Finset.mem_univ, true_and]
@ -95,7 +95,7 @@ def uncontractedFromInvolution : {φs : List 𝓕} →
conv_rhs =>
rw [Finset.card_filter]
erw [Fin.sum_univ_succ, if_neg (Option.isSome_dite'.mp hkIsSome)]
simp only [Nat.succ_eq_add_one, Mathlib.Vector.length_val, List.length_cons,
simp only [Nat.succ_eq_add_one, List.Vector.length_val, List.length_cons,
Nat.cast_id, zero_add]
conv_rhs => lhs; rw [Eq.symm (Fintype.sum_ite_eq' k fun j => 1)]
rw [← Finset.sum_add_distrib, Finset.card_filter]
@ -275,7 +275,7 @@ lemma toInvolution_length_uncontracted {φs φsᵤₙ : List 𝓕} {c : Contract
have h2 := (toInvolution ⟨φsᵤₙ, c⟩).2
simp only at h2
conv_lhs => rw [← h2]
exact Mathlib.Vector.length_val (uncontractedFromInvolution (toInvolution ⟨φsᵤₙ, c⟩).1)
exact List.Vector.length_val (uncontractedFromInvolution (toInvolution ⟨φsᵤₙ, c⟩).1)
lemma toInvolution_cons {φs φsᵤₙ : List 𝓕} {φ : 𝓕}
(c : ContractionsAux φs φsᵤₙ) (n : Option (Fin (φsᵤₙ.length))) :

View file

@ -213,7 +213,7 @@ instance preimageVertexMapFintype [IsFinitePreFeynmanRule P] {𝓔 𝓥 : Type}
[DecidableEq 𝓥] (v : 𝓥) (f : 𝓥 ⟶ P.VertexLabel) (F : Over (P.HalfEdgeLabel × 𝓔 × 𝓥))
[Fintype F.left] :
Fintype ((P.vertexLabelMap (f v)).left → ((P.preimageVertex v).obj F).left) :=
Pi.fintype
Pi.instFintype
/-- Given an edge, there is a finite number of maps between the indexing set of the
expected half-edges corresponding to that edges label, and the actual indexing
@ -222,7 +222,7 @@ instance preimageEdgeMapFintype [IsFinitePreFeynmanRule P] {𝓔 𝓥 : Type}
[DecidableEq 𝓔] (v : 𝓔) (f : 𝓔 ⟶ P.EdgeLabel) (F : Over (P.HalfEdgeLabel × 𝓔 × 𝓥))
[Fintype F.left] :
Fintype ((P.edgeLabelMap (f v)).left → ((P.preimageEdge v).obj F).left) :=
Pi.fintype
Pi.instFintype
/-!

View file

@ -30,7 +30,7 @@ variable {𝓕 : Type} {f : 𝓕 → Type} [∀ i, Fintype (f i)] {φs : List
(a : CreateAnnihilateSect f φs)
/-- The type `CreateAnnihilateSect f φs` is finite. -/
instance fintype : Fintype (CreateAnnihilateSect f φs) := Pi.fintype
instance fintype : Fintype (CreateAnnihilateSect f φs) := Pi.instFintype
/-- The section got by dropping the first element of `φs` if it exists. -/
def tail : {φs : List 𝓕} → (a : CreateAnnihilateSect f φs) → CreateAnnihilateSect f φs.tail

View file

@ -82,8 +82,8 @@ lemma koszulSign_insertIdx [IsTotal 𝓕 le] [IsTrans 𝓕 le] (φ : 𝓕) :
(φs : List 𝓕) → (n : ) → (hn : n ≤ φs.length) →
koszulSign q le (List.insertIdx n φ φs) = insertSign q n φ φs * koszulSign q le φs *
insertSign q (insertionSortEquiv le (List.insertIdx n φ φs) ⟨n, by
rw [List.length_insertIdx _ _ hn]
omega⟩) φ (List.insertionSort le (List.insertIdx n φ φs))
rw [List.length_insertIdx, if_pos hn]
exact Nat.succ_le_succ hn⟩) φ (List.insertionSort le (List.insertIdx n φ φs))
| [], 0, h => by
simp [koszulSign, insertSign, superCommuteCoef, koszulSignInsert]
| [], n + 1, h => by
@ -138,8 +138,8 @@ lemma koszulSign_insertIdx [IsTotal 𝓕 le] [IsTrans 𝓕 le] (φ : 𝓕) :
have hnsL : n < (List.insertIdx n φ φs).length := by
rw [List.length_insertIdx _ _]
simp only [List.length_cons, add_le_add_iff_right] at h
omega
exact Nat.le_of_lt_succ h
rw [if_pos h]
exact Nat.succ_le_succ h
let ni : Fin rs.length := (insertionSortEquiv le (List.insertIdx n φ φs))
⟨n, hnsL⟩
let nro : Fin (rs.length + 1) :=

View file

@ -83,11 +83,11 @@ lemma koszulSignInsert_eq_filter (φ : 𝓕) : (φs : List 𝓕) →
dsimp only [koszulSignInsert, Fin.isValue]
simp only [Fin.isValue, List.filter, decide_not]
by_cases h : le φ φ1
· simp only [h, ↓reduceIte, decide_True, Bool.not_true]
· simp only [h, ↓reduceIte, decide_true, Bool.not_true]
rw [koszulSignInsert_eq_filter]
congr
simp
· simp only [h, ↓reduceIte, Fin.isValue, decide_False, Bool.not_false]
· simp only [h, ↓reduceIte, Fin.isValue, decide_false, Bool.not_false]
dsimp only [Fin.isValue, koszulSignInsert]
simp only [Fin.isValue, h, ↓reduceIte]
rw [koszulSignInsert_eq_filter]
@ -176,7 +176,7 @@ lemma koszulSignInsert_eq_insertSign [IsTotal 𝓕 le] [IsTrans 𝓕 le] (φ :
simp_all
rw [h1]
rw [List.filter_cons]
simp only [decide_not, (IsTotal.to_isRefl le).refl φ, not_true_eq_false, decide_False,
simp only [decide_not, (IsTotal.to_isRefl le).refl φ, not_true_eq_false, decide_false,
Bool.false_eq_true, ↓reduceIte]
rw [orderedInsertPos_take]
simp only [decide_not, List.append_right_eq_self, List.filter_eq_nil_iff, Bool.not_eq_eq_eq_not,