refactor: More simps
This commit is contained in:
parent
4a396783ab
commit
269f4d53a7
16 changed files with 139 additions and 62 deletions
|
@ -121,7 +121,7 @@ theorem linesInPlane_constAbs {S : (PureU1 (n.succ.succ.succ.succ.succ)).LinSols
|
|||
|
||||
lemma linesInPlane_four (S : (PureU1 4).Sols) (hS : LineInPlaneCond S.1.1) :
|
||||
ConstAbsProp (S.val (0 : Fin 4), S.val (1 : Fin 4)) := by
|
||||
simp [ConstAbsProp]
|
||||
simp only [ConstAbsProp, Fin.isValue]
|
||||
by_contra hn
|
||||
have hLin := pureU1_linear S.1.1
|
||||
have hcube := pureU1_cube S
|
||||
|
|
|
@ -93,13 +93,15 @@ def accGrav : (SMνCharges n).Charges →ₗ[ℚ] ℚ where
|
|||
map_add' S T := by
|
||||
simp only
|
||||
repeat rw [map_add]
|
||||
simp [Pi.add_apply, mul_add]
|
||||
simp only [SMνSpecies_numberCharges, ACCSystemCharges.chargesAddCommMonoid_add, toSpecies_apply,
|
||||
Fin.isValue, mul_add]
|
||||
repeat erw [Finset.sum_add_distrib]
|
||||
ring
|
||||
map_smul' a S := by
|
||||
simp only
|
||||
repeat erw [map_smul]
|
||||
simp [HSMul.hSMul, SMul.smul]
|
||||
simp only [SMνSpecies_numberCharges, HSMul.hSMul, SMul.smul, toSpecies_apply, Fin.isValue,
|
||||
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]
|
||||
|
@ -127,13 +129,15 @@ def accSU2 : (SMνCharges n).Charges →ₗ[ℚ] ℚ where
|
|||
map_add' S T := by
|
||||
simp only
|
||||
repeat rw [map_add]
|
||||
simp [Pi.add_apply, mul_add]
|
||||
simp only [SMνSpecies_numberCharges, ACCSystemCharges.chargesAddCommMonoid_add, toSpecies_apply,
|
||||
Fin.isValue, mul_add]
|
||||
repeat erw [Finset.sum_add_distrib]
|
||||
ring
|
||||
map_smul' a S := by
|
||||
simp only
|
||||
repeat erw [map_smul]
|
||||
simp [HSMul.hSMul, SMul.smul]
|
||||
simp only [SMνSpecies_numberCharges, HSMul.hSMul, SMul.smul, toSpecies_apply, Fin.isValue,
|
||||
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]
|
||||
|
@ -160,13 +164,15 @@ def accSU3 : (SMνCharges n).Charges →ₗ[ℚ] ℚ where
|
|||
map_add' S T := by
|
||||
simp only
|
||||
repeat rw [map_add]
|
||||
simp [Pi.add_apply, mul_add]
|
||||
simp only [SMνSpecies_numberCharges, ACCSystemCharges.chargesAddCommMonoid_add, toSpecies_apply,
|
||||
Fin.isValue, mul_add]
|
||||
repeat erw [Finset.sum_add_distrib]
|
||||
ring
|
||||
map_smul' a S := by
|
||||
simp only
|
||||
repeat erw [map_smul]
|
||||
simp [HSMul.hSMul, SMul.smul]
|
||||
simp only [SMνSpecies_numberCharges, HSMul.hSMul, SMul.smul, toSpecies_apply, Fin.isValue,
|
||||
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]
|
||||
|
@ -201,7 +207,8 @@ def accYY : (SMνCharges n).Charges →ₗ[ℚ] ℚ where
|
|||
map_smul' a S := by
|
||||
simp only
|
||||
repeat erw [map_smul]
|
||||
simp [HSMul.hSMul, SMul.smul]
|
||||
simp only [SMνSpecies_numberCharges, HSMul.hSMul, SMul.smul, toSpecies_apply, Fin.isValue,
|
||||
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]
|
||||
|
@ -235,7 +242,7 @@ def quadBiLin : BiLinearSymm (SMνCharges n).Charges := BiLinearSymm.mk₂
|
|||
rw [Finset.mul_sum]
|
||||
refine Fintype.sum_congr _ _ fun i ↦ ?_
|
||||
repeat erw [map_smul]
|
||||
simp [HSMul.hSMul, SMul.smul]
|
||||
simp only [HSMul.hSMul, SMul.smul, toSpecies_apply, Fin.isValue, neg_mul, one_mul]
|
||||
ring)
|
||||
(by
|
||||
intro S T R
|
||||
|
@ -298,7 +305,7 @@ def cubeTriLin : TriLinearSymm (SMνCharges n).Charges := TriLinearSymm.mk₃
|
|||
rw [Finset.mul_sum]
|
||||
refine Fintype.sum_congr _ _ fun i ↦ ?_
|
||||
repeat erw [map_smul]
|
||||
simp [HSMul.hSMul, SMul.smul]
|
||||
simp only [HSMul.hSMul, SMul.smul, toSpecies_apply, Fin.isValue]
|
||||
ring)
|
||||
(by
|
||||
intro S T R L
|
||||
|
|
|
@ -69,7 +69,8 @@ def speciesEmbed (m n : ℕ) :
|
|||
rfl
|
||||
map_smul' a S := by
|
||||
funext i
|
||||
simp [HSMul.hSMul]
|
||||
simp only [SMνSpecies_numberCharges, HSMul.hSMul, ACCSystemCharges.chargesModule_smul,
|
||||
eq_ratCast, Rat.cast_eq_id, id_eq]
|
||||
by_cases hi : i.val < m
|
||||
· erw [dif_pos hi, dif_pos hi]
|
||||
· erw [dif_neg hi, dif_neg hi]
|
||||
|
|
|
@ -38,12 +38,12 @@ variable {n : ℕ}
|
|||
|
||||
lemma SU2Sol (S : (SMNoGrav n).LinSols) : accSU2 S.val = 0 := by
|
||||
have hS := S.linearSol
|
||||
simp at hS
|
||||
simp only [SMNoGrav_numberLinear, SMNoGrav_linearACCs, Fin.isValue] at hS
|
||||
exact hS 0
|
||||
|
||||
lemma SU3Sol (S : (SMNoGrav n).LinSols) : accSU3 S.val = 0 := by
|
||||
have hS := S.linearSol
|
||||
simp at hS
|
||||
simp only [SMNoGrav_numberLinear, SMNoGrav_linearACCs, Fin.isValue] at hS
|
||||
exact hS 1
|
||||
|
||||
lemma cubeSol (S : (SMNoGrav n).Sols) : accCube S.val = 0 := S.cubicSol
|
||||
|
@ -54,7 +54,7 @@ def chargeToLinear (S : (SMNoGrav n).Charges) (hSU2 : accSU2 S = 0) (hSU3 : accS
|
|||
(SMNoGrav n).LinSols :=
|
||||
⟨S, by
|
||||
intro i
|
||||
simp at i
|
||||
simp only [SMNoGrav_numberLinear] at i
|
||||
match i with
|
||||
| 0 => exact hSU2
|
||||
| 1 => exact hSU3⟩
|
||||
|
|
|
@ -130,7 +130,11 @@ def B₆ : (SM 3).Charges := toSpeciesEquiv.invFun (fun s => fun i =>
|
|||
|
||||
lemma B₆_cubic (S T : (SM 3).Charges) : cubeTriLin B₆ S T =
|
||||
3 * (S (5 : Fin 18) * T (5 : Fin 18) - S (8 : Fin 18) * T (8 : Fin 18)) := by
|
||||
simp [Fin.sum_univ_three, B₆, Fin.divNat, Fin.modNat, finProdFinEquiv]
|
||||
simp only [B₆, Equiv.invFun_as_coe, cubeTriLin_toFun_apply_apply, Nat.reduceMul, finProdFinEquiv,
|
||||
Fin.divNat, Fin.modNat, Fin.isValue, Equiv.coe_fn_mk, Fin.val_zero, mul_zero, add_zero,
|
||||
toSpeciesEquiv_symm_apply, Fin.val_one, mul_one, Nat.ofNat_pos, Nat.add_div_right,
|
||||
Nat.add_mod_right, Fin.val_two, Nat.add_mul_mod_self_left, Fin.sum_univ_three, Fin.zero_eta,
|
||||
zero_mul, zero_add, Fin.reduceFinMk, Fin.mk_one, Nat.reduceAdd, one_mul, neg_mul, mul_neg]
|
||||
ring_nf
|
||||
|
||||
/-- The charge assignments forming a basis of the plane. -/
|
||||
|
@ -150,7 +154,8 @@ lemma B₀_Bi_cubic {i : Fin 7} (hi : 0 ≠ i) (S : (SM 3).Charges) :
|
|||
change cubeTriLin B₀ (B i) S = 0
|
||||
rw [B₀_cubic]
|
||||
fin_cases i <;>
|
||||
simp at hi <;>
|
||||
simp only [Fin.isValue, Fin.zero_eta, ne_eq, Fin.reduceEq, not_false_eq_true, Fin.mk_one,
|
||||
Fin.reduceFinMk, not_true_eq_false] at hi <;>
|
||||
simp [B₀, B₁, B₂, B₃, B₄, B₅, B₆, Fin.divNat, Fin.modNat, finProdFinEquiv]
|
||||
|
||||
lemma B₁_Bi_cubic {i : Fin 7} (hi : 1 ≠ i) (S : (SM 3).Charges) :
|
||||
|
@ -158,7 +163,8 @@ lemma B₁_Bi_cubic {i : Fin 7} (hi : 1 ≠ i) (S : (SM 3).Charges) :
|
|||
change cubeTriLin B₁ (B i) S = 0
|
||||
rw [B₁_cubic]
|
||||
fin_cases i <;>
|
||||
simp at hi <;>
|
||||
simp only [Fin.isValue, Fin.zero_eta, ne_eq, Fin.reduceEq, not_false_eq_true, Fin.mk_one,
|
||||
Fin.reduceFinMk, not_true_eq_false] at hi <;>
|
||||
simp [B₀, B₁, B₂, B₃, B₄, B₅, B₆, Fin.divNat, Fin.modNat, finProdFinEquiv]
|
||||
|
||||
lemma B₂_Bi_cubic {i : Fin 7} (hi : 2 ≠ i) (S : (SM 3).Charges) :
|
||||
|
@ -166,7 +172,8 @@ lemma B₂_Bi_cubic {i : Fin 7} (hi : 2 ≠ i) (S : (SM 3).Charges) :
|
|||
change cubeTriLin B₂ (B i) S = 0
|
||||
rw [B₂_cubic]
|
||||
fin_cases i <;>
|
||||
simp at hi <;>
|
||||
simp only [Fin.isValue, Fin.zero_eta, ne_eq, Fin.reduceEq, not_false_eq_true, Fin.mk_one,
|
||||
Fin.reduceFinMk, not_true_eq_false] at hi <;>
|
||||
simp [B₀, B₁, B₂, B₃, B₄, B₅, B₆, Fin.divNat, Fin.modNat, finProdFinEquiv]
|
||||
|
||||
lemma B₃_Bi_cubic {i : Fin 7} (hi : 3 ≠ i) (S : (SM 3).Charges) :
|
||||
|
@ -174,7 +181,8 @@ lemma B₃_Bi_cubic {i : Fin 7} (hi : 3 ≠ i) (S : (SM 3).Charges) :
|
|||
change cubeTriLin (B₃) (B i) S = 0
|
||||
rw [B₃_cubic]
|
||||
fin_cases i <;>
|
||||
simp at hi <;>
|
||||
simp only [Fin.isValue, Fin.zero_eta, ne_eq, Fin.reduceEq, not_false_eq_true, Fin.mk_one,
|
||||
Fin.reduceFinMk, not_true_eq_false] at hi <;>
|
||||
simp [B₀, B₁, B₂, B₃, B₄, B₅, B₆, Fin.divNat, Fin.modNat, finProdFinEquiv]
|
||||
|
||||
lemma B₄_Bi_cubic {i : Fin 7} (hi : 4 ≠ i) (S : (SM 3).Charges) :
|
||||
|
@ -182,7 +190,8 @@ lemma B₄_Bi_cubic {i : Fin 7} (hi : 4 ≠ i) (S : (SM 3).Charges) :
|
|||
change cubeTriLin (B₄) (B i) S = 0
|
||||
rw [B₄_cubic]
|
||||
fin_cases i <;>
|
||||
simp at hi <;>
|
||||
simp only [Fin.isValue, Fin.zero_eta, ne_eq, Fin.reduceEq, not_false_eq_true, Fin.mk_one,
|
||||
Fin.reduceFinMk, not_true_eq_false] at hi <;>
|
||||
simp [B₀, B₁, B₂, B₃, B₄, B₅, B₆, Fin.divNat, Fin.modNat, finProdFinEquiv]
|
||||
|
||||
lemma B₅_Bi_cubic {i : Fin 7} (hi : 5 ≠ i) (S : (SM 3).Charges) :
|
||||
|
@ -190,7 +199,8 @@ lemma B₅_Bi_cubic {i : Fin 7} (hi : 5 ≠ i) (S : (SM 3).Charges) :
|
|||
change cubeTriLin (B₅) (B i) S = 0
|
||||
rw [B₅_cubic]
|
||||
fin_cases i <;>
|
||||
simp at hi <;>
|
||||
simp only [Fin.isValue, Fin.zero_eta, ne_eq, Fin.reduceEq, not_false_eq_true, Fin.mk_one,
|
||||
Fin.reduceFinMk, not_true_eq_false] at hi <;>
|
||||
simp [B₀, B₁, B₂, B₃, B₄, B₅, B₆, Fin.divNat, Fin.modNat, finProdFinEquiv]
|
||||
|
||||
lemma B₆_Bi_cubic {i : Fin 7} (hi : 6 ≠ i) (S : (SM 3).Charges) :
|
||||
|
@ -198,7 +208,8 @@ lemma B₆_Bi_cubic {i : Fin 7} (hi : 6 ≠ i) (S : (SM 3).Charges) :
|
|||
change cubeTriLin (B₆) (B i) S = 0
|
||||
rw [B₆_cubic]
|
||||
fin_cases i <;>
|
||||
simp at hi <;>
|
||||
simp only [Fin.isValue, Fin.zero_eta, ne_eq, Fin.reduceEq, not_false_eq_true, Fin.mk_one,
|
||||
Fin.reduceFinMk, not_true_eq_false] at hi <;>
|
||||
simp [B₀, B₁, B₂, B₃, B₄, B₅, B₆, Fin.divNat, Fin.modNat, finProdFinEquiv]
|
||||
|
||||
lemma Bi_Bj_ne_cubic {i j : Fin 7} (h : i ≠ j) (S : (SM 3).Charges) :
|
||||
|
|
|
@ -35,7 +35,7 @@ def BL₁ : (PlusU1 1).Sols where
|
|||
| (5 : Fin 6) => 3
|
||||
linearSol := by
|
||||
intro i
|
||||
simp at i
|
||||
simp only [PlusU1_numberLinear] at i
|
||||
match i with
|
||||
| 0 => rfl
|
||||
| 1 => rfl
|
||||
|
@ -43,7 +43,7 @@ def BL₁ : (PlusU1 1).Sols where
|
|||
| 3 => rfl
|
||||
quadSol := by
|
||||
intro i
|
||||
simp at i
|
||||
simp only [PlusU1_numberQuadratic] at i
|
||||
match i with
|
||||
| 0 => rfl
|
||||
cubicSol := by rfl
|
||||
|
|
|
@ -40,27 +40,27 @@ variable {n : ℕ}
|
|||
|
||||
lemma gravSol (S : (PlusU1 n).LinSols) : accGrav S.val = 0 := by
|
||||
have hS := S.linearSol
|
||||
simp at hS
|
||||
simp only [PlusU1_numberLinear, PlusU1_linearACCs, Fin.isValue] at hS
|
||||
exact hS 0
|
||||
|
||||
lemma SU2Sol (S : (PlusU1 n).LinSols) : accSU2 S.val = 0 := by
|
||||
have hS := S.linearSol
|
||||
simp at hS
|
||||
simp only [PlusU1_numberLinear, PlusU1_linearACCs, Fin.isValue] at hS
|
||||
exact hS 1
|
||||
|
||||
lemma SU3Sol (S : (PlusU1 n).LinSols) : accSU3 S.val = 0 := by
|
||||
have hS := S.linearSol
|
||||
simp at hS
|
||||
simp only [PlusU1_numberLinear, PlusU1_linearACCs, Fin.isValue] at hS
|
||||
exact hS 2
|
||||
|
||||
lemma YYsol (S : (PlusU1 n).LinSols) : accYY S.val = 0 := by
|
||||
have hS := S.linearSol
|
||||
simp at hS
|
||||
simp only [PlusU1_numberLinear, PlusU1_linearACCs, Fin.isValue] at hS
|
||||
exact hS 3
|
||||
|
||||
lemma quadSol (S : (PlusU1 n).QuadSols) : accQuad S.val = 0 := by
|
||||
have hS := S.quadSol
|
||||
simp at hS
|
||||
simp only [PlusU1_numberQuadratic, HomogeneousQuadratic.eq_1, PlusU1_quadraticACCs] at hS
|
||||
exact hS 0
|
||||
|
||||
lemma cubeSol (S : (PlusU1 n).Sols) : accCube S.val = 0 := by
|
||||
|
@ -73,7 +73,7 @@ def chargeToLinear (S : (PlusU1 n).Charges) (hGrav : accGrav S = 0)
|
|||
(PlusU1 n).LinSols :=
|
||||
⟨S, by
|
||||
intro i
|
||||
simp at i
|
||||
simp only [PlusU1_numberLinear] at i
|
||||
match i with
|
||||
| 0 => exact hGrav
|
||||
| 1 => exact hSU2
|
||||
|
@ -86,7 +86,7 @@ def linearToQuad (S : (PlusU1 n).LinSols) (hQ : accQuad S.val = 0) :
|
|||
(PlusU1 n).QuadSols :=
|
||||
⟨S, by
|
||||
intro i
|
||||
simp at i
|
||||
simp only [PlusU1_numberQuadratic] at i
|
||||
match i with
|
||||
| 0 => exact hQ⟩
|
||||
|
||||
|
@ -122,7 +122,7 @@ def perm (n : ℕ) : ACCSystemGroupAction (PlusU1 n) where
|
|||
rep := repCharges
|
||||
linearInvariant := by
|
||||
intro i
|
||||
simp at i
|
||||
simp only [PlusU1_numberLinear] at i
|
||||
match i with
|
||||
| 0 => exact accGrav_invariant
|
||||
| 1 => exact accSU2_invariant
|
||||
|
@ -130,7 +130,7 @@ def perm (n : ℕ) : ACCSystemGroupAction (PlusU1 n) where
|
|||
| 3 => exact accYY_invariant
|
||||
quadInvariant := by
|
||||
intro i
|
||||
simp at i
|
||||
simp only [PlusU1_numberQuadratic] at i
|
||||
match i with
|
||||
| 0 => exact accQuad_invariant
|
||||
cubicInvariant := accCube_invariant
|
||||
|
|
|
@ -162,7 +162,7 @@ lemma isSolution_sum_part (f : Fin 11 → ℚ) (hS : (PlusU1 3).IsSolution (∑
|
|||
rw [isSolution_f0 f hS, isSolution_f1 f hS, isSolution_f2 f hS, isSolution_f3 f hS,
|
||||
isSolution_f4 f hS, isSolution_f5 f hS,
|
||||
isSolution_f6 f hS, isSolution_f7 f hS, isSolution_f8 f hS]
|
||||
simp
|
||||
simp only [Fin.isValue, zero_smul, add_zero, zero_add]
|
||||
rfl
|
||||
|
||||
lemma isSolution_grav (f : Fin 11 → ℚ) (hS : (PlusU1 3).IsSolution (∑ i, f i • B i)) :
|
||||
|
@ -172,7 +172,7 @@ lemma isSolution_grav (f : Fin 11 → ℚ) (hS : (PlusU1 3).IsSolution (∑ i, f
|
|||
have hg := gravSol S.toLinSols
|
||||
rw [hS', hx, accGrav.map_add, accGrav.map_smul, accGrav.map_smul, show accGrav B₉ = 3 by rfl,
|
||||
show accGrav B₁₀ = 1 by rfl] at hg
|
||||
simp at hg
|
||||
simp only [Fin.isValue, smul_eq_mul, mul_one] at hg
|
||||
linear_combination hg
|
||||
|
||||
lemma isSolution_sum_part' (f : Fin 11 → ℚ) (hS : (PlusU1 3).IsSolution (∑ i, f i • B i)) :
|
||||
|
@ -193,7 +193,7 @@ lemma isSolution_f9 (f : Fin 11 → ℚ) (hS : (PlusU1 3).IsSolution (∑ i, f i
|
|||
cubeTriLin.map_smul₃, cubeTriLin.map_smul₃] at hc
|
||||
rw [show accCube B₉ = 9 by rfl, show accCube B₁₀ = 1 by rfl, show cubeTriLin B₉ B₉ B₁₀ = 0 by rfl,
|
||||
show cubeTriLin B₁₀ B₁₀ B₉ = 0 by rfl] at hc
|
||||
simp at hc
|
||||
simp only [Fin.isValue, neg_mul, mul_one, mul_zero, add_zero] at hc
|
||||
have h1 : f 9 ^ 3 * 9 + (-(3 * f 9)) ^ 3 = - 18 * f 9 ^ 3 := by ring
|
||||
rw [h1] at hc
|
||||
simpa using hc
|
||||
|
|
|
@ -55,7 +55,9 @@ lemma accQuad_α₁_α₂ (S : (PlusU1 n).LinSols) :
|
|||
lemma accQuad_α₁_α₂_zero (S : (PlusU1 n).LinSols) (h1 : α₁ C S = 0)
|
||||
(h2 : α₂ S = 0) (a b : ℚ) : accQuad (a • S + b • C.1).val = 0 := by
|
||||
erw [add_AFL_quad]
|
||||
simp [α₁, α₂] at h1 h2
|
||||
simp only [α₁, quadBiLin_toFun_apply, Fin.isValue, neg_mul, neg_eq_zero, mul_eq_zero,
|
||||
OfNat.ofNat_ne_zero, false_or, α₂, HomogeneousQuadratic.eq_1, accQuad,
|
||||
BiLinearSymm.toHomogeneousQuad_apply] at h1 h2
|
||||
field_simp [h1, h2]
|
||||
|
||||
/-- The construction of a `QuadSol` from a `LinSols` in the generic case. -/
|
||||
|
|
|
@ -118,14 +118,14 @@ lemma quadSolToSolInv_α₁_α₂_neq_zero (S : (PlusU1 n).Sols) (h : α₁ S.1
|
|||
lemma quadSolToSolInv_special (S : (PlusU1 n).Sols) (h : α₁ S.1 = 0) :
|
||||
special (quadSolToSolInv S).1 (quadSolToSolInv S).2.1 (quadSolToSolInv S).2.2
|
||||
(quadSolToSolInv_α₁_α₂_zero S h).1 (quadSolToSolInv_α₁_α₂_zero S h).2 = S := by
|
||||
simp [quadSolToSolInv_1]
|
||||
simp only [quadSolToSolInv_1]
|
||||
rw [show (quadSolToSolInv S).2.1 = 1 by rw [quadSolToSolInv, if_pos h]]
|
||||
rw [show (quadSolToSolInv S).2.2 = 0 by rw [quadSolToSolInv, if_pos h]]
|
||||
rw [special_on_AF]
|
||||
|
||||
lemma quadSolToSolInv_generic (S : (PlusU1 n).Sols) (h : α₁ S.1 ≠ 0) :
|
||||
(quadSolToSolInv S).2.1 • generic (quadSolToSolInv S).1 = S := by
|
||||
simp [quadSolToSolInv_1]
|
||||
simp only [quadSolToSolInv_1]
|
||||
rw [show (quadSolToSolInv S).2.1 = (α₁ S.1)⁻¹ by rw [quadSolToSolInv, if_neg h]]
|
||||
rw [generic_on_AF_α₁_ne_zero S h]
|
||||
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue