Merge pull request #330 from HEPLean/bump

chore: bump to lean v4.16.0
This commit is contained in:
Joseph Tooby-Smith 2025-02-12 15:35:22 +00:00 committed by GitHub
commit 3b00f963bf
No known key found for this signature in database
GPG key ID: B5690EEEBB952194
36 changed files with 204 additions and 148 deletions

View file

@ -3,7 +3,7 @@ Copyright (c) 2025 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import Mathlib.Algebra.BigOperators.Group.Finset
import Mathlib.Algebra.BigOperators.Group.Finset.Basic
/-!
# Creation and annihilation parts of fields
@ -66,7 +66,7 @@ lemma not_normalOrder_annihilate_iff_false (a : CreateAnnihilate) :
lemma sum_eq {M : Type} [AddCommMonoid M] (f : CreateAnnihilate → M) :
∑ i, f i = f create + f annihilate := by
change ∑ i in {create, annihilate}, f i = f create + f annihilate
change ∑ i {create, annihilate}, f i = f create + f annihilate
simp
@[simp]

View file

@ -84,7 +84,7 @@ lemma koszulSignInsert_append_annihilate (φ' φ : 𝓕.CrAnFieldOp)
Wick.koszulSignInsert 𝓕.crAnStatistics normalOrderRel φ' (φs ++ [φ]) =
Wick.koszulSignInsert 𝓕.crAnStatistics normalOrderRel φ' φs
| [] => by
simp only [Wick.koszulSignInsert, normalOrderRel, hφ, ite_eq_left_iff,
simp only [List.nil_append, Wick.koszulSignInsert, normalOrderRel, hφ, ite_eq_left_iff,
CreateAnnihilate.not_normalOrder_annihilate_iff_false, ite_eq_right_iff, and_imp,
IsEmpty.forall_iff]
| φ'' :: φs => by
@ -281,6 +281,7 @@ lemma normalOrderList_append_annihilate (φ : 𝓕.CrAnFieldOp)
simp only [normalOrderList, List.insertionSort, List.append_eq]
have hi := normalOrderList_append_annihilate φ hφ φs
dsimp only [normalOrderList] at hi
simp only [List.cons_append, List.insertionSort]
rw [hi, orderedInsert_append_annihilate φ' φ hφ]
lemma normalOrder_swap_create_annihilate_fst (φc φa : 𝓕.CrAnFieldOp)

View file

@ -308,15 +308,15 @@ lemma orderedInsert_in_swap_eq_time {φ ψ φ': 𝓕.CrAnFieldOp} (h1 : crAnTime
have h1 (b : 𝓕.CrAnFieldOp) : (crAnTimeOrderRel b φ) ↔ (crAnTimeOrderRel b ψ) :=
Iff.intro (fun h => IsTrans.trans _ _ _ h h1) (fun h => IsTrans.trans _ _ _ h h2)
by_cases h : crAnTimeOrderRel φ' φ
· simp only [List.orderedInsert, h, ↓reduceIte, ← h1 φ']
· simp only [List.nil_append, List.orderedInsert, h, ↓reduceIte, ← h1 φ']
use [φ'], φs'
simp
· simp only [List.orderedInsert, h, ↓reduceIte, ← h1 φ']
· simp only [List.nil_append, List.orderedInsert, h, ↓reduceIte, ← h1 φ']
use [], List.orderedInsert crAnTimeOrderRel φ' φs'
simp
| φ'' :: φs, φs' => by
obtain ⟨l1, l2, hl⟩ := orderedInsert_in_swap_eq_time (φ' := φ') h1 h2 φs φs'
simp only [List.orderedInsert, List.append_eq]
simp only [List.cons_append, List.orderedInsert]
rw [hl.1, hl.2]
by_cases h : crAnTimeOrderRel φ' φ''
· simp only [h, ↓reduceIte]
@ -334,7 +334,7 @@ lemma crAnTimeOrderList_swap_eq_time {φ ψ : 𝓕.CrAnFieldOp}
crAnTimeOrderList (φs ++ ψ :: φ :: φs') = l1 ++ ψ :: φ :: l2
| [], φs' => by
simp only [crAnTimeOrderList]
simp only [List.insertionSort]
simp only [List.nil_append, List.insertionSort]
use List.takeWhile (fun b => ¬ crAnTimeOrderRel ψ b) (List.insertionSort crAnTimeOrderRel φs'),
List.dropWhile (fun b => ¬ crAnTimeOrderRel ψ b) (List.insertionSort crAnTimeOrderRel φs')
apply And.intro
@ -345,7 +345,7 @@ lemma crAnTimeOrderList_swap_eq_time {φ ψ : 𝓕.CrAnFieldOp}
simpa using orderedInsert_swap_eq_time h2 h1 _
| φ'' :: φs, φs' => by
rw [crAnTimeOrderList, crAnTimeOrderList]
simp only [List.insertionSort, List.append_eq]
simp only [List.cons_append, List.insertionSort]
obtain ⟨l1, l2, hl⟩ := crAnTimeOrderList_swap_eq_time h1 h2 φs φs'
simp only [crAnTimeOrderList] at hl
rw [hl.1, hl.2]

View file

@ -184,7 +184,7 @@ lemma ofList_append (s : 𝓕 → FieldStatistic) (φs φs' : List 𝓕) :
(if a = (if b = c then bosonic else fermionic) then bosonic else fermionic) =
if (if a = b then bosonic else fermionic) = c then bosonic else fermionic := by
fin_cases a <;> fin_cases b <;> fin_cases c <;> rfl
simp only [ofList, List.append_eq, Fin.isValue, ih, hab]
simp only [List.cons_append, ofList, ih, hab]
lemma ofList_append_eq_mul (s : 𝓕 → FieldStatistic) (φs φs' : List 𝓕) :
ofList s (φs ++ φs') = ofList s φs * ofList s φs' := by

View file

@ -42,6 +42,7 @@ lemma koszulSign_mul_self (l : List 𝓕) : koszulSign q le l * koszulSign q le
@[simp]
lemma koszulSign_freeMonoid_of (φ : 𝓕) : koszulSign q le (FreeMonoid.of φ) = 1 := by
change koszulSign q le [φ] = 1
simp only [koszulSign, mul_one]
rfl
@ -235,6 +236,9 @@ lemma koszulSign_eraseIdx [IsTotal 𝓕 le] [IsTrans 𝓕 le] (φs : List 𝓕)
𝓢(q φs[↑n], ofList q (List.take (↑((insertionSortEquiv le φs) n)) (List.insertionSort le φs))))
swap
· simp only [Fin.getElem_fin]
rw [Equiv.trans_apply, Equiv.trans_apply]
simp only [instCommGroup.eq_1, mul_one, Fin.castOrderIso,
Equiv.coe_fn_mk, Fin.cast_mk, Fin.eta, Fin.coe_cast]
ring
conv_rhs =>
rhs
@ -274,7 +278,7 @@ lemma koszulSign_swap_eq_rel {ψ φ : 𝓕} (h1 : le φ ψ) (h2 : le ψ φ) : (
simp only [List.nil_append]
exact koszulSign_swap_eq_rel_cons q le h1 h2 φs'
| φ'' :: φs, φs' => by
simp only [Wick.koszulSign, List.append_eq]
simp only [List.cons_append, koszulSign]
rw [koszulSign_swap_eq_rel h1 h2]
congr 1
apply Wick.koszulSignInsert_eq_perm
@ -301,7 +305,7 @@ lemma koszulSign_eq_rel_eq_stat {ψ φ : 𝓕} [IsTrans 𝓕 le]
simp only [List.nil_append]
exact koszulSign_eq_rel_eq_stat_append q le h1 h2 hq φs
| φ'' :: φs', φs => by
simp only [koszulSign, List.append_eq]
simp only [List.cons_append, koszulSign]
rw [koszulSign_eq_rel_eq_stat h1 h2 hq φs' φs]
simp only [mul_eq_mul_right_iff]
left
@ -381,7 +385,7 @@ lemma koszulSign_of_append_eq_insertionSort [IsTotal 𝓕 le] [IsTrans 𝓕 le]
simp only [List.nil_append]
exact koszulSign_of_append_eq_insertionSort_left q le φs φs'
| φ'' :: φs'', φs, φs' => by
simp only [koszulSign, List.append_eq]
simp only [List.cons_append, koszulSign]
rw [koszulSign_of_append_eq_insertionSort φs'' φs φs', ← mul_assoc]
congr 2
apply koszulSignInsert_eq_perm
@ -430,7 +434,7 @@ lemma koszulSign_perm_eq [IsTrans 𝓕 le] (φ : 𝓕) : (φs1 φs φs' φs2 : L
simp only [List.nil_append]
exact koszulSign_perm_eq_append q le φ φs φs' φs2 hp h
| φ1 :: φs1, φs, φs', φs2, h, hp => by
simp only [koszulSign, List.append_eq]
simp only [List.cons_append, koszulSign]
have ih := koszulSign_perm_eq φ φs1 φs φs' φs2 h hp
rw [ih]
congr 1

View file

@ -69,7 +69,7 @@ lemma koszulSignInsert_ge_forall_append (φs : List 𝓕) (φ' φ : 𝓕) (hi :
induction φs with
| nil => simp [koszulSignInsert, hi]
| cons φ'' φs ih =>
simp only [koszulSignInsert, Fin.isValue, List.append_eq]
simp only [koszulSignInsert, List.cons_append]
by_cases hr : le φ' φ''
· rw [if_pos hr, if_pos hr, ih]
· rw [if_neg hr, if_neg hr, ih]

View file

@ -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⟩ :

View file

@ -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]

View file

@ -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) _