refactor: Replace some simp with simp?
This commit is contained in:
parent
03edf78931
commit
fccb283a5c
14 changed files with 110 additions and 85 deletions
|
@ -158,7 +158,7 @@ lemma mul_minkowskiMatrix_mul_transpose :
|
|||
(Subtype.val Λ) * minkowskiMatrix * (Subtype.val Λ).transpose = minkowskiMatrix := by
|
||||
have h2 := Λ.prop
|
||||
rw [LorentzGroup.mem_iff_self_mul_dual] at h2
|
||||
simp [minkowskiMetric.dual] at h2
|
||||
simp only [dual] at h2
|
||||
refine (right_inv_eq_left_inv minkowskiMatrix.sq ?_).symm
|
||||
rw [← h2]
|
||||
noncomm_ring
|
||||
|
@ -168,7 +168,7 @@ lemma transpose_mul_minkowskiMatrix_mul_self :
|
|||
(Subtype.val Λ).transpose * minkowskiMatrix * (Subtype.val Λ) = minkowskiMatrix := by
|
||||
have h2 := Λ.prop
|
||||
rw [LorentzGroup.mem_iff_dual_mul_self] at h2
|
||||
simp [minkowskiMetric.dual] at h2
|
||||
simp only [dual] at h2
|
||||
refine right_inv_eq_left_inv ?_ minkowskiMatrix.sq
|
||||
rw [← h2]
|
||||
noncomm_ring
|
||||
|
|
|
@ -72,7 +72,7 @@ namespace genBoost
|
|||
-/
|
||||
lemma self (u : FuturePointing d) : genBoost u u = LinearMap.id := by
|
||||
ext x
|
||||
simp [genBoost]
|
||||
simp only [genBoost, LinearMap.add_apply, LinearMap.id_coe, id_eq]
|
||||
rw [add_assoc, add_right_eq_self, add_eq_zero_iff_eq_neg, genBoostAux₁, genBoostAux₂]
|
||||
simp only [LinearMap.coe_mk, AddHom.coe_mk, map_add, smul_add, neg_smul, neg_add_rev, neg_neg]
|
||||
rw [← add_smul]
|
||||
|
|
|
@ -76,7 +76,7 @@ lemma stepFunction_continuous : Continuous stepFunction := by
|
|||
<;> intro a ha
|
||||
· rw [@Set.Iic_def, @frontier_Iic, @Set.mem_singleton_iff] at ha
|
||||
rw [ha]
|
||||
simp [neg_lt_self_iff, zero_lt_one, ↓reduceIte]
|
||||
simp only [le_neg_self_iff, id_eq]
|
||||
have h1 : ¬ (1 : ℝ) ≤ 0 := by simp
|
||||
exact Eq.symm (if_neg h1)
|
||||
· rw [Set.Ici_def, @frontier_Ici, @Set.mem_singleton_iff] at ha
|
||||
|
|
|
@ -176,7 +176,7 @@ lemma id_IsProper : @IsProper d 1 := by
|
|||
|
||||
lemma isProper_on_connected_component {Λ Λ' : LorentzGroup d} (h : Λ' ∈ connectedComponent Λ) :
|
||||
IsProper Λ ↔ IsProper Λ' := by
|
||||
simp [detRep_apply, detRep_apply, detContinuous]
|
||||
simp only [IsProper]
|
||||
rw [det_on_connected_component h]
|
||||
|
||||
end LorentzGroup
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue