PhysLean/HepLean/AnomalyCancellation/PureU1/Sort.lean

79 lines
2.2 KiB
Text
Raw Normal View History

2024-04-18 10:09:08 -04:00
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Authors: Joseph Tooby-Smith
-/
import HepLean.AnomalyCancellation.PureU1.Permutations
2024-06-25 07:06:32 -04:00
import Mathlib.Data.Fin.Tuple.Sort
2024-04-18 10:09:08 -04:00
/-!
# Sort for Pure U(1) charges
We define the sort function for Pure U(1) charges, and prove some basic properties.
-/
universe v u
open Nat
2024-07-12 11:23:02 -04:00
open Finset
2024-04-18 10:09:08 -04:00
namespace PureU1
variable {n : }
/-- We say a charge is shorted if for all `i ≤ j`, then `S i ≤ S j`. -/
@[simp]
2024-07-12 11:23:02 -04:00
def Sorted {n : } (S : (PureU1 n).Charges) : Prop :=
∀ i j (_ : i ≤ j), S i ≤ S j
2024-04-18 10:09:08 -04:00
/-- Given a charge assignment `S`, the corresponding sorted charge assignment. -/
@[simp]
def sort {n : } (S : (PureU1 n).Charges) : (PureU1 n).Charges :=
2024-04-18 10:09:08 -04:00
((FamilyPermutations n).rep (Tuple.sort S).symm S)
2024-07-12 11:23:02 -04:00
lemma sort_sorted {n : } (S : (PureU1 n).Charges) : Sorted (sort S) := by
simp only [Sorted, PureU1_numberCharges, sort, FamilyPermutations, PermGroup, permCharges,
2024-04-18 10:12:55 -04:00
MonoidHom.coe_mk, OneHom.coe_mk, chargeMap_apply]
2024-04-18 10:09:08 -04:00
intro i j hij
exact Tuple.monotone_sort S hij
lemma sort_perm {n : } (S : (PureU1 n).Charges) (M :(FamilyPermutations n).group) :
2024-04-18 10:09:08 -04:00
sort ((FamilyPermutations n).rep M S) = sort S :=
@Tuple.comp_perm_comp_sort_eq_comp_sort n _ S M⁻¹
lemma sort_apply {n : } (S : (PureU1 n).Charges) (j : Fin n) :
2024-04-18 10:09:08 -04:00
sort S j = S ((Tuple.sort S) j) := by
rfl
2024-07-12 11:23:02 -04:00
lemma sort_zero {n : } (S : (PureU1 n).Charges) (hS : sort S = 0) : S = 0 := by
2024-04-18 10:09:08 -04:00
funext i
have hj : ∀ j, sort S j = 0 := by
rw [hS]
intro j
rfl
have hi := hj ((Tuple.sort S).invFun i)
rw [sort_apply] at hi
simp at hi
rw [hi]
rfl
lemma sort_projection {n : } (S : (PureU1 n).Charges) : sort (sort S) = sort S :=
2024-04-18 10:09:08 -04:00
sort_perm S (Tuple.sort S).symm
/-- The sort function acting on `LinSols`. -/
2024-07-12 11:23:02 -04:00
def sortAFL {n : } (S : (PureU1 n).LinSols) : (PureU1 n).LinSols :=
2024-04-18 10:09:08 -04:00
((FamilyPermutations n).linSolRep (Tuple.sort S.val).symm S)
2024-07-12 11:23:02 -04:00
lemma sortAFL_val {n : } (S : (PureU1 n).LinSols) : (sortAFL S).val = sort S.val := by
2024-04-18 10:09:08 -04:00
rfl
2024-07-12 11:23:02 -04:00
lemma sortAFL_zero {n : } (S : (PureU1 n).LinSols) (hS : sortAFL S = 0) : S = 0 := by
2024-04-18 10:09:08 -04:00
apply ACCSystemLinear.LinSols.ext
have h1 : sort S.val = 0 := by
rw [← sortAFL_val]
rw [hS]
rfl
exact sort_zero S.val h1
end PureU1