refactor: Replace some simp with simp only

This commit is contained in:
jstoobysmith 2024-09-04 15:33:54 -04:00
parent da5e0e3f00
commit 49d089d4cd
17 changed files with 56 additions and 41 deletions

View file

@ -71,9 +71,11 @@ lemma doublePoint_Y₃_Y₃ (R : MSSMACC.LinSols) :
simp only [mul_one, Fin.isValue, toSMSpecies_apply, one_mul, mul_neg, neg_mul, neg_neg, mul_zero, simp only [mul_one, Fin.isValue, toSMSpecies_apply, one_mul, mul_neg, neg_mul, neg_neg, mul_zero,
zero_mul, add_zero, Hd_apply, Fin.reduceFinMk, Hu_apply] zero_mul, add_zero, Hd_apply, Fin.reduceFinMk, Hu_apply]
have hLin := R.linearSol have hLin := R.linearSol
simp at hLin simp only [MSSMACC_numberLinear, MSSMACC_linearACCs, Nat.reduceMul, Fin.isValue,
Fin.reduceFinMk] at hLin
have h3 := hLin 3 have h3 := hLin 3
simp [Fin.sum_univ_three] at h3 simp only [Fin.isValue, Fin.sum_univ_three, Prod.mk_zero_zero, Prod.mk_one_one, LinearMap.coe_mk,
AddHom.coe_mk] at h3
linear_combination (norm := ring_nf) 6 * h3 linear_combination (norm := ring_nf) 6 * h3
simp only [Fin.isValue, Prod.mk_zero_zero, Prod.mk_one_one, add_add_sub_cancel, add_neg_cancel] simp only [Fin.isValue, Prod.mk_zero_zero, Prod.mk_one_one, add_add_sub_cancel, add_neg_cancel]

View file

