feat: Remove more erws

This commit is contained in:
jstoobysmith 2025-03-24 11:50:41 -04:00
parent becf9d1841
commit 70ace79ccc
9 changed files with 47 additions and 29 deletions

View file

@ -158,7 +158,7 @@ lemma succAbove_succAbove_predAboveI (i : Fin n.succ.succ) (j : Fin n.succ) (x :
/-- The equivalence between `Fin n.succ` and `Fin 1 ⊕ Fin n` extracting the
`i`th component. -/
def finExtractOne {n : } (i : Fin n.succ) : Fin n.succ ≃ Fin 1 ⊕ Fin n :=
def finExtractOne {n : } (i : Fin (n + 1)) : Fin (n + 1) ≃ Fin 1 ⊕ Fin n :=
(finCongr (by omega : n.succ = i + 1 + (n - i))).trans <|
finSumFinEquiv.symm.trans <|
(Equiv.sumCongr (finSumFinEquiv.symm.trans (Equiv.sumComm (Fin i) (Fin 1)))
@ -231,7 +231,7 @@ lemma finExtractOne_symm_inl_apply {n : } (i : Fin n.succ) :
(finExtractOne i).symm (Sum.inl 0) = i := by
rfl
lemma finExtractOne_apply_neq {n : } (i j : Fin n.succ.succ) (hij : i ≠ j) :
lemma finExtractOne_apply_neq {n : } (i j : Fin (n + 1 + 1)) (hij : i ≠ j) :
finExtractOne i j = Sum.inr (predAboveI i j) := by
symm
apply (Equiv.symm_apply_eq _).mp ?_

View file

@ -344,7 +344,7 @@ lemma orderedInsert_eraseIdx_orderedInsertPos_le {I : Type} (le1 : I → I → P
/-- The equivalence between `Fin (r0 :: r).length` and `Fin (List.orderedInsert le1 r0 r).length`
according to where the elements map, i.e. `0` is taken to `orderedInsertPos le1 r r0`. -/
def orderedInsertEquiv {I : Type} (le1 : I → I → Prop) [DecidableRel le1] (r : List I) (r0 : I) :
Fin (r0 :: r).length ≃ Fin (List.orderedInsert le1 r0 r).length := by
Fin (r.length + 1) ≃ Fin (List.orderedInsert le1 r0 r).length := by
let e2 : Fin (List.orderedInsert le1 r0 r).length ≃ Fin (r0 :: r).length :=
(Fin.castOrderIso (List.orderedInsert_length le1 r r0)).toEquiv
let e3 : Fin (r0 :: r).length ≃ Fin 1 ⊕ Fin (r).length := finExtractOne 0
@ -352,8 +352,9 @@ def orderedInsertEquiv {I : Type} (le1 : I → I → Prop) [DecidableRel le1] (r
finExtractOne ⟨orderedInsertPos le1 r r0, orderedInsertPos_lt_length le1 r r0⟩
exact e3.trans (e4.symm.trans e2.symm)
@[simp]
lemma orderedInsertEquiv_zero {I : Type} (le1 : I → I → Prop) [DecidableRel le1] (r : List I)
(r0 : I) : orderedInsertEquiv le1 r r0 ⟨0, by simp⟩ = orderedInsertPos le1 r r0 := by
(r0 : I) : orderedInsertEquiv le1 r r0 0 = orderedInsertPos le1 r r0 := by
simp [orderedInsertEquiv]
lemma orderedInsertEquiv_succ {I : Type} (le1 : I → I → Prop) [DecidableRel le1] (r : List I)
@ -368,7 +369,8 @@ lemma orderedInsertEquiv_succ {I : Type} (le1 : I → I → Prop) [DecidableRel
| [] =>
simp
| r1 :: r =>
erw [finExtractOne_apply_neq]
simp only [List.orderedInsert.eq_2, List.length_cons, Fin.cast_inj]
rw [finExtractOne_apply_neq]
simp only [List.orderedInsert.eq_2, List.length_cons, orderedInsertPos, decide_not,
Nat.succ_eq_add_one, finExtractOne_symm_inr_apply]
rfl
@ -384,9 +386,12 @@ lemma orderedInsertEquiv_fin_succ {I : Type} (le1 : I → I → Prop) [Decidable
| [] =>
simp
| r1 :: r =>
erw [finExtractOne_apply_neq]
simp only [List.orderedInsert.eq_2, List.length_cons, orderedInsertPos, decide_not,
Nat.succ_eq_add_one, finExtractOne_symm_inr_apply]
simp only [List.orderedInsert.eq_2, List.length_cons, OrderIso.toEquiv_symm,
Fin.symm_castOrderIso, Nat.succ_eq_add_one, RelIso.coe_fn_toEquiv, Fin.castOrderIso_apply,
Fin.eta, Fin.cast_inj]
rw [finExtractOne_apply_neq]
simp only [orderedInsertPos, List.orderedInsert.eq_2, decide_not, Nat.succ_eq_add_one,
finExtractOne_symm_inr_apply]
rfl
exact ne_of_beq_false rfl
@ -413,14 +418,11 @@ lemma get_eq_orderedInsertEquiv {I : Type} (le1 : I → I → Prop) [DecidableRe
funext x
match x with
| ⟨0, h⟩ =>
simp only [List.length_cons, Fin.zero_eta, List.get_eq_getElem, Fin.val_zero,
List.getElem_cons_zero, Function.comp_apply]
erw [orderedInsertEquiv_zero]
simp
| ⟨Nat.succ n, h⟩ =>
simp only [List.length_cons, Nat.succ_eq_add_one, List.get_eq_getElem, List.getElem_cons_succ,
Function.comp_apply]
erw [orderedInsertEquiv_succ]
rw [orderedInsertEquiv_succ]
simp only [Fin.succAbove, Fin.castSucc_mk, Fin.mk_lt_mk, Fin.succ_mk, Fin.coe_cast]
by_cases hn : n < ↑(orderedInsertPos le1 r r0)
· simp [hn, orderedInsert_get_lt]
@ -501,7 +503,7 @@ lemma orderedInsertEquiv_sigma {I : Type} {f : I → Type}
| ⟨0, h0⟩ =>
simp only [List.length_cons, Fin.zero_eta, Equiv.trans_apply, RelIso.coe_fn_toEquiv,
Fin.castOrderIso_apply, Fin.cast_zero, Fin.coe_cast]
erw [orderedInsertEquiv_zero, orderedInsertEquiv_zero]
rw [orderedInsertEquiv_zero, orderedInsertEquiv_zero]
simp [orderedInsertPos_sigma]
| ⟨Nat.succ n, h0⟩ =>
simp only [List.length_cons, Nat.succ_eq_add_one, Equiv.trans_apply, RelIso.coe_fn_toEquiv,

View file

@ -138,7 +138,7 @@ def accGrav : MSSMCharges.Charges →ₗ[] where
repeat rw [map_add]
simp only [MSSMSpecies_numberCharges, ACCSystemCharges.chargesAddCommMonoid_add,
toSMSpecies_apply, reduceMul, Fin.isValue, mul_add, Hd_apply, Fin.reduceFinMk, Hu_apply]
repeat erw [Finset.sum_add_distrib]
repeat rw [Finset.sum_add_distrib]
ring
map_smul' a S := by
simp only
@ -146,9 +146,8 @@ def accGrav : MSSMCharges.Charges →ₗ[] where
erw [Hd.map_smul, Hu.map_smul]
simp only [MSSMSpecies_numberCharges, HSMul.hSMul, SMul.smul, Fin.isValue, toSMSpecies_apply,
reduceMul, Hd_apply, Fin.reduceFinMk, Hu_apply, eq_ratCast, Rat.cast_eq_id, id_eq]
repeat erw [Finset.sum_add_distrib]
repeat erw [← Finset.mul_sum]
--rw [show Rat.cast a = a from rfl]
repeat rw [Finset.sum_add_distrib]
repeat rw [← Finset.mul_sum]
ring
/-- Extensionality lemma for `accGrav`. -/

View file

@ -119,7 +119,10 @@ lemma cube_proj_proj_Y₃ (T : MSSMACC.LinSols) :
(dot Y₃.val B₃.val)^2 * cubeTriLin T.val T.val Y₃.val := by
rw [proj_val]
rw [cubeTriLin.map_add₁, cubeTriLin.map_add₂]
erw [lineY₃B₃_doublePoint]
conv_lhs =>
enter [1, 1]
change ((cubeTriLin (lineY₃B₃ _ _).val) (lineY₃B₃ _ _).val) _
rw [lineY₃B₃_doublePoint]
rw [cubeTriLin.map_add₂]
rw [cubeTriLin.swap₂]
rw [cubeTriLin.map_add₁, cubeTriLin.map_smul₁, cubeTriLin.map_smul₃]

View file

@ -104,7 +104,7 @@ open BigOperators
/-- A solution to the pure U(1) accs satisfies the linear ACCs. -/
lemma pureU1_linear {n : } (S : (PureU1 n).LinSols) :
i, S.val i = 0 := by
(i : Fin n), S.val i = 0 := by
have hS := S.linearSol
simp only [succ_eq_add_one, PureU1_numberLinear, PureU1_linearACCs] at hS
exact hS 0

View file

@ -176,7 +176,7 @@ lemma AFL_hasBoundary (h : A.val (0 : Fin n.succ) ≠ 0) : HasBoundary A.val :=
by_contra hn
have h0 := not_hasBoundary_grav hA hn
simp only [succ_eq_add_one, accGrav, LinearMap.coe_mk, AddHom.coe_mk, cast_add, cast_one] at h0
erw [pureU1_linear A] at h0
rw [pureU1_linear A] at h0
simp only [zero_eq_mul] at h0
cases' h0
· linarith
@ -188,7 +188,7 @@ lemma AFL_odd_noBoundary {A : (PureU1 (2 * n + 1)).LinSols} (h : ConstAbsSorted
obtain ⟨k, hk⟩ := hn
have h0 := boundary_accGrav'' h k hk
simp only [succ_eq_add_one, accGrav, LinearMap.coe_mk, AddHom.coe_mk, cast_mul, cast_ofNat] at h0
erw [pureU1_linear A] at h0
rw [pureU1_linear A] at h0
simp only [zero_eq_mul, hA, or_false] at h0
have h1 : 2 * n = 2 * k.val + 1 := by
rw [← @Nat.cast_inj ]
@ -210,8 +210,10 @@ lemma AFL_even_Boundary {A : (PureU1 (2 * n.succ)).LinSols} (h : ConstAbsSorted
(hA : A.val (0 : Fin (2 * n.succ)) ≠ 0) {k : Fin (2 * n + 1)} (hk : Boundary A.val k) :
k.val = n := by
have h0 := boundary_accGrav'' h k hk
change ∑ i : Fin (succ (Nat.mul 2 n + 1)), A.val i = _ at h0
erw [pureU1_linear A] at h0
change ∑ i, A.val i = _ at h0
simp only [succ_eq_add_one, PureU1_numberCharges, mul_eq, cast_add, cast_mul, cast_ofNat,
cast_one, add_sub_add_right_eq_sub] at h0
rw [pureU1_linear A] at h0
simp only [mul_eq, cast_add, cast_mul, cast_ofNat, cast_one, add_sub_add_right_eq_sub,
succ_eq_add_one, zero_eq_mul, hA, or_false] at h0
rw [← @Nat.cast_inj ]

View file

@ -45,7 +45,14 @@ lemma lineInCubic_expand {S : (PureU1 (2 * n.succ)).LinSols} (h : LineInCubic S)
simp only [TriLinearSymm.toCubic_add] at h1
simp only [HomogeneousCubic.map_smul,
accCubeTriLinSymm.map_smul₁, accCubeTriLinSymm.map_smul₂, accCubeTriLinSymm.map_smul₃] at h1
erw [P_accCube, P!_accCube] at h1
conv_lhs at h1 =>
enter [1, 1, 1, 2]
change accCube _ _
rw [P_accCube]
conv_lhs at h1 =>
enter [1, 1, 2, 2]
change accCube _ _
rw [P!_accCube]
rw [← h1]
ring

View file

@ -48,7 +48,7 @@ def pairIso (c : C) : (pair F).obj (Discrete.mk c) ≅ (lift.obj F).obj (OverCol
/-- The isomorphism between `F.obj (Discrete.mk c1) ⊗ F.obj (Discrete.mk c2)` and
`(lift.obj F).obj (OverColor.mk ![c1,c2])` formed by the tensor. -/
def pairIsoSep {c1 c2 : C} : F.obj (Discrete.mk c1) ⊗ F.obj (Discrete.mk c2) ≅
(lift.obj F).obj (OverColor.mk ![c1,c2]) := by
(lift.obj F).obj (OverColor.mk (![c1 ,c2] : Fin 2 → C)) := by
symm
apply ((lift.obj F).mapIso fin2Iso).trans
apply (Functor.Monoidal.μIso (lift.obj F).toFunctor _ _).symm.trans
@ -67,8 +67,12 @@ def pairIsoSep {c1 c2 : C} : F.obj (Discrete.mk c1) ⊗ F.obj (Discrete.mk c2)
rfl
lemma pairIsoSep_tmul {c1 c2 : C} (x : F.obj (Discrete.mk c1)) (y : F.obj (Discrete.mk c2)) :
(pairIsoSep F).hom.hom (x ⊗ₜ[k] y) =
((pairIsoSep F).hom.hom : (F.obj (Discrete.mk c1)).V ⊗ (F.obj (Discrete.mk c2)).V ⟶ _ ) (x ⊗ₜ[k] y) =
PiTensorProduct.tprod k (Fin.cases x (Fin.cases y (fun i => Fin.elim0 i))) := by
conv_lhs =>
simp only [Action.instMonoidalCategory_tensorObj_V, Nat.succ_eq_add_one, Nat.reduceAdd,
Equivalence.symm_inverse, Action.functorCategoryEquivalence_functor,
Action.FunctorCategoryEquivalence.functor_obj_obj]
simp only [Nat.succ_eq_add_one, Nat.reduceAdd, Action.instMonoidalCategory_tensorObj_V,
pairIsoSep, Fin.isValue, Matrix.cons_val_zero, Matrix.cons_val_one, Matrix.head_cons,
forgetLiftApp, Iso.trans_symm, Iso.symm_symm_eq, Iso.trans_assoc, Iso.trans_hom, Iso.symm_hom,

View file

@ -80,13 +80,14 @@ def coordinateMultiLinear {n : } (c : Fin n → S.C) :
/-- The linear map from tensors to coordinates. -/
def coordinate {n : } (c : Fin n → S.C) :
S.F.obj (OverColor.mk c) →ₗ[k] ((Π j, Fin (S.repDim (c j))) → k) :=
(S.liftTensor (c := c)).toFun (S.coordinateMultiLinear c)
PiTensorProduct.lift (S.coordinateMultiLinear c)
@[simp]
lemma coordinate_tprod {n : } (c : Fin n → S.C) (x : (i : Fin n) → S.FD.obj (Discrete.mk (c i))) :
S.coordinate c (PiTensorProduct.tprod k x) = S.coordinateMultiLinear c x := by
simp only [coordinate, liftTensor, AddHom.toFun_eq_coe, LinearMap.coe_toAddHom,
LinearEquiv.coe_coe, mk_left, Functor.id_obj, mk_hom]
rw [coordinate]
erw [PiTensorProduct.lift.tprod]
rfl
/-- The basis vector of tensors given a `b : Π j, Fin (S.repDim (c j))` . -/
def basisVector {n : } (c : Fin n → S.C) (b : Π j, Fin (S.repDim (c j))) :