diff --git a/HepLean/Mathematics/SO3/Basic.lean b/HepLean/Mathematics/SO3/Basic.lean index 1c527c6..f8c5a39 100644 --- a/HepLean/Mathematics/SO3/Basic.lean +++ b/HepLean/Mathematics/SO3/Basic.lean @@ -25,12 +25,13 @@ instance SO3Group : Group SO3 where by simp only [det_mul, A.2.1, B.2.1, mul_one], by - simp [A.2.2, B.2.2, ← Matrix.mul_assoc, Matrix.mul_assoc]⟩ + simp only [transpose_mul, ← Matrix.mul_assoc, Matrix.mul_assoc, B.2.2, mul_one, A.2.2]⟩ mul_assoc A B C := Subtype.eq (Matrix.mul_assoc A.1 B.1 C.1) one := ⟨1, det_one, by rw [transpose_one, mul_one]⟩ one_mul A := Subtype.eq (Matrix.one_mul A.1) mul_one A := Subtype.eq (Matrix.mul_one A.1) - inv A := ⟨A.1ᵀ, by simp [A.2], by simp [mul_eq_one_comm.mpr A.2.2]⟩ + inv A := ⟨A.1ᵀ, by simp only [det_transpose, A.2], + by simp only [transpose_transpose, mul_eq_one_comm.mpr A.2.2]⟩ inv_mul_cancel A := Subtype.eq (mul_eq_one_comm.mpr A.2.2) /-- Notation for the group `SO3`. -/ @@ -99,7 +100,7 @@ lemma toGL_embedding : Embedding toGL.toFun where rw [isOpen_induced_iff] at hU1 obtain ⟨V, hV1, hV2⟩ := hU1 use V - simp [hV1] + simp only [hV1, true_and] rw [← hU2, ← hV2] rfl · intro h diff --git a/HepLean/SpaceTime/LorentzGroup/Basic.lean b/HepLean/SpaceTime/LorentzGroup/Basic.lean index 60dc761..cbfbd6d 100644 --- a/HepLean/SpaceTime/LorentzGroup/Basic.lean +++ b/HepLean/SpaceTime/LorentzGroup/Basic.lean @@ -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 diff --git a/HepLean/SpaceTime/LorentzGroup/Boosts.lean b/HepLean/SpaceTime/LorentzGroup/Boosts.lean index 5d08bf5..d3bbf34 100644 --- a/HepLean/SpaceTime/LorentzGroup/Boosts.lean +++ b/HepLean/SpaceTime/LorentzGroup/Boosts.lean @@ -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] diff --git a/HepLean/SpaceTime/LorentzGroup/Orthochronous.lean b/HepLean/SpaceTime/LorentzGroup/Orthochronous.lean index 260c99e..e900b41 100644 --- a/HepLean/SpaceTime/LorentzGroup/Orthochronous.lean +++ b/HepLean/SpaceTime/LorentzGroup/Orthochronous.lean @@ -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 diff --git a/HepLean/SpaceTime/LorentzGroup/Proper.lean b/HepLean/SpaceTime/LorentzGroup/Proper.lean index 2688228..32768de 100644 --- a/HepLean/SpaceTime/LorentzGroup/Proper.lean +++ b/HepLean/SpaceTime/LorentzGroup/Proper.lean @@ -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 diff --git a/HepLean/SpaceTime/LorentzTensor/Basic.lean b/HepLean/SpaceTime/LorentzTensor/Basic.lean index 2f4c9d8..3eb49c6 100644 --- a/HepLean/SpaceTime/LorentzTensor/Basic.lean +++ b/HepLean/SpaceTime/LorentzTensor/Basic.lean @@ -204,7 +204,9 @@ lemma contrRightAux_comp {V1 V2 V3 V4 V5 : Type} [AddCommMonoid V1] [AddCommMono refine TensorProduct.induction_on y (by simp) ?_ (fun x z h1 h2 => by simp [add_tmul, tmul_add, LinearMap.map_add, h1, h2]) intro x3 x4 - simp [contrRightAux, contrMidAux, contrLeftAux] + simp only [contrRightAux, LinearEquiv.refl_toLinearMap, LinearMap.coe_comp, LinearEquiv.coe_coe, + Function.comp_apply, map_tmul, LinearMap.id_coe, id_eq, assoc_tmul, rid_tmul, tmul_smul, + map_smul, contrMidAux, contrLeftAux, assoc_symm_tmul, lid_tmul] rfl end TensorStructure @@ -293,7 +295,7 @@ lemma tensorProd_piTensorProd_ext {M : Type} [AddCommMonoid M] [Module R M] refine fun y ↦ PiTensorProduct.induction_on' y ?_ (by intro a b hx hy - simp at hx hy + simp only [PiTensorProduct.tprodCoeff_eq_smul_tprod] at hx hy simp [map_add, tmul_add, hx, hy]) intro ry fy simp only [PiTensorProduct.tprodCoeff_eq_smul_tprod, tmul_smul, LinearMapClass.map_smul] @@ -355,8 +357,8 @@ lemma mapIso_symm (e : X ≃ Y) (h : cX.MapIso e cY) : apply PiTensorProduct.ext apply MultilinearMap.ext intro x - simp [mapIso, LinearMap.compMultilinearMap_apply, LinearEquiv.coe_coe, - LinearEquiv.symm_apply_apply, PiTensorProduct.reindex_tprod] + simp only [mapIso, LinearEquiv.trans_symm, LinearMap.compMultilinearMap_apply, + LinearEquiv.coe_coe, LinearEquiv.trans_apply, Equiv.symm_symm] change (PiTensorProduct.reindex R (fun x => 𝓣.ColorModule (cX x)) e).symm ((PiTensorProduct.congr fun y => 𝓣.colorModuleCast _).symm ((PiTensorProduct.tprod R) x)) = (PiTensorProduct.congr fun y => 𝓣.colorModuleCast _) @@ -469,7 +471,7 @@ lemma inlPureTensor_update_left [DecidableEq (X ⊕ Y)] (f : 𝓣.PureTensor (Su 𝓣.inlPureTensor (Function.update f (Sum.inl x) v1) = Function.update (𝓣.inlPureTensor f) x v1 := by funext y - simp [inlPureTensor, Function.update, Sum.inl.injEq, Sum.elim_inl] + simp only [inlPureTensor, Function.update, Sum.inl.injEq, Sum.elim_inl] split · rename_i h subst h @@ -491,7 +493,7 @@ lemma inrPureTensor_update_right [DecidableEq (X ⊕ Y)] (f : 𝓣.PureTensor (S 𝓣.inrPureTensor (Function.update f (Sum.inr y) v1) = Function.update (𝓣.inrPureTensor f) y v1 := by funext y - simp [inrPureTensor, Function.update, Sum.inl.injEq, Sum.elim_inl] + simp only [inrPureTensor, Function.update, Sum.inr.injEq, Sum.elim_inr] split · rename_i h subst h @@ -570,11 +572,13 @@ def tensoratorEquiv (c : X → 𝓣.Color) (d : Y → 𝓣.Color) : apply PiTensorProduct.ext apply MultilinearMap.ext intro p - simp [tensorator, tensoratorSymm, domCoprod] + simp only [tensoratorSymm, tensorator, domCoprod, LinearMap.compMultilinearMap_apply, + LinearMap.coe_comp, Function.comp_apply, PiTensorProduct.lift.tprod, MultilinearMap.coe_mk, + lift.tmul, LinearMap.coe_mk, AddHom.coe_mk] change (PiTensorProduct.lift _) ((PiTensorProduct.tprod R) _) = LinearMap.id ((PiTensorProduct.tprod R) p) rw [PiTensorProduct.lift.tprod] - simp [elimPureTensorMulLin, elimPureTensor] + simp only [elimPureTensorMulLin, MultilinearMap.coe_mk, LinearMap.id_coe, id_eq] change (PiTensorProduct.tprod R) _ = _ apply congrArg funext x @@ -584,11 +588,12 @@ def tensoratorEquiv (c : X → 𝓣.Color) (d : Y → 𝓣.Color) : (by apply tensorProd_piTensorProd_ext intro p q - simp [tensorator, tensoratorSymm] + simp only [tensorator, tensoratorSymm, LinearMap.coe_comp, Function.comp_apply, lift.tmul, + LinearMap.coe_mk, AddHom.coe_mk, PiTensorProduct.lift.tprod, LinearMap.id_coe, id_eq] change (PiTensorProduct.lift 𝓣.domCoprod) ((PiTensorProduct.lift (𝓣.elimPureTensorMulLin p)) ((PiTensorProduct.tprod R) q)) =_ rw [PiTensorProduct.lift.tprod] - simp [elimPureTensorMulLin] + simp only [elimPureTensorMulLin, MultilinearMap.coe_mk, PiTensorProduct.lift.tprod] rfl) omit [Fintype X] [Fintype Y] in @@ -606,7 +611,7 @@ lemma tensoratorEquiv_symm_tprod (f : 𝓣.PureTensor (Sum.elim cX cY)) : (𝓣.tensoratorEquiv cX cY).symm ((PiTensorProduct.tprod R) f) = (PiTensorProduct.tprod R) (𝓣.inlPureTensor f) ⊗ₜ[R] (PiTensorProduct.tprod R) (𝓣.inrPureTensor f) := by - simp [tensoratorEquiv, tensorator] + simp only [tensoratorEquiv, tensorator, LinearEquiv.ofLinear_symm_apply] change (PiTensorProduct.lift 𝓣.domCoprod) ((PiTensorProduct.tprod R) f) = _ simp [domCoprod] @@ -684,7 +689,9 @@ lemma contrDual_symm_contrRightAux (h : ν = η) : intro x y refine TensorProduct.induction_on x (by simp) ?_ ?_ · intro x z - simp [contrRightAux] + simp only [contrRightAux, LinearEquiv.refl_toLinearMap, LinearMap.coe_comp, LinearEquiv.coe_coe, + Function.comp_apply, assoc_tmul, map_tmul, LinearMap.id_coe, id_eq, rid_tmul, map_smul, + congr_tmul, contrDual_symm'] congr · exact (LinearEquiv.symm_apply_eq (𝓣.colorModuleCast (𝓣.τ_involutive μ))).mp rfl · exact (LinearEquiv.symm_apply_eq (𝓣.colorModuleCast (𝓣.τ_involutive (𝓣.τ μ)))).mp rfl diff --git a/HepLean/SpaceTime/LorentzTensor/Contraction.lean b/HepLean/SpaceTime/LorentzTensor/Contraction.lean index 30c88f4..f267574 100644 --- a/HepLean/SpaceTime/LorentzTensor/Contraction.lean +++ b/HepLean/SpaceTime/LorentzTensor/Contraction.lean @@ -177,7 +177,7 @@ lemma pairProd_tmul_tprod_tprod (fx : (i : X) → 𝓣.ColorModule (cX i)) (fx2 : (i : X) → 𝓣.ColorModule (cX2 i)) : 𝓣.pairProd (PiTensorProduct.tprod R fx ⊗ₜ[R] PiTensorProduct.tprod R fx2) = PiTensorProduct.tprod R (fun x => fx x ⊗ₜ[R] fx2 x) := by - simp [pairProd] + simp only [pairProd, lift.tmul] erw [PiTensorProduct.map₂_tprod_tprod] rfl @@ -220,15 +220,17 @@ lemma contrAll'_isEmpty_tmul [IsEmpty X] (x : 𝓣.Tensor cX) (y : 𝓣.Tensor ( intro rx fx refine PiTensorProduct.induction_on' y ?_ (by intro a b hx hy - simp at hx hy + simp only [PiTensorProduct.tprodCoeff_eq_smul_tprod, map_smul, isEmptyEquiv_tprod, + smul_eq_mul, mul_one] at hx hy simp [map_add, tmul_add, mul_add, hx, hy]) intro ry fy - simp [smul_tmul] + simp only [PiTensorProduct.tprodCoeff_eq_smul_tprod, Function.comp_apply, tmul_smul, smul_tmul, + map_smul, smul_eq_mul, isEmptyEquiv_tprod, mul_one] ring_nf rw [mul_assoc, mul_assoc] apply congrArg apply congrArg - simp [contrAll'] + simp only [contrAll', LinearMap.coe_comp, Function.comp_apply] erw [pairProd_tmul_tprod_tprod] simp only [Function.comp_apply, PiTensorProduct.map_tprod, PiTensorProduct.lift.tprod, MultilinearMap.mkPiAlgebra_apply, Finset.univ_eq_empty, Finset.prod_empty] @@ -278,7 +280,8 @@ lemma contrAll'_mapIso (e : X ≃ Y) (h : cX.MapIso e cY) : (𝓣.colorModuleCast (congrFun (TensorColor.ColorMap.MapIso.symm h) y))).mp rfl · symm apply cast_eq_iff_heq.mpr - simp [colorModuleCast, Equiv.apply_symm_apply] + simp only [Function.comp_apply, colorModuleCast, Equiv.cast_symm, LinearEquiv.coe_mk, + Equiv.cast_apply] rw [Equiv.apply_symm_apply] exact HEq.symm (cast_heq _ _) @@ -385,7 +388,8 @@ lemma contrAll_rep (e : X ≃ Y) (h : cX.ContrAll e cY) (g : G) : intro rx fx refine fun y ↦ PiTensorProduct.induction_on' y ?_ (by intro a b hx hy - simp at hx hy + simp only [PiTensorProduct.tprodCoeff_eq_smul_tprod, LinearMap.coe_comp, Function.comp_apply, + map_tmul, map_smul, rep_tprod] at hx hy simp [map_add, tmul_add, hx, hy]) intro ry fy simp only [contrAll_tmul, PiTensorProduct.tprodCoeff_eq_smul_tprod, tmul_smul, smul_tmul, diff --git a/HepLean/SpaceTime/LorentzTensor/IndexNotation/ColorIndexList/Append.lean b/HepLean/SpaceTime/LorentzTensor/IndexNotation/ColorIndexList/Append.lean index 754bcf5..d750dad 100644 --- a/HepLean/SpaceTime/LorentzTensor/IndexNotation/ColorIndexList/Append.lean +++ b/HepLean/SpaceTime/LorentzTensor/IndexNotation/ColorIndexList/Append.lean @@ -69,11 +69,10 @@ lemma inr (h : AppendCond l l2) (h' : AppendCond (l ++[h] l2) l3) : AppendCond l2 l3 := by apply And.intro · have h1 := h'.1 - simp at h1 - rw [append_assoc] at h1 + rw [append_toIndexList, append_assoc] at h1 exact OnlyUniqueDuals.inr h1 · have h1 := h'.2 - simp at h1 + simp only [append_toIndexList] at h1 rw [append_assoc] at h1 refine ColorCond.inr ?_ h1 rw [← append_assoc] @@ -82,10 +81,10 @@ lemma inr (h : AppendCond l l2) (h' : AppendCond (l ++[h] l2) l3) : lemma assoc (h : AppendCond l l2) (h' : AppendCond (l ++[h] l2) l3) : AppendCond l (l2 ++[h.inr h'] l3) := by apply And.intro - · simp + · simp only [append_toIndexList] rw [← append_assoc] simpa using h'.1 - · simp + · simp only [append_toIndexList] rw [← append_assoc] exact h'.2 @@ -132,7 +131,7 @@ lemma countId_contr_fst_eq_zero_mem_snd (h : AppendCond l l2) {I : Index 𝓒.Co have h1I := h.1 rw [OnlyUniqueDuals.iff_countId_leq_two'] at h1I have h1I' := h1I I - simp at h1I' + simp only [countId_append] at h1I' omega lemma countId_contr_snd_eq_zero_mem_fst (h : AppendCond l l2) {I : Index 𝓒.Color} @@ -151,13 +150,13 @@ lemma append_contr_left (h : AppendCond l l2) : apply List.filter_congr intro I hI simp only [decide_eq_decide] - simp [contrIndexList] at hI + simp only [contrIndexList, List.mem_filter, decide_eq_true_eq] at hI exact AppendCond.countId_contr_fst_eq_zero_mem_snd h hI.1 lemma append_contr_right (h : AppendCond l l2) : (l ++[h.contr_right] l2.contr).contr = (l ++[h] l2).contr := by apply ext - simp [contr] + simp only [contr, append_toIndexList] rw [contrIndexList_append_eq_filter, contrIndexList_append_eq_filter, contrIndexList_contrIndexList] apply congrFun diff --git a/HepLean/SpaceTime/LorentzTensor/IndexNotation/ColorIndexList/ContrPerm.lean b/HepLean/SpaceTime/LorentzTensor/IndexNotation/ColorIndexList/ContrPerm.lean index ff62a2f..c6f8385 100644 --- a/HepLean/SpaceTime/LorentzTensor/IndexNotation/ColorIndexList/ContrPerm.lean +++ b/HepLean/SpaceTime/LorentzTensor/IndexNotation/ColorIndexList/ContrPerm.lean @@ -59,7 +59,7 @@ lemma getDualInOtherEquiv_eq (h : l.ContrPerm l2) (i : Fin l.contr.length) : rw [h.2.1] exact Finset.mem_univ i⟩).1 · rfl - · simp [getDualInOtherEquiv] + · simp only [getDualInOtherEquiv, Equiv.coe_fn_mk, getDualInOther?_id, List.get_eq_getElem] rfl lemma mem_snd_of_mem_snd (h : l.ContrPerm l2) {I : Index 𝓒.Color} (hI : I ∈ l.contr.val) : @@ -96,11 +96,11 @@ lemma getDualInOtherEquiv_eq_of_countSelf rw [← List.mem_singleton, ← filter_id_of_countId_eq_one_mem l2.contr.toIndexList h1' h1] simp only [List.get_eq_getElem, List.mem_filter, decide_eq_true_eq] apply And.intro (List.getElem_mem _ _ _) - simp [getDualInOtherEquiv] + simp only [getDualInOtherEquiv, Equiv.coe_fn_mk] change _ = l2.contr.idMap (l.contr.getDualInOtherEquiv l2.contr ⟨i, by rw [hn] exact Finset.mem_univ i⟩).1 - simp [getDualInOtherEquiv] + simp only [getDualInOtherEquiv, Equiv.coe_fn_mk, getDualInOther?_id] rfl lemma colorMap_eq_of_countSelf (hn : IndexList.Subperm l.contr l2.contr.toIndexList) @@ -109,7 +109,7 @@ lemma colorMap_eq_of_countSelf (hn : IndexList.Subperm l.contr l2.contr.toIndexL l2.contr.colorMap' ∘ Subtype.val ∘ (l.contr.getDualInOtherEquiv l2.contr) = l.contr.colorMap' ∘ Subtype.val := by funext a - simp [colorMap', colorMap] + simp only [colorMap', Function.comp_apply, colorMap, List.get_eq_getElem] change _ = (l.contr.val.get a.1).toColor rw [← getDualInOtherEquiv_eq_of_countSelf hn a.1] · rfl @@ -119,7 +119,7 @@ lemma colorMap_eq_of_countSelf (hn : IndexList.Subperm l.contr l2.contr.toIndexL (countSelf_eq_one_of_countId_eq_one l.contr.toIndexList (l.contr.val.get a.1) ?h1 ?hme) · rw [Subperm.iff_countId_fin] at hn exact (hn a.1).1 - · simp + · simp only [List.get_eq_getElem] exact List.getElem_mem l.contr.val (↑↑a) a.1.isLt lemma iff_count_fin : l.ContrPerm l2 ↔ @@ -264,9 +264,11 @@ def contrPermEquiv {l l' : ColorIndexList 𝓒} (h : ContrPerm l l') : lemma contrPermEquiv_colorMap_iso {l l' : ColorIndexList 𝓒} (h : ContrPerm l l') : ColorMap.MapIso (contrPermEquiv h).symm l'.contr.colorMap' l.contr.colorMap' := by - simp [ColorMap.MapIso] + simp only [ColorMap.MapIso] funext i - simp [contrPermEquiv, getDualInOtherEquiv] + simp only [contrPermEquiv, getDualInOtherEquiv, Function.comp_apply, Equiv.symm_trans_apply, + Equiv.symm_symm, Equiv.subtypeUnivEquiv_symm_apply, Equiv.coe_fn_symm_mk, + Equiv.subtypeUnivEquiv_apply] have h' := h.symm.2.2 have hi : i ∈ (l'.contr.withUniqueDualInOther l.contr.toIndexList) := by rw [h.symm.2.1] @@ -291,7 +293,7 @@ lemma contrPermEquiv_symm {l l' : ColorIndexList 𝓒} (h : ContrPerm l l') : lemma contrPermEquiv_trans {l l2 l3 : ColorIndexList 𝓒} (h1 : ContrPerm l l2) (h2 : ContrPerm l2 l3) : (contrPermEquiv h1).trans (contrPermEquiv h2) = contrPermEquiv (h1.trans h2) := by - simp [contrPermEquiv] + simp only [contrPermEquiv] ext x simp only [getDualInOtherEquiv, Equiv.trans_apply, Equiv.subtypeUnivEquiv_symm_apply, Equiv.coe_fn_mk, Equiv.subtypeUnivEquiv_apply] @@ -315,7 +317,7 @@ lemma contrPermEquiv_trans {l l2 l3 : ColorIndexList 𝓒} lemma contrPermEquiv_self_contr {l : ColorIndexList 𝓒} : contrPermEquiv (contr_self : ContrPerm l l.contr) = (Fin.castOrderIso (by simp)).toEquiv := by - simp [contrPermEquiv] + simp only [contrPermEquiv] ext1 x simp only [getDualInOtherEquiv, Equiv.trans_apply, Equiv.subtypeUnivEquiv_symm_apply, Equiv.coe_fn_mk, Equiv.subtypeUnivEquiv_apply, RelIso.coe_fn_toEquiv, Fin.castOrderIso_apply, diff --git a/HepLean/SpaceTime/LorentzTensor/IndexNotation/ColorIndexList/Contraction.lean b/HepLean/SpaceTime/LorentzTensor/IndexNotation/ColorIndexList/Contraction.lean index 3d33d25..41ce875 100644 --- a/HepLean/SpaceTime/LorentzTensor/IndexNotation/ColorIndexList/Contraction.lean +++ b/HepLean/SpaceTime/LorentzTensor/IndexNotation/ColorIndexList/Contraction.lean @@ -81,7 +81,7 @@ lemma contr_contr_colorMap (i : Fin l.contr.contr.length) : @[simp] lemma contr_of_withDual_empty (h : l.withDual = ∅) : l.contr = l := by - simp [contr] + simp only [contr] apply ext simp [l.contrIndexList_of_withDual_empty h] @@ -97,14 +97,14 @@ lemma contr_areDualInSelf (i j : Fin l.contr.length) : lemma contr_countId_eq_zero_of_countId_zero (I : Index 𝓒.Color) (h : l.countId I = 0) : l.contr.countId I = 0 := by - simp [contr] + simp only [contr] exact countId_contrIndexList_zero_of_countId l.toIndexList I h lemma contr_countId_eq_filter (I : Index 𝓒.Color) : l.contr.countId I = (l.val.filter (fun J => I.id = J.id)).countP (fun i => l.val.countP (fun j => i.id = j.id) = 1) := by - simp [contr, contrIndexList, countId] + simp only [countId, contr, contrIndexList] rw [List.countP_filter, List.countP_filter] congr funext J @@ -185,13 +185,16 @@ lemma contrEquiv_contrCond : ColorMap.ContrCond l.contrEquiv l.colorMap := by @[simp] lemma contrEquiv_on_withDual_empty (i : Fin l.contr.length) (h : l.withDual = ∅) : l.contrEquiv (Sum.inr i) = Fin.cast (by simp [h]) i := by - simp [contrEquiv] + simp only [contrEquiv, Equiv.trans_apply, Equiv.sumCongr_apply, Equiv.coe_refl, Sum.map_inr, + id_eq] change l.dualEquiv (Sum.inr ((Fin.castOrderIso _).toEquiv i)) = _ - simp [dualEquiv, withoutDualEquiv] + simp only [dualEquiv, withoutDualEquiv, RelIso.coe_fn_toEquiv, Fin.castOrderIso_apply, + Equiv.trans_apply, Equiv.sumCongr_apply, Equiv.coe_refl, Sum.map_inr, + Equiv.Finset.union_symm_inr, Finset.coe_orderIsoOfFin_apply, Equiv.subtypeUnivEquiv_apply] have h : l.withoutDual = Finset.univ := by have hx := l.withDual_union_withoutDual simp_all - simp [h] + simp only [h] rw [orderEmbOfFin_univ] · rfl · rw [h] diff --git a/HepLean/SpaceTime/LorentzTensor/IndexNotation/IndexList/CountId.lean b/HepLean/SpaceTime/LorentzTensor/IndexNotation/IndexList/CountId.lean index 4560fe4..083aa8d 100644 --- a/HepLean/SpaceTime/LorentzTensor/IndexNotation/IndexList/CountId.lean +++ b/HepLean/SpaceTime/LorentzTensor/IndexNotation/IndexList/CountId.lean @@ -40,8 +40,7 @@ lemma countId_append (I : Index X) : (l ++ l2).countId I = l.countId I + l2.coun omit [IndexNotation X] [Fintype X] [DecidableEq X] in lemma countId_eq_length_filter (I : Index X) : l.countId I = (l.val.filter (fun J => I.id = J.id)).length := by - simp [countId] - rw [List.countP_eq_length_filter] + rw [countId, List.countP_eq_length_filter] omit [IndexNotation X] [Fintype X] [DecidableEq X] in lemma countId_index_neq_zero (i : Fin l.length) : l.countId (l.val.get i) ≠ 0 := by @@ -58,7 +57,7 @@ lemma countId_append_symm (I : Index X) : (l ++ l2).countId I = (l2 ++ l).countI omit [IndexNotation X] [Fintype X] [DecidableEq X] in lemma countId_eq_one_append_mem_right_self_eq_one {I : Index X} (hI : I ∈ l2.val) (h : (l ++ l2).countId I = 1) : l2.countId I = 1 := by - simp at h + simp only [countId_append] at h have hmem : I ∈ l2.val.filter (fun J => I.id = J.id) := by simp [List.mem_filter, decide_True, and_true, hI] have h1 : l2.countId I ≠ 0 := by @@ -72,7 +71,7 @@ lemma countId_eq_one_append_mem_right_self_eq_one {I : Index X} (hI : I ∈ l2.v omit [IndexNotation X] [Fintype X] [DecidableEq X] in lemma countId_eq_one_append_mem_right_other_eq_zero {I : Index X} (hI : I ∈ l2.val) (h : (l ++ l2).countId I = 1) : l.countId I = 0 := by - simp at h + simp only [countId_append] at h have hmem : I ∈ l2.val.filter (fun J => I.id = J.id) := by simp [List.mem_filter, decide_True, and_true, hI] have h1 : l2.countId I ≠ 0 := by @@ -143,7 +142,7 @@ lemma filter_finRange (i : Fin l.length) : rw [@List.length_eq_one] at h3 obtain ⟨a, ha⟩ := h3 rw [ha] at h4 - simp at h4 + simp only [List.mem_singleton] at h4 subst h4 exact ha @@ -185,7 +184,7 @@ lemma countId_gt_zero_of_mem_withDual (i : Fin l.length) (h : i ∈ l.withDual) 1 < l.countId (l.val.get i) := by rw [countId_get] by_contra hn - simp at hn + simp only [Nat.lt_add_left_iff_pos, not_lt, Nat.le_zero_eq] at hn rw [List.countP_eq_length_filter, List.length_eq_zero] at hn rw [mem_withDual_iff_exists] at h obtain ⟨j, hj⟩ := h @@ -244,8 +243,8 @@ lemma countId_of_not_mem_withDualInOther (i : Fin l.length) (h : i ∉ l.withDua rw [mem_withInDualOther_iff_exists] at h simp at h have hj' := h ⟨l2.val.indexOf j, hj'⟩ - simp [AreDualInOther, idMap] at hj' - simp at hj + simp only [AreDualInOther, idMap, List.get_eq_getElem, List.getElem_indexOf] at hj' + simp only [List.get_eq_getElem, List.mem_filter, decide_eq_true_eq] at hj simp_all only [List.get_eq_getElem, List.isEmpty_eq_true, List.getElem_indexOf, not_true_eq_false] omit [IndexNotation X] [Fintype X] in @@ -264,7 +263,8 @@ lemma mem_withoutDual_iff_countId_eq_one (i : Fin l.length) : · exact countId_of_not_mem_withDual l i (l.not_mem_withDual_of_mem_withoutDual i h) · by_contra hn have h : i ∈ l.withDual := by - simp [withoutDual] at hn + simp only [withoutDual, Option.isNone_iff_eq_none, Finset.mem_filter, Finset.mem_univ, + true_and] at hn simpa using Option.ne_none_iff_isSome.mp hn rw [mem_withDual_iff_countId_gt_one] at h omega @@ -291,7 +291,8 @@ lemma countId_eq_two_of_mem_withUniqueDual (i : Fin l.length) (h : i ∈ l.withU · simp only [List.filter_filter, decide_eq_true_eq, Bool.decide_and] refine List.filter_congr (fun j _ => ?_) simp only [Bool.and_iff_right_iff_imp, decide_eq_true_eq] - simp [withUniqueDual] at h + simp only [withUniqueDual, mem_withDual_iff_isSome, Finset.mem_filter, Finset.mem_univ, + true_and] at h intro hj have hj' := h.2 j hj apply Option.some_injective @@ -306,11 +307,11 @@ lemma mem_withUniqueDual_of_countId_eq_two (i : Fin l.length) have hw : i ∈ l.withDual := by rw [mem_withDual_iff_countId_gt_one, h] exact Nat.one_lt_two - simp [withUniqueDual] + simp only [withUniqueDual, mem_withDual_iff_isSome, Finset.mem_filter, Finset.mem_univ, true_and] apply And.intro ((mem_withDual_iff_isSome l i).mp hw) intro j hj rw [@countId_get] at h - simp [List.countP_eq_length_filter] at h + simp only [List.countP_eq_length_filter, Nat.reduceEqDiff] at h rw [List.length_eq_one] at h obtain ⟨a, ha⟩ := h have hj : j ∈ List.filter (fun j => decide (l.AreDualInSelf i j)) (List.finRange l.length) := by @@ -322,7 +323,7 @@ lemma mem_withUniqueDual_of_countId_eq_two (i : Fin l.length) (List.finRange l.length).filter (fun j => decide (l.AreDualInSelf i j)) := by simp rw [ha] at ht - simp at ht + simp only [List.mem_singleton] at ht subst ht simp @@ -348,12 +349,14 @@ lemma countId_eq_one_of_mem_withUniqueDualInOther (i : Fin l.length) simpa using eq_comm trans List.filter (fun j => j = i') ((List.finRange l2.length).filter (fun j => (l.AreDualInOther l2 i j))) - · simp + · simp only [List.filter_filter, decide_eq_true_eq, Bool.decide_and] apply List.filter_congr (fun j _ => ?_) exact Bool.and_comm (decide (l.AreDualInOther l2 i j)) (decide (j = i')) - · simp + · simp only [List.filter_filter, decide_eq_true_eq, Bool.decide_and] refine List.filter_congr (fun j _ => ?_) - simp [withUniqueDualInOther] at h + simp only [withUniqueDualInOther, mem_withDual_iff_isSome, Bool.not_eq_true, + Option.not_isSome, Option.isNone_iff_eq_none, mem_withInDualOther_iff_isSome, + Finset.mem_filter, Finset.mem_univ, true_and] at h simp only [Bool.and_iff_right_iff_imp, decide_eq_true_eq] intro hj have hj' := h.2.2 j hj @@ -390,13 +393,13 @@ lemma mem_withUniqueDualInOther_of_countId_eq_one (i : Fin l.length) (List.finRange l2.length) := by simpa using hj rw [ha] at hj - simp at hj + simp only [List.mem_singleton] at hj subst hj have ht : (l.getDualInOther? l2 i).get ((mem_withInDualOther_iff_isSome l l2 i).mp hw) ∈ (List.finRange l2.length).filter (fun j => decide (l.AreDualInOther l2 i j)) := by simp rw [ha] at ht - simp at ht + simp only [List.mem_singleton] at ht subst ht exact Option.some_get ((mem_withInDualOther_iff_isSome l l2 i).mp hw) @@ -419,7 +422,7 @@ lemma getDual?_countId (i : Fin l.length) (h : (l.getDual? i).isSome) : l.countId (l.val[Fin.val ((l.getDual? i).get h)]) = l.countId (l.val.get i) := by apply countId_congr change l.idMap ((l.getDual? i).get h) = _ - simp + simp only [getDual?_get_id, List.get_eq_getElem] rfl @[simp] @@ -427,7 +430,7 @@ lemma getDualInOther?_countId_right (i : Fin l.length) (h : (l.getDualInOther? l l2.countId (l2.val[Fin.val ((l.getDualInOther? l2 i).get h)]) = l2.countId (l.val.get i) := by apply countId_congr change l2.idMap ((l.getDualInOther? l2 i).get h) = _ - simp + simp only [getDualInOther?_id, List.get_eq_getElem] rfl @[simp] @@ -435,7 +438,7 @@ lemma getDualInOther?_countId_left (i : Fin l.length) (h : (l.getDualInOther? l2 l.countId (l2.val[Fin.val ((l.getDualInOther? l2 i).get h)]) = l.countId (l.val.get i) := by apply countId_congr change l2.idMap ((l.getDualInOther? l2 i).get h) = _ - simp + simp only [getDualInOther?_id, List.get_eq_getElem] rfl lemma getDual?_isSome_of_countId_eq_two {i : Fin l.length} @@ -542,8 +545,14 @@ lemma filter_id_of_countId_eq_two {i : Fin l.length} exact Fin.ne_of_lt hi' · intro a b fin_cases a, b - <;> simp [hi'] - exact Fin.le_of_lt hi' + · simp [hi'] + · simp only [List.get_eq_getElem, List.length_cons, List.length_singleton, Nat.reduceAdd, + List.length_nil, Fin.zero_eta, Fin.isValue, Function.Embedding.coeFn_mk, + Matrix.cons_val_zero, Fin.mk_one, Matrix.cons_val_one, Matrix.head_cons, Fin.zero_le, + iff_true] + exact Fin.le_of_lt hi' + · simp [hi'] + · simp [hi'] · intro a fin_cases a <;> rfl · rw [hc] diff --git a/HepLean/SpaceTime/LorentzTensor/MulActionTensor.lean b/HepLean/SpaceTime/LorentzTensor/MulActionTensor.lean index f8ba924..052b4dc 100644 --- a/HepLean/SpaceTime/LorentzTensor/MulActionTensor.lean +++ b/HepLean/SpaceTime/LorentzTensor/MulActionTensor.lean @@ -148,7 +148,7 @@ lemma rep_mapIso (e : X ≃ Y) (h : cX = cY ∘ e) (g : G) : simp only [LinearMap.compMultilinearMap_apply, LinearMap.coe_comp, LinearEquiv.coe_coe, Function.comp_apply] erw [mapIso_tprod] - simp [rep, colorModuleCast_equivariant_apply] + simp only [rep, MonoidHom.coe_mk, OneHom.coe_mk] change (PiTensorProduct.map fun x => (repColorModule (cY x)) g) ((PiTensorProduct.tprod R) fun i => (𝓣.colorModuleCast _) (x (e.symm i))) = (𝓣.mapIso e h) ((PiTensorProduct.map _) ((PiTensorProduct.tprod R) x)) @@ -171,7 +171,7 @@ omit [Fintype X] [DecidableEq X] [Fintype Y] [DecidableEq Y] in lemma rep_tprod (g : G) (f : (i : X) → 𝓣.ColorModule (cX i)) : g • (PiTensorProduct.tprod R f) = PiTensorProduct.tprod R (fun x => repColorModule (cX x) g (f x)) := by - simp [rep] + simp only [rep, MonoidHom.coe_mk, OneHom.coe_mk] change (PiTensorProduct.map _) ((PiTensorProduct.tprod R) f) = _ rw [PiTensorProduct.map_tprod] diff --git a/HepLean/SpaceTime/LorentzTensor/Real/IndexNotation.lean b/HepLean/SpaceTime/LorentzTensor/Real/IndexNotation.lean index 8b3d531..f3906fd 100644 --- a/HepLean/SpaceTime/LorentzTensor/Real/IndexNotation.lean +++ b/HepLean/SpaceTime/LorentzTensor/Real/IndexNotation.lean @@ -33,13 +33,13 @@ instance : IndexNotation realTensorColor.Color where by intro c by_cases hc : c = 'ᵘ' - · simp [hc] + · simp only [↓Char.isValue, hc, ↓reduceIte] exact SetCoe.ext (id (Eq.symm hc)) · have hc' : c = 'ᵤ' := by have hc2 := c.2 - simp at hc2 + simp only [↓Char.isValue, Finset.mem_insert, Finset.mem_singleton] at hc2 simp_all - simp [hc'] + simp only [↓Char.isValue, hc', Char.reduceEq, ↓reduceIte] exact SetCoe.ext (id (Eq.symm hc'))⟩ namespace realLorentzTensor diff --git a/HepLean/SpaceTime/LorentzTensor/RisingLowering.lean b/HepLean/SpaceTime/LorentzTensor/RisingLowering.lean index 2652cb4..949e310 100644 --- a/HepLean/SpaceTime/LorentzTensor/RisingLowering.lean +++ b/HepLean/SpaceTime/LorentzTensor/RisingLowering.lean @@ -58,8 +58,8 @@ def boolFin (c₁ c₂ : 𝓒.ColorMap (Fin n)) : Bool := omit [Fintype 𝓒.Color] in lemma boolFin_DualMap {c₁ c₂ : 𝓒.ColorMap (Fin n)} (h : boolFin c₁ c₂ = true) : DualMap c₁ c₂ := by - simp [boolFin] at h - simp [DualMap] + simp only [boolFin, Bool.if_false_right, Bool.and_true, List.all_eq_true, decide_eq_true_eq] at h + simp only [DualMap] funext x have h2 {n : ℕ} (m : Fin n) : m ∈ Fin.list n := by have h1' : (Fin.list n)[m] = m := by @@ -76,8 +76,8 @@ def boolFin' (c₁ c₂ : 𝓒.ColorMap (Fin n)) : Bool := omit [Fintype 𝓒.Color] lemma boolFin'_DualMap {c₁ c₂ : 𝓒.ColorMap (Fin n)} (h : boolFin' c₁ c₂ = true) : DualMap c₁ c₂ := by - simp [boolFin'] at h - simp [DualMap] + simp only [boolFin', decide_eq_true_eq] at h + simp only [DualMap] funext x exact h x @@ -103,7 +103,7 @@ lemma dual_eq_of_neq (h : DualMap c₁ c₂) {x : X} (h' : c₁ x ≠ c₂ x) : 𝓒.τ (c₁ x) = c₂ x := by rw [DualMap] at h have h1 := congrFun h x - simp [colorQuot, HasEquiv.Equiv, Setoid.r, colorRel] at h1 + simp only [Function.comp_apply, colorQuot, Quotient.eq, HasEquiv.Equiv, Setoid.r, colorRel] at h1 simp_all only [ne_eq, false_or] exact 𝓒.τ_involutive (c₂ x) @@ -270,7 +270,7 @@ def dualizeAll : 𝓣.Tensor cX ≃ₗ[R] 𝓣.Tensor (𝓣.τ ∘ cX) := by apply LinearMap.ext refine fun x ↦ PiTensorProduct.induction_on' x ?_ (by intro a b hx a - simp [map_add, add_tmul, hx] + simp only [Function.comp_apply, map_add, hx, LinearMap.id_coe, id_eq, LinearMap.coe_comp] simp_all only [Function.comp_apply, LinearMap.coe_comp, LinearMap.id_coe, id_eq]) intro rx fx simp only [Function.comp_apply, PiTensorProduct.tprodCoeff_eq_smul_tprod, @@ -287,10 +287,10 @@ lemma dualizeAll_equivariant (g : G) : (𝓣.dualizeAll.toLinearMap) ∘ₗ (@re = 𝓣.rep g ∘ₗ (𝓣.dualizeAll.toLinearMap) := by apply LinearMap.ext intro x - simp [dualizeAll] + simp only [dualizeAll, Function.comp_apply, LinearEquiv.ofLinear_toLinearMap, LinearMap.coe_comp] refine PiTensorProduct.induction_on' x ?_ (by intro a b hx a - simp [map_add, add_tmul, hx] + simp only [map_add, hx] simp_all only [Function.comp_apply, LinearMap.coe_comp, LinearMap.id_coe, id_eq]) intro rx fx simp only [PiTensorProduct.tprodCoeff_eq_smul_tprod, LinearMapClass.map_smul, rep_tprod]