@ -51,7 +51,8 @@ def permCharges {n : } : Representation (PermGroup n) (PureU1 n).Charges
lemma accGrav_invariant {n : } (f : (PermGroup n)) (S : (PureU1 n).Charges) : lemma accGrav_invariant {n : } (f : (PermGroup n)) (S : (PureU1 n).Charges) :
PureU1.accGrav n (permCharges f S) = accGrav n S := by PureU1.accGrav n (permCharges f S) = accGrav n S := by
simp [accGrav, permCharges] simp only [accGrav, PermGroup, permCharges, MonoidHom.coe_mk, OneHom.coe_mk, LinearMap.coe_mk,
AddHom.coe_mk, chargeMap_apply]
apply (Equiv.Perm.sum_comp _ _ _ ?_) apply (Equiv.Perm.sum_comp _ _ _ ?_)
simp simp
@ -144,7 +145,7 @@ lemma pairSwap_other {n : } (i j k : Fin n) (hik : i ≠ k) (hjk : j ≠ k) :
lemma pairSwap_inv_other {n : } {i j k : Fin n} (hik : i ≠ k) (hjk : j ≠ k) : lemma pairSwap_inv_other {n : } {i j k : Fin n} (hik : i ≠ k) (hjk : j ≠ k) :
(pairSwap i j).invFun k = k := by (pairSwap i j).invFun k = k := by
simp [pairSwap] simp only [pairSwap, Equiv.invFun_as_coe, Equiv.coe_fn_symm_mk]
split split
· rename_i h · rename_i h
exact False.elim (hik (id (Eq.symm h))) exact False.elim (hik (id (Eq.symm h)))
@ -197,8 +198,8 @@ lemma permTwo_fst : (permTwo hij hij').toFun i' = i := by
have ht := Equiv.extendSubtype_apply_of_mem have ht := Equiv.extendSubtype_apply_of_mem
((permTwoInj hij').toEquivRange.symm.trans ((permTwoInj hij').toEquivRange.symm.trans
(permTwoInj hij).toEquivRange) i' (permTwoInj_fst hij') (permTwoInj hij).toEquivRange) i' (permTwoInj_fst hij')
simp at ht simp only [Equiv.trans_apply, Function.Embedding.toEquivRange_apply] at ht
simp [ht, permTwoInj_fst_apply] simp only [Equiv.toFun_as_coe, ht, permTwoInj_fst_apply, Fin.isValue]
rfl rfl
lemma permTwo_snd : (permTwo hij hij').toFun j' = j := by lemma permTwo_snd : (permTwo hij hij').toFun j' = j := by

View file

@ -269,7 +269,7 @@ def cubeTriLin : TriLinearSymm (SMCharges n).Charges := TriLinearSymm.mk₃
apply Fintype.sum_congr apply Fintype.sum_congr
intro i intro i
repeat erw [map_smul] repeat erw [map_smul]
simp [HSMul.hSMul, SMul.smul] simp only [HSMul.hSMul, SMul.smul, toSpecies_apply, Fin.isValue]
ring) ring)
(by (by
intro S T R L intro S T R L

View file

@ -38,12 +38,12 @@ variable {n : }
lemma SU2Sol (S : (SMNoGrav n).LinSols) : accSU2 S.val = 0 := by lemma SU2Sol (S : (SMNoGrav n).LinSols) : accSU2 S.val = 0 := by
have hS := S.linearSol have hS := S.linearSol
simp at hS simp only [SMNoGrav_numberLinear, SMNoGrav_linearACCs, Fin.isValue] at hS
exact hS 0 exact hS 0
lemma SU3Sol (S : (SMNoGrav n).LinSols) : accSU3 S.val = 0 := by lemma SU3Sol (S : (SMNoGrav n).LinSols) : accSU3 S.val = 0 := by
have hS := S.linearSol have hS := S.linearSol
simp at hS simp only [SMNoGrav_numberLinear, SMNoGrav_linearACCs, Fin.isValue] at hS
exact hS 1 exact hS 1
lemma cubeSol (S : (SMNoGrav n).Sols) : accCube S.val = 0 := by lemma cubeSol (S : (SMNoGrav n).Sols) : accCube S.val = 0 := by
@ -55,7 +55,7 @@ def chargeToLinear (S : (SMNoGrav n).Charges) (hSU2 : accSU2 S = 0) (hSU3 : accS
(SMNoGrav n).LinSols := (SMNoGrav n).LinSols :=
⟨S, by ⟨S, by
intro i intro i
simp at i simp only [SMNoGrav_numberLinear] at i
match i with match i with
| 0 => exact hSU2 | 0 => exact hSU2
| 1 => exact hSU3⟩ | 1 => exact hSU3⟩

View file

@ -37,17 +37,17 @@ variable {n : }
lemma gravSol (S : (SM n).LinSols) : accGrav S.val = 0 := by lemma gravSol (S : (SM n).LinSols) : accGrav S.val = 0 := by
have hS := S.linearSol have hS := S.linearSol
simp at hS simp only [SM_numberLinear, SM_linearACCs, Fin.isValue] at hS
exact hS 0 exact hS 0
lemma SU2Sol (S : (SM n).LinSols) : accSU2 S.val = 0 := by lemma SU2Sol (S : (SM n).LinSols) : accSU2 S.val = 0 := by
have hS := S.linearSol have hS := S.linearSol
simp at hS simp only [SM_numberLinear, SM_linearACCs, Fin.isValue] at hS
exact hS 1 exact hS 1
lemma SU3Sol (S : (SM n).LinSols) : accSU3 S.val = 0 := by lemma SU3Sol (S : (SM n).LinSols) : accSU3 S.val = 0 := by
have hS := S.linearSol have hS := S.linearSol
simp at hS simp only [SM_numberLinear, SM_linearACCs, Fin.isValue] at hS
exact hS 2 exact hS 2
lemma cubeSol (S : (SM n).Sols) : accCube S.val = 0 := S.cubicSol lemma cubeSol (S : (SM n).Sols) : accCube S.val = 0 := S.cubicSol
@ -58,7 +58,7 @@ def chargeToLinear (S : (SM n).Charges) (hGrav : accGrav S = 0)
(hSU2 : accSU2 S = 0) (hSU3 : accSU3 S = 0) : (SM n).LinSols := (hSU2 : accSU2 S = 0) (hSU3 : accSU3 S = 0) : (SM n).LinSols :=
⟨S, by ⟨S, by
intro i intro i
simp at i simp only [SM_numberLinear] at i
match i with match i with
| 0 => exact hGrav | 0 => exact hGrav
| 1 => exact hSU2 | 1 => exact hSU2
@ -100,14 +100,14 @@ def perm (n : ) : ACCSystemGroupAction (SM n) where
rep := repCharges rep := repCharges
linearInvariant := by linearInvariant := by
intro i intro i
simp at i simp only [SM_numberLinear] at i
match i with match i with
| 0 => exact accGrav_invariant | 0 => exact accGrav_invariant
| 1 => exact accSU2_invariant | 1 => exact accSU2_invariant
| 2 => exact accSU3_invariant | 2 => exact accSU3_invariant
quadInvariant := by quadInvariant := by
intro i intro i
simp at i simp only [SM_numberQuadratic] at i
exact Fin.elim0 i exact Fin.elim0 i
cubicInvariant := accCube_invariant cubicInvariant := accCube_invariant

View file

@ -88,6 +88,7 @@ lemma left_zero : potential m₁₁2 m₂₂2 𝓵₁ 𝓵₂ 𝓵₃ 𝓵₄ m
def IsBounded (m₁₁2 m₂₂2 𝓵₁ 𝓵₂ 𝓵₃ 𝓵₄ : ) (m₁₂2 𝓵₅ 𝓵₆ 𝓵₇ : ) : Prop := def IsBounded (m₁₁2 m₂₂2 𝓵₁ 𝓵₂ 𝓵₃ 𝓵₄ : ) (m₁₂2 𝓵₅ 𝓵₆ 𝓵₇ : ) : Prop :=
∃ c, ∀ Φ1 Φ2 x, c ≤ potential m₁₁2 m₂₂2 𝓵₁ 𝓵₂ 𝓵₃ 𝓵₄ m₁₂2 𝓵₅ 𝓵₆ 𝓵₇ Φ1 Φ2 x ∃ c, ∀ Φ1 Φ2 x, c ≤ potential m₁₁2 m₂₂2 𝓵₁ 𝓵₂ 𝓵₃ 𝓵₄ m₁₂2 𝓵₅ 𝓵₆ 𝓵₇ Φ1 Φ2 x
/-! TODO: Show that if the potential is bounded then `0 ≤ 𝓵₁` and `0 ≤ 𝓵₂`. -/
/-! /-!
## Smoothness of the potential ## Smoothness of the potential

View file

@ -131,7 +131,7 @@ lemma toHomogeneousQuad_add {V : Type} [AddCommMonoid V] [Module V]
(τ : BiLinearSymm V) (S T : V) : (τ : BiLinearSymm V) (S T : V) :
τ.toHomogeneousQuad (S + T) = τ.toHomogeneousQuad S + τ.toHomogeneousQuad (S + T) = τ.toHomogeneousQuad S +
τ.toHomogeneousQuad T + 2 * τ S T := by τ.toHomogeneousQuad T + 2 * τ S T := by
simp [toHomogeneousQuad_apply] simp only [HomogeneousQuadratic, toHomogeneousQuad_apply, map_add]
rw [τ.map_add₁, τ.map_add₁, τ.swap T S] rw [τ.map_add₁, τ.map_add₁, τ.swap T S]
ring ring

View file

@ -99,7 +99,7 @@ instance spaceTimeAsLieModule : LieModule lorentzAlgebra (LorentzVector 3) w
smul_lie r Λ x := by smul_lie r Λ x := by
simp [Bracket.bracket, smul_mulVec_assoc] simp [Bracket.bracket, smul_mulVec_assoc]
lie_smul r Λ x := by lie_smul r Λ x := by
simp [Bracket.bracket] simp only [Bracket.bracket]
rw [mulVec_smul] rw [mulVec_smul]
end SpaceTime end SpaceTime

View file

@ -59,14 +59,23 @@ noncomputable def toSelfAdjointMatrix :
funext i funext i
match i with match i with
| Sum.inl 0 => | Sum.inl 0 =>
simp [fromSelfAdjointMatrix', toSelfAdjointMatrix', toMatrix, toMatrix_isSelfAdjoint] simp only [fromSelfAdjointMatrix', one_div, toSelfAdjointMatrix', toMatrix,
LorentzVector.time, Fin.isValue, LorentzVector.space, Function.comp_apply, of_apply,
cons_val', cons_val_zero, empty_val', cons_val_fin_one, cons_val_one, head_cons,
head_fin_const, add_add_sub_cancel, add_re, ofReal_re]
ring_nf ring_nf
| Sum.inr 0 => | Sum.inr 0 =>
simp [fromSelfAdjointMatrix', toSelfAdjointMatrix', toMatrix, toMatrix_isSelfAdjoint] simp [fromSelfAdjointMatrix', toSelfAdjointMatrix', toMatrix, toMatrix_isSelfAdjoint]
| Sum.inr 1 => | Sum.inr 1 =>
simp [fromSelfAdjointMatrix', toSelfAdjointMatrix', toMatrix, toMatrix_isSelfAdjoint] simp only [fromSelfAdjointMatrix', toSelfAdjointMatrix', toMatrix, LorentzVector.time,
Fin.isValue, LorentzVector.space, Function.comp_apply, of_apply, cons_val', cons_val_zero,
empty_val', cons_val_fin_one, cons_val_one, head_fin_const, add_im, ofReal_im, mul_im,
ofReal_re, I_im, mul_one, I_re, mul_zero, add_zero, zero_add]
| Sum.inr 2 => | Sum.inr 2 =>
simp [fromSelfAdjointMatrix', toSelfAdjointMatrix', toMatrix, toMatrix_isSelfAdjoint] simp only [fromSelfAdjointMatrix', one_div, toSelfAdjointMatrix', toMatrix,
LorentzVector.time, Fin.isValue, LorentzVector.space, Function.comp_apply, of_apply,
cons_val', cons_val_zero, empty_val', cons_val_fin_one, cons_val_one, head_cons,
head_fin_const, add_sub_sub_cancel, add_re, ofReal_re]
ring ring
right_inv x := by right_inv x := by
simp only [toSelfAdjointMatrix', toMatrix, LorentzVector.time, fromSelfAdjointMatrix', one_div, simp only [toSelfAdjointMatrix', toMatrix, LorentzVector.time, fromSelfAdjointMatrix', one_div,

View file

@ -74,12 +74,12 @@ lemma decomp_stdBasis (v : LorentzVector d) : ∑ i, v i • e i = v := by
funext ν funext ν
rw [Finset.sum_apply] rw [Finset.sum_apply]
rw [Finset.sum_eq_single_of_mem ν] rw [Finset.sum_eq_single_of_mem ν]
· simp [HSMul.hSMul, SMul.smul, stdBasis, Pi.basisFun_apply] · simp only [HSMul.hSMul, SMul.smul, stdBasis]
erw [Pi.basisFun_apply] erw [Pi.basisFun_apply]
simp only [Pi.single_eq_same, mul_one] simp only [Pi.single_eq_same, mul_one]
· exact Finset.mem_univ ν · exact Finset.mem_univ ν
· intros b _ hbi · intros b _ hbi
simp [HSMul.hSMul, SMul.smul, stdBasis, Pi.basisFun_apply] simp only [HSMul.hSMul, SMul.smul, stdBasis, mul_eq_zero]
erw [Pi.basisFun_apply] erw [Pi.basisFun_apply]
simp only [Pi.single] simp only [Pi.single]
apply Or.inr $ Function.update_noteq (id (Ne.symm hbi)) 1 0 apply Or.inr $ Function.update_noteq (id (Ne.symm hbi)) 1 0

View file

@ -32,7 +32,7 @@ lemma rep_apply_stdBasis (g : LorentzGroup d) (μ : Fin 1 ⊕ Fin d) :
simp only [rep_apply, Fintype.sum_sum_type, Finset.univ_unique, Fin.default_eq_zero, Fin.isValue, simp only [rep_apply, Fintype.sum_sum_type, Finset.univ_unique, Fin.default_eq_zero, Fin.isValue,
Finset.sum_singleton, decomp_stdBasis'] Finset.sum_singleton, decomp_stdBasis']
funext ν funext ν
simp [LorentzVector.stdBasis, Pi.basisFun_apply] simp only [stdBasis, Matrix.transpose_apply]
erw [Pi.basisFun_apply, Matrix.mulVec_single_one] erw [Pi.basisFun_apply, Matrix.mulVec_single_one]
rfl rfl

View file

@ -114,7 +114,7 @@ lemma not_mem_iff : v ∉ FuturePointing d ↔ v.1.time ≤ 0 := by
refine Iff.intro (fun h => ?_) (fun h => ?_) refine Iff.intro (fun h => ?_) (fun h => ?_)
· exact le_of_not_lt ((mem_iff v).mp.mt h) · exact le_of_not_lt ((mem_iff v).mp.mt h)
· have h1 := (mem_iff v).mp.mt · have h1 := (mem_iff v).mp.mt
simp at h1 simp only [LorentzVector.time, Fin.isValue, not_lt] at h1
exact h1 h exact h1 h
lemma not_mem_iff_neg : v ∉ FuturePointing d ↔ neg v ∈ FuturePointing d := by lemma not_mem_iff_neg : v ∉ FuturePointing d ↔ neg v ∈ FuturePointing d := by

View file

@ -36,7 +36,7 @@ scoped[minkowskiMatrix] notation "η" => minkowskiMatrix
@[simp] @[simp]
lemma sq : @minkowskiMatrix d * minkowskiMatrix = 1 := by lemma sq : @minkowskiMatrix d * minkowskiMatrix = 1 := by
simp [minkowskiMatrix, LieAlgebra.Orthogonal.indefiniteDiagonal] simp only [minkowskiMatrix, LieAlgebra.Orthogonal.indefiniteDiagonal, diagonal_mul_diagonal]
ext1 i j ext1 i j
rcases i with i | i <;> rcases j with j | j rcases i with i | i <;> rcases j with j | j
· simp only [diagonal, of_apply, Sum.inl.injEq, Sum.elim_inl, mul_one] · simp only [diagonal, of_apply, Sum.inl.injEq, Sum.elim_inl, mul_one]
@ -70,9 +70,7 @@ lemma η_apply_mul_η_apply_diag (μ : Fin 1 ⊕ Fin d) : η μ μ * η μ μ =
lemma as_block : @minkowskiMatrix d = ( lemma as_block : @minkowskiMatrix d = (
Matrix.fromBlocks (1 : Matrix (Fin 1) (Fin 1) ) 0 0 (-1 : Matrix (Fin d) (Fin d) )) := by Matrix.fromBlocks (1 : Matrix (Fin 1) (Fin 1) ) 0 0 (-1 : Matrix (Fin d) (Fin d) )) := by
rw [minkowskiMatrix] rw [minkowskiMatrix, LieAlgebra.Orthogonal.indefiniteDiagonal, ← fromBlocks_diagonal]
simp [LieAlgebra.Orthogonal.indefiniteDiagonal]
rw [← fromBlocks_diagonal]
refine fromBlocks_inj.mpr ?_ refine fromBlocks_inj.mpr ?_
simp only [diagonal_one, true_and] simp only [diagonal_one, true_and]
funext i j funext i j
@ -291,7 +289,7 @@ lemma matrix_eq_iff_eq_forall' : (∀ v, Λ *ᵥ v= Λ' *ᵥ v) ↔ ∀ w v, ⟪
refine sub_eq_zero.1 ?_ refine sub_eq_zero.1 ?_
have h1 := h v have h1 := h v
rw [nondegenerate] at h1 rw [nondegenerate] at h1
simp [sub_mulVec] at h1 simp only [sub_mulVec] at h1
exact h1 exact h1
lemma matrix_eq_iff_eq_forall : Λ = Λ' ↔ ∀ w v, ⟪v, Λ *ᵥ w⟫ₘ = ⟪v, Λ' *ᵥ w⟫ₘ := by lemma matrix_eq_iff_eq_forall : Λ = Λ' ↔ ∀ w v, ⟪v, Λ *ᵥ w⟫ₘ = ⟪v, Λ' *ᵥ w⟫ₘ := by
@ -334,7 +332,8 @@ lemma on_basis (μ ν : Fin 1 ⊕ Fin d) : ⟪e μ, e ν⟫ₘ = η μ ν := by
rw [basis_left, stdBasis_apply] rw [basis_left, stdBasis_apply]
by_cases h : μ = ν by_cases h : μ = ν
· simp [h] · simp [h]
· simp [h, LieAlgebra.Orthogonal.indefiniteDiagonal, minkowskiMatrix] · simp only [minkowskiMatrix, LieAlgebra.Orthogonal.indefiniteDiagonal, diagonal_apply_eq,
mul_ite, mul_one, mul_zero, ne_eq, h, not_false_eq_true, diagonal_apply_ne, ite_eq_right_iff]
exact fun a => False.elim (h (id (Eq.symm a))) exact fun a => False.elim (h (id (Eq.symm a)))
lemma matrix_apply_stdBasis (ν μ : Fin 1 ⊕ Fin d) : lemma matrix_apply_stdBasis (ν μ : Fin 1 ⊕ Fin d) :

View file

@ -212,10 +212,8 @@ def appendInl : Fin l.length ↪ Fin (l ++ l2).length where
/-- The inclusion of the indices of `l2` into the indices of `l ++ l2`. -/ /-- The inclusion of the indices of `l2` into the indices of `l ++ l2`. -/
def appendInr : Fin l2.length ↪ Fin (l ++ l2).length where def appendInr : Fin l2.length ↪ Fin (l ++ l2).length where
toFun := appendEquiv ∘ Sum.inr toFun := appendEquiv ∘ Sum.inr
inj' := by inj' i j h := by
intro i j h simpa only [Function.comp, EmbeddingLike.apply_eq_iff_eq, Sum.inr.injEq] using h
simp [Function.comp] at h
exact h
@[simp] @[simp]
lemma appendInl_appendEquiv : lemma appendInl_appendEquiv :
@ -315,7 +313,7 @@ lemma filter_sort_comm {n : } (s : Finset (Fin n)) (p : Fin n → Prop) [Deci
intro i j h1 _ _ intro i j h1 _ _
have hs : List.Sorted (fun i j => i ≤ j) (List.mergeSort (fun i j => i ≤ j) m) := by have hs : List.Sorted (fun i j => i ≤ j) (List.mergeSort (fun i j => i ≤ j) m) := by
exact List.sorted_mergeSort (fun i j => i ≤ j) m exact List.sorted_mergeSort (fun i j => i ≤ j) m
simp [List.Sorted] at hs simp only [List.Sorted] at hs
rw [List.pairwise_iff_get] at hs rw [List.pairwise_iff_get] at hs
exact hs i j h1 exact hs i j h1
have hp1 : (List.mergeSort (fun i j => i ≤ j) m).Perm m := by have hp1 : (List.mergeSort (fun i j => i ≤ j) m).Perm m := by

View file

@ -105,7 +105,8 @@ lemma countColorQuot_eq_countId_iff_of_isSome (hl : l.OnlyUniqueDuals) (i : Fin
rcases l.filter_id_of_countId_eq_two hi1 with hf | hf rcases l.filter_id_of_countId_eq_two hi1 with hf | hf
all_goals all_goals
erw [hf] erw [hf]
simp [List.countP, List.countP.go] simp only [List.countP, List.countP.go, List.get_eq_getElem, true_or, decide_True,
Bool.decide_or, zero_add, Nat.reduceAdd, cond_true, List.length_cons, List.length_singleton]
refine Iff.intro (fun h => ?_) (fun h => ?_) refine Iff.intro (fun h => ?_) (fun h => ?_)
· by_contra hn · by_contra hn
have hn' : (decide (l.val[↑i].toColor = l.val[↑((l.getDual? i).get hi)].toColor) || have hn' : (decide (l.val[↑i].toColor = l.val[↑((l.getDual? i).get hi)].toColor) ||

View file

@ -67,7 +67,7 @@ lemma trans (h1 : l.Subperm l2) (h2 : l2.Subperm l3) : l.Subperm l3 := by
lemma fst_eq_contrIndexList (h : l.Subperm l2) : l.contrIndexList = l := by lemma fst_eq_contrIndexList (h : l.Subperm l2) : l.contrIndexList = l := by
rw [iff_countId] at * rw [iff_countId] at *
apply ext apply ext
simp [contrIndexList] simp only [contrIndexList, List.filter_eq_self, decide_eq_true_eq]
intro I hI intro I hI
have h' := countId_mem l I hI have h' := countId_mem l I hI
have hI' := h I have hI' := h I

View file

@ -63,13 +63,17 @@ unsafe def processAllFiles : IO Unit := do
((IO.asTask $ IO.Process.run ((IO.asTask $ IO.Process.run
{cmd := "lake", args := #["exe", "free_simps", f.toString]}), f) {cmd := "lake", args := #["exe", "free_simps", f.toString]}), f)
tasks.toList.enum.forM fun (n, (t, path)) => do tasks.toList.enum.forM fun (n, (t, path)) => do
println! "{n} of {tasks.toList.length}: {path}"
let tn ← IO.wait (← t) let tn ← IO.wait (← t)
match tn with match tn with
| .ok x => println! x | .ok x =>
if x ≠ "" then
println! "{n} of {tasks.toList.length}: {path}"
println! x
| .error _ => println! "Error" | .error _ => println! "Error"
unsafe def main (args : List String) : IO Unit := do unsafe def main (args : List String) : IO Unit := do
match args with match args with
| [path] => transverseTactics path visitTacticInfo | [path] => transverseTactics path visitTacticInfo
| _ => processAllFiles | _ =>
processAllFiles
IO.println "Finished"