diff --git a/Archive/Imo/Imo1994Q1.lean b/Archive/Imo/Imo1994Q1.lean index e8d62532d713d..6ea9e04e8862f 100644 --- a/Archive/Imo/Imo1994Q1.lean +++ b/Archive/Imo/Imo1994Q1.lean @@ -77,7 +77,7 @@ theorem imo1994_q1 (n : ℕ) (m : ℕ) (A : Finset ℕ) (hm : A.card = m + 1) have hf : map f (Icc 0 k) ⊆ map a.toEmbedding (Ioc (rev k) (Fin.last m)) := by intro x hx simp only [Equiv.subLeft_apply, a, rev] at h - simp only [mem_map, mem_Icc, mem_Ioc, Fin.zero_le, true_and_iff, Equiv.subLeft_apply, + simp only [mem_map, mem_Icc, mem_Ioc, Fin.zero_le, true_and, Equiv.subLeft_apply, Function.Embedding.coeFn_mk, exists_prop, RelEmbedding.coe_toEmbedding, f, rev] at hx ⊢ rcases hx with ⟨i, ⟨hi, rfl⟩⟩ have h1 : a i + a (Fin.last m - k) ≤ n := by unfold_let; linarith only [h, a.monotone hi] diff --git a/Archive/Imo/Imo1998Q2.lean b/Archive/Imo/Imo1998Q2.lean index ac4e0dfb0f196..367aec35fd5fc 100644 --- a/Archive/Imo/Imo1998Q2.lean +++ b/Archive/Imo/Imo1998Q2.lean @@ -104,9 +104,9 @@ theorem A_fibre_over_contestant (c : C) : (Finset.univ.filter fun p : JudgePair J => p.Agree r c ∧ p.Distinct) = ((A r).filter fun a : AgreedTriple C J => a.contestant = c).image Prod.snd := by ext p - simp only [A, Finset.mem_univ, Finset.mem_filter, Finset.mem_image, true_and_iff, exists_prop] + simp only [A, Finset.mem_univ, Finset.mem_filter, Finset.mem_image, exists_prop] constructor - · rintro ⟨h₁, h₂⟩; refine ⟨(c, p), ?_⟩; tauto + · rintro ⟨_, h₂⟩; refine ⟨(c, p), ?_⟩; tauto · intro h; aesop theorem A_fibre_over_contestant_card (c : C) : diff --git a/Archive/MiuLanguage/DecisionSuf.lean b/Archive/MiuLanguage/DecisionSuf.lean index 161e777bc7f34..2f1f8bbf87750 100644 --- a/Archive/MiuLanguage/DecisionSuf.lean +++ b/Archive/MiuLanguage/DecisionSuf.lean @@ -265,7 +265,7 @@ theorem base_case_suf (en : Miustr) (h : Decstr en) (hu : count U en = 0) : Deri rcases h with ⟨⟨mhead, nmtail⟩, hi⟩ have : en ≠ nil := by intro k - simp only [k, count, countP, countP.go, if_false, zero_mod, zero_ne_one, false_or_iff, + simp only [k, count, countP, countP.go, if_false, zero_mod, zero_ne_one, false_or, reduceCtorEq] at hi rcases exists_cons_of_ne_nil this with ⟨y, ys, rfl⟩ rcases mhead diff --git a/Archive/Wiedijk100Theorems/AscendingDescendingSequences.lean b/Archive/Wiedijk100Theorems/AscendingDescendingSequences.lean index 19e2a5834b237..f4dbb1df5644d 100644 --- a/Archive/Wiedijk100Theorems/AscendingDescendingSequences.lean +++ b/Archive/Wiedijk100Theorems/AscendingDescendingSequences.lean @@ -147,7 +147,7 @@ theorem erdos_szekeres {r s n : ℕ} {f : Fin n → α} (hn : r * s < n) (hf : I have : image ab univ ⊆ ran := by -- First some logical shuffling rintro ⟨x₁, x₂⟩ - simp only [ran, mem_image, exists_prop, mem_range, mem_univ, mem_product, true_and_iff, + simp only [ran, mem_image, exists_prop, mem_range, mem_univ, mem_product, true_and, Prod.ext_iff] rintro ⟨i, rfl, rfl⟩ specialize q i diff --git a/Archive/Wiedijk100Theorems/BallotProblem.lean b/Archive/Wiedijk100Theorems/BallotProblem.lean index 79a0cf487315f..5d22c4ad829b2 100644 --- a/Archive/Wiedijk100Theorems/BallotProblem.lean +++ b/Archive/Wiedijk100Theorems/BallotProblem.lean @@ -206,7 +206,7 @@ theorem first_vote_pos : ((countedSequence_nonempty _ _).image _)] · have : List.cons (-1) '' countedSequence (p + 1) q ∩ {l : List ℤ | l.headI = 1} = ∅ := by ext - simp only [mem_inter_iff, mem_image, mem_setOf_eq, mem_empty_iff_false, iff_false_iff, + simp only [mem_inter_iff, mem_image, mem_setOf_eq, mem_empty_iff_false, iff_false, not_and, forall_exists_index, and_imp] rintro l _ rfl norm_num @@ -262,7 +262,7 @@ theorem countedSequence_int_pos_counted_succ_succ (p q : ℕ) : (_ : List.cons (-1) '' countedSequence (p + 1) q ∩ {l | l.headI = 1} = ∅), union_empty] <;> · ext simp only [mem_inter_iff, mem_image, mem_setOf_eq, and_iff_left_iff_imp, mem_empty_iff_false, - iff_false_iff, not_and, forall_exists_index, and_imp] + iff_false, not_and, forall_exists_index, and_imp] rintro y _ rfl norm_num @@ -289,7 +289,7 @@ theorem countedSequence_int_neg_counted_succ_succ (p q : ℕ) : empty_union] <;> · ext simp only [mem_inter_iff, mem_image, mem_setOf_eq, and_iff_left_iff_imp, mem_empty_iff_false, - iff_false_iff, not_and, forall_exists_index, and_imp] + iff_false, not_and, forall_exists_index, and_imp] rintro y _ rfl norm_num @@ -363,7 +363,7 @@ theorem ballot_problem : rwa [ENNReal.toReal_eq_toReal (measure_lt_top _ _).ne] at this simp only [Ne, ENNReal.div_eq_top, tsub_eq_zero_iff_le, Nat.cast_le, not_le, add_eq_zero, Nat.cast_eq_zero, ENNReal.add_eq_top, ENNReal.natCast_ne_top, or_self_iff, - not_false_iff, and_true_iff] + not_false_iff, and_true] push_neg exact ⟨fun _ _ => by linarith, (tsub_le_self.trans_lt (ENNReal.natCast_ne_top p).lt_top).ne⟩ diff --git a/Archive/Wiedijk100Theorems/CubingACube.lean b/Archive/Wiedijk100Theorems/CubingACube.lean index 05b1712152b67..d7acf6143b041 100644 --- a/Archive/Wiedijk100Theorems/CubingACube.lean +++ b/Archive/Wiedijk100Theorems/CubingACube.lean @@ -71,7 +71,7 @@ theorem univ_pi_side (c : Cube n) : pi univ (side c) = c.toSet := theorem toSet_subset {c c' : Cube n} : c.toSet ⊆ c'.toSet ↔ ∀ j, c.side j ⊆ c'.side j := by simp only [← univ_pi_side, univ_pi_subset_univ_pi_iff, (c.side_nonempty _).ne_empty, exists_false, - or_false_iff] + or_false] theorem toSet_disjoint {c c' : Cube n} : Disjoint c.toSet c'.toSet ↔ ∃ j, Disjoint (c.side j) (c'.side j) := by @@ -361,7 +361,7 @@ theorem smallest_onBoundary {j} (bi : OnBoundary (mi_mem_bcubes : mi h v ∈ _) dsimp only [x]; rw [← bi, add_sub_assoc, add_lt_iff_neg_left, sub_lt_zero] apply mi_strict_minimal (Ne.symm h2i') hi' refine ⟨x, ⟨?_, ?_⟩, ?_⟩ - · simp only [side, neg_lt_zero, hw, add_lt_iff_neg_left, and_true_iff, mem_Ico, sub_eq_add_neg, x] + · simp only [side, neg_lt_zero, hw, add_lt_iff_neg_left, and_true, mem_Ico, sub_eq_add_neg, x] rw [add_assoc, le_add_iff_nonneg_right, ← sub_eq_add_neg, sub_nonneg] apply le_of_lt (w_lt_w h v hi') · simp only [side, not_and_or, not_lt, not_le, mem_Ico]; left; exact hx @@ -408,7 +408,7 @@ theorem mi_not_onBoundary (j : Fin n) : ¬OnBoundary (mi_mem_bcubes : mi h v ∈ have i'_i'' : i' ≠ i'' := by rintro ⟨⟩ have : (cs i).b ∈ (cs i').toSet := by - simp only [toSet, forall_iff_succ, hi.1, bottom_mem_side h2i', true_and_iff, mem_setOf_eq] + simp only [toSet, forall_iff_succ, hi.1, bottom_mem_side h2i', true_and, mem_setOf_eq] intro j₂; by_cases hj₂ : j₂ = j · simpa [p', side_tail, hj'.symm, hj₂] using hi''.2 j · simpa [p, hj₂] using hi'.2 j₂ diff --git a/Archive/Wiedijk100Theorems/FriendshipGraphs.lean b/Archive/Wiedijk100Theorems/FriendshipGraphs.lean index 30ac6b66643fd..860afa94c8e2e 100644 --- a/Archive/Wiedijk100Theorems/FriendshipGraphs.lean +++ b/Archive/Wiedijk100Theorems/FriendshipGraphs.lean @@ -179,7 +179,7 @@ theorem card_of_regular (hd : G.IsRegularOfDegree d) : d + (Fintype.card V - 1) trans ((G.adjMatrix ℕ ^ 2) *ᵥ (fun _ => 1)) v · rw [adjMatrix_sq_of_regular hG hd, mulVec, dotProduct, ← insert_erase (mem_univ v)] simp only [sum_insert, mul_one, if_true, Nat.cast_id, eq_self_iff_true, mem_erase, not_true, - Ne, not_false_iff, add_right_inj, false_and_iff, of_apply] + Ne, not_false_iff, add_right_inj, false_and, of_apply] rw [Finset.sum_const_nat, card_erase_of_mem (mem_univ v), mul_one]; · rfl intro x hx; simp [(ne_of_mem_erase hx).symm] · rw [sq, ← mulVec_mulVec] diff --git a/Archive/Wiedijk100Theorems/Partition.lean b/Archive/Wiedijk100Theorems/Partition.lean index 5895ebcb7fdd8..0ced772d9f68f 100644 --- a/Archive/Wiedijk100Theorems/Partition.lean +++ b/Archive/Wiedijk100Theorems/Partition.lean @@ -206,7 +206,7 @@ theorem partialGF_prop (α : Type*) [CommSemiring α] (n : ℕ) (s : Finset ℕ) · dsimp only intro p₁ hp₁ p₂ hp₂ h apply Nat.Partition.ext - simp only [true_and_iff, mem_univ, mem_filter] at hp₁ hp₂ + simp only [true_and, mem_univ, mem_filter] at hp₁ hp₂ ext i simp only [φ, ne_eq, Multiset.mem_toFinset, not_not, smul_eq_mul, Finsupp.mk.injEq] at h by_cases hi : i = 0 @@ -218,7 +218,7 @@ theorem partialGF_prop (α : Type*) [CommSemiring α] (n : ℕ) (s : Finset ℕ) · rw [← mul_left_inj' hi] rw [Function.funext_iff] at h exact h.2 i - · simp only [φ, mem_filter, mem_finsuppAntidiag, mem_univ, exists_prop, true_and_iff, and_assoc] + · simp only [φ, mem_filter, mem_finsuppAntidiag, mem_univ, exists_prop, true_and, and_assoc] rintro f ⟨hf, hf₃, hf₄⟩ have hf' : f ∈ finsuppAntidiag s n := mem_finsuppAntidiag.mpr ⟨hf, hf₃⟩ simp only [mem_finsuppAntidiag] at hf' @@ -266,7 +266,7 @@ theorem partialOddGF_prop [Field α] (n m : ℕ) : convert partialGF_prop α n ((range m).map mkOdd) _ (fun _ => Set.univ) (fun _ _ => trivial) using 2 · congr - simp only [true_and_iff, forall_const, Set.mem_univ] + simp only [true_and, forall_const, Set.mem_univ] · rw [Finset.prod_map] simp_rw [num_series'] congr! 2 with x diff --git a/Counterexamples/Phillips.lean b/Counterexamples/Phillips.lean index 755cb6b94dd69..86aeacfb5f059 100644 --- a/Counterexamples/Phillips.lean +++ b/Counterexamples/Phillips.lean @@ -467,14 +467,14 @@ theorem sierpinski_pathological_family (Hcont : #ℝ = aleph 1) : refine ⟨fun x => {y | r x y}, fun x => ?_, fun y => ?_⟩ · have : univ \ {y | r x y} = {y | r y x} ∪ {x} := by ext y - simp only [true_and_iff, mem_univ, mem_setOf_eq, mem_insert_iff, union_singleton, mem_diff] + simp only [true_and, mem_univ, mem_setOf_eq, mem_insert_iff, union_singleton, mem_diff] rcases trichotomous_of r x y with (h | rfl | h) - · simp only [h, not_or, false_iff_iff, not_true] + · simp only [h, not_or, false_iff, not_true] constructor · rintro rfl; exact irrefl_of r y h · exact asymm h - · simp only [true_or_iff, eq_self_iff_true, iff_true_iff]; exact irrefl x - · simp only [h, iff_true_iff, or_true_iff]; exact asymm h + · simp only [true_or, eq_self_iff_true, iff_true]; exact irrefl x + · simp only [h, iff_true, or_true]; exact asymm h rw [this] apply Countable.union _ (countable_singleton _) rw [Cardinal.countable_iff_lt_aleph_one, ← Hcont] diff --git a/Mathlib/Algebra/BigOperators/Group/Finset.lean b/Mathlib/Algebra/BigOperators/Group/Finset.lean index 84b2a9a1ee80b..33e51269d348b 100644 --- a/Mathlib/Algebra/BigOperators/Group/Finset.lean +++ b/Mathlib/Algebra/BigOperators/Group/Finset.lean @@ -2074,7 +2074,7 @@ theorem disjoint_list_sum_left {a : Multiset α} {l : List (Multiset α)} : simp only [zero_disjoint, List.not_mem_nil, IsEmpty.forall_iff, forall_const, List.sum_nil] | cons b bs ih => simp_rw [List.sum_cons, disjoint_add_left, List.mem_cons, forall_eq_or_imp] - simp [and_congr_left_iff, iff_self_iff, ih] + simp [and_congr_left_iff, ih] theorem disjoint_list_sum_right {a : Multiset α} {l : List (Multiset α)} : Multiset.Disjoint a l.sum ↔ ∀ b ∈ l, Multiset.Disjoint a b := by @@ -2093,7 +2093,7 @@ theorem disjoint_sum_right {a : Multiset α} {i : Multiset (Multiset α)} : theorem disjoint_finset_sum_left {β : Type*} {i : Finset β} {f : β → Multiset α} {a : Multiset α} : Multiset.Disjoint (i.sum f) a ↔ ∀ b ∈ i, Multiset.Disjoint (f b) a := by convert @disjoint_sum_left _ a (map f i.val) - simp [and_congr_left_iff, iff_self_iff] + simp [and_congr_left_iff] theorem disjoint_finset_sum_right {β : Type*} {i : Finset β} {f : β → Multiset α} {a : Multiset α} : Multiset.Disjoint a (i.sum f) ↔ ∀ b ∈ i, Multiset.Disjoint a (f b) := by diff --git a/Mathlib/Algebra/CharZero/Lemmas.lean b/Mathlib/Algebra/CharZero/Lemmas.lean index 4a3176488e297..15bdc16d9f97e 100644 --- a/Mathlib/Algebra/CharZero/Lemmas.lean +++ b/Mathlib/Algebra/CharZero/Lemmas.lean @@ -79,7 +79,7 @@ variable {R : Type*} [NonAssocSemiring R] [NoZeroDivisors R] [CharZero R] {a : R @[simp] theorem add_self_eq_zero {a : R} : a + a = 0 ↔ a = 0 := by - simp only [(two_mul a).symm, mul_eq_zero, two_ne_zero, false_or_iff] + simp only [(two_mul a).symm, mul_eq_zero, two_ne_zero, false_or] end diff --git a/Mathlib/Algebra/GCDMonoid/Basic.lean b/Mathlib/Algebra/GCDMonoid/Basic.lean index eb193ba73400a..f578796cbf0ff 100644 --- a/Mathlib/Algebra/GCDMonoid/Basic.lean +++ b/Mathlib/Algebra/GCDMonoid/Basic.lean @@ -666,7 +666,7 @@ theorem lcm_dvd_iff [GCDMonoid α] {a b c : α} : lcm a b ∣ c ↔ a ∣ c ∧ by_cases h : a = 0 ∨ b = 0 · rcases h with (rfl | rfl) <;> simp (config := { contextual := true }) only [iff_def, lcm_zero_left, lcm_zero_right, - zero_dvd_iff, dvd_zero, eq_self_iff_true, and_true_iff, imp_true_iff] + zero_dvd_iff, dvd_zero, eq_self_iff_true, and_true, imp_true_iff] · obtain ⟨h1, h2⟩ := not_or.1 h have h : gcd a b ≠ 0 := fun H => h1 ((gcd_eq_zero_iff _ _).1 H).1 rw [← mul_dvd_mul_iff_left h, (gcd_mul_lcm a b).dvd_iff_dvd_left, ← @@ -1306,7 +1306,7 @@ instance (priority := 100) : NormalizedGCDMonoid G₀ where exact Associated.refl _ -- Porting note(#12129): additional beta reduction needed · beta_reduce - rw [if_neg (not_and_of_not_left _ ha), one_mul, if_neg (not_or_of_not ha hb)] + rw [if_neg (not_and_of_not_left _ ha), one_mul, if_neg (not_or_intro ha hb)] exact (associated_one_iff_isUnit.mpr ((IsUnit.mk0 _ ha).mul (IsUnit.mk0 _ hb))).symm lcm_zero_left b := if_pos (Or.inl rfl) lcm_zero_right a := if_pos (Or.inr rfl) diff --git a/Mathlib/Algebra/Group/Pointwise/Set.lean b/Mathlib/Algebra/Group/Pointwise/Set.lean index b2ed19084371d..a1ac52d71f441 100644 --- a/Mathlib/Algebra/Group/Pointwise/Set.lean +++ b/Mathlib/Algebra/Group/Pointwise/Set.lean @@ -1146,7 +1146,7 @@ theorem isUnit_singleton (a : α) : IsUnit ({a} : Set α) := @[to_additive (attr := simp)] theorem isUnit_iff_singleton : IsUnit s ↔ ∃ a, s = {a} := by - simp only [isUnit_iff, Group.isUnit, and_true_iff] + simp only [isUnit_iff, Group.isUnit, and_true] @[to_additive (attr := simp)] theorem image_mul_left : (a * ·) '' t = (a⁻¹ * ·) ⁻¹' t := by diff --git a/Mathlib/Algebra/Group/Subgroup/Basic.lean b/Mathlib/Algebra/Group/Subgroup/Basic.lean index aa8d59fc8d74c..ba423e5b8a907 100644 --- a/Mathlib/Algebra/Group/Subgroup/Basic.lean +++ b/Mathlib/Algebra/Group/Subgroup/Basic.lean @@ -1717,7 +1717,7 @@ This may be easier to work with, as it avoids inequalities and negations. -/ theorem _root_.normalizerCondition_iff_only_full_group_self_normalizing : NormalizerCondition G ↔ ∀ H : Subgroup G, H.normalizer = H → H = ⊤ := by apply forall_congr'; intro H - simp only [lt_iff_le_and_ne, le_normalizer, true_and_iff, le_top, Ne] + simp only [lt_iff_le_and_ne, le_normalizer, le_top, Ne] tauto variable (H) @@ -2219,10 +2219,10 @@ instance (priority := 100) normal_ker (f : G →* M) : f.ker.Normal := rw [mem_ker, map_mul, map_mul, f.mem_ker.1 hx, mul_one, map_mul_eq_one f (mul_inv_cancel y)]⟩ @[to_additive (attr := simp)] -lemma ker_fst : ker (fst G G') = .prod ⊥ ⊤ := SetLike.ext fun _ => (and_true_iff _).symm +lemma ker_fst : ker (fst G G') = .prod ⊥ ⊤ := SetLike.ext fun _ => (iff_of_eq (and_true _)).symm @[to_additive (attr := simp)] -lemma ker_snd : ker (snd G G') = .prod ⊤ ⊥ := SetLike.ext fun _ => (true_and_iff _).symm +lemma ker_snd : ker (snd G G') = .prod ⊤ ⊥ := SetLike.ext fun _ => (iff_of_eq (true_and _)).symm @[simp] theorem coe_toAdditive_ker (f : G →* G') : @@ -2384,7 +2384,7 @@ theorem map_le_map_iff {f : G →* N} {H K : Subgroup G} : H.map f ≤ K.map f @[to_additive] theorem map_le_map_iff' {f : G →* N} {H K : Subgroup G} : H.map f ≤ K.map f ↔ H ⊔ f.ker ≤ K ⊔ f.ker := by - simp only [map_le_map_iff, sup_le_iff, le_sup_right, and_true_iff] + simp only [map_le_map_iff, sup_le_iff, le_sup_right, and_true] @[to_additive] theorem map_eq_map_iff {f : G →* N} {H K : Subgroup G} : diff --git a/Mathlib/Algebra/Group/Subgroup/Pointwise.lean b/Mathlib/Algebra/Group/Subgroup/Pointwise.lean index 1f026608f9626..17dc2929357ae 100644 --- a/Mathlib/Algebra/Group/Subgroup/Pointwise.lean +++ b/Mathlib/Algebra/Group/Subgroup/Pointwise.lean @@ -72,7 +72,7 @@ theorem closure_toSubmonoid (S : Set G) : (fun x hx => Submonoid.closure_mono subset_union_left (Submonoid.subset_closure hx)) (Submonoid.one_mem _) (fun x y hx hy => Submonoid.mul_mem _ hx hy) fun x hx => ?_ rwa [← Submonoid.mem_closure_inv, Set.union_inv, inv_inv, Set.union_comm] - · simp only [true_and_iff, coe_toSubmonoid, union_subset_iff, subset_closure, inv_subset_closure] + · simp only [true_and, coe_toSubmonoid, union_subset_iff, subset_closure, inv_subset_closure] /-- For subgroups generated by a single element, see the simpler `zpow_induction_left`. -/ @[to_additive (attr := elab_as_elim) diff --git a/Mathlib/Algebra/Group/Submonoid/Operations.lean b/Mathlib/Algebra/Group/Submonoid/Operations.lean index 04c31e0b2686f..e410115baf4ec 100644 --- a/Mathlib/Algebra/Group/Submonoid/Operations.lean +++ b/Mathlib/Algebra/Group/Submonoid/Operations.lean @@ -914,10 +914,10 @@ theorem mker_inr : mker (inr M N) = ⊥ := by simp [mem_mker] @[to_additive (attr := simp)] -lemma mker_fst : mker (fst M N) = .prod ⊥ ⊤ := SetLike.ext fun _ => (and_true_iff _).symm +lemma mker_fst : mker (fst M N) = .prod ⊥ ⊤ := SetLike.ext fun _ => (iff_of_eq (and_true _)).symm @[to_additive (attr := simp)] -lemma mker_snd : mker (snd M N) = .prod ⊤ ⊥ := SetLike.ext fun _ => (true_and_iff _).symm +lemma mker_snd : mker (snd M N) = .prod ⊤ ⊥ := SetLike.ext fun _ => (iff_of_eq (true_and _)).symm /-- The `MonoidHom` from the preimage of a submonoid to itself. -/ @[to_additive (attr := simps) diff --git a/Mathlib/Algebra/GroupWithZero/Basic.lean b/Mathlib/Algebra/GroupWithZero/Basic.lean index 78525d46a0706..8b4b08535c105 100644 --- a/Mathlib/Algebra/GroupWithZero/Basic.lean +++ b/Mathlib/Algebra/GroupWithZero/Basic.lean @@ -411,8 +411,7 @@ lemma zpow_add' {m n : ℤ} (h : a ≠ 0 ∨ m + n ≠ 0 ∨ m = 0 ∧ n = 0) : · simp [hn] by_cases ha : a = 0 · subst a - simp only [false_or_iff, eq_self_iff_true, not_true, Ne, hm, hn, false_and_iff, - or_false_iff] at h + simp only [false_or, eq_self_iff_true, not_true, Ne, hm, hn, false_and, or_false] at h rw [zero_zpow _ h, zero_zpow _ hm, zero_mul] · exact zpow_add₀ ha m n diff --git a/Mathlib/Algebra/Lie/CartanSubalgebra.lean b/Mathlib/Algebra/Lie/CartanSubalgebra.lean index f5d08788b2163..c904d1921c59f 100644 --- a/Mathlib/Algebra/Lie/CartanSubalgebra.lean +++ b/Mathlib/Algebra/Lie/CartanSubalgebra.lean @@ -107,7 +107,7 @@ end LieSubalgebra theorem LieIdeal.normalizer_eq_top {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] (I : LieIdeal R L) : (I : LieSubalgebra R L).normalizer = ⊤ := by ext x - simpa only [LieSubalgebra.mem_normalizer_iff, LieSubalgebra.mem_top, iff_true_iff] using + simpa only [LieSubalgebra.mem_normalizer_iff, LieSubalgebra.mem_top, iff_true] using fun y hy => I.lie_mem hy open LieIdeal diff --git a/Mathlib/Algebra/Lie/Submodule.lean b/Mathlib/Algebra/Lie/Submodule.lean index 9a496ce1a3f1e..3c1f32a790e1d 100644 --- a/Mathlib/Algebra/Lie/Submodule.lean +++ b/Mathlib/Algebra/Lie/Submodule.lean @@ -550,7 +550,7 @@ nonrec theorem eq_bot_iff : N = ⊥ ↔ ∀ m : M, m ∈ N → m = 0 := by rw [e instance subsingleton_of_bot : Subsingleton (LieSubmodule R L ↑(⊥ : LieSubmodule R L M)) := by apply subsingleton_of_bot_eq_top ext ⟨x, hx⟩; change x ∈ ⊥ at hx; rw [Submodule.mem_bot] at hx; subst hx - simp only [true_iff_iff, eq_self_iff_true, Submodule.mk_eq_zero, LieSubmodule.mem_bot, mem_top] + simp only [eq_self_iff_true, Submodule.mk_eq_zero, LieSubmodule.mem_bot, mem_top] instance : IsModularLattice (LieSubmodule R L M) where sup_inf_le_assoc_of_le _ _ := by diff --git a/Mathlib/Algebra/Module/Basic.lean b/Mathlib/Algebra/Module/Basic.lean index 2e3fb20b33d35..c8f18e8483148 100644 --- a/Mathlib/Algebra/Module/Basic.lean +++ b/Mathlib/Algebra/Module/Basic.lean @@ -135,7 +135,7 @@ lemma support_smul_subset_right [Zero M] [SMulZeroClass R M] (f : α → R) (g : lemma support_const_smul_of_ne_zero [Zero R] [Zero M] [SMulWithZero R M] [NoZeroSMulDivisors R M] (c : R) (g : α → M) (hc : c ≠ 0) : support (c • g) = support g := - ext fun x ↦ by simp only [hc, mem_support, Pi.smul_apply, Ne, smul_eq_zero, false_or_iff] + ext fun x ↦ by simp only [hc, mem_support, Pi.smul_apply, Ne, smul_eq_zero, false_or] lemma support_smul [Zero R] [Zero M] [SMulWithZero R M] [NoZeroSMulDivisors R M] (f : α → R) (g : α → M) : support (f • g) = support f ∩ support g := diff --git a/Mathlib/Algebra/Module/Submodule/Map.lean b/Mathlib/Algebra/Module/Submodule/Map.lean index eaf5c011557e5..54e88c4e936d3 100644 --- a/Mathlib/Algebra/Module/Submodule/Map.lean +++ b/Mathlib/Algebra/Module/Submodule/Map.lean @@ -624,7 +624,7 @@ This is `LinearEquiv.ofSubmodule'` but with `map` on the right instead of `comap def submoduleMap (p : Submodule R M) : p ≃ₛₗ[σ₁₂] ↥(p.map (e : M →ₛₗ[σ₁₂] M₂) : Submodule R₂ M₂) := { ((e : M →ₛₗ[σ₁₂] M₂).domRestrict p).codRestrict (p.map (e : M →ₛₗ[σ₁₂] M₂)) fun x => ⟨x, by - simp only [LinearMap.domRestrict_apply, eq_self_iff_true, and_true_iff, SetLike.coe_mem, + simp only [LinearMap.domRestrict_apply, eq_self_iff_true, and_true, SetLike.coe_mem, SetLike.mem_coe]⟩ with invFun := fun y => ⟨(e.symm : M₂ →ₛₗ[σ₂₁] M) y, by diff --git a/Mathlib/Algebra/MvPolynomial/Division.lean b/Mathlib/Algebra/MvPolynomial/Division.lean index 73946aa78f7a6..7d505e900d7d2 100644 --- a/Mathlib/Algebra/MvPolynomial/Division.lean +++ b/Mathlib/Algebra/MvPolynomial/Division.lean @@ -206,7 +206,7 @@ theorem monomial_dvd_monomial {r s : R} {i j : σ →₀ ℕ} : theorem monomial_one_dvd_monomial_one [Nontrivial R] {i j : σ →₀ ℕ} : monomial i (1 : R) ∣ monomial j 1 ↔ i ≤ j := by rw [monomial_dvd_monomial] - simp_rw [one_ne_zero, false_or_iff, dvd_rfl, and_true_iff] + simp_rw [one_ne_zero, false_or, dvd_rfl, and_true] @[simp] theorem X_dvd_X [Nontrivial R] {i j : σ} : @@ -219,6 +219,6 @@ theorem X_dvd_X [Nontrivial R] {i j : σ} : theorem X_dvd_monomial {i : σ} {j : σ →₀ ℕ} {r : R} : (X i : MvPolynomial σ R) ∣ monomial j r ↔ r = 0 ∨ j i ≠ 0 := by refine monomial_dvd_monomial.trans ?_ - simp_rw [one_dvd, and_true_iff, Finsupp.single_le_iff, Nat.one_le_iff_ne_zero] + simp_rw [one_dvd, and_true, Finsupp.single_le_iff, Nat.one_le_iff_ne_zero] end MvPolynomial diff --git a/Mathlib/Algebra/MvPolynomial/Rename.lean b/Mathlib/Algebra/MvPolynomial/Rename.lean index eda39a1e51a4b..48bfa47e642c2 100644 --- a/Mathlib/Algebra/MvPolynomial/Rename.lean +++ b/Mathlib/Algebra/MvPolynomial/Rename.lean @@ -206,15 +206,14 @@ theorem exists_finset_rename (p : MvPolynomial σ R) : · rintro p q ⟨s, p, rfl⟩ ⟨t, q, rfl⟩ refine ⟨s ∪ t, ⟨?_, ?_⟩⟩ · refine rename (Subtype.map id ?_) p + rename (Subtype.map id ?_) q <;> - simp (config := { contextual := true }) only [id, true_or_iff, or_true_iff, + simp (config := { contextual := true }) only [id, true_or, or_true, Finset.mem_union, forall_true_iff] · simp only [rename_rename, map_add] rfl · rintro p n ⟨s, p, rfl⟩ refine ⟨insert n s, ⟨?_, ?_⟩⟩ · refine rename (Subtype.map id ?_) p * X ⟨n, s.mem_insert_self n⟩ - simp (config := { contextual := true }) only [id, or_true_iff, Finset.mem_insert, - forall_true_iff] + simp (config := { contextual := true }) only [id, or_true, Finset.mem_insert, forall_true_iff] · simp only [rename_rename, rename_X, Subtype.coe_mk, map_mul] rfl diff --git a/Mathlib/Algebra/MvPolynomial/Variables.lean b/Mathlib/Algebra/MvPolynomial/Variables.lean index 7f43b2f3d2866..730a6443dcfec 100644 --- a/Mathlib/Algebra/MvPolynomial/Variables.lean +++ b/Mathlib/Algebra/MvPolynomial/Variables.lean @@ -149,7 +149,7 @@ theorem vars_C_mul (a : A) (ha : a ≠ 0) (φ : MvPolynomial σ A) : apply exists_congr intro d apply and_congr _ Iff.rfl - rw [coeff_C_mul, mul_ne_zero_iff, eq_true ha, true_and_iff] + rw [coeff_C_mul, mul_ne_zero_iff, eq_true ha, true_and] end IsDomain diff --git a/Mathlib/Algebra/Order/BigOperators/Group/Finset.lean b/Mathlib/Algebra/Order/BigOperators/Group/Finset.lean index 93535400087e6..1d59c749630b2 100644 --- a/Mathlib/Algebra/Order/BigOperators/Group/Finset.lean +++ b/Mathlib/Algebra/Order/BigOperators/Group/Finset.lean @@ -545,7 +545,7 @@ theorem finset_sum_eq_sup_iff_disjoint [DecidableEq α] {β : Type*} {i : Finset · simp only [Finset.not_mem_empty, IsEmpty.forall_iff, imp_true_iff, Finset.sum_empty, Finset.sup_empty, bot_eq_zero, eq_self_iff_true] · simp_rw [Finset.sum_cons hz, Finset.sup_cons, Finset.mem_cons, Multiset.sup_eq_union, - forall_eq_or_imp, Ne, not_true_eq_false, IsEmpty.forall_iff, true_and_iff, + forall_eq_or_imp, Ne, not_true_eq_false, IsEmpty.forall_iff, true_and, imp_and, forall_and, ← hr, @eq_comm _ z] have := fun x (H : x ∈ i) => ne_of_mem_of_not_mem H hz simp (config := { contextual := true }) only [this, not_false_iff, true_imp_iff] diff --git a/Mathlib/Algebra/Order/CauSeq/Basic.lean b/Mathlib/Algebra/Order/CauSeq/Basic.lean index 461ef07e12d62..b390f7b8dceaf 100644 --- a/Mathlib/Algebra/Order/CauSeq/Basic.lean +++ b/Mathlib/Algebra/Order/CauSeq/Basic.lean @@ -678,7 +678,7 @@ instance : Preorder (CauSeq α abs) where | Or.inr fg, Or.inl gh => Or.inl <| lt_of_eq_of_lt fg gh | Or.inr fg, Or.inr gh => Or.inr <| Setoid.trans fg gh lt_iff_le_not_le _ _ := - ⟨fun h => ⟨Or.inl h, not_or_of_not (mt (lt_trans h) lt_irrefl) (not_limZero_of_pos h)⟩, + ⟨fun h => ⟨Or.inl h, not_or_intro (mt (lt_trans h) lt_irrefl) (not_limZero_of_pos h)⟩, fun ⟨h₁, h₂⟩ => h₁.resolve_right (mt (fun h => Or.inr (Setoid.symm h)) h₂)⟩ theorem le_antisymm {f g : CauSeq α abs} (fg : f ≤ g) (gf : g ≤ f) : f ≈ g := diff --git a/Mathlib/Algebra/Order/Field/Basic.lean b/Mathlib/Algebra/Order/Field/Basic.lean index 10cb8d1525691..2385c27aac032 100644 --- a/Mathlib/Algebra/Order/Field/Basic.lean +++ b/Mathlib/Algebra/Order/Field/Basic.lean @@ -174,7 +174,7 @@ theorem inv_lt_one_iff_of_pos (h₀ : 0 < a) : a⁻¹ < 1 ↔ 1 < a := theorem inv_lt_one_iff : a⁻¹ < 1 ↔ a ≤ 0 ∨ 1 < a := by rcases le_or_lt a 0 with ha | ha · simp [ha, (inv_nonpos.2 ha).trans_lt zero_lt_one] - · simp only [ha.not_le, false_or_iff, inv_lt_one_iff_of_pos ha] + · simp only [ha.not_le, false_or, inv_lt_one_iff_of_pos ha] theorem one_lt_inv_iff : 1 < a⁻¹ ↔ 0 < a ∧ a < 1 := ⟨fun h => ⟨inv_pos.1 (zero_lt_one.trans h), diff --git a/Mathlib/Algebra/Order/Floor.lean b/Mathlib/Algebra/Order/Floor.lean index 0431b57d759fd..d909f10092190 100644 --- a/Mathlib/Algebra/Order/Floor.lean +++ b/Mathlib/Algebra/Order/Floor.lean @@ -936,7 +936,7 @@ theorem fract_neg {x : α} (hx : fract x ≠ 0) : fract (-x) = 1 - fract x := by @[simp] theorem fract_neg_eq_zero {x : α} : fract (-x) = 0 ↔ fract x = 0 := by - simp only [fract_eq_iff, le_refl, zero_lt_one, tsub_zero, true_and_iff] + simp only [fract_eq_iff, le_refl, zero_lt_one, tsub_zero, true_and] constructor <;> rintro ⟨z, hz⟩ <;> use -z <;> simp [← hz] theorem fract_mul_nat (a : α) (b : ℕ) : ∃ z : ℤ, fract a * b - fract (a * b) = z := by diff --git a/Mathlib/Algebra/Order/Ring/Abs.lean b/Mathlib/Algebra/Order/Ring/Abs.lean index 7d11531e4a537..e61b8d4297a31 100644 --- a/Mathlib/Algebra/Order/Ring/Abs.lean +++ b/Mathlib/Algebra/Order/Ring/Abs.lean @@ -43,7 +43,7 @@ lemma abs_two : |(2 : α)| = 2 := abs_of_pos zero_lt_two lemma abs_mul (a b : α) : |a * b| = |a| * |b| := by rw [abs_eq (mul_nonneg (abs_nonneg a) (abs_nonneg b))] rcases le_total a 0 with ha | ha <;> rcases le_total b 0 with hb | hb <;> - simp only [abs_of_nonpos, abs_of_nonneg, true_or_iff, or_true_iff, eq_self_iff_true, neg_mul, + simp only [abs_of_nonpos, abs_of_nonneg, true_or, or_true, eq_self_iff_true, neg_mul, mul_neg, neg_neg, *] /-- `abs` as a `MonoidWithZeroHom`. -/ diff --git a/Mathlib/Algebra/Polynomial/AlgebraMap.lean b/Mathlib/Algebra/Polynomial/AlgebraMap.lean index 9f4745493a802..ae287ad87cf7f 100644 --- a/Mathlib/Algebra/Polynomial/AlgebraMap.lean +++ b/Mathlib/Algebra/Polynomial/AlgebraMap.lean @@ -102,7 +102,7 @@ instance subalgebraNontrivial [Nontrivial A] : Nontrivial (Subalgebra R A[X]) := ⟨⟨⊥, ⊤, by rw [Ne, SetLike.ext_iff, not_forall] refine ⟨X, ?_⟩ - simp only [Algebra.mem_bot, not_exists, Set.mem_range, iff_true_iff, Algebra.mem_top, + simp only [Algebra.mem_bot, not_exists, Set.mem_range, iff_true, Algebra.mem_top, algebraMap_apply, not_forall] intro x rw [ext_iff, not_forall] diff --git a/Mathlib/Algebra/Polynomial/BigOperators.lean b/Mathlib/Algebra/Polynomial/BigOperators.lean index ffa4e0a8e9d60..a94d1ed2ce7df 100644 --- a/Mathlib/Algebra/Polynomial/BigOperators.lean +++ b/Mathlib/Algebra/Polynomial/BigOperators.lean @@ -186,7 +186,7 @@ theorem natDegree_multiset_prod_of_monic (h : ∀ f ∈ t, Monic f) : rw [this] simp convert prod_replicate (Multiset.card t) (1 : R) - · simp only [eq_replicate, Multiset.card_map, eq_self_iff_true, true_and_iff] + · simp only [eq_replicate, Multiset.card_map, eq_self_iff_true, true_and] rintro i hi obtain ⟨i, hi, rfl⟩ := Multiset.mem_map.mp hi apply h diff --git a/Mathlib/Algebra/Polynomial/Coeff.lean b/Mathlib/Algebra/Polynomial/Coeff.lean index b2747d4680f9f..fe2f8eca5f560 100644 --- a/Mathlib/Algebra/Polynomial/Coeff.lean +++ b/Mathlib/Algebra/Polynomial/Coeff.lean @@ -220,7 +220,7 @@ theorem card_support_trinomial {k m n : ℕ} (hkm : k < m) (hmn : m < n) {x y z (hy : y ≠ 0) (hz : z ≠ 0) : card (support (C x * X ^ k + C y * X ^ m + C z * X ^ n)) = 3 := by rw [support_trinomial hkm hmn hx hy hz, card_insert_of_not_mem - (mt mem_insert.mp (not_or_of_not hkm.ne (mt mem_singleton.mp (hkm.trans hmn).ne))), + (mt mem_insert.mp (not_or_intro hkm.ne (mt mem_singleton.mp (hkm.trans hmn).ne))), card_insert_of_not_mem (mt mem_singleton.mp hmn.ne), card_singleton] end Fewnomials diff --git a/Mathlib/Algebra/Polynomial/Derivative.lean b/Mathlib/Algebra/Polynomial/Derivative.lean index e4eb94abf7e30..d29f194ce5d73 100644 --- a/Mathlib/Algebra/Polynomial/Derivative.lean +++ b/Mathlib/Algebra/Polynomial/Derivative.lean @@ -295,7 +295,7 @@ theorem mem_support_derivative [NoZeroSMulDivisors ℕ R] (p : R[X]) (n : ℕ) : suffices ¬p.coeff (n + 1) * (n + 1 : ℕ) = 0 ↔ coeff p (n + 1) ≠ 0 by simpa only [mem_support_iff, coeff_derivative, Ne, Nat.cast_succ] rw [← nsmul_eq_mul', smul_eq_zero] - simp only [Nat.succ_ne_zero, false_or_iff] + simp only [Nat.succ_ne_zero, false_or] @[simp] theorem degree_derivative_eq [NoZeroSMulDivisors ℕ R] (p : R[X]) (hp : 0 < natDegree p) : diff --git a/Mathlib/Algebra/Polynomial/Div.lean b/Mathlib/Algebra/Polynomial/Div.lean index 5ef9e03f9a99f..5dacd0350992a 100644 --- a/Mathlib/Algebra/Polynomial/Div.lean +++ b/Mathlib/Algebra/Polynomial/Div.lean @@ -281,7 +281,7 @@ theorem degree_divByMonic_le (p q : R[X]) : degree (p /ₘ q) ≤ degree p := exact WithBot.coe_le_coe.2 (Nat.le_add_left _ _) else by unfold divByMonic divModByMonicAux - simp [dif_pos hq, h, false_and_iff, if_false, degree_zero, bot_le] + simp [dif_pos hq, h, if_false, degree_zero, bot_le] else (divByMonic_eq_of_not_monic p hq).symm ▸ bot_le theorem degree_divByMonic_lt (p : R[X]) {q : R[X]} (hq : Monic q) (hp0 : p ≠ 0) diff --git a/Mathlib/Algebra/Polynomial/FieldDivision.lean b/Mathlib/Algebra/Polynomial/FieldDivision.lean index 0460864bfceb5..0df84dde52989 100644 --- a/Mathlib/Algebra/Polynomial/FieldDivision.lean +++ b/Mathlib/Algebra/Polynomial/FieldDivision.lean @@ -334,7 +334,7 @@ theorem mod_eq_self_iff (hq0 : q ≠ 0) : p % q = p ↔ degree p < degree q := rw [mod_def, modByMonic, dif_pos (monic_mul_leadingCoeff_inv hq0)] unfold divModByMonicAux dsimp - simp only [this, false_and_iff, if_false]⟩ + simp only [this, false_and, if_false]⟩ theorem div_eq_zero_iff (hq0 : q ≠ 0) : p / q = 0 ↔ degree p < degree q := ⟨fun h => by diff --git a/Mathlib/Algebra/Polynomial/HasseDeriv.lean b/Mathlib/Algebra/Polynomial/HasseDeriv.lean index 526b25659bb93..7f0f31d75a921 100644 --- a/Mathlib/Algebra/Polynomial/HasseDeriv.lean +++ b/Mathlib/Algebra/Polynomial/HasseDeriv.lean @@ -179,7 +179,7 @@ theorem natDegree_hasseDeriv_le (p : R[X]) (n : ℕ) : refine (natDegree_sum_le _ _).trans ?_ simp_rw [Function.comp, natDegree_monomial] rw [Finset.fold_ite, Finset.fold_const] - · simp only [ite_self, max_eq_right, zero_le', Finset.fold_max_le, true_and_iff, and_imp, + · simp only [ite_self, max_eq_right, zero_le', Finset.fold_max_le, true_and, and_imp, tsub_le_iff_right, mem_support_iff, Ne, Finset.mem_filter] intro x hx hx' have hxp : x ≤ p.natDegree := le_natDegree_of_ne_zero hx diff --git a/Mathlib/Algebra/Polynomial/UnitTrinomial.lean b/Mathlib/Algebra/Polynomial/UnitTrinomial.lean index f22387feb2113..c21059979dbcf 100644 --- a/Mathlib/Algebra/Polynomial/UnitTrinomial.lean +++ b/Mathlib/Algebra/Polynomial/UnitTrinomial.lean @@ -178,7 +178,7 @@ theorem isUnitTrinomial_iff' : refine ⟨?_, fun hp => ?_⟩ · rintro ⟨k, m, n, hkm, hmn, u, v, w, rfl⟩ rw [sum_def, trinomial_support hkm hmn u.ne_zero v.ne_zero w.ne_zero, - sum_insert (mt mem_insert.mp (not_or_of_not hkm.ne (mt mem_singleton.mp (hkm.trans hmn).ne))), + sum_insert (mt mem_insert.mp (not_or_intro hkm.ne (mt mem_singleton.mp (hkm.trans hmn).ne))), sum_insert (mt mem_singleton.mp hmn.ne), sum_singleton, trinomial_leading_coeff' hkm hmn, trinomial_middle_coeff hkm hmn, trinomial_trailing_coeff' hkm hmn] simp_rw [← Units.val_pow_eq_pow_val, Int.units_sq, Units.val_one] diff --git a/Mathlib/Algebra/Ring/Int.lean b/Mathlib/Algebra/Ring/Int.lean index 68823fef425a9..8848a57ffc4d7 100644 --- a/Mathlib/Algebra/Ring/Int.lean +++ b/Mathlib/Algebra/Ring/Int.lean @@ -118,9 +118,9 @@ lemma even_xor'_odd (n : ℤ) : Xor' (Even n) (Odd n) := by lemma even_xor'_odd' (n : ℤ) : ∃ k, Xor' (n = 2 * k) (n = 2 * k + 1) := by rcases even_or_odd n with (⟨k, rfl⟩ | ⟨k, rfl⟩) <;> use k - · simpa only [← two_mul, Xor', true_and_iff, eq_self_iff_true, not_true, or_false_iff, - and_false_iff] using (succ_ne_self (2 * k)).symm - · simp only [Xor', add_right_eq_self, false_or_iff, eq_self_iff_true, not_true, not_false_iff, + · simpa only [← two_mul, Xor', true_and, eq_self_iff_true, not_true, or_false, + and_false] using (succ_ne_self (2 * k)).symm + · simp only [Xor', add_right_eq_self, false_or, eq_self_iff_true, not_true, not_false_iff, one_ne_zero, and_self_iff] instance : DecidablePred (Odd : ℤ → Prop) := fun _ => decidable_of_iff _ not_even_iff_odd diff --git a/Mathlib/Algebra/Squarefree/Basic.lean b/Mathlib/Algebra/Squarefree/Basic.lean index 4ea689f18ad28..2ae3fd080cb05 100644 --- a/Mathlib/Algebra/Squarefree/Basic.lean +++ b/Mathlib/Algebra/Squarefree/Basic.lean @@ -153,7 +153,7 @@ theorem squarefree_iff_irreducible_sq_not_dvd_of_ne_zero {r : R} (hr : r ≠ 0) theorem squarefree_iff_irreducible_sq_not_dvd_of_exists_irreducible {r : R} (hr : ∃ x : R, Irreducible x) : Squarefree r ↔ ∀ x : R, Irreducible x → ¬x * x ∣ r := by rw [irreducible_sq_not_dvd_iff_eq_zero_and_no_irreducibles_or_squarefree, ← not_exists] - simp only [hr, not_true, false_or_iff, and_false_iff] + simp only [hr, not_true, false_or, and_false] end Irreducible diff --git a/Mathlib/Algebra/Star/NonUnitalSubalgebra.lean b/Mathlib/Algebra/Star/NonUnitalSubalgebra.lean index 4cb61d4429883..88ae6784d80a8 100644 --- a/Mathlib/Algebra/Star/NonUnitalSubalgebra.lean +++ b/Mathlib/Algebra/Star/NonUnitalSubalgebra.lean @@ -786,7 +786,7 @@ theorem toNonUnitalSubalgebra_bot : @[simp] theorem coe_bot : ((⊥ : NonUnitalStarSubalgebra R A) : Set A) = {0} := by simp only [Set.ext_iff, NonUnitalStarAlgebra.mem_bot, SetLike.mem_coe, Set.mem_singleton_iff, - iff_self_iff, forall_const] + forall_const] theorem eq_top_iff {S : NonUnitalStarSubalgebra R A} : S = ⊤ ↔ ∀ x : A, x ∈ S := ⟨fun h x => by rw [h]; exact mem_top, diff --git a/Mathlib/AlgebraicTopology/AlternatingFaceMapComplex.lean b/Mathlib/AlgebraicTopology/AlternatingFaceMapComplex.lean index 4a3ffcf186da8..de7b73ca2ffc3 100644 --- a/Mathlib/AlgebraicTopology/AlternatingFaceMapComplex.lean +++ b/Mathlib/AlgebraicTopology/AlternatingFaceMapComplex.lean @@ -80,7 +80,7 @@ theorem d_squared (n : ℕ) : objD X (n + 1) ≫ objD X n = 0 := by apply Finset.sum_bij φ · -- φ(S) is contained in Sᶜ intro ij hij - simp only [S, Finset.mem_univ, Finset.compl_filter, Finset.mem_filter, true_and_iff, + simp only [S, Finset.mem_univ, Finset.compl_filter, Finset.mem_filter, true_and, Fin.val_succ, Fin.coe_castLT] at hij ⊢ linarith · -- φ : S → Sᶜ is injective diff --git a/Mathlib/AlgebraicTopology/DoldKan/Degeneracies.lean b/Mathlib/AlgebraicTopology/DoldKan/Degeneracies.lean index ac8cbcf733f76..dde1b3dea55c9 100644 --- a/Mathlib/AlgebraicTopology/DoldKan/Degeneracies.lean +++ b/Mathlib/AlgebraicTopology/DoldKan/Degeneracies.lean @@ -95,7 +95,7 @@ theorem σ_comp_P_eq_zero (X : SimplicialObject C) {n q : ℕ} (i : Fin (n + 1)) comp_id, v.comp_Hσ_eq hi, assoc, SimplicialObject.δ_comp_σ_succ_assoc, Fin.eta, decomposition_Q n q, sum_comp, sum_comp, Finset.sum_eq_zero, add_zero, add_neg_eq_zero] intro j hj - simp only [true_and_iff, Finset.mem_univ, Finset.mem_filter] at hj + simp only [Finset.mem_univ, Finset.mem_filter] at hj obtain ⟨k, hk⟩ := Nat.le.dest (Nat.lt_succ_iff.mp (Fin.is_lt j)) rw [add_comm] at hk have hi' : i = Fin.castSucc ⟨i, by omega⟩ := by diff --git a/Mathlib/AlgebraicTopology/SimplexCategory.lean b/Mathlib/AlgebraicTopology/SimplexCategory.lean index 1f52cf7037c0b..d5bc35e4c0a4f 100644 --- a/Mathlib/AlgebraicTopology/SimplexCategory.lean +++ b/Mathlib/AlgebraicTopology/SimplexCategory.lean @@ -725,7 +725,7 @@ theorem eq_id_of_mono {x : SimplexCategory} (i : x ⟶ x) [Mono i] : i = 𝟙 _ apply isIso_of_bijective dsimp rw [Fintype.bijective_iff_injective_and_card i.toOrderHom, ← mono_iff_injective, - eq_self_iff_true, and_true_iff] + eq_self_iff_true, and_true] infer_instance theorem eq_id_of_epi {x : SimplexCategory} (i : x ⟶ x) [Epi i] : i = 𝟙 _ := by @@ -735,7 +735,7 @@ theorem eq_id_of_epi {x : SimplexCategory} (i : x ⟶ x) [Epi i] : i = 𝟙 _ := apply isIso_of_bijective dsimp rw [Fintype.bijective_iff_surjective_and_card i.toOrderHom, ← epi_iff_surjective, - eq_self_iff_true, and_true_iff] + eq_self_iff_true, and_true] infer_instance theorem eq_σ_of_epi {n : ℕ} (θ : mk (n + 1) ⟶ mk n) [Epi θ] : ∃ i : Fin (n + 1), θ = σ i := by diff --git a/Mathlib/Analysis/Analytic/Basic.lean b/Mathlib/Analysis/Analytic/Basic.lean index 9caab40add1e5..26093e42f5ff8 100644 --- a/Mathlib/Analysis/Analytic/Basic.lean +++ b/Mathlib/Analysis/Analytic/Basic.lean @@ -1347,7 +1347,7 @@ theorem hasFPowerSeriesAt_iff : simp only [Metric.eventually_nhds_iff] rintro ⟨r, r_pos, h⟩ refine ⟨p.radius ⊓ r.toNNReal, by simp, ?_, ?_⟩ - · simp only [r_pos.lt, lt_inf_iff, ENNReal.coe_pos, Real.toNNReal_pos, and_true_iff] + · simp only [r_pos.lt, lt_inf_iff, ENNReal.coe_pos, Real.toNNReal_pos, and_true] obtain ⟨z, z_pos, le_z⟩ := NormedField.exists_norm_lt 𝕜 r_pos.lt have : (‖z‖₊ : ENNReal) ≤ p.radius := by simp only [dist_zero_right] at h diff --git a/Mathlib/Analysis/Analytic/Composition.lean b/Mathlib/Analysis/Analytic/Composition.lean index 2267a001aa658..658ad80e3df11 100644 --- a/Mathlib/Analysis/Analytic/Composition.lean +++ b/Mathlib/Analysis/Analytic/Composition.lean @@ -527,8 +527,7 @@ def compPartialSumSource (m M N : ℕ) : Finset (Σ n, Fin n → ℕ) := theorem mem_compPartialSumSource_iff (m M N : ℕ) (i : Σ n, Fin n → ℕ) : i ∈ compPartialSumSource m M N ↔ (m ≤ i.1 ∧ i.1 < M) ∧ ∀ a : Fin i.1, 1 ≤ i.2 a ∧ i.2 a < N := by - simp only [compPartialSumSource, Finset.mem_Ico, Fintype.mem_piFinset, Finset.mem_sigma, - iff_self_iff] + simp only [compPartialSumSource, Finset.mem_Ico, Fintype.mem_piFinset, Finset.mem_sigma] /-- Change of variables appearing to compute the composition of partial sums of formal power series -/ @@ -569,7 +568,7 @@ theorem compPartialSumTargetSet_image_compPartialSumSource (m M N : ℕ) rcases i with ⟨n, c⟩ refine ⟨⟨c.length, c.blocksFun⟩, ?_, ?_⟩ · simp only [compPartialSumTargetSet, Set.mem_setOf_eq] at hi - simp only [mem_compPartialSumSource_iff, hi.left, hi.right, true_and_iff, and_true_iff] + simp only [mem_compPartialSumSource_iff, hi.left, hi.right, true_and, and_true] exact fun a => c.one_le_blocks' _ · dsimp [compChangeOfVariables] rw [Composition.sigma_eq_iff_blocks_eq] @@ -608,7 +607,7 @@ theorem compChangeOfVariables_sum {α : Type*} [AddCommMonoid α] (m M N : ℕ) -- Porting note: added simp only at H simp only [mem_compPartialSumTarget_iff, Composition.length, Composition.blocks, H.left, - map_ofFn, length_ofFn, true_and_iff, compChangeOfVariables] + map_ofFn, length_ofFn, true_and, compChangeOfVariables] intro j simp only [Composition.blocksFun, (H.right _).right, List.get_ofFn] -- 2 - show that the map is injective diff --git a/Mathlib/Analysis/Asymptotics/Asymptotics.lean b/Mathlib/Analysis/Asymptotics/Asymptotics.lean index 47518d245fd96..fda3788c247e8 100644 --- a/Mathlib/Analysis/Asymptotics/Asymptotics.lean +++ b/Mathlib/Analysis/Asymptotics/Asymptotics.lean @@ -593,7 +593,7 @@ theorem isLittleO_sup : f =o[l ⊔ l'] g ↔ f =o[l] g ∧ f =o[l'] g := theorem isBigOWith_insert [TopologicalSpace α] {x : α} {s : Set α} {C : ℝ} {g : α → E} {g' : α → F} (h : ‖g x‖ ≤ C * ‖g' x‖) : IsBigOWith C (𝓝[insert x s] x) g g' ↔ IsBigOWith C (𝓝[s] x) g g' := by - simp_rw [IsBigOWith_def, nhdsWithin_insert, eventually_sup, eventually_pure, h, true_and_iff] + simp_rw [IsBigOWith_def, nhdsWithin_insert, eventually_sup, eventually_pure, h, true_and] protected theorem IsBigOWith.insert [TopologicalSpace α] {x : α} {s : Set α} {C : ℝ} {g : α → E} {g' : α → F} (h1 : IsBigOWith C (𝓝[s] x) g g') (h2 : ‖g x‖ ≤ C * ‖g' x‖) : @@ -1059,7 +1059,7 @@ variable {g g' l} @[simp] theorem isBigOWith_zero_right_iff : (IsBigOWith c l f'' fun _x => (0 : F')) ↔ f'' =ᶠ[l] 0 := by - simp only [IsBigOWith_def, exists_prop, true_and_iff, norm_zero, mul_zero, + simp only [IsBigOWith_def, exists_prop, norm_zero, mul_zero, norm_le_zero_iff, EventuallyEq, Pi.zero_apply] @[simp] @@ -1622,8 +1622,8 @@ theorem isLittleO_const_left_of_ne {c : E''} (hc : c ≠ 0) : theorem isLittleO_const_left {c : E''} : (fun _x => c) =o[l] g'' ↔ c = 0 ∨ Tendsto (norm ∘ g'') l atTop := by rcases eq_or_ne c 0 with (rfl | hc) - · simp only [isLittleO_zero, eq_self_iff_true, true_or_iff] - · simp only [hc, false_or_iff, isLittleO_const_left_of_ne hc]; rfl + · simp only [isLittleO_zero, eq_self_iff_true, true_or] + · simp only [hc, false_or, isLittleO_const_left_of_ne hc]; rfl @[simp 1001] -- Porting note: increase priority so that this triggers before `isLittleO_const_left` theorem isLittleO_const_const_iff [NeBot l] {d : E''} {c : F''} : diff --git a/Mathlib/Analysis/Asymptotics/Theta.lean b/Mathlib/Analysis/Asymptotics/Theta.lean index 7b78a91770b61..bf570aafef4e4 100644 --- a/Mathlib/Analysis/Asymptotics/Theta.lean +++ b/Mathlib/Analysis/Asymptotics/Theta.lean @@ -238,7 +238,7 @@ theorem isTheta_const_const_iff [NeBot l] {c₁ : E''} {c₂ : F''} : @[simp] theorem isTheta_zero_left : (fun _ ↦ (0 : E')) =Θ[l] g'' ↔ g'' =ᶠ[l] 0 := by - simp only [IsTheta, isBigO_zero, isBigO_zero_right_iff, true_and_iff] + simp only [IsTheta, isBigO_zero, isBigO_zero_right_iff, true_and] @[simp] theorem isTheta_zero_right : (f'' =Θ[l] fun _ ↦ (0 : F')) ↔ f'' =ᶠ[l] 0 := diff --git a/Mathlib/Analysis/BoundedVariation.lean b/Mathlib/Analysis/BoundedVariation.lean index dd6942a44733b..4f281d9bfca6d 100644 --- a/Mathlib/Analysis/BoundedVariation.lean +++ b/Mathlib/Analysis/BoundedVariation.lean @@ -292,7 +292,7 @@ theorem add_point (f : α → E) {s : Set α} {x : α} (hx : x ∈ s) (u : ℕ apply Finset.sum_congr rfl fun i _hi => ?_ dsimp only [w] simp only [← Npos, Nat.not_lt_zero, Nat.add_succ_sub_one, add_zero, if_false, - add_eq_zero, Nat.one_ne_zero, false_and_iff, Nat.succ_add_sub_one, zero_add] + add_eq_zero, Nat.one_ne_zero, false_and, Nat.succ_add_sub_one, zero_add] rw [add_comm 1 i] _ = ∑ i ∈ Finset.Ico 1 (n + 1), edist (f (w (i + 1))) (f (w i)) := by rw [Finset.range_eq_Ico] @@ -317,7 +317,7 @@ theorem add_point (f : α → E) {s : Set α} {x : α} (hx : x ∈ s) (u : ℕ congr 1 · congr 1 · apply Finset.sum_congr rfl fun i hi => ?_ - simp only [Finset.mem_Ico, zero_le', true_and_iff] at hi + simp only [Finset.mem_Ico, zero_le', true_and] at hi dsimp only [w] have A : i + 1 < N := Nat.lt_pred_iff.1 hi have B : i < N := Nat.lt_of_succ_lt A diff --git a/Mathlib/Analysis/BoxIntegral/Box/SubboxInduction.lean b/Mathlib/Analysis/BoxIntegral/Box/SubboxInduction.lean index c516b55504b50..f70f3e2cb58e1 100644 --- a/Mathlib/Analysis/BoxIntegral/Box/SubboxInduction.lean +++ b/Mathlib/Analysis/BoxIntegral/Box/SubboxInduction.lean @@ -52,7 +52,7 @@ theorem mem_splitCenterBox {s : Set ι} {y : ι → ℝ} : simp only [splitCenterBox, mem_def, ← forall_and] refine forall_congr' fun i ↦ ?_ dsimp only [Set.piecewise] - split_ifs with hs <;> simp only [hs, iff_true_iff, iff_false_iff, not_lt] + split_ifs with hs <;> simp only [hs, iff_true, iff_false, not_lt] exacts [⟨fun H ↦ ⟨⟨(left_lt_add_div_two.2 (I.lower_lt_upper i)).trans H.1, H.2⟩, H.1⟩, fun H ↦ ⟨H.2, H.1.2⟩⟩, ⟨fun H ↦ ⟨⟨H.1, H.2.trans (add_div_two_lt_right.2 (I.lower_lt_upper i)).le⟩, H.2⟩, diff --git a/Mathlib/Analysis/BoxIntegral/Partition/Basic.lean b/Mathlib/Analysis/BoxIntegral/Partition/Basic.lean index a596f5d3c17d9..788c55499703d 100644 --- a/Mathlib/Analysis/BoxIntegral/Partition/Basic.lean +++ b/Mathlib/Analysis/BoxIntegral/Partition/Basic.lean @@ -623,7 +623,7 @@ def IsPartition (π : Prepartition I) := ∀ x ∈ I, ∃ J ∈ π, x ∈ J theorem isPartition_iff_iUnion_eq {π : Prepartition I} : π.IsPartition ↔ π.iUnion = I := by - simp_rw [IsPartition, Set.Subset.antisymm_iff, π.iUnion_subset, true_and_iff, Set.subset_def, + simp_rw [IsPartition, Set.Subset.antisymm_iff, π.iUnion_subset, true_and, Set.subset_def, mem_iUnion, Box.mem_coe] @[simp] diff --git a/Mathlib/Analysis/BoxIntegral/Partition/Split.lean b/Mathlib/Analysis/BoxIntegral/Partition/Split.lean index fa8553041f1e2..06363982a7538 100644 --- a/Mathlib/Analysis/BoxIntegral/Partition/Split.lean +++ b/Mathlib/Analysis/BoxIntegral/Partition/Split.lean @@ -149,7 +149,7 @@ def split (I : Box ι) (i : ι) (x : ℝ) : Prepartition I := rintro J (rfl | rfl) exacts [Box.splitLower_le, Box.splitUpper_le]) (by - simp only [Finset.coe_insert, Finset.coe_singleton, true_and_iff, Set.mem_singleton_iff, + simp only [Finset.coe_insert, Finset.coe_singleton, true_and, Set.mem_singleton_iff, pairwise_insert_of_symmetric symmetric_disjoint, pairwise_singleton] rintro J rfl - exact I.disjoint_splitLower_splitUpper i x) diff --git a/Mathlib/Analysis/Calculus/BumpFunction/FiniteDimension.lean b/Mathlib/Analysis/Calculus/BumpFunction/FiniteDimension.lean index 314fbbe3fa861..a246268057e33 100644 --- a/Mathlib/Analysis/Calculus/BumpFunction/FiniteDimension.lean +++ b/Mathlib/Analysis/Calculus/BumpFunction/FiniteDimension.lean @@ -407,7 +407,7 @@ theorem y_pos_of_mem_ball {D : ℝ} {x : E} (Dpos : 0 < D) (D_lt_one : D < 1) intro y hy simp only [support_mul, w_support E Dpos] simp only [φ, mem_inter_iff, mem_support, Ne, indicator_apply_eq_zero, - mem_closedBall_zero_iff, one_ne_zero, not_forall, not_false_iff, exists_prop, and_true_iff] + mem_closedBall_zero_iff, one_ne_zero, not_forall, not_false_iff, exists_prop, and_true] constructor · apply ball_subset_ball' _ hy simp only [hz, norm_smul, abs_of_nonneg Dpos.le, abs_of_nonneg B.le, dist_zero_right, @@ -512,7 +512,7 @@ instance (priority := 100) {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E · rintro ⟨R, x⟩ ⟨hR : 1 < R, _⟩ have A : 0 < (R - 1) / (R + 1) := by apply div_pos <;> linarith have B : (R - 1) / (R + 1) < 1 := by apply (div_lt_one _).2 <;> linarith - simp only [mem_preimage, prod_mk_mem_set_prod_eq, mem_Ioo, mem_univ, and_true_iff, A, B] + simp only [mem_preimage, prod_mk_mem_set_prod_eq, mem_Ioo, mem_univ, and_true, A, B] eq_one := fun R hR x hx => by have A : 0 < R + 1 := by linarith simp only [hR, if_true] diff --git a/Mathlib/Analysis/Calculus/ContDiff/Basic.lean b/Mathlib/Analysis/Calculus/ContDiff/Basic.lean index 4eebb545531aa..78cf41f061212 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/Basic.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/Basic.lean @@ -892,7 +892,7 @@ theorem ContDiffWithinAt.hasFDerivWithinAt_nhds {f : E → F → G} {g : E → F have hst : insert x₀ s ×ˢ t ∈ 𝓝[(fun x => (x, g x)) '' s] (x₀, g x₀) := by refine nhdsWithin_mono _ ?_ (nhdsWithin_prod self_mem_nhdsWithin hgt) simp_rw [image_subset_iff, mk_preimage_prod, preimage_id', subset_inter_iff, subset_insert, - true_and_iff, subset_preimage_image] + true_and, subset_preimage_image] obtain ⟨v, hv, hvs, f', hvf', hf'⟩ := contDiffWithinAt_succ_iff_hasFDerivWithinAt'.mp hf refine ⟨(fun z => (z, g z)) ⁻¹' v ∩ insert x₀ s, ?_, inter_subset_right, fun z => diff --git a/Mathlib/Analysis/Calculus/ContDiff/Defs.lean b/Mathlib/Analysis/Calculus/ContDiff/Defs.lean index 501132040acea..e58a5dda3a781 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/Defs.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/Defs.lean @@ -323,7 +323,7 @@ theorem HasFTaylorSeriesUpToOn.contDiffOn {f' : E → FormalMultilinearSeries (hf : HasFTaylorSeriesUpToOn n f f' s) : ContDiffOn 𝕜 n f s := by intro x hx m hm use s - simp only [Set.insert_eq_of_mem hx, self_mem_nhdsWithin, true_and_iff] + simp only [Set.insert_eq_of_mem hx, self_mem_nhdsWithin, true_and] exact ⟨f', hf.of_le hm⟩ theorem ContDiffOn.contDiffWithinAt (h : ContDiffOn 𝕜 n f s) (hx : x ∈ s) : diff --git a/Mathlib/Analysis/Calculus/FDeriv/Measurable.lean b/Mathlib/Analysis/Calculus/FDeriv/Measurable.lean index 5623962121150..3066b3ab3c333 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Measurable.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Measurable.lean @@ -661,7 +661,7 @@ theorem D_subset_differentiable_set {K : Set F} (hK : IsComplete K) : · simp only [one_div, inv_pow, left_mem_Icc, le_add_iff_nonneg_right] positivity · simp only [pow_add, tsub_le_iff_left] at h'k - simpa only [hy.1, mem_Icc, true_and_iff, one_div, pow_one] using h'k + simpa only [hy.1, mem_Icc, true_and, one_div, pow_one] using h'k _ = 4 * (1 / 2) ^ e * (1 / 2) ^ (m + 2) := by field_simp; ring _ ≤ 4 * (1 / 2) ^ e * (y - x) := by gcongr _ = 4 * (1 / 2) ^ e * ‖y - x‖ := by rw [Real.norm_of_nonneg yzero.le] diff --git a/Mathlib/Analysis/Calculus/InverseFunctionTheorem/ApproximatesLinearOn.lean b/Mathlib/Analysis/Calculus/InverseFunctionTheorem/ApproximatesLinearOn.lean index 56145cc2bf428..ea6f8405c48bf 100644 --- a/Mathlib/Analysis/Calculus/InverseFunctionTheorem/ApproximatesLinearOn.lean +++ b/Mathlib/Analysis/Calculus/InverseFunctionTheorem/ApproximatesLinearOn.lean @@ -257,7 +257,7 @@ theorem surjOn_closedBall_of_nonlinearRightInverse -- It remains to check that `f x = y`. This follows from continuity of `f` on `closedBall b ε` -- and from the fact that `f uₙ` is converging to `y` by construction. have hx' : Tendsto u atTop (𝓝[closedBall b ε] x) := by - simp only [nhdsWithin, tendsto_inf, hx, true_and_iff, tendsto_principal] + simp only [nhdsWithin, tendsto_inf, hx, true_and, tendsto_principal] exact Eventually.of_forall fun n => C n _ (D n).2 have T1 : Tendsto (f ∘ u) atTop (𝓝 (f x)) := (hf.continuousOn.mono hε x xmem).tendsto.comp hx' diff --git a/Mathlib/Analysis/Complex/Arg.lean b/Mathlib/Analysis/Complex/Arg.lean index ec78377f02673..13b4096755966 100644 --- a/Mathlib/Analysis/Complex/Arg.lean +++ b/Mathlib/Analysis/Complex/Arg.lean @@ -31,7 +31,7 @@ theorem sameRay_iff : SameRay ℝ x y ↔ x = 0 ∨ y = 0 ∨ x.arg = y.arg := b · simp rcases eq_or_ne y 0 with (rfl | hy) · simp - simp only [hx, hy, false_or_iff, sameRay_iff_norm_smul_eq, arg_eq_arg_iff hx hy] + simp only [hx, hy, sameRay_iff_norm_smul_eq, arg_eq_arg_iff hx hy] field_simp [hx, hy] rw [mul_comm, eq_comm] diff --git a/Mathlib/Analysis/Complex/UpperHalfPlane/Basic.lean b/Mathlib/Analysis/Complex/UpperHalfPlane/Basic.lean index 919cb65d7a8e2..f9c084804d861 100644 --- a/Mathlib/Analysis/Complex/UpperHalfPlane/Basic.lean +++ b/Mathlib/Analysis/Complex/UpperHalfPlane/Basic.lean @@ -178,7 +178,7 @@ theorem linear_ne_zero (cd : Fin 2 → ℝ) (z : ℍ) (h : cd ≠ 0) : (cd 0 : have : cd 0 = 0 := by -- we will need this twice apply_fun Complex.im at h - simpa only [z.im_ne_zero, Complex.add_im, add_zero, coe_im, zero_mul, or_false_iff, + simpa only [z.im_ne_zero, Complex.add_im, add_zero, coe_im, zero_mul, or_false, Complex.ofReal_im, Complex.zero_im, Complex.mul_im, mul_eq_zero] using h simp only [this, zero_mul, Complex.ofReal_zero, zero_add, Complex.ofReal_eq_zero] at h diff --git a/Mathlib/Analysis/Complex/UpperHalfPlane/Metric.lean b/Mathlib/Analysis/Complex/UpperHalfPlane/Metric.lean index dec00ac0e5457..f4c7f5c6d05f3 100644 --- a/Mathlib/Analysis/Complex/UpperHalfPlane/Metric.lean +++ b/Mathlib/Analysis/Complex/UpperHalfPlane/Metric.lean @@ -327,7 +327,7 @@ instance : IsometricSMul SL(2, ℝ) ℍ := have h₂ : Complex.abs (y₁ * y₂) ≠ 0 := by simp [y₁.ne_zero, y₂.ne_zero] simp only [dist_eq, modular_S_smul, inv_neg, neg_div, div_mul_div_comm, coe_mk, mk_im, div_one, Complex.inv_im, Complex.neg_im, coe_im, neg_neg, Complex.normSq_neg, - mul_eq_mul_left_iff, Real.arsinh_inj, one_ne_zero, or_false_iff, + mul_eq_mul_left_iff, Real.arsinh_inj, one_ne_zero, dist_neg_neg, mul_neg, neg_mul, dist_inv_inv₀ y₁.ne_zero y₂.ne_zero, ← AbsoluteValue.map_mul, ← Complex.normSq_mul, Real.sqrt_div h₁, ← Complex.abs_apply, mul_div (2 : ℝ), div_div_div_comm, div_self h₂, Complex.norm_eq_abs] diff --git a/Mathlib/Analysis/Convex/Caratheodory.lean b/Mathlib/Analysis/Convex/Caratheodory.lean index 75f0b5f68dfeb..4d14dfff94413 100644 --- a/Mathlib/Analysis/Convex/Caratheodory.lean +++ b/Mathlib/Analysis/Convex/Caratheodory.lean @@ -85,7 +85,7 @@ theorem mem_convexHull_erase [DecidableEq E] {t : Finset E} (h : ¬AffineIndepen _ ≤ 0 := by apply mul_nonpos_of_nonneg_of_nonpos · apply div_nonneg (fpos i₀ (mem_of_subset (filter_subset _ t) mem)) (le_of_lt hg) - · simpa only [s, mem_filter, het, true_and_iff, not_lt] using hes + · simpa only [s, mem_filter, het, true_and, not_lt] using hes _ ≤ f e := fpos e het · rw [Subtype.coe_mk, centerMass_eq_of_sum_1 _ id ksum] calc diff --git a/Mathlib/Analysis/Convex/Combination.lean b/Mathlib/Analysis/Convex/Combination.lean index b851e544f9998..1fc98bcd1e819 100644 --- a/Mathlib/Analysis/Convex/Combination.lean +++ b/Mathlib/Analysis/Convex/Combination.lean @@ -124,7 +124,7 @@ theorem Finset.centerMass_subset {t' : Finset ι} (ht : t ⊆ t') (h : ∀ i ∈ theorem Finset.centerMass_filter_ne_zero : (t.filter fun i => w i ≠ 0).centerMass w z = t.centerMass w z := Finset.centerMass_subset z (filter_subset _ _) fun i hit hit' => by - simpa only [hit, mem_filter, true_and_iff, Ne, Classical.not_not] using hit' + simpa only [hit, mem_filter, true_and, Ne, Classical.not_not] using hit' namespace Finset diff --git a/Mathlib/Analysis/Convex/Segment.lean b/Mathlib/Analysis/Convex/Segment.lean index 87adc0a70ac2e..9f5026dbeb8cc 100644 --- a/Mathlib/Analysis/Convex/Segment.lean +++ b/Mathlib/Analysis/Convex/Segment.lean @@ -122,7 +122,7 @@ theorem segment_same (x : E) : [x -[𝕜] x] = {x} := theorem insert_endpoints_openSegment (x y : E) : insert x (insert y (openSegment 𝕜 x y)) = [x -[𝕜] y] := by simp only [subset_antisymm_iff, insert_subset_iff, left_mem_segment, right_mem_segment, - openSegment_subset_segment, true_and_iff] + openSegment_subset_segment, true_and] rintro z ⟨a, b, ha, hb, hab, rfl⟩ refine hb.eq_or_gt.imp ?_ fun hb' => ha.eq_or_gt.imp ?_ fun ha' => ?_ · rintro rfl @@ -140,7 +140,7 @@ theorem mem_openSegment_of_ne_left_right (hx : x ≠ z) (hy : y ≠ z) (hz : z theorem openSegment_subset_iff_segment_subset (hx : x ∈ s) (hy : y ∈ s) : openSegment 𝕜 x y ⊆ s ↔ [x -[𝕜] y] ⊆ s := by - simp only [← insert_endpoints_openSegment, insert_subset_iff, *, true_and_iff] + simp only [← insert_endpoints_openSegment, insert_subset_iff, *, true_and] end Module diff --git a/Mathlib/Analysis/Convex/Side.lean b/Mathlib/Analysis/Convex/Side.lean index 13933585a56e0..4bedc05dc363e 100644 --- a/Mathlib/Analysis/Convex/Side.lean +++ b/Mathlib/Analysis/Convex/Side.lean @@ -776,7 +776,7 @@ theorem isPreconnected_setOf_sSameSide (s : AffineSubspace ℝ P) (x : P) : simp only [h, not_sSameSide_bot] exact isPreconnected_empty · by_cases hx : x ∈ s - · simp only [hx, SSameSide, not_true, false_and_iff, and_false_iff] + · simp only [hx, SSameSide, not_true, false_and, and_false] exact isPreconnected_empty · exact (isConnected_setOf_sSameSide hx h).isPreconnected @@ -817,7 +817,7 @@ theorem isPreconnected_setOf_sOppSide (s : AffineSubspace ℝ P) (x : P) : simp only [h, not_sOppSide_bot] exact isPreconnected_empty · by_cases hx : x ∈ s - · simp only [hx, SOppSide, not_true, false_and_iff, and_false_iff] + · simp only [hx, SOppSide, not_true, false_and, and_false] exact isPreconnected_empty · exact (isConnected_setOf_sOppSide hx h).isPreconnected diff --git a/Mathlib/Analysis/Convex/StrictConvexBetween.lean b/Mathlib/Analysis/Convex/StrictConvexBetween.lean index c3bb34bd7fb30..08a6eed3ee484 100644 --- a/Mathlib/Analysis/Convex/StrictConvexBetween.lean +++ b/Mathlib/Analysis/Convex/StrictConvexBetween.lean @@ -58,13 +58,13 @@ theorem Collinear.wbtw_of_dist_eq_of_dist_le {p p₁ p₂ p₃ : P} {r : ℝ} · simp [hp₃p₂] have hs : Sbtw ℝ p₂ p₃ p₁ := ⟨hw, hp₃p₂, hp₁p₃.symm⟩ have hs' := hs.dist_lt_max_dist p - rw [hp₁, hp₃, lt_max_iff, lt_self_iff_false, or_false_iff] at hs' + rw [hp₁, hp₃, lt_max_iff, lt_self_iff_false, or_false] at hs' exact False.elim (hp₂.not_lt hs') · by_cases hp₁p₂ : p₁ = p₂ · simp [hp₁p₂] have hs : Sbtw ℝ p₃ p₁ p₂ := ⟨hw, hp₁p₃, hp₁p₂⟩ have hs' := hs.dist_lt_max_dist p - rw [hp₁, hp₃, lt_max_iff, lt_self_iff_false, false_or_iff] at hs' + rw [hp₁, hp₃, lt_max_iff, lt_self_iff_false, false_or] at hs' exact False.elim (hp₂.not_lt hs') /-- Given three collinear points, two (not equal) with distance `r` from `p` and one with diff --git a/Mathlib/Analysis/Convolution.lean b/Mathlib/Analysis/Convolution.lean index 4b2168a084049..55806cb2f24fa 100644 --- a/Mathlib/Analysis/Convolution.lean +++ b/Mathlib/Analysis/Convolution.lean @@ -1031,10 +1031,10 @@ theorem hasFDerivAt_convolution_right_with_param {g : P → G → E'} {s : Set P let g' := fderiv 𝕜 ↿g have A : ∀ p ∈ s, Continuous (g p) := fun p hp ↦ by refine hg.continuousOn.comp_continuous (continuous_const.prod_mk continuous_id') fun x => ?_ - simpa only [prod_mk_mem_set_prod_eq, mem_univ, and_true_iff] using hp + simpa only [prod_mk_mem_set_prod_eq, mem_univ, and_true] using hp have A' : ∀ q : P × G, q.1 ∈ s → s ×ˢ univ ∈ 𝓝 q := fun q hq ↦ by apply (hs.prod isOpen_univ).mem_nhds - simpa only [mem_prod, mem_univ, and_true_iff] using hq + simpa only [mem_prod, mem_univ, and_true] using hq -- The derivative of `g` vanishes away from `k`. have g'_zero : ∀ p x, p ∈ s → x ∉ k → g' (p, x) = 0 := by intro p x hp hx @@ -1055,7 +1055,7 @@ theorem hasFDerivAt_convolution_right_with_param {g : P → G → E'} {s : Set P hg.continuousOn_fderiv_of_isOpen (hs.prod isOpen_univ) le_rfl apply exists_isOpen_isBounded_image_of_isCompact_of_continuousOn A (hs.prod isOpen_univ) _ B simp only [prod_subset_prod_iff, hq₀, singleton_subset_iff, subset_univ, and_self_iff, - true_or_iff] + true_or] obtain ⟨ε, εpos, hε, h'ε⟩ : ∃ ε : ℝ, 0 < ε ∧ thickening ε ({q₀.fst} ×ˢ k) ⊆ t ∧ ball q₀.1 ε ⊆ s := by obtain ⟨ε, εpos, hε⟩ : ∃ ε : ℝ, 0 < ε ∧ thickening ε (({q₀.fst} : Set P) ×ˢ k) ⊆ t := @@ -1071,10 +1071,10 @@ theorem hasFDerivAt_convolution_right_with_param {g : P → G → E'} {s : Set P · have H : (p, x) ∈ t := by apply hε refine mem_thickening_iff.2 ⟨(q₀.1, x), ?_, ?_⟩ - · simp only [hx, singleton_prod, mem_image, Prod.mk.inj_iff, eq_self_iff_true, true_and_iff, + · simp only [hx, singleton_prod, mem_image, Prod.mk.inj_iff, eq_self_iff_true, true_and, exists_eq_right] · rw [← dist_eq_norm] at hp - simpa only [Prod.dist_eq, εpos, dist_self, max_lt_iff, and_true_iff] using hp + simpa only [Prod.dist_eq, εpos, dist_self, max_lt_iff, and_true] using hp have : g' (p, x) ∈ closedBall (0 : P × G →L[𝕜] E') C := hC (mem_image_of_mem _ H) rwa [mem_closedBall_zero_iff] at this · have : g' (p, x) = 0 := g'_zero _ _ hps hx @@ -1102,7 +1102,7 @@ theorem hasFDerivAt_convolution_right_with_param {g : P → G → E'} {s : Set P hg.continuousOn_fderiv_of_isOpen (hs.prod isOpen_univ) le_rfl apply this.comp_continuous (continuous_const.prod_mk continuous_id') intro x - simpa only [prod_mk_mem_set_prod_eq, mem_univ, and_true_iff] using hq₀ + simpa only [prod_mk_mem_set_prod_eq, mem_univ, and_true] using hq₀ set K' := (-k + {q₀.2} : Set G) with K'_def have hK' : IsCompact K' := hk.neg.add isCompact_singleton obtain ⟨U, U_open, K'U, hU⟩ : ∃ U, IsOpen U ∧ K' ⊆ U ∧ IntegrableOn f U μ := @@ -1251,12 +1251,12 @@ theorem contDiffOn_convolution_right_with_param {f : G → E} {n : ℕ∞} (L : apply hg.comp (isoP.prod isoG).contDiff.contDiffOn rintro ⟨p, x⟩ ⟨hp, -⟩ simpa only [mem_preimage, ContinuousLinearEquiv.prod_apply, prod_mk_mem_set_prod_eq, mem_univ, - and_true_iff] using hp + and_true] using hp have A : ContDiffOn 𝕜 n (isoF ∘ R ∘ (isoP.prod isoG).symm) (s ×ˢ univ) := by apply isoF.contDiff.comp_contDiffOn apply R_contdiff.comp (ContinuousLinearEquiv.contDiff _).contDiffOn rintro ⟨p, x⟩ ⟨hp, -⟩ - simpa only [mem_preimage, mem_prod, mem_univ, and_true_iff, ContinuousLinearEquiv.prod_symm, + simpa only [mem_preimage, mem_prod, mem_univ, and_true, ContinuousLinearEquiv.prod_symm, ContinuousLinearEquiv.prod_apply, ContinuousLinearEquiv.apply_symm_apply] using hp have : isoF ∘ R ∘ (isoP.prod isoG).symm = fun q : P × G => (f ⋆[L, μ] g q.1) q.2 := by apply funext diff --git a/Mathlib/Analysis/Distribution/SchwartzSpace.lean b/Mathlib/Analysis/Distribution/SchwartzSpace.lean index 73c009395f7d1..310970a55e4ef 100644 --- a/Mathlib/Analysis/Distribution/SchwartzSpace.lean +++ b/Mathlib/Analysis/Distribution/SchwartzSpace.lean @@ -533,7 +533,7 @@ theorem _root_.Function.HasTemperateGrowth.norm_iteratedFDeriv_le_uniform_aux {f choose k C f using hf_temperate.2 use (Finset.range (n + 1)).sup k let C' := max (0 : ℝ) ((Finset.range (n + 1)).sup' (by simp) C) - have hC' : 0 ≤ C' := by simp only [C', le_refl, Finset.le_sup'_iff, true_or_iff, le_max_iff] + have hC' : 0 ≤ C' := by simp only [C', le_refl, Finset.le_sup'_iff, true_or, le_max_iff] use C', hC' intro N hN x rw [← Finset.mem_range_succ_iff] at hN diff --git a/Mathlib/Analysis/InnerProductSpace/Basic.lean b/Mathlib/Analysis/InnerProductSpace/Basic.lean index 5ee958f6175b3..2e4adbdbee16f 100644 --- a/Mathlib/Analysis/InnerProductSpace/Basic.lean +++ b/Mathlib/Analysis/InnerProductSpace/Basic.lean @@ -447,7 +447,7 @@ theorem inner_self_eq_zero {x : F} : ⟪x, x⟫ = 0 ↔ x = 0 := theorem normSq_eq_zero {x : F} : normSqF x = 0 ↔ x = 0 := Iff.trans - (by simp only [normSq, ext_iff, map_zero, inner_self_im, eq_self_iff_true, and_true_iff]) + (by simp only [normSq, ext_iff, map_zero, inner_self_im, eq_self_iff_true, and_true]) (@inner_self_eq_zero 𝕜 _ _ _ _ _ x) theorem inner_self_ne_zero {x : F} : ⟪x, x⟫ ≠ 0 ↔ x ≠ 0 := diff --git a/Mathlib/Analysis/InnerProductSpace/LinearPMap.lean b/Mathlib/Analysis/InnerProductSpace/LinearPMap.lean index 9019b62655b0c..4fa2bc6e39428 100644 --- a/Mathlib/Analysis/InnerProductSpace/LinearPMap.lean +++ b/Mathlib/Analysis/InnerProductSpace/LinearPMap.lean @@ -202,7 +202,7 @@ as taking the `ContinuousLinearMap.adjoint` interpreted as a `LinearPMap`. -/ theorem toPMap_adjoint_eq_adjoint_toPMap_of_dense (hp : Dense (p : Set E)) : (A.toPMap p).adjoint = A.adjoint.toPMap ⊤ := by ext x y hxy - · simp only [LinearMap.toPMap_domain, Submodule.mem_top, iff_true_iff, + · simp only [LinearMap.toPMap_domain, Submodule.mem_top, iff_true, LinearPMap.mem_adjoint_domain_iff, LinearMap.coe_comp, innerₛₗ_apply_coe] exact ((innerSL 𝕜 x).comp <| A.comp <| Submodule.subtypeL _).cont refine LinearPMap.adjoint_apply_eq ?_ _ fun v => ?_ diff --git a/Mathlib/Analysis/InnerProductSpace/OfNorm.lean b/Mathlib/Analysis/InnerProductSpace/OfNorm.lean index ad51b9d339160..2616c07526d0b 100644 --- a/Mathlib/Analysis/InnerProductSpace/OfNorm.lean +++ b/Mathlib/Analysis/InnerProductSpace/OfNorm.lean @@ -247,13 +247,13 @@ private theorem int_prop (n : ℤ) : innerProp' E (n : 𝕜) := by simp only [Int.cast_natCast, map_natCast, map_intCast, Int.cast_mul, map_mul, mul_smul] obtain hn | rfl | hn := lt_trichotomy n 0 · rw [Int.sign_eq_neg_one_of_neg hn, innerProp_neg_one ((n.natAbs : 𝕜) • x), nat] - simp only [map_neg, neg_mul, one_mul, mul_eq_mul_left_iff, true_or_iff, Int.natAbs_eq_zero, + simp only [map_neg, neg_mul, one_mul, mul_eq_mul_left_iff, Int.natAbs_eq_zero, eq_self_iff_true, Int.cast_one, map_one, neg_inj, Nat.cast_eq_zero, Int.cast_neg] · simp only [inner_, Int.cast_zero, zero_sub, Nat.cast_zero, zero_mul, eq_self_iff_true, Int.sign_zero, zero_smul, zero_add, mul_zero, smul_zero, sub_self, norm_neg, Int.natAbs_zero] · rw [Int.sign_eq_one_of_pos hn] - simp only [one_mul, mul_eq_mul_left_iff, true_or_iff, Int.natAbs_eq_zero, eq_self_iff_true, + simp only [one_mul, mul_eq_mul_left_iff, Int.natAbs_eq_zero, eq_self_iff_true, Int.cast_one, one_smul, Nat.cast_eq_zero, nat] private theorem rat_prop (r : ℚ) : innerProp' E (r : 𝕜) := by diff --git a/Mathlib/Analysis/InnerProductSpace/PiL2.lean b/Mathlib/Analysis/InnerProductSpace/PiL2.lean index faf543a3bcbef..12ea8605df263 100644 --- a/Mathlib/Analysis/InnerProductSpace/PiL2.lean +++ b/Mathlib/Analysis/InnerProductSpace/PiL2.lean @@ -916,7 +916,7 @@ noncomputable def LinearIsometry.extend (L : S →ₗᵢ[𝕜] V) : V →ₗᵢ[ simp only [sq, Mx_decomp] rw [norm_add_sq_eq_norm_sq_add_norm_sq_of_inner_eq_zero (L (p1 x)) (L3 (p2 x)) Mx_orth] simp only [p1, p2, LinearIsometry.norm_map, _root_.add_left_inj, mul_eq_mul_left_iff, - norm_eq_zero, true_or_iff, eq_self_iff_true, ContinuousLinearMap.coe_coe, Submodule.coe_norm, + norm_eq_zero, eq_self_iff_true, ContinuousLinearMap.coe_coe, Submodule.coe_norm, Submodule.coe_eq_zero] exact { toLinearMap := M diff --git a/Mathlib/Analysis/InnerProductSpace/Projection.lean b/Mathlib/Analysis/InnerProductSpace/Projection.lean index 0cc0c1494c616..dff8bfa3ecd3d 100644 --- a/Mathlib/Analysis/InnerProductSpace/Projection.lean +++ b/Mathlib/Analysis/InnerProductSpace/Projection.lean @@ -1188,7 +1188,7 @@ theorem OrthogonalFamily.isInternal_iff_of_isComplete [DecidableEq ι] {V : ι (hc : IsComplete (↑(iSup V) : Set E)) : DirectSum.IsInternal V ↔ (iSup V)ᗮ = ⊥ := by haveI : CompleteSpace (↥(iSup V)) := hc.completeSpace_coe simp only [DirectSum.isInternal_submodule_iff_independent_and_iSup_eq_top, hV.independent, - true_and_iff, Submodule.orthogonal_eq_bot_iff] + true_and, Submodule.orthogonal_eq_bot_iff] /-- An orthogonal family of subspaces of `E` satisfies `DirectSum.IsInternal` (that is, they provide an internal direct sum decomposition of `E`) if and only if their span has trivial diff --git a/Mathlib/Analysis/InnerProductSpace/TwoDim.lean b/Mathlib/Analysis/InnerProductSpace/TwoDim.lean index 5f43d060381e7..1f3cbc87585d7 100644 --- a/Mathlib/Analysis/InnerProductSpace/TwoDim.lean +++ b/Mathlib/Analysis/InnerProductSpace/TwoDim.lean @@ -422,8 +422,7 @@ theorem nonneg_inner_and_areaForm_eq_zero_iff_sameRay (x y : E) : · intro h obtain ⟨r, hr, rfl⟩ := h.exists_nonneg_left hx simp only [inner_smul_right, real_inner_self_eq_norm_sq, LinearMap.map_smulₛₗ, - areaForm_apply_self, Algebra.id.smul_eq_mul, mul_zero, eq_self_iff_true, - and_true_iff] + areaForm_apply_self, Algebra.id.smul_eq_mul, mul_zero, eq_self_iff_true, and_true] positivity /-- A complex-valued real-bilinear map on an oriented real inner product space of dimension 2. Its diff --git a/Mathlib/Analysis/LocallyConvex/ContinuousOfBounded.lean b/Mathlib/Analysis/LocallyConvex/ContinuousOfBounded.lean index a58db852ed74d..a61d1ac3cdcf5 100644 --- a/Mathlib/Analysis/LocallyConvex/ContinuousOfBounded.lean +++ b/Mathlib/Analysis/LocallyConvex/ContinuousOfBounded.lean @@ -99,7 +99,7 @@ theorem LinearMap.continuousAt_zero_of_locally_bounded (f : E →ₛₗ[σ] F) refine bE.1.to_hasBasis ?_ ?_ · intro n _ use n + 1 - simp only [Ne, Nat.succ_ne_zero, not_false_iff, Nat.cast_add, Nat.cast_one, true_and_iff] + simp only [Ne, Nat.succ_ne_zero, not_false_iff, Nat.cast_add, Nat.cast_one, true_and] -- `b (n + 1) ⊆ b n` follows from `Antitone`. have h : b (n + 1) ⊆ b n := bE.2 (by simp) refine _root_.trans ?_ h diff --git a/Mathlib/Analysis/LocallyConvex/WeakDual.lean b/Mathlib/Analysis/LocallyConvex/WeakDual.lean index 455453490b735..80adb1f3e3bbd 100644 --- a/Mathlib/Analysis/LocallyConvex/WeakDual.lean +++ b/Mathlib/Analysis/LocallyConvex/WeakDual.lean @@ -112,7 +112,7 @@ theorem LinearMap.hasBasis_weakBilin (B : E →ₗ[𝕜] F →ₗ[𝕜] 𝕜) : refine lt_of_le_of_lt (hp x) (lt_of_lt_of_le hx ?_) exact Finset.inf'_le _ hyU' rw [Set.not_nonempty_iff_eq_empty.mp hU₃] - simp only [Set.empty_pi, Set.preimage_univ, Set.subset_univ, and_true_iff] + simp only [Set.empty_pi, Set.preimage_univ, Set.subset_univ, and_true] exact Exists.intro ((p 0).ball 0 1) (p.basisSets_singleton_mem 0 one_pos) rintro U (hU : U ∈ p.basisSets) rw [SeminormFamily.basisSets_iff] at hU diff --git a/Mathlib/Analysis/MeanInequalities.lean b/Mathlib/Analysis/MeanInequalities.lean index e871352abb1cc..2426536d2bcae 100644 --- a/Mathlib/Analysis/MeanInequalities.lean +++ b/Mathlib/Analysis/MeanInequalities.lean @@ -520,7 +520,7 @@ theorem isGreatest_Lp (f : ι → ℝ≥0) {p q : ℝ} (hpq : p.IsConjExponent q simp [h, hpq.ne_zero] simp only [Set.mem_setOf_eq, div_rpow, ← sum_div, ← rpow_mul, div_mul_cancel₀ _ hpq.symm.ne_zero, rpow_one, div_le_iff₀ hf, one_mul, hpq.mul_eq_add, ← - rpow_sub' A, add_sub_cancel_right, le_refl, true_and_iff, ← mul_div_assoc, B] + rpow_sub' A, add_sub_cancel_right, le_refl, true_and, ← mul_div_assoc, B] rw [div_eq_iff, ← rpow_add hf.ne', one_div, one_div, hpq.inv_add_inv_conj, rpow_one] simpa [hpq.symm.ne_zero] using hf.ne' · rintro _ ⟨g, hg, rfl⟩ diff --git a/Mathlib/Analysis/MeanInequalitiesPow.lean b/Mathlib/Analysis/MeanInequalitiesPow.lean index 2c7f60fd4576c..daa489d39347c 100644 --- a/Mathlib/Analysis/MeanInequalitiesPow.lean +++ b/Mathlib/Analysis/MeanInequalitiesPow.lean @@ -221,7 +221,7 @@ theorem rpow_arith_mean_le_arith_mean_rpow (w z : ι → ℝ≥0∞) (hw' : ∑ · -- first, prove `(∑ i ∈ s, w i * z i) ^ p = ⊤ → ∑ i ∈ s, (w i * z i ^ p) = ⊤` rw [rpow_eq_top_iff, sum_eq_top, sum_eq_top] intro h - simp only [and_false_iff, hp_not_neg, false_or_iff] at h + simp only [and_false, hp_not_neg, false_or] at h rcases h.left with ⟨a, H, ha⟩ use a, H rwa [← h_top_iff_rpow_top a H] diff --git a/Mathlib/Analysis/Normed/Algebra/Spectrum.lean b/Mathlib/Analysis/Normed/Algebra/Spectrum.lean index e6e984888bc72..6bdbec8adca03 100644 --- a/Mathlib/Analysis/Normed/Algebra/Spectrum.lean +++ b/Mathlib/Analysis/Normed/Algebra/Spectrum.lean @@ -280,7 +280,7 @@ theorem hasFPowerSeriesOnBall_inverse_one_sub_smul [CompleteSpace A] (a : A) : le_radius_of_bound_nnreal _ (max 1 ‖(1 : A)‖₊) fun n => ?_ rw [← norm_toNNReal, norm_mkPiRing, norm_toNNReal] cases' n with n - · simp only [le_refl, mul_one, or_true_iff, le_max_iff, pow_zero] + · simp only [le_refl, mul_one, or_true, le_max_iff, pow_zero] · refine le_trans (le_trans (mul_le_mul_right' (nnnorm_pow_le' a n.succ_pos) (r ^ n.succ)) ?_) (le_max_left _ _) diff --git a/Mathlib/Analysis/Normed/Module/Ray.lean b/Mathlib/Analysis/Normed/Module/Ray.lean index 93b860a1c1924..f01da59497153 100644 --- a/Mathlib/Analysis/Normed/Module/Ray.lean +++ b/Mathlib/Analysis/Normed/Module/Ray.lean @@ -81,7 +81,7 @@ the unit vectors `‖x‖⁻¹ • x` and `‖y‖⁻¹ • y` are equal. -/ theorem sameRay_iff_inv_norm_smul_eq : SameRay ℝ x y ↔ x = 0 ∨ y = 0 ∨ ‖x‖⁻¹ • x = ‖y‖⁻¹ • y := by rcases eq_or_ne x 0 with (rfl | hx); · simp [SameRay.zero_left] rcases eq_or_ne y 0 with (rfl | hy); · simp [SameRay.zero_right] - simp only [sameRay_iff_inv_norm_smul_eq_of_ne hx hy, *, false_or_iff] + simp only [sameRay_iff_inv_norm_smul_eq_of_ne hx hy, *, false_or] /-- Two vectors of the same norm are on the same ray if and only if they are equal. -/ theorem sameRay_iff_of_norm_eq (h : ‖x‖ = ‖y‖) : SameRay ℝ x y ↔ x = y := by diff --git a/Mathlib/Analysis/PSeries.lean b/Mathlib/Analysis/PSeries.lean index 6b1fde6df52ee..29941c56c0472 100644 --- a/Mathlib/Analysis/PSeries.lean +++ b/Mathlib/Analysis/PSeries.lean @@ -401,7 +401,7 @@ theorem sum_Ioo_inv_sq_le (k n : ℕ) : (∑ i ∈ Ioo k n, (i ^ 2 : α)⁻¹) apply sum_le_sum_of_subset_of_nonneg · intro x hx simp only [mem_Ioo] at hx - simp only [hx, hx.2.le, mem_Ioc, le_max_iff, or_true_iff, and_self_iff] + simp only [hx, hx.2.le, mem_Ioc, le_max_iff, or_true, and_self_iff] · intro i _hi _hident positivity _ ≤ ((k + 1 : α) ^ 2)⁻¹ + ∑ i ∈ Ioc k.succ (max (k + 1) n), ((i : α) ^ 2)⁻¹ := by diff --git a/Mathlib/Analysis/Seminorm.lean b/Mathlib/Analysis/Seminorm.lean index 46659a1b853cf..6ca315beb04f3 100644 --- a/Mathlib/Analysis/Seminorm.lean +++ b/Mathlib/Analysis/Seminorm.lean @@ -815,14 +815,14 @@ theorem closedBall_finset_sup (p : ι → Seminorm 𝕜 E) (s : Finset ι) (x : @[simp] theorem ball_eq_emptyset (p : Seminorm 𝕜 E) {x : E} {r : ℝ} (hr : r ≤ 0) : p.ball x r = ∅ := by ext - rw [Seminorm.mem_ball, Set.mem_empty_iff_false, iff_false_iff, not_lt] + rw [Seminorm.mem_ball, Set.mem_empty_iff_false, iff_false, not_lt] exact hr.trans (apply_nonneg p _) @[simp] theorem closedBall_eq_emptyset (p : Seminorm 𝕜 E) {x : E} {r : ℝ} (hr : r < 0) : p.closedBall x r = ∅ := by ext - rw [Seminorm.mem_closedBall, Set.mem_empty_iff_false, iff_false_iff, not_le] + rw [Seminorm.mem_closedBall, Set.mem_empty_iff_false, iff_false, not_le] exact hr.trans_le (apply_nonneg _ _) theorem closedBall_smul_ball (p : Seminorm 𝕜 E) {r₁ : ℝ} (hr₁ : r₁ ≠ 0) (r₂ : ℝ) : diff --git a/Mathlib/Analysis/SpecialFunctions/Complex/Arg.lean b/Mathlib/Analysis/SpecialFunctions/Complex/Arg.lean index 94901e3540e4f..f34699a4c7f97 100644 --- a/Mathlib/Analysis/SpecialFunctions/Complex/Arg.lean +++ b/Mathlib/Analysis/SpecialFunctions/Complex/Arg.lean @@ -172,7 +172,7 @@ theorem arg_mul_real {r : ℝ} (hr : 0 < r) (x : ℂ) : arg (x * r) = arg x := theorem arg_eq_arg_iff {x y : ℂ} (hx : x ≠ 0) (hy : y ≠ 0) : arg x = arg y ↔ (abs y / abs x : ℂ) * x = y := by simp only [ext_abs_arg_iff, map_mul, map_div₀, abs_ofReal, abs_abs, - div_mul_cancel₀ _ (abs.ne_zero hx), eq_self_iff_true, true_and_iff] + div_mul_cancel₀ _ (abs.ne_zero hx), eq_self_iff_true, true_and] rw [← ofReal_div, arg_real_mul] exact div_pos (abs.pos hy) (abs.pos hx) @@ -322,28 +322,28 @@ lemma image_exp_Ioc_eq_sphere : (fun θ : ℝ ↦ exp (θ * I)) '' Set.Ioc (-π) theorem arg_le_pi_div_two_iff {z : ℂ} : arg z ≤ π / 2 ↔ 0 ≤ re z ∨ im z < 0 := by rcases le_or_lt 0 (re z) with hre | hre - · simp only [hre, arg_of_re_nonneg hre, Real.arcsin_le_pi_div_two, true_or_iff] - simp only [hre.not_le, false_or_iff] + · simp only [hre, arg_of_re_nonneg hre, Real.arcsin_le_pi_div_two, true_or] + simp only [hre.not_le, false_or] rcases le_or_lt 0 (im z) with him | him · simp only [him.not_lt] - rw [iff_false_iff, not_le, arg_of_re_neg_of_im_nonneg hre him, ← sub_lt_iff_lt_add, half_sub, + rw [iff_false, not_le, arg_of_re_neg_of_im_nonneg hre him, ← sub_lt_iff_lt_add, half_sub, Real.neg_pi_div_two_lt_arcsin, neg_im, neg_div, neg_lt_neg_iff, div_lt_one, ← _root_.abs_of_nonneg him, abs_im_lt_abs] exacts [hre.ne, abs.pos <| ne_of_apply_ne re hre.ne] · simp only [him] - rw [iff_true_iff, arg_of_re_neg_of_im_neg hre him] + rw [iff_true, arg_of_re_neg_of_im_neg hre him] exact (sub_le_self _ Real.pi_pos.le).trans (Real.arcsin_le_pi_div_two _) theorem neg_pi_div_two_le_arg_iff {z : ℂ} : -(π / 2) ≤ arg z ↔ 0 ≤ re z ∨ 0 ≤ im z := by rcases le_or_lt 0 (re z) with hre | hre - · simp only [hre, arg_of_re_nonneg hre, Real.neg_pi_div_two_le_arcsin, true_or_iff] - simp only [hre.not_le, false_or_iff] + · simp only [hre, arg_of_re_nonneg hre, Real.neg_pi_div_two_le_arcsin, true_or] + simp only [hre.not_le, false_or] rcases le_or_lt 0 (im z) with him | him · simp only [him] - rw [iff_true_iff, arg_of_re_neg_of_im_nonneg hre him] + rw [iff_true, arg_of_re_neg_of_im_nonneg hre him] exact (Real.neg_pi_div_two_le_arcsin _).trans (le_add_of_nonneg_right Real.pi_pos.le) · simp only [him.not_le] - rw [iff_false_iff, not_le, arg_of_re_neg_of_im_neg hre him, sub_lt_iff_lt_add', ← + rw [iff_false, not_le, arg_of_re_neg_of_im_neg hre him, sub_lt_iff_lt_add', ← sub_eq_add_neg, sub_half, Real.arcsin_lt_pi_div_two, div_lt_one, neg_im, ← abs_of_neg him, abs_im_lt_abs] exacts [hre.ne, abs.pos <| ne_of_apply_ne re hre.ne] @@ -367,7 +367,7 @@ lemma arg_lt_pi_div_two_iff {z : ℂ} : arg z < π / 2 ↔ 0 < re z ∨ im z < 0 @[simp] theorem abs_arg_le_pi_div_two_iff {z : ℂ} : |arg z| ≤ π / 2 ↔ 0 ≤ re z := by rw [abs_le, arg_le_pi_div_two_iff, neg_pi_div_two_le_arg_iff, ← or_and_left, ← not_le, - and_not_self_iff, or_false_iff] + and_not_self_iff, or_false] @[simp] theorem abs_arg_lt_pi_div_two_iff {z : ℂ} : |arg z| < π / 2 ↔ 0 < re z ∨ z = 0 := by diff --git a/Mathlib/Analysis/SpecialFunctions/Gamma/Beta.lean b/Mathlib/Analysis/SpecialFunctions/Gamma/Beta.lean index 5b9cac12afe16..8f72e2ec290e0 100644 --- a/Mathlib/Analysis/SpecialFunctions/Gamma/Beta.lean +++ b/Mathlib/Analysis/SpecialFunctions/Gamma/Beta.lean @@ -183,11 +183,11 @@ theorem betaIntegral_recurrence {u v : ℂ} (hu : 0 < re u) (hv : 0 < re v) : have int_ev := intervalIntegral.integral_eq_sub_of_hasDerivAt_of_le zero_le_one hc hder h_int have hF0 : F 0 = 0 := by simp only [F, mul_eq_zero, ofReal_zero, cpow_eq_zero_iff, eq_self_iff_true, Ne, - true_and_iff, sub_zero, one_cpow, one_ne_zero, or_false_iff] + true_and, sub_zero, one_cpow, one_ne_zero, or_false] contrapose! hu; rw [hu, zero_re] have hF1 : F 1 = 0 := by simp only [F, mul_eq_zero, ofReal_one, one_cpow, one_ne_zero, sub_self, cpow_eq_zero_iff, - eq_self_iff_true, Ne, true_and_iff, false_or_iff] + eq_self_iff_true, Ne, true_and, false_or] contrapose! hv; rw [hv, zero_re] rw [hF0, hF1, sub_zero, intervalIntegral.integral_sub, intervalIntegral.integral_const_mul, intervalIntegral.integral_const_mul] at int_ev @@ -407,7 +407,7 @@ theorem Gamma_mul_Gamma_one_sub (z : ℂ) : Gamma z * Gamma (1 - z) = π / sin ( rw [hs, div_zero] rw [← neg_eq_zero, ← Complex.sin_neg, ← mul_neg, Complex.sin_eq_zero_iff, mul_comm] at hs obtain ⟨k, hk⟩ := hs - rw [mul_eq_mul_right_iff, eq_false (ofReal_ne_zero.mpr pi_pos.ne'), or_false_iff, + rw [mul_eq_mul_right_iff, eq_false (ofReal_ne_zero.mpr pi_pos.ne'), or_false, neg_eq_iff_eq_neg] at hk rw [hk] cases k diff --git a/Mathlib/Analysis/SpecialFunctions/Integrals.lean b/Mathlib/Analysis/SpecialFunctions/Integrals.lean index ca08532bc9736..506735693822c 100644 --- a/Mathlib/Analysis/SpecialFunctions/Integrals.lean +++ b/Mathlib/Analysis/SpecialFunctions/Integrals.lean @@ -117,7 +117,7 @@ theorem intervalIntegrable_cpow {r : ℂ} (h : 0 ≤ r.re ∨ (0 : ℝ) ∉ [[a, · -- Easy case #1: 0 ∉ [a, b] -- use continuity. refine (ContinuousAt.continuousOn fun x hx => ?_).intervalIntegrable exact Complex.continuousAt_ofReal_cpow_const _ _ (Or.inr <| ne_of_mem_of_not_mem hx h2) - rw [eq_false h2, or_false_iff] at h + rw [eq_false h2, or_false] at h rcases lt_or_eq_of_le h with (h' | h') · -- Easy case #2: 0 < re r -- again use continuity exact (Complex.continuous_ofReal_cpow_const h').intervalIntegrable _ _ diff --git a/Mathlib/Analysis/SpecialFunctions/PolarCoord.lean b/Mathlib/Analysis/SpecialFunctions/PolarCoord.lean index e11dde25a0bad..62dafde2c1433 100644 --- a/Mathlib/Analysis/SpecialFunctions/PolarCoord.lean +++ b/Mathlib/Analysis/SpecialFunctions/PolarCoord.lean @@ -39,12 +39,12 @@ def polarCoord : PartialHomeomorph (ℝ × ℝ) (ℝ × ℝ) where · simpa using hr · right simp at hr - simpa only [ne_of_gt hr, Ne, mem_setOf_eq, mul_eq_zero, false_or_iff, + simpa only [ne_of_gt hr, Ne, mem_setOf_eq, mul_eq_zero, false_or, sin_eq_zero_iff_of_lt_of_lt hθ.1 hθ.2] using h'θ map_source' := by rintro ⟨x, y⟩ hxy simp only [prod_mk_mem_set_prod_eq, mem_Ioi, sqrt_pos, mem_Ioo, Complex.neg_pi_lt_arg, - true_and_iff, Complex.arg_lt_pi_iff] + true_and, Complex.arg_lt_pi_iff] constructor · cases' hxy with hxy hxy · dsimp at hxy; linarith [sq_pos_of_ne_zero hxy.ne', sq_nonneg y] diff --git a/Mathlib/Analysis/SpecialFunctions/Pow/NNReal.lean b/Mathlib/Analysis/SpecialFunctions/Pow/NNReal.lean index e826c5461f294..d2306464d57ff 100644 --- a/Mathlib/Analysis/SpecialFunctions/Pow/NNReal.lean +++ b/Mathlib/Analysis/SpecialFunctions/Pow/NNReal.lean @@ -666,7 +666,7 @@ theorem mul_rpow_eq_ite (x y : ℝ≥0∞) (z : ℝ) : induction y · rw [ne_eq, coe_eq_zero] at hx0 cases' hz with hz hz <;> simp [*] - simp only [*, false_and_iff, and_false_iff, false_or_iff, if_false] + simp only [*, if_false] norm_cast at * rw [← coe_rpow_of_ne_zero (mul_ne_zero hx0 hy0), NNReal.mul_rpow] norm_cast diff --git a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Angle.lean b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Angle.lean index 806f844aeaaeb..364f5e72076c6 100644 --- a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Angle.lean +++ b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Angle.lean @@ -204,7 +204,7 @@ theorem cos_eq_iff_coe_eq_or_eq_neg {θ ψ : ℝ} : constructor · intro Hcos rw [← sub_eq_zero, cos_sub_cos, mul_eq_zero, mul_eq_zero, neg_eq_zero, - eq_false (two_ne_zero' ℝ), false_or_iff, sin_eq_zero_iff, sin_eq_zero_iff] at Hcos + eq_false (two_ne_zero' ℝ), false_or, sin_eq_zero_iff, sin_eq_zero_iff] at Hcos rcases Hcos with (⟨n, hn⟩ | ⟨n, hn⟩) · right rw [eq_div_iff_mul_eq (two_ne_zero' ℝ), ← sub_eq_iff_eq_add] at hn @@ -252,7 +252,7 @@ theorem cos_sin_inj {θ ψ : ℝ} (Hcos : cos θ = cos ψ) (Hsin : sin θ = sin rw [eq_neg_iff_add_eq_zero, hs] at hc obtain ⟨n, hn⟩ : ∃ n, n • _ = _ := QuotientAddGroup.leftRel_apply.mp (Quotient.exact' hc) rw [← neg_one_mul, add_zero, ← sub_eq_zero, zsmul_eq_mul, ← mul_assoc, ← sub_mul, mul_eq_zero, - eq_false (ne_of_gt pi_pos), or_false_iff, sub_neg_eq_add, ← Int.cast_zero, ← Int.cast_one, + eq_false (ne_of_gt pi_pos), or_false, sub_neg_eq_add, ← Int.cast_zero, ← Int.cast_one, ← Int.cast_ofNat, ← Int.cast_mul, ← Int.cast_add, Int.cast_inj] at hn have : (n * 2 + 1) % (2 : ℤ) = 0 % (2 : ℤ) := congr_arg (· % (2 : ℤ)) hn rw [add_comm, Int.add_mul_emod_self] at this diff --git a/Mathlib/Analysis/SpecificLimits/FloorPow.lean b/Mathlib/Analysis/SpecificLimits/FloorPow.lean index 0916d04a949a1..fe6f17c4e3162 100644 --- a/Mathlib/Analysis/SpecificLimits/FloorPow.lean +++ b/Mathlib/Analysis/SpecificLimits/FloorPow.lean @@ -234,7 +234,7 @@ theorem sum_div_pow_sq_le_div_sq (N : ℕ) {j : ℝ} (hj : 0 < j) {c : ℝ} (hc ∑ i ∈ Ico ⌊Real.log j / Real.log c⌋₊ N, (1 : ℝ) / (c ^ i) ^ 2 := by refine sum_le_sum_of_subset_of_nonneg (fun i hi ↦ ?_) (by intros; positivity) simp only [mem_filter, mem_range] at hi - simp only [hi.1, mem_Ico, and_true_iff] + simp only [hi.1, mem_Ico, and_true] apply Nat.floor_le_of_le apply le_of_lt rw [div_lt_iff (Real.log_pos hc), ← Real.log_pow] diff --git a/Mathlib/Analysis/SpecificLimits/Normed.lean b/Mathlib/Analysis/SpecificLimits/Normed.lean index 2fa99368ef2a5..c25802286eeec 100644 --- a/Mathlib/Analysis/SpecificLimits/Normed.lean +++ b/Mathlib/Analysis/SpecificLimits/Normed.lean @@ -157,7 +157,7 @@ theorem TFAE_exists_lt_isLittleO_pow (f : ℕ → ℝ) (R : ℝ) : · obtain rfl : f = 0 := by ext n simpa using H n - simp only [lt_irrefl, false_or_iff] at h₀ + simp only [lt_irrefl, false_or] at h₀ exact ⟨0, ⟨neg_lt_zero.2 h₀, h₀⟩, isBigO_zero _ _⟩ exact ⟨a, A ⟨ha₀, ha⟩, isBigO_of_le' _ fun n ↦ (H n).trans <| mul_le_mul_of_nonneg_left (le_abs_self _) hC₀.le⟩ diff --git a/Mathlib/CategoryTheory/Limits/FilteredColimitCommutesFiniteLimit.lean b/Mathlib/CategoryTheory/Limits/FilteredColimitCommutesFiniteLimit.lean index 4d3665564222e..ffb3e847a42f9 100644 --- a/Mathlib/CategoryTheory/Limits/FilteredColimitCommutesFiniteLimit.lean +++ b/Mathlib/CategoryTheory/Limits/FilteredColimitCommutesFiniteLimit.lean @@ -112,7 +112,7 @@ theorem colimitLimitToLimitColimit_injective : Finset.mem_union.mpr (Or.inl (by - simp only [true_and_iff, Finset.mem_univ, eq_self_iff_true, exists_prop_of_true, + simp only [true_and, Finset.mem_univ, eq_self_iff_true, exists_prop_of_true, Finset.mem_image, heq_iff_eq] refine ⟨j, ?_⟩ simp only [heq_iff_eq] )) @@ -122,7 +122,7 @@ theorem colimitLimitToLimitColimit_injective : Finset.mem_union.mpr (Or.inr (by - simp only [true_and_iff, Finset.mem_univ, eq_self_iff_true, exists_prop_of_true, + simp only [true_and, Finset.mem_univ, eq_self_iff_true, exists_prop_of_true, Finset.mem_image, heq_iff_eq] refine ⟨j, ?_⟩ simp only [heq_iff_eq])) diff --git a/Mathlib/CategoryTheory/Sites/LocallySurjective.lean b/Mathlib/CategoryTheory/Sites/LocallySurjective.lean index 5d9ab1ddb8a9c..06a7da7304b54 100644 --- a/Mathlib/CategoryTheory/Sites/LocallySurjective.lean +++ b/Mathlib/CategoryTheory/Sites/LocallySurjective.lean @@ -60,7 +60,7 @@ theorem imageSieve_whisker_forget {F G : Cᵒᵖ ⥤ A} (f : F ⟶ G) {U : C} (s theorem imageSieve_app {F G : Cᵒᵖ ⥤ A} (f : F ⟶ G) {U : C} (s : F.obj (op U)) : imageSieve f (f.app _ s) = ⊤ := by ext V i - simp only [Sieve.top_apply, iff_true_iff, imageSieve_apply] + simp only [Sieve.top_apply, iff_true, imageSieve_apply] have := elementwise_of% (f.naturality i.op) exact ⟨F.map i.op s, this s⟩ @@ -94,7 +94,7 @@ instance {F G : Cᵒᵖ ⥤ A} (f : F ⟶ G) [IsLocallySurjective J f] : theorem isLocallySurjective_iff_imagePresheaf_sheafify_eq_top {F G : Cᵒᵖ ⥤ A} (f : F ⟶ G) : IsLocallySurjective J f ↔ (imagePresheaf (whiskerRight f (forget A))).sheafify J = ⊤ := by simp only [Subpresheaf.ext_iff, Function.funext_iff, Set.ext_iff, top_subpresheaf_obj, - Set.top_eq_univ, Set.mem_univ, iff_true_iff] + Set.top_eq_univ, Set.mem_univ, iff_true] exact ⟨fun H _ => H.imageSieve_mem, fun H => ⟨H _⟩⟩ theorem isLocallySurjective_iff_imagePresheaf_sheafify_eq_top' {F G : Cᵒᵖ ⥤ Type w} (f : F ⟶ G) : diff --git a/Mathlib/CategoryTheory/Sites/Subsheaf.lean b/Mathlib/CategoryTheory/Sites/Subsheaf.lean index 3731eadbee58d..198726f465d9b 100644 --- a/Mathlib/CategoryTheory/Sites/Subsheaf.lean +++ b/Mathlib/CategoryTheory/Sites/Subsheaf.lean @@ -118,7 +118,7 @@ theorem Subpresheaf.eq_top_iff_isIso : G = ⊤ ↔ IsIso G.ι := by infer_instance · intro H ext U x - apply iff_true_iff.mpr + apply (iff_of_eq (iff_true _)).mpr rw [← IsIso.inv_hom_id_apply (G.ι.app U) x] exact ((inv (G.ι.app U)) x).2 diff --git a/Mathlib/Combinatorics/Additive/Energy.lean b/Mathlib/Combinatorics/Additive/Energy.lean index 02b6197b5c3cb..a6a99d5f15e0d 100644 --- a/Mathlib/Combinatorics/Additive/Energy.lean +++ b/Mathlib/Combinatorics/Additive/Energy.lean @@ -183,7 +183,7 @@ lemma mulEnergy_univ_left : Eₘ[univ, t] = Fintype.card α * t.card ^ 2 := by rw [mul_right_cancel h.1] rw [← card_image_of_injOn this] congr with a - simp only [mem_filter, mem_product, mem_univ, true_and_iff, mem_image, exists_prop, + simp only [mem_filter, mem_product, mem_univ, true_and, mem_image, exists_prop, Prod.exists] refine ⟨fun h => ⟨a.1.1 * a.2.2⁻¹, _, _, h.1, by simp [f, mul_right_comm, h.2]⟩, ?_⟩ rintro ⟨b, c, d, hcd, rfl⟩ diff --git a/Mathlib/Combinatorics/Configuration.lean b/Mathlib/Combinatorics/Configuration.lean index fb13f1517300b..aabfcf9052b58 100644 --- a/Mathlib/Combinatorics/Configuration.lean +++ b/Mathlib/Combinatorics/Configuration.lean @@ -220,7 +220,7 @@ theorem HasLines.card_le [HasLines P L] [Fintype P] [Fintype L] : _ < ∑ p, lineCount L p := by obtain ⟨p, hp⟩ := not_forall.mp (mt (Fintype.card_le_of_surjective f) hc₂) refine sum_lt_sum_of_subset (subset_univ _) (mem_univ p) ?_ ?_ fun p _ _ ↦ zero_le _ - · simpa only [Finset.mem_map, exists_prop, Finset.mem_univ, true_and_iff] + · simpa only [Finset.mem_map, exists_prop, Finset.mem_univ, true_and] · rw [lineCount, Nat.card_eq_fintype_card, Fintype.card_pos_iff] obtain ⟨l, _⟩ := @exists_line P L _ _ p exact @@ -261,7 +261,7 @@ theorem HasLines.lineCount_eq_pointCount [HasLines P L] [Fintype P] [Fintype L] simp_rw [hf2, sum_const, Set.toFinset_card, ← Nat.card_eq_fintype_card] change pointCount P l • _ = lineCount L (f l) • _ rw [hf2] - all_goals simp_rw [s, Finset.mem_univ, true_and_iff, Set.mem_toFinset]; exact fun p => Iff.rfl + all_goals simp_rw [s, Finset.mem_univ, true_and, Set.mem_toFinset]; exact fun p => Iff.rfl have step3 : ∑ i ∈ sᶜ, lineCount L i.1 = ∑ i ∈ sᶜ, pointCount P i.2 := by rwa [← s.sum_add_sum_compl, ← s.sum_add_sum_compl, step2, add_left_cancel_iff] at step1 rw [← Set.toFinset_compl] at step3 diff --git a/Mathlib/Combinatorics/SetFamily/KruskalKatona.lean b/Mathlib/Combinatorics/SetFamily/KruskalKatona.lean index c20ab83898964..8b5ce0135fec7 100644 --- a/Mathlib/Combinatorics/SetFamily/KruskalKatona.lean +++ b/Mathlib/Combinatorics/SetFamily/KruskalKatona.lean @@ -86,7 +86,7 @@ lemma shadow_initSeg [Fintype α] (hs : s.Nonempty) : -- if j < k, k is our colex witness for t ∪ {j} < s · refine Or.inr ⟨k, mem_of_mem_erase ‹_›, fun hk ↦ hkt <| mem_of_mem_insert_of_ne hk hjk.ne', fun x hx ↦ ?_⟩ - simpa only [mem_insert, z hx, (hjk.trans hx).ne', mem_erase, Ne, false_or_iff, + simpa only [mem_insert, z hx, (hjk.trans hx).ne', mem_erase, Ne, false_or, and_iff_right_iff_imp] using fun _ ↦ ((min'_le _ _ <| mem_of_mem_erase hks).trans_lt hx).ne' -- if j = k, all of range k is in t so by sizes t ∪ {j} = s refine Or.inl (eq_of_subset_of_card_le (fun a ha ↦ ?_) hcard.ge).symm diff --git a/Mathlib/Combinatorics/SimpleGraph/Acyclic.lean b/Mathlib/Combinatorics/SimpleGraph/Acyclic.lean index 2cb9bcfec2287..a93671aaf9864 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Acyclic.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Acyclic.lean @@ -114,7 +114,7 @@ theorem isAcyclic_of_path_unique (h : ∀ (v w : V) (p q : G.Path v w), p = q) : cases c with | nil => cases hc.2.1 rfl | cons ha c' => - simp only [Walk.cons_isTrail_iff, Walk.support_cons, List.tail_cons, true_and_iff] at hc + simp only [Walk.cons_isTrail_iff, Walk.support_cons, List.tail_cons] at hc specialize h _ _ ⟨c', by simp only [Walk.isPath_def, hc.2]⟩ (Path.singleton ha.symm) rw [Path.singleton, Subtype.mk.injEq] at h simp [h] at hc @@ -132,7 +132,7 @@ theorem isTree_iff_existsUnique_path : intro v w let q := (hc v w).some.toPath use q - simp only [true_and_iff, Path.isPath] + simp only [true_and, Path.isPath] intro p hp specialize hu ⟨p, hp⟩ q exact Subtype.ext_iff.mp hu diff --git a/Mathlib/Combinatorics/SimpleGraph/Basic.lean b/Mathlib/Combinatorics/SimpleGraph/Basic.lean index 6aaf4e25f4463..c2be3bbe6e5c2 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Basic.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Basic.lean @@ -573,12 +573,12 @@ theorem fromEdgeSet_edgeSet : fromEdgeSet G.edgeSet = G := by @[simp] theorem fromEdgeSet_empty : fromEdgeSet (∅ : Set (Sym2 V)) = ⊥ := by ext v w - simp only [fromEdgeSet_adj, Set.mem_empty_iff_false, false_and_iff, bot_adj] + simp only [fromEdgeSet_adj, Set.mem_empty_iff_false, false_and, bot_adj] @[simp] theorem fromEdgeSet_univ : fromEdgeSet (Set.univ : Set (Sym2 V)) = ⊤ := by ext v w - simp only [fromEdgeSet_adj, Set.mem_univ, true_and_iff, top_adj] + simp only [fromEdgeSet_adj, Set.mem_univ, true_and, top_adj] @[simp] theorem fromEdgeSet_inter (s t : Set (Sym2 V)) : @@ -603,7 +603,7 @@ theorem fromEdgeSet_sdiff (s t : Set (Sym2 V)) : theorem fromEdgeSet_mono {s t : Set (Sym2 V)} (h : s ⊆ t) : fromEdgeSet s ≤ fromEdgeSet t := by rintro v w simp (config := { contextual := true }) only [fromEdgeSet_adj, Ne, not_false_iff, - and_true_iff, and_imp] + and_true, and_imp] exact fun vws _ => h vws @[simp] lemma disjoint_fromEdgeSet : Disjoint G (fromEdgeSet s) ↔ Disjoint G.edgeSet s := by diff --git a/Mathlib/Combinatorics/SimpleGraph/Clique.lean b/Mathlib/Combinatorics/SimpleGraph/Clique.lean index 8de6f50777720..cd44c839664d1 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Clique.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Clique.lean @@ -270,9 +270,9 @@ theorem not_cliqueFree_of_top_embedding {n : ℕ} (f : (⊤ : SimpleGraph (Fin n ¬G.CliqueFree n := by simp only [CliqueFree, isNClique_iff, isClique_iff_induce_eq, not_forall, Classical.not_not] use Finset.univ.map f.toEmbedding - simp only [card_map, Finset.card_fin, eq_self_iff_true, and_true_iff] + simp only [card_map, Finset.card_fin, eq_self_iff_true, and_true] ext ⟨v, hv⟩ ⟨w, hw⟩ - simp only [coe_map, Set.mem_image, coe_univ, Set.mem_univ, true_and_iff] at hv hw + simp only [coe_map, Set.mem_image, coe_univ, Set.mem_univ, true_and] at hv hw obtain ⟨v', rfl⟩ := hv obtain ⟨w', rfl⟩ := hw simp only [coe_sort_coe, RelEmbedding.coe_toEmbedding, comap_adj, Function.Embedding.coe_subtype, diff --git a/Mathlib/Combinatorics/SimpleGraph/Connectivity/WalkCounting.lean b/Mathlib/Combinatorics/SimpleGraph/Connectivity/WalkCounting.lean index 15302f6c68c45..bd09d692268b5 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Connectivity/WalkCounting.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Connectivity/WalkCounting.lean @@ -41,7 +41,7 @@ theorem set_walk_self_length_zero_eq (u : V) : {p : G.Walk u u | p.length = 0} = theorem set_walk_length_zero_eq_of_ne {u v : V} (h : u ≠ v) : {p : G.Walk u v | p.length = 0} = ∅ := by ext p - simp only [Set.mem_setOf_eq, Set.mem_empty_iff_false, iff_false_iff] + simp only [Set.mem_setOf_eq, Set.mem_empty_iff_false, iff_false] exact fun h' => absurd (Walk.eq_of_length_eq_zero h') h theorem set_walk_length_succ_eq (u v : V) (n : ℕ) : diff --git a/Mathlib/Combinatorics/SimpleGraph/DegreeSum.lean b/Mathlib/Combinatorics/SimpleGraph/DegreeSum.lean index 7a78571cb9edc..868f25e9ef0db 100644 --- a/Mathlib/Combinatorics/SimpleGraph/DegreeSum.lean +++ b/Mathlib/Combinatorics/SimpleGraph/DegreeSum.lean @@ -51,7 +51,7 @@ variable [Fintype V] [DecidableRel G.Adj] theorem dart_fst_fiber [DecidableEq V] (v : V) : (univ.filter fun d : G.Dart => d.fst = v) = univ.image (G.dartOfNeighborSet v) := by ext d - simp only [mem_image, true_and_iff, mem_filter, SetCoe.exists, mem_univ, exists_prop_of_true] + simp only [mem_image, true_and, mem_filter, SetCoe.exists, mem_univ, exists_prop_of_true] constructor · rintro rfl exact ⟨_, d.adj, by ext <;> rfl⟩ @@ -121,7 +121,7 @@ theorem even_card_odd_degree_vertices [Fintype V] [DecidableRel G.Adj] : convert h exact ZMod.ne_zero_iff_odd.symm · intro v - simp only [true_and_iff, mem_filter, mem_univ, Ne] + simp only [true_and, mem_filter, mem_univ, Ne] rw [ZMod.eq_zero_iff_even, ZMod.eq_one_iff_odd, ← Nat.not_even_iff_odd, imp_self] trivial @@ -131,7 +131,7 @@ theorem odd_card_odd_degree_vertices_ne [Fintype V] [DecidableEq V] [DecidableRe have hk : 0 < k := by have hh : (filter (fun v : V => Odd (G.degree v)) univ).Nonempty := by use v - simp only [true_and_iff, mem_filter, mem_univ] + simp only [true_and, mem_filter, mem_univ] exact h rwa [← card_pos, hg, ← two_mul, mul_pos_iff_of_pos_left] at hh exact zero_lt_two @@ -144,7 +144,7 @@ theorem odd_card_odd_degree_vertices_ne [Fintype V] [DecidableEq V] [DecidableRe rw [add_assoc, one_add_one_eq_two, ← Nat.mul_succ, ← two_mul] congr omega - · simpa only [true_and_iff, mem_filter, mem_univ] + · simpa only [true_and, mem_filter, mem_univ] theorem exists_ne_odd_degree_of_exists_odd_degree [Fintype V] [DecidableRel G.Adj] (v : V) (h : Odd (G.degree v)) : ∃ w : V, w ≠ v ∧ Odd (G.degree w) := by @@ -154,7 +154,7 @@ theorem exists_ne_odd_degree_of_exists_odd_degree [Fintype V] [DecidableRel G.Ad rw [hg] apply Nat.succ_pos rcases card_pos.mp hg' with ⟨w, hw⟩ - simp only [true_and_iff, mem_filter, mem_univ, Ne] at hw + simp only [true_and, mem_filter, mem_univ, Ne] at hw exact ⟨w, hw⟩ end SimpleGraph diff --git a/Mathlib/Combinatorics/SimpleGraph/IncMatrix.lean b/Mathlib/Combinatorics/SimpleGraph/IncMatrix.lean index ec09c234563e3..8a8ce3e22216a 100644 --- a/Mathlib/Combinatorics/SimpleGraph/IncMatrix.lean +++ b/Mathlib/Combinatorics/SimpleGraph/IncMatrix.lean @@ -125,10 +125,10 @@ theorem sum_incMatrix_apply_of_mem_edgeSet [Fintype α] : intro a b h rw [mem_edgeSet] at h rw [← Nat.cast_two, ← card_pair h.ne] - simp only [incMatrix_apply', sum_boole, mk'_mem_incidenceSet_iff, h, true_and_iff] + simp only [incMatrix_apply', sum_boole, mk'_mem_incidenceSet_iff, h] congr 2 ext e - simp only [mem_filter, mem_univ, true_and_iff, mem_insert, mem_singleton] + simp only [mem_filter, mem_univ, true_and, mem_insert, mem_singleton] theorem sum_incMatrix_apply_of_not_mem_edgeSet [Fintype α] (h : e ∉ G.edgeSet) : ∑ a, G.incMatrix R a e = 0 := diff --git a/Mathlib/Combinatorics/SimpleGraph/Path.lean b/Mathlib/Combinatorics/SimpleGraph/Path.lean index a094d7bce9012..9a5f20e639064 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Path.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Path.lean @@ -1124,7 +1124,7 @@ theorem adj_and_reachable_delete_edges_iff_exists_cycle {v w : V} : rw [Sym2.eq_swap] intro h cases hp (Walk.edges_toPath_subset p h) - · simp only [Sym2.eq_swap, Walk.edges_cons, List.mem_cons, eq_self_iff_true, true_or_iff] + · simp only [Sym2.eq_swap, Walk.edges_cons, List.mem_cons, eq_self_iff_true, true_or] · rintro ⟨u, c, hc, he⟩ refine ⟨c.adj_of_mem_edges he, ?_⟩ by_contra! hb @@ -1143,7 +1143,7 @@ theorem isBridge_iff_adj_and_forall_cycle_not_mem {v w : V} : G.IsBridge s(v, w) rw [← not_iff_not] push_neg rw [← adj_and_reachable_delete_edges_iff_exists_cycle] - simp only [h, true_and_iff] + simp only [h, true_and] theorem isBridge_iff_mem_and_forall_cycle_not_mem {e : Sym2 V} : G.IsBridge e ↔ e ∈ G.edgeSet ∧ ∀ ⦃u : V⦄ (p : G.Walk u u), p.IsCycle → e ∉ p.edges := diff --git a/Mathlib/Combinatorics/SimpleGraph/Regularity/Increment.lean b/Mathlib/Combinatorics/SimpleGraph/Regularity/Increment.lean index 9e7da73f754a5..4fe6e82a487fa 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Regularity/Increment.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Regularity/Increment.lean @@ -93,7 +93,7 @@ private theorem distinctPairs_increment : P.parts.offDiag.attach.biUnion (distinctPairs hP G ε) ⊆ (increment hP G ε).parts.offDiag := by rintro ⟨Ui, Vj⟩ simp only [distinctPairs, increment, mem_offDiag, bind_parts, mem_biUnion, Prod.exists, - exists_and_left, exists_prop, mem_product, mem_attach, true_and_iff, Subtype.exists, and_imp, + exists_and_left, exists_prop, mem_product, mem_attach, true_and, Subtype.exists, and_imp, mem_offDiag, forall_exists_index, exists₂_imp, Ne] refine fun U V hUV hUi hVj => ⟨⟨_, hUV.1, hUi⟩, ⟨_, hUV.2.1, hVj⟩, ?_⟩ rintro rfl diff --git a/Mathlib/Combinatorics/SimpleGraph/Subgraph.lean b/Mathlib/Combinatorics/SimpleGraph/Subgraph.lean index 68777353d2638..9d64a2270e2d0 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Subgraph.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Subgraph.lean @@ -606,7 +606,7 @@ theorem comap_monotone {G' : SimpleGraph W} (f : G →g G') : Monotone (Subgraph simp only [comap_verts, Set.mem_preimage] apply h.1 · intro v w - simp (config := { contextual := true }) only [comap_adj, and_imp, true_and_iff] + simp (config := { contextual := true }) only [comap_adj, and_imp, true_and] intro apply h.2 @@ -615,7 +615,7 @@ theorem map_le_iff_le_comap {G' : SimpleGraph W} (f : G →g G') (H : G.Subgraph refine ⟨fun h ↦ ⟨fun v hv ↦ ?_, fun v w hvw ↦ ?_⟩, fun h ↦ ⟨fun v ↦ ?_, fun v w ↦ ?_⟩⟩ · simp only [comap_verts, Set.mem_preimage] exact h.1 ⟨v, hv, rfl⟩ - · simp only [H.adj_sub hvw, comap_adj, true_and_iff] + · simp only [H.adj_sub hvw, comap_adj, true_and] exact h.2 ⟨v, w, hvw, rfl, rfl⟩ · simp only [map_verts, Set.mem_image, forall_exists_index, and_imp] rintro w hw rfl @@ -767,7 +767,7 @@ theorem eq_singletonSubgraph_iff_verts_eq (H : G.Subgraph) {v : V} : refine ⟨fun h ↦ by rw [h, singletonSubgraph_verts], fun h ↦ ?_⟩ ext · rw [h, singletonSubgraph_verts] - · simp only [Prop.bot_eq_false, singletonSubgraph_adj, Pi.bot_apply, iff_false_iff] + · simp only [Prop.bot_eq_false, singletonSubgraph_adj, Pi.bot_apply, iff_false] intro ha have ha1 := ha.fst_mem have ha2 := ha.snd_mem @@ -784,7 +784,7 @@ theorem edgeSet_subgraphOfAdj {v w : V} (hvw : G.Adj v w) : (G.subgraphOfAdj hvw).edgeSet = {s(v, w)} := by ext e refine e.ind ?_ - simp only [eq_comm, Set.mem_singleton_iff, Subgraph.mem_edgeSet, subgraphOfAdj_adj, iff_self_iff, + simp only [eq_comm, Set.mem_singleton_iff, Subgraph.mem_edgeSet, subgraphOfAdj_adj, forall₂_true_iff] lemma subgraphOfAdj_le_of_adj {v w : V} (H : G.Subgraph) (h : H.Adj v w) : @@ -992,7 +992,7 @@ theorem deleteEdges_le : G'.deleteEdges s ≤ G' := by theorem deleteEdges_le_of_le {s s' : Set (Sym2 V)} (h : s ⊆ s') : G'.deleteEdges s' ≤ G'.deleteEdges s := by constructor <;> simp (config := { contextual := true }) only [deleteEdges_verts, deleteEdges_adj, - true_and_iff, and_imp, subset_rfl] + true_and, and_imp, subset_rfl] exact fun _ _ _ hs' hs ↦ hs' (h hs) @[simp] @@ -1044,7 +1044,7 @@ variable {G' G'' : G.Subgraph} {s s' : Set V} theorem induce_mono (hg : G' ≤ G'') (hs : s ⊆ s') : G'.induce s ≤ G''.induce s' := by constructor · simp [hs] - · simp (config := { contextual := true }) only [induce_adj, true_and_iff, and_imp] + · simp (config := { contextual := true }) only [induce_adj, and_imp] intro v w hv hw ha exact ⟨hs hv, hs hw, hg.2 ha⟩ @@ -1065,7 +1065,7 @@ theorem induce_self_verts : G'.induce G'.verts = G' := by ext · simp · constructor <;> - simp (config := { contextual := true }) only [induce_adj, imp_true_iff, and_true_iff] + simp (config := { contextual := true }) only [induce_adj, imp_true_iff, and_true] exact fun ha ↦ ⟨G'.edge_vert ha, G'.edge_vert ha.symm⟩ lemma le_induce_top_verts : G' ≤ (⊤ : G.Subgraph).induce G'.verts := diff --git a/Mathlib/Combinatorics/SimpleGraph/Trails.lean b/Mathlib/Combinatorics/SimpleGraph/Trails.lean index bf530ee6beb46..174a67d7b2b8d 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Trails.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Trails.lean @@ -58,18 +58,18 @@ theorem IsTrail.even_countP_edges_iff {u v : V} {p : G.Walk u v} (ht : p.IsTrail · rw [decide_eq_true_eq] at h obtain (rfl | rfl) := h · rw [Nat.even_add_one, ih] - simp only [huv.ne, imp_false, Ne, not_false_iff, true_and_iff, not_forall, - Classical.not_not, exists_prop, eq_self_iff_true, not_true, false_and_iff, + simp only [huv.ne, imp_false, Ne, not_false_iff, true_and, not_forall, + Classical.not_not, exists_prop, eq_self_iff_true, not_true, false_and, and_iff_right_iff_imp] rintro rfl rfl exact G.loopless _ huv · rw [Nat.even_add_one, ih, ← not_iff_not] - simp only [huv.ne.symm, Ne, eq_self_iff_true, not_true, false_and_iff, not_forall, - not_false_iff, exists_prop, and_true_iff, Classical.not_not, true_and_iff, iff_and_self] + simp only [huv.ne.symm, Ne, eq_self_iff_true, not_true, false_and, not_forall, + not_false_iff, exists_prop, and_true, Classical.not_not, true_and, iff_and_self] rintro rfl exact huv.ne · rw [decide_eq_true_eq, not_or] at h - simp only [h.1, h.2, not_false_iff, true_and_iff, add_zero, Ne] at ih ⊢ + simp only [h.1, h.2, not_false_iff, true_and, add_zero, Ne] at ih ⊢ rw [ih] constructor <;> · rintro h' h'' rfl diff --git a/Mathlib/Combinatorics/SimpleGraph/Walk.lean b/Mathlib/Combinatorics/SimpleGraph/Walk.lean index c86f6abe53427..9fd198d9d716f 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Walk.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Walk.lean @@ -560,7 +560,7 @@ theorem subset_support_append_left {V : Type u} {G : SimpleGraph V} {u v w : V} theorem subset_support_append_right {V : Type u} {G : SimpleGraph V} {u v w : V} (p : G.Walk u v) (q : G.Walk v w) : q.support ⊆ (p.append q).support := by intro h - simp (config := { contextual := true }) only [mem_support_append_iff, or_true_iff, imp_true_iff] + simp (config := { contextual := true }) only [mem_support_append_iff, or_true, imp_true_iff] theorem coe_support {u v : V} (p : G.Walk u v) : (p.support : Multiset V) = {u} + p.support.tail := by cases p <;> rfl @@ -1201,7 +1201,7 @@ theorem map_injective_of_injective {f : G →g G'} (hinj : Function.Injective f) | cons _ _ => simp only [map_cons, cons.injEq] at h cases hinj h.1 - simp only [cons.injEq, heq_iff_eq, true_and_iff] + simp only [cons.injEq, heq_iff_eq, true_and] apply ih simpa using h.2 diff --git a/Mathlib/Computability/Reduce.lean b/Mathlib/Computability/Reduce.lean index 32ee465296600..753ba658feba9 100644 --- a/Mathlib/Computability/Reduce.lean +++ b/Mathlib/Computability/Reduce.lean @@ -370,7 +370,7 @@ private theorem le_antisymm {d₁ d₂ : ManyOneDegree} : d₁ ≤ d₂ → d₂ induction d₁ using ManyOneDegree.ind_on induction d₂ using ManyOneDegree.ind_on intro hp hq - simp_all only [ManyOneEquiv, of_le_of, of_eq_of, true_and_iff] + simp_all only [ManyOneEquiv, of_le_of, of_eq_of, true_and] private theorem le_trans {d₁ d₂ d₃ : ManyOneDegree} : d₁ ≤ d₂ → d₂ ≤ d₃ → d₁ ≤ d₃ := by induction d₁ using ManyOneDegree.ind_on diff --git a/Mathlib/Computability/RegularExpressions.lean b/Mathlib/Computability/RegularExpressions.lean index 04ff068cce5f2..cbcdb86be9877 100644 --- a/Mathlib/Computability/RegularExpressions.lean +++ b/Mathlib/Computability/RegularExpressions.lean @@ -306,7 +306,7 @@ theorem star_rmatch_iff (P : RegularExpression α) : · exact ⟨[], [], by tauto⟩ · cases' t' with b t · simp only [forall_eq_or_imp, List.mem_cons] at helem - simp only [eq_self_iff_true, not_true, Ne, false_and_iff] at helem + simp only [eq_self_iff_true, not_true, Ne, false_and] at helem simp only [List.join, List.cons_append, List.cons_eq_cons] at hsum refine ⟨t, U.join, hsum.2, ?_, ?_⟩ · specialize helem (b :: t) (by simp) diff --git a/Mathlib/Computability/TMToPartrec.lean b/Mathlib/Computability/TMToPartrec.lean index 59fb6a5cf5aed..68249b2638fe3 100644 --- a/Mathlib/Computability/TMToPartrec.lean +++ b/Mathlib/Computability/TMToPartrec.lean @@ -340,8 +340,8 @@ theorem exists_code {n} {f : Vector ℕ n →. ℕ} (hf : Nat.Partrec' f) : have := PFun.mem_fix_iff.1 h2 simp only [hf, Part.bind_some] at this split_ifs at this with h - · simp only [List.headI_nil, List.headI_cons, exists_false, or_false_iff, Part.mem_some_iff, - List.tail_cons, false_and_iff, Sum.inl.injEq, reduceCtorEq] at this + · simp only [List.headI_nil, List.headI_cons, exists_false, or_false, Part.mem_some_iff, + List.tail_cons, false_and, Sum.inl.injEq, reduceCtorEq] at this subst this exact ⟨_, ⟨h, @(hm)⟩, rfl⟩ · refine IH (n.succ::v.val) (by simp_all) _ rfl fun m h' => ?_ @@ -1607,7 +1607,7 @@ theorem trStmts₁_trans {q q'} : q' ∈ trStmts₁ q → trStmts₁ q' ⊆ trSt | read q q_ih => _ | succ q q_ih => _ | pred q₁ q₂ q₁_ih q₂_ih => _ | ret => _ all_goals simp (config := { contextual := true }) only [trStmts₁, Finset.mem_insert, Finset.mem_union, - or_imp, Finset.mem_singleton, Finset.Subset.refl, imp_true_iff, true_and_iff] + or_imp, Finset.mem_singleton, Finset.Subset.refl, imp_true_iff, true_and] repeat exact fun h => Finset.Subset.trans (q_ih h) (Finset.subset_insert _ _) · simp intro s h x h' diff --git a/Mathlib/Computability/TuringMachine.lean b/Mathlib/Computability/TuringMachine.lean index 8ebbf88168bcd..ea36cdcbb6053 100644 --- a/Mathlib/Computability/TuringMachine.lean +++ b/Mathlib/Computability/TuringMachine.lean @@ -1383,7 +1383,7 @@ theorem tr_supports {S : Finset Λ} (ss : TM1.Supports M S) : cases' q' with q' v' simp only [trStmts, Finset.mem_coe] at h₂ ⊢ rw [Finset.mem_product] at h₂ ⊢ - simp only [Finset.mem_univ, and_true_iff] at h₂ ⊢ + simp only [Finset.mem_univ, and_true] at h₂ ⊢ cases q'; · exact Multiset.mem_cons_self _ _ simp only [tr, Option.mem_def] at h₁ have := TM1.stmts_supportsStmt ss h₂ @@ -1734,7 +1734,7 @@ theorem tr_supports [Inhabited Λ] {S : Finset Λ} (ss : Supports M S) : cases d <;> simp only [trNormal, iterate, supportsStmt_move, IH] | write f q IH => unfold writes at hw ⊢ - simp only [Finset.mem_image, Finset.mem_union, Finset.mem_univ, exists_prop, true_and_iff] + simp only [Finset.mem_image, Finset.mem_union, Finset.mem_univ, exists_prop, true_and] at hw ⊢ replace IH := IH hs fun q hq ↦ hw q (Or.inr hq) refine ⟨supportsStmt_read _ fun a _ s ↦ hw _ (Or.inl ⟨_, rfl⟩), fun q' hq ↦ ?_⟩ diff --git a/Mathlib/Control/LawfulFix.lean b/Mathlib/Control/LawfulFix.lean index 90e5b963198d1..485cc3ade329c 100644 --- a/Mathlib/Control/LawfulFix.lean +++ b/Mathlib/Control/LawfulFix.lean @@ -80,7 +80,7 @@ theorem mem_iff (a : α) (b : β a) : b ∈ Part.fix f a ↔ ∃ i, b ∈ approx · rcases le_total i j with H | H <;> [skip; symm] <;> apply_assumption <;> assumption replace hh := approx_mono f case _ _ hh apply Part.mem_unique h₁ hh - · simp only [fix_def' (⇑f) h₀, not_exists, false_iff_iff, not_mem_none] + · simp only [fix_def' (⇑f) h₀, not_exists, false_iff, not_mem_none] simp only [dom_iff_mem, not_exists] at h₀ intro; apply h₀ diff --git a/Mathlib/Data/Fin/Tuple/NatAntidiagonal.lean b/Mathlib/Data/Fin/Tuple/NatAntidiagonal.lean index 4cef85d6ad0d3..7f72be2b670cb 100644 --- a/Mathlib/Data/Fin/Tuple/NatAntidiagonal.lean +++ b/Mathlib/Data/Fin/Tuple/NatAntidiagonal.lean @@ -143,8 +143,8 @@ theorem antidiagonalTuple_pairwise_pi_lex : simp_rw [antidiagonalTuple, List.pairwise_bind, List.pairwise_map, List.mem_map, forall_exists_index, and_imp, forall_apply_eq_imp_iff₂] simp only [mem_antidiagonal, Prod.forall, and_imp, forall_apply_eq_imp_iff₂] - simp only [Fin.pi_lex_lt_cons_cons, eq_self_iff_true, true_and_iff, lt_self_iff_false, - false_or_iff] + simp only [Fin.pi_lex_lt_cons_cons, eq_self_iff_true, true_and, lt_self_iff_false, + false_or] refine ⟨fun _ _ _ => antidiagonalTuple_pairwise_pi_lex k _, ?_⟩ induction' n with n n_ih · rw [antidiagonal_zero] diff --git a/Mathlib/Data/FinEnum.lean b/Mathlib/Data/FinEnum.lean index 82688911d453d..7e93b826d44d3 100644 --- a/Mathlib/Data/FinEnum.lean +++ b/Mathlib/Data/FinEnum.lean @@ -76,7 +76,7 @@ noncomputable def ofInjective {α β} (f : α → β) [DecidableEq α] [FinEnum ofList ((toList β).filterMap (partialInv f)) (by intro x - simp only [mem_toList, true_and_iff, List.mem_filterMap] + simp only [mem_toList, true_and, List.mem_filterMap] use f x simp only [h, Function.partialInv_left]) @@ -133,7 +133,7 @@ theorem Finset.mem_enum [DecidableEq α] (s : Finset α) (xs : List α) : by_cases h : xs_hd ∈ s · have : {xs_hd} ⊆ s := by simp only [HasSubset.Subset, *, forall_eq, mem_singleton] - simp only [union_sdiff_of_subset this, or_true_iff, Finset.union_sdiff_of_subset, + simp only [union_sdiff_of_subset this, or_true, Finset.union_sdiff_of_subset, eq_self_iff_true] · left symm diff --git a/Mathlib/Data/Finmap.lean b/Mathlib/Data/Finmap.lean index ed5e6a411c850..3c10573ce6ef0 100644 --- a/Mathlib/Data/Finmap.lean +++ b/Mathlib/Data/Finmap.lean @@ -213,7 +213,7 @@ theorem keys_singleton (a : α) (b : β a) : (singleton a b).keys = {a} := @[simp] theorem mem_singleton (x y : α) (b : β y) : x ∈ singleton y b ↔ x = y := by - simp only [singleton]; erw [mem_cons, mem_nil_iff, or_false_iff] + simp only [singleton]; erw [mem_cons, mem_nil_iff, or_false] section diff --git a/Mathlib/Data/Finset/Basic.lean b/Mathlib/Data/Finset/Basic.lean index 07d86ddb5847a..4015ac7c6e26d 100644 --- a/Mathlib/Data/Finset/Basic.lean +++ b/Mathlib/Data/Finset/Basic.lean @@ -1427,7 +1427,7 @@ theorem insert_inter_of_not_mem {s₁ s₂ : Finset α} {a : α} (h : a ∉ s₂ insert a s₁ ∩ s₂ = s₁ ∩ s₂ := ext fun x => by have : ¬(x = a ∧ x ∈ s₂) := by rintro ⟨rfl, H⟩; exact h H - simp only [mem_inter, mem_insert, or_and_right, this, false_or_iff] + simp only [mem_inter, mem_insert, or_and_right, this, false_or] @[simp] theorem inter_insert_of_not_mem {s₁ s₂ : Finset α} {a : α} (h : a ∉ s₁) : @@ -1479,7 +1479,7 @@ instance : DistribLattice (Finset α) := { le_sup_inf := fun a b c => by simp (config := { contextual := true }) only [sup_eq_union, inf_eq_inter, le_eq_subset, subset_iff, mem_inter, mem_union, and_imp, - or_imp, true_or_iff, imp_true_iff, true_and_iff, or_true_iff] } + or_imp, true_or, imp_true_iff, true_and, or_true] } @[simp] theorem union_left_idem (s t : Finset α) : s ∪ (s ∪ t) = s ∪ t := sup_left_idem _ _ @@ -1637,7 +1637,7 @@ theorem erase_eq_self : s.erase a = s ↔ a ∉ s := theorem erase_insert_eq_erase (s : Finset α) (a : α) : (insert a s).erase a = s.erase a := ext fun x => by simp (config := { contextual := true }) only [mem_erase, mem_insert, and_congr_right_iff, - false_or_iff, iff_self_iff, imp_true_iff] + false_or, iff_self, imp_true_iff] theorem erase_insert {a : α} {s : Finset α} (h : a ∉ s) : erase (insert a s) a = s := by rw [erase_insert_eq_erase, erase_eq_of_not_mem h] @@ -1654,7 +1654,7 @@ theorem erase_cons_of_ne {a b : α} {s : Finset α} (ha : a ∉ s) (hb : a ≠ b @[simp] theorem insert_erase (h : a ∈ s) : insert a (erase s a) = s := ext fun x => by - simp only [mem_insert, mem_erase, or_and_left, dec_em, true_and_iff] + simp only [mem_insert, mem_erase, or_and_left, dec_em, true_and] apply or_iff_right_of_imp rintro rfl exact h @@ -1772,11 +1772,11 @@ instance : GeneralizedBooleanAlgebra (Finset α) := simp only [Finset.ext_iff, mem_union, mem_sdiff, inf_eq_inter, sup_eq_union, mem_inter, ← and_or_left, em, and_true, implies_true] inf_inf_sdiff := fun x y => by - simp only [Finset.ext_iff, inter_sdiff_self, inter_empty, inter_assoc, false_iff_iff, + simp only [Finset.ext_iff, inter_sdiff_self, inter_empty, inter_assoc, false_iff, inf_eq_inter, not_mem_empty, bot_eq_empty, not_false_iff, implies_true] } theorem not_mem_sdiff_of_mem_right (h : a ∈ t) : a ∉ s \ t := by - simp only [mem_sdiff, h, not_true, not_false_iff, and_false_iff] + simp only [mem_sdiff, h, not_true, not_false_iff, and_false] theorem not_mem_sdiff_of_not_mem_left (h : a ∉ s) : a ∉ s \ t := by simp [h] @@ -2393,7 +2393,7 @@ theorem filter_insert (a : α) (s : Finset α) : theorem filter_erase (a : α) (s : Finset α) : filter p (erase s a) = erase (filter p s) a := by ext x - simp only [and_assoc, mem_filter, iff_self_iff, mem_erase] + simp only [and_assoc, mem_filter, iff_self, mem_erase] theorem filter_or (s : Finset α) : (s.filter fun a => p a ∨ q a) = s.filter p ∪ s.filter q := ext fun _ => by simp [mem_filter, mem_union, and_or_left] @@ -2463,7 +2463,7 @@ theorem filter_eq [DecidableEq β] (s : Finset β) (b : β) : rintro rfl exact ⟨h, rfl⟩ · ext - simp only [mem_filter, not_and, iff_false_iff, not_mem_empty, decide_eq_true_eq] + simp only [mem_filter, not_and, iff_false, not_mem_empty, decide_eq_true_eq] rintro m rfl exact h m @@ -2595,7 +2595,7 @@ end Range -- useful rules for calculations with quantifiers theorem exists_mem_empty_iff (p : α → Prop) : (∃ x, x ∈ (∅ : Finset α) ∧ p x) ↔ False := by - simp only [not_mem_empty, false_and_iff, exists_false] + simp only [not_mem_empty, false_and, exists_false] theorem exists_mem_insert [DecidableEq α] (a : α) (s : Finset α) (p : α → Prop) : (∃ x, x ∈ insert a s ∧ p x) ↔ p a ∨ ∃ x, x ∈ s ∧ p x := by diff --git a/Mathlib/Data/Finset/Card.lean b/Mathlib/Data/Finset/Card.lean index ee90c07098ad8..c143030e91559 100644 --- a/Mathlib/Data/Finset/Card.lean +++ b/Mathlib/Data/Finset/Card.lean @@ -296,7 +296,7 @@ theorem card_eq_of_bijective (f : ∀ i, i < n → α) (hf : ∀ a ∈ s, ∃ i, have : s = (range n).attach.image fun i => f i.1 (mem_range.1 i.2) := by ext a suffices _ : a ∈ s ↔ ∃ (i : _) (hi : i ∈ range n), f i (mem_range.1 hi) = a by - simpa only [mem_image, mem_attach, true_and_iff, Subtype.exists] + simpa only [mem_image, mem_attach, true_and, Subtype.exists] constructor · intro ha; obtain ⟨i, hi, rfl⟩ := hf a ha; use i, mem_range.2 hi · rintro ⟨i, hi, rfl⟩; apply hf' diff --git a/Mathlib/Data/Finset/Lattice.lean b/Mathlib/Data/Finset/Lattice.lean index 452204c7d625f..d577bf61dda82 100644 --- a/Mathlib/Data/Finset/Lattice.lean +++ b/Mathlib/Data/Finset/Lattice.lean @@ -220,7 +220,7 @@ theorem sup_le_of_le_directed {α : Type*} [SemilatticeSup α] [OrderBot α] (s (∀ x ∈ t, ∃ y ∈ s, x ≤ y) → ∃ x ∈ s, t.sup id ≤ x := by classical induction' t using Finset.induction_on with a r _ ih h - · simpa only [forall_prop_of_true, and_true_iff, forall_prop_of_false, bot_le, not_false_iff, + · simpa only [forall_prop_of_true, and_true, forall_prop_of_false, bot_le, not_false_iff, sup_empty, forall_true_iff, not_mem_empty] · intro h have incs : (r : Set α) ⊆ ↑(insert a r) := by diff --git a/Mathlib/Data/Finset/Pointwise/Basic.lean b/Mathlib/Data/Finset/Pointwise/Basic.lean index f9133dd51a7aa..95b5171decef7 100644 --- a/Mathlib/Data/Finset/Pointwise/Basic.lean +++ b/Mathlib/Data/Finset/Pointwise/Basic.lean @@ -988,12 +988,12 @@ to ∃ a, s = {a} ∧ IsUnit a -/ -- @[simp] theorem isUnit_iff_singleton : IsUnit s ↔ ∃ a, s = {a} := by - simp only [isUnit_iff, Group.isUnit, and_true_iff] + simp only [isUnit_iff, Group.isUnit, and_true] @[simp] theorem isUnit_iff_singleton_aux {α} [Group α] {s : Finset α} : (∃ a, s = {a} ∧ IsUnit a) ↔ ∃ a, s = {a} := by - simp only [Group.isUnit, and_true_iff] + simp only [Group.isUnit, and_true] @[to_additive (attr := simp)] theorem image_mul_left : diff --git a/Mathlib/Data/Finset/Union.lean b/Mathlib/Data/Finset/Union.lean index b54499d3e58d9..3df81fb5d2be1 100644 --- a/Mathlib/Data/Finset/Union.lean +++ b/Mathlib/Data/Finset/Union.lean @@ -52,7 +52,7 @@ lemma disjiUnion_val (s : Finset α) (t : α → Finset β) (h) : @[simp, norm_cast] lemma coe_disjiUnion {h} : (s.disjiUnion t h : Set β) = ⋃ x ∈ (s : Set α), t x := by - simp [Set.ext_iff, mem_disjiUnion, Set.mem_iUnion, iff_self_iff, mem_coe, imp_true_iff] + simp [Set.ext_iff, mem_disjiUnion, Set.mem_iUnion, mem_coe, imp_true_iff] @[simp] lemma disjiUnion_cons (a : α) (s : Finset α) (ha : a ∉ s) (f : α → Finset β) (H) : disjiUnion (cons a s ha) f H = @@ -119,7 +119,7 @@ protected def biUnion (s : Finset α) (t : α → Finset β) : Finset β := @[simp, norm_cast] lemma coe_biUnion : (s.biUnion t : Set β) = ⋃ x ∈ (s : Set α), t x := by - simp [Set.ext_iff, mem_biUnion, Set.mem_iUnion, iff_self_iff, mem_coe, imp_true_iff] + simp [Set.ext_iff, mem_biUnion, Set.mem_iUnion, mem_coe, imp_true_iff] @[simp] lemma biUnion_insert [DecidableEq α] {a : α} : (insert a s).biUnion t = t a ∪ s.biUnion t := diff --git a/Mathlib/Data/Finset/Update.lean b/Mathlib/Data/Finset/Update.lean index a0a7aa87054bf..d4dfb055c18ea 100644 --- a/Mathlib/Data/Finset/Update.lean +++ b/Mathlib/Data/Finset/Update.lean @@ -56,8 +56,7 @@ theorem updateFinset_updateFinset {s t : Finset ι} (hst : Disjoint s t) set e := Equiv.Finset.union s t hst congr with i by_cases his : i ∈ s <;> by_cases hit : i ∈ t <;> - simp only [updateFinset, his, hit, dif_pos, dif_neg, Finset.mem_union, true_or_iff, - false_or_iff, not_false_iff] + simp only [updateFinset, his, hit, dif_pos, dif_neg, Finset.mem_union, false_or, not_false_iff] · exfalso; exact Finset.disjoint_left.mp hst his hit · exact piCongrLeft_sum_inl (fun b : ↥(s ∪ t) => π b) e y z ⟨i, his⟩ |>.symm · exact piCongrLeft_sum_inr (fun b : ↥(s ∪ t) => π b) e y z ⟨i, hit⟩ |>.symm diff --git a/Mathlib/Data/Finsupp/Basic.lean b/Mathlib/Data/Finsupp/Basic.lean index 17123acc6aff6..e6107678fa4e7 100644 --- a/Mathlib/Data/Finsupp/Basic.lean +++ b/Mathlib/Data/Finsupp/Basic.lean @@ -622,7 +622,7 @@ theorem sum_comapDomain [Zero M] [AddCommMonoid N] (f : α → β) (l : β → theorem eq_zero_of_comapDomain_eq_zero [AddCommMonoid M] (f : α → β) (l : β →₀ M) (hf : Set.BijOn f (f ⁻¹' ↑l.support) ↑l.support) : comapDomain f l hf.injOn = 0 → l = 0 := by rw [← support_eq_empty, ← support_eq_empty, comapDomain] - simp only [Finset.ext_iff, Finset.not_mem_empty, iff_false_iff, mem_preimage] + simp only [Finset.ext_iff, Finset.not_mem_empty, iff_false, mem_preimage] intro h a ha cases' hf.2.2 ha with b hb exact h b (hb.2.symm ▸ ha) diff --git a/Mathlib/Data/Finsupp/Defs.lean b/Mathlib/Data/Finsupp/Defs.lean index 4e65162454d8b..a41978fa0311d 100644 --- a/Mathlib/Data/Finsupp/Defs.lean +++ b/Mathlib/Data/Finsupp/Defs.lean @@ -756,7 +756,7 @@ def embDomain (f : α ↪ β) (v : α →₀ M) : β →₀ M where mem_support_toFun a₂ := by dsimp split_ifs with h - · simp only [h, true_iff_iff, Ne] + · simp only [h, true_iff, Ne] rw [← not_mem_support_iff, not_not] classical apply Finset.choose_mem · simp only [h, Ne, ne_self_iff_false, not_true_eq_false] diff --git a/Mathlib/Data/Fintype/Basic.lean b/Mathlib/Data/Fintype/Basic.lean index 64062b2403702..065dfecfbeb4e 100644 --- a/Mathlib/Data/Fintype/Basic.lean +++ b/Mathlib/Data/Fintype/Basic.lean @@ -270,7 +270,7 @@ theorem compl_erase : (s.erase a)ᶜ = insert a sᶜ := by @[simp] theorem compl_insert : (insert a s)ᶜ = sᶜ.erase a := by ext - simp only [not_or, mem_insert, iff_self_iff, mem_compl, mem_erase] + simp only [not_or, mem_insert, mem_compl, mem_erase] theorem insert_compl_insert (ha : a ∉ s) : insert a (insert a s)ᶜ = sᶜ := by simp_rw [compl_insert, insert_erase (mem_compl.2 ha)] diff --git a/Mathlib/Data/Fintype/Card.lean b/Mathlib/Data/Fintype/Card.lean index c2ac99212e5f5..8920878fe8f7e 100644 --- a/Mathlib/Data/Fintype/Card.lean +++ b/Mathlib/Data/Fintype/Card.lean @@ -542,7 +542,7 @@ theorem one_lt_card_iff : 1 < card α ↔ ∃ a b : α, a ≠ b := one_lt_card_iff_nontrivial.trans nontrivial_iff nonrec theorem two_lt_card_iff : 2 < card α ↔ ∃ a b c : α, a ≠ b ∧ a ≠ c ∧ b ≠ c := by - simp_rw [← Finset.card_univ, two_lt_card_iff, mem_univ, true_and_iff] + simp_rw [← Finset.card_univ, two_lt_card_iff, mem_univ, true_and] theorem card_of_bijective {f : α → β} (hf : Bijective f) : card α = card β := card_congr (Equiv.ofBijective f hf) @@ -815,7 +815,7 @@ theorem wellFounded_of_trans_of_irrefl (r : α → α → Prop) [IsTrans α r] [ fun x y hxy => Finset.card_lt_card <| by simp only [Finset.lt_iff_ssubset.symm, lt_iff_le_not_le, Finset.le_iff_subset, - Finset.subset_iff, mem_filter, true_and_iff, mem_univ, hxy] + Finset.subset_iff, mem_filter, true_and, mem_univ, hxy] exact ⟨fun z hzx => _root_.trans hzx hxy, not_forall_of_exists_not ⟨x, Classical.not_imp.2 ⟨hxy, irrefl x⟩⟩⟩ diff --git a/Mathlib/Data/Fintype/Fin.lean b/Mathlib/Data/Fintype/Fin.lean index 1c108db32e13b..79f88ae3dc5ed 100644 --- a/Mathlib/Data/Fintype/Fin.lean +++ b/Mathlib/Data/Fintype/Fin.lean @@ -37,8 +37,7 @@ theorem Iio_last_eq_map : Iio (Fin.last n) = Finset.univ.map Fin.castSuccEmb := @[simp] theorem Ioi_succ (i : Fin n) : Ioi i.succ = (Ioi i).map (Fin.succEmb _) := by ext i - simp only [mem_filter, mem_Ioi, mem_map, mem_univ, true_and_iff, Function.Embedding.coeFn_mk, - exists_true_left] + simp only [mem_filter, mem_Ioi, mem_map, mem_univ, Function.Embedding.coeFn_mk, exists_true_left] constructor · refine cases ?_ ?_ i · rintro ⟨⟨⟩⟩ diff --git a/Mathlib/Data/Int/GCD.lean b/Mathlib/Data/Int/GCD.lean index 038b354fda39a..6c4e2b4ad9ed9 100644 --- a/Mathlib/Data/Int/GCD.lean +++ b/Mathlib/Data/Int/GCD.lean @@ -322,7 +322,7 @@ theorem gcd_least_linear {a b : ℤ} (ha : a ≠ 0) : IsLeast { n : ℕ | 0 < n ∧ ∃ x y : ℤ, ↑n = a * x + b * y } (a.gcd b) := by simp_rw [← gcd_dvd_iff] constructor - · simpa [and_true_iff, dvd_refl, Set.mem_setOf_eq] using gcd_pos_of_ne_zero_left b ha + · simpa [and_true, dvd_refl, Set.mem_setOf_eq] using gcd_pos_of_ne_zero_left b ha · simp only [lowerBounds, and_imp, Set.mem_setOf_eq] exact fun n hn_pos hn => Nat.le_of_dvd hn_pos hn diff --git a/Mathlib/Data/List/Basic.lean b/Mathlib/Data/List/Basic.lean index ff5b94381bca9..a19aa53596821 100644 --- a/Mathlib/Data/List/Basic.lean +++ b/Mathlib/Data/List/Basic.lean @@ -718,7 +718,7 @@ theorem indexOf_eq_length {a : α} {l : List α} : indexOf a l = length l ↔ a rw [cond_eq_if] split_ifs with h <;> simp at h · exact iff_of_false (by rintro ⟨⟩) fun H => H <| Or.inl h.symm - · simp only [Ne.symm h, false_or_iff] + · simp only [Ne.symm h, false_or] rw [← ih] exact succ_inj' @@ -2197,7 +2197,7 @@ variable {p q : α → Prop} {l : List α} @[simp] theorem forall_cons (p : α → Prop) (x : α) : ∀ l : List α, Forall p (x :: l) ↔ p x ∧ Forall p l - | [] => (and_true_iff _).symm + | [] => (and_iff_left_of_imp fun _ ↦ trivial).symm | _ :: _ => Iff.rfl theorem forall_iff_forall_mem : ∀ {l : List α}, Forall p l ↔ ∀ x ∈ l, p x diff --git a/Mathlib/Data/List/Chain.lean b/Mathlib/Data/List/Chain.lean index b73a2aa957cd3..a1953d0529518 100644 --- a/Mathlib/Data/List/Chain.lean +++ b/Mathlib/Data/List/Chain.lean @@ -45,12 +45,12 @@ theorem Chain.iff_mem {a : α} {l : List α} : Chain.imp fun a b h => h.2.2⟩ theorem chain_singleton {a b : α} : Chain R a [b] ↔ R a b := by - simp only [chain_cons, Chain.nil, and_true_iff] + simp only [chain_cons, Chain.nil, and_true] theorem chain_split {a b : α} {l₁ l₂ : List α} : Chain R a (l₁ ++ b :: l₂) ↔ Chain R a (l₁ ++ [b]) ∧ Chain R b l₂ := by induction' l₁ with x l₁ IH generalizing a <;> - simp only [*, nil_append, cons_append, Chain.nil, chain_cons, and_true_iff, and_assoc] + simp only [*, nil_append, cons_append, Chain.nil, chain_cons, and_true, and_assoc] @[simp] theorem chain_append_cons_cons {a b c : α} {l₁ l₂ : List α} : @@ -99,7 +99,7 @@ protected theorem Chain.pairwise [IsTrans α R] : | a, _, @Chain.cons _ _ _ b l h hb => hb.pairwise.cons (by - simp only [mem_cons, forall_eq_or_imp, h, true_and_iff] + simp only [mem_cons, forall_eq_or_imp, h, true_and] exact fun c hc => _root_.trans h (rel_of_pairwise_cons hb.pairwise hc)) theorem chain_iff_pairwise [IsTrans α R] {a : α} {l : List α} : Chain R a l ↔ Pairwise R (a :: l) := @@ -266,7 +266,7 @@ theorem Chain'.take (h : Chain' R l) (n : ℕ) : Chain' R (take n l) := h.prefix (take_prefix _ _) theorem chain'_pair {x y} : Chain' R [x, y] ↔ R x y := by - simp only [chain'_singleton, chain'_cons, and_true_iff] + simp only [chain'_singleton, chain'_cons, and_true] theorem Chain'.imp_head {x y} (h : ∀ {z}, R x z → R y z) {l} (hl : Chain' R (x :: l)) : Chain' R (y :: l) := diff --git a/Mathlib/Data/List/Cycle.lean b/Mathlib/Data/List/Cycle.lean index d68df32b94451..5f038e2ae1e94 100644 --- a/Mathlib/Data/List/Cycle.lean +++ b/Mathlib/Data/List/Cycle.lean @@ -528,7 +528,7 @@ theorem Subsingleton.congr {s : Cycle α} (h : Subsingleton s) : ∀ ⦃x⦄ (_hx : x ∈ s) ⦃y⦄ (_hy : y ∈ s), x = y := by induction' s using Quot.inductionOn with l simp only [length_subsingleton_iff, length_coe, mk_eq_coe, le_iff_lt_or_eq, Nat.lt_add_one_iff, - length_eq_zero, length_eq_one, Nat.not_lt_zero, false_or_iff] at h + length_eq_zero, length_eq_one, Nat.not_lt_zero, false_or] at h rcases h with (rfl | ⟨z, rfl⟩) <;> simp /-- A `s : Cycle α` that is made up of at least two unique elements. -/ @@ -543,7 +543,7 @@ theorem nontrivial_coe_nodup_iff {l : List α} (hl : l.Nodup) : · simp · simp · simp only [mem_cons, exists_prop, mem_coe_iff, List.length, Ne, Nat.succ_le_succ_iff, - Nat.zero_le, iff_true_iff] + Nat.zero_le, iff_true] refine ⟨hd, hd', ?_, by simp⟩ simp only [not_or, mem_cons, nodup_cons] at hl exact hl.left.left @@ -880,7 +880,7 @@ theorem chain_iff_pairwise [IsTrans α r] : Chain r s ↔ ∀ a ∈ s, ∀ b ∈ intro hs b hb c hc rw [Cycle.chain_coe_cons, List.chain_iff_pairwise] at hs simp only [pairwise_append, pairwise_cons, mem_append, mem_singleton, List.not_mem_nil, - IsEmpty.forall_iff, imp_true_iff, Pairwise.nil, forall_eq, true_and_iff] at hs + IsEmpty.forall_iff, imp_true_iff, Pairwise.nil, forall_eq, true_and] at hs simp only [mem_coe_iff, mem_cons] at hb hc rcases hb with (rfl | hb) <;> rcases hc with (rfl | hc) · exact hs.1 c (Or.inr rfl) diff --git a/Mathlib/Data/List/FinRange.lean b/Mathlib/Data/List/FinRange.lean index bb4b35d88e9a0..22b38b33b6259 100644 --- a/Mathlib/Data/List/FinRange.lean +++ b/Mathlib/Data/List/FinRange.lean @@ -72,7 +72,7 @@ open List theorem Equiv.Perm.map_finRange_perm {n : ℕ} (σ : Equiv.Perm (Fin n)) : map σ (finRange n) ~ finRange n := by rw [perm_ext_iff_of_nodup ((nodup_finRange n).map σ.injective) <| nodup_finRange n] - simpa [mem_map, mem_finRange, true_and_iff, iff_true_iff] using σ.surjective + simpa [mem_map, mem_finRange] using σ.surjective /-- The list obtained from a permutation of a tuple `f` is permutation equivalent to the list obtained from `f`. -/ diff --git a/Mathlib/Data/List/Forall2.lean b/Mathlib/Data/List/Forall2.lean index f2e78d52a7c08..79f5bf6c1cb0a 100644 --- a/Mathlib/Data/List/Forall2.lean +++ b/Mathlib/Data/List/Forall2.lean @@ -87,8 +87,7 @@ theorem forall₂_cons_right_iff {b l u} : theorem forall₂_and_left {p : α → Prop} : ∀ l u, Forall₂ (fun a b => p a ∧ R a b) l u ↔ (∀ a ∈ l, p a) ∧ Forall₂ R l u | [], u => by - simp only [forall₂_nil_left_iff, forall_prop_of_false (not_mem_nil _), imp_true_iff, - true_and_iff] + simp only [forall₂_nil_left_iff, forall_prop_of_false (not_mem_nil _), imp_true_iff, true_and] | a :: l, u => by simp only [forall₂_and_left l, forall₂_cons_left_iff, forall_mem_cons, and_assoc, @and_comm _ (p a), @and_left_comm _ (p a), exists_and_left] @@ -244,7 +243,7 @@ theorem rel_filter {p : α → Bool} {q : β → Bool} dsimp [LiftFun] at hpq by_cases h : p a · have : q b := by rwa [← hpq h₁] - simp only [filter_cons_of_pos h, filter_cons_of_pos this, forall₂_cons, h₁, true_and_iff, + simp only [filter_cons_of_pos h, filter_cons_of_pos this, forall₂_cons, h₁, true_and, rel_filter hpq h₂] · have : ¬q b := by rwa [← hpq h₁] simp only [filter_cons_of_neg h, filter_cons_of_neg this, rel_filter hpq h₂] diff --git a/Mathlib/Data/List/Lattice.lean b/Mathlib/Data/List/Lattice.lean index 181f14ce99536..93192b359af50 100644 --- a/Mathlib/Data/List/Lattice.lean +++ b/Mathlib/Data/List/Lattice.lean @@ -178,13 +178,13 @@ theorem cons_bagInter_of_neg (l₁ : List α) (h : a ∉ l₂) : @[simp] theorem mem_bagInter {a : α} : ∀ {l₁ l₂ : List α}, a ∈ l₁.bagInter l₂ ↔ a ∈ l₁ ∧ a ∈ l₂ - | [], l₂ => by simp only [nil_bagInter, not_mem_nil, false_and_iff] + | [], l₂ => by simp only [nil_bagInter, not_mem_nil, false_and] | b :: l₁, l₂ => by by_cases h : b ∈ l₂ · rw [cons_bagInter_of_pos _ h, mem_cons, mem_cons, mem_bagInter] by_cases ba : a = b - · simp only [ba, h, eq_self_iff_true, true_or_iff, true_and_iff] - · simp only [mem_erase_of_ne ba, ba, false_or_iff] + · simp only [ba, h, eq_self_iff_true, true_or, true_and] + · simp only [mem_erase_of_ne ba, ba, false_or] · rw [cons_bagInter_of_neg _ h, mem_bagInter, mem_cons, or_and_right] symm apply or_iff_right_of_imp diff --git a/Mathlib/Data/List/Lemmas.lean b/Mathlib/Data/List/Lemmas.lean index 72e4ddac0a2a7..0d7c5ffd0b7ff 100644 --- a/Mathlib/Data/List/Lemmas.lean +++ b/Mathlib/Data/List/Lemmas.lean @@ -56,7 +56,7 @@ theorem injOn_insertNth_index_of_not_mem (l : List α) (x : α) (hx : x ∉ l) : · rfl · simp [hx.left] at h · simp [Ne.symm hx.left] at h - · simp only [true_and_iff, eq_self_iff_true, insertNth_succ_cons] at h + · simp only [true_and, eq_self_iff_true, insertNth_succ_cons] at h rw [Nat.succ_inj'] refine IH hx.right ?_ ?_ (by injection h) · simpa [Nat.succ_le_succ_iff] using hn diff --git a/Mathlib/Data/List/NatAntidiagonal.lean b/Mathlib/Data/List/NatAntidiagonal.lean index fbbf06a9c47bb..3d11150f01d9e 100644 --- a/Mathlib/Data/List/NatAntidiagonal.lean +++ b/Mathlib/Data/List/NatAntidiagonal.lean @@ -59,7 +59,7 @@ theorem nodup_antidiagonal (n : ℕ) : Nodup (antidiagonal n) := @[simp] theorem antidiagonal_succ {n : ℕ} : antidiagonal (n + 1) = (0, n + 1) :: (antidiagonal n).map (Prod.map Nat.succ id) := by - simp only [antidiagonal, range_succ_eq_map, map_cons, true_and_iff, Nat.add_succ_sub_one, + simp only [antidiagonal, range_succ_eq_map, map_cons, Nat.add_succ_sub_one, Nat.add_zero, id, eq_self_iff_true, Nat.sub_zero, map_map, Prod.map_mk] apply congr rfl (congr rfl _) ext; simp diff --git a/Mathlib/Data/List/Perm.lean b/Mathlib/Data/List/Perm.lean index 565fc6ade6048..71f0f7c142106 100644 --- a/Mathlib/Data/List/Perm.lean +++ b/Mathlib/Data/List/Perm.lean @@ -266,7 +266,7 @@ theorem Perm.inter_append {l t₁ t₂ : List α} (h : Disjoint t₁ t₂) : · have h₂ : x ∉ t₂ := h h₁ simp [*] by_cases h₂ : x ∈ t₂ - · simp only [*, inter_cons_of_not_mem, false_or_iff, mem_append, inter_cons_of_mem, + · simp only [*, inter_cons_of_not_mem, false_or, mem_append, inter_cons_of_mem, not_false_iff] refine Perm.trans (Perm.cons _ l_ih) ?_ change [x] ++ xs ∩ t₁ ++ xs ∩ t₂ ~ xs ∩ t₁ ++ ([x] ++ xs ∩ t₂) diff --git a/Mathlib/Data/List/Range.lean b/Mathlib/Data/List/Range.lean index 66dd6c10fdb9b..66f26df43fc84 100644 --- a/Mathlib/Data/List/Range.lean +++ b/Mathlib/Data/List/Range.lean @@ -34,8 +34,7 @@ theorem chain'_range_succ (r : ℕ → ℕ → Prop) (n : ℕ) : induction' n with n hn · simp · rw [range_succ] - simp only [append_assoc, singleton_append, chain'_append_cons_cons, chain'_singleton, - and_true_iff] + simp only [append_assoc, singleton_append, chain'_append_cons_cons, chain'_singleton, and_true] rw [hn, forall_lt_succ] theorem chain_range_succ (r : ℕ → ℕ → Prop) (n a : ℕ) : diff --git a/Mathlib/Data/List/ReduceOption.lean b/Mathlib/Data/List/ReduceOption.lean index f86d28a812369..8994967357edb 100644 --- a/Mathlib/Data/List/ReduceOption.lean +++ b/Mathlib/Data/List/ReduceOption.lean @@ -34,8 +34,7 @@ theorem reduceOption_map {l : List (Option α)} {f : α → β} : induction' l with hd tl hl · simp only [reduceOption_nil, map_nil] · cases hd <;> - simpa [true_and_iff, Option.map_some', map, eq_self_iff_true, - reduceOption_cons_of_some] using hl + simpa [Option.map_some', map, eq_self_iff_true, reduceOption_cons_of_some] using hl theorem reduceOption_append (l l' : List (Option α)) : (l ++ l').reduceOption = l.reduceOption ++ l'.reduceOption := diff --git a/Mathlib/Data/List/Sections.lean b/Mathlib/Data/List/Sections.lean index 5fc1305b9cafa..c604e4d75e47c 100644 --- a/Mathlib/Data/List/Sections.lean +++ b/Mathlib/Data/List/Sections.lean @@ -25,7 +25,7 @@ theorem mem_sections {L : List (List α)} {f} : f ∈ sections L ↔ Forall₂ ( exact Forall₂.nil simp only [sections, bind_eq_bind, mem_bind, mem_map] at h rcases h with ⟨_, _, _, _, rfl⟩ - simp only [*, forall₂_cons, true_and_iff] + simp only [*, forall₂_cons, true_and] · induction' h with a l f L al fL fs · simp only [sections, mem_singleton] simp only [sections, bind_eq_bind, mem_bind, mem_map] diff --git a/Mathlib/Data/Multiset/Basic.lean b/Mathlib/Data/Multiset/Basic.lean index 94e09d5bf9e63..3de385213638a 100644 --- a/Mathlib/Data/Multiset/Basic.lean +++ b/Mathlib/Data/Multiset/Basic.lean @@ -306,7 +306,7 @@ theorem coe_singleton (a : α) : ([a] : Multiset α) = {a} := @[simp] theorem mem_singleton {a b : α} : b ∈ ({a} : Multiset α) ↔ b = a := by - simp only [← cons_zero, mem_cons, iff_self_iff, or_false_iff, not_mem_zero] + simp only [← cons_zero, mem_cons, iff_self, or_false, not_mem_zero] theorem mem_singleton_self (a : α) : a ∈ ({a} : Multiset α) := by rw [← cons_zero] diff --git a/Mathlib/Data/Multiset/Fintype.lean b/Mathlib/Data/Multiset/Fintype.lean index eef7f8002adbf..ee4962cb23b82 100644 --- a/Mathlib/Data/Multiset/Fintype.lean +++ b/Mathlib/Data/Multiset/Fintype.lean @@ -188,8 +188,7 @@ theorem Multiset.map_univ_coeEmbedding (m : Multiset α) : ext ⟨x, i⟩ simp only [Fin.exists_iff, Finset.mem_map, Finset.mem_univ, Multiset.coeEmbedding_apply, Prod.mk.inj_iff, exists_true_left, Multiset.exists_coe, Multiset.coe_mk, Fin.val_mk, - exists_prop, exists_eq_right_right, exists_eq_right, Multiset.mem_toEnumFinset, iff_self_iff, - true_and_iff] + exists_prop, exists_eq_right_right, exists_eq_right, Multiset.mem_toEnumFinset, true_and] @[simp] theorem Multiset.map_univ_coe (m : Multiset α) : diff --git a/Mathlib/Data/Nat/Bits.lean b/Mathlib/Data/Nat/Bits.lean index 61e802b8c4e17..feb4a5b86056f 100644 --- a/Mathlib/Data/Nat/Bits.lean +++ b/Mathlib/Data/Nat/Bits.lean @@ -340,7 +340,7 @@ theorem binaryRec_eq' {C : ℕ → Sort*} {z : C 0} {f : ∀ b n, C n → C (bit split_ifs with h' · rcases bit_eq_zero_iff.mp h' with ⟨rfl, rfl⟩ rw [binaryRec_zero] - simp only [imp_false, or_false_iff, eq_self_iff_true, not_true, reduceCtorEq] at h + simp only [imp_false, or_false, eq_self_iff_true, not_true, reduceCtorEq] at h exact h.symm · dsimp only [] generalize_proofs e diff --git a/Mathlib/Data/Nat/Factorization/Basic.lean b/Mathlib/Data/Nat/Factorization/Basic.lean index 6e116bcf98273..a00644b704354 100644 --- a/Mathlib/Data/Nat/Factorization/Basic.lean +++ b/Mathlib/Data/Nat/Factorization/Basic.lean @@ -309,7 +309,7 @@ theorem dvd_iff_prime_pow_dvd_dvd (n d : ℕ) : rcases eq_or_ne n 0 with (rfl | hn) · simp rcases eq_or_ne d 0 with (rfl | hd) - · simp only [zero_dvd_iff, hn, false_iff_iff, not_forall] + · simp only [zero_dvd_iff, hn, false_iff, not_forall] exact ⟨2, n, prime_two, dvd_zero _, mt (le_of_dvd hn.bot_lt) (lt_two_pow n).not_le⟩ refine ⟨fun h p k _ hpkd => dvd_trans hpkd h, ?_⟩ rw [← factorization_prime_le_iff_dvd hd hn] diff --git a/Mathlib/Data/Nat/Factorization/PrimePow.lean b/Mathlib/Data/Nat/Factorization/PrimePow.lean index 5f35804b888b8..37333d9e42ecf 100644 --- a/Mathlib/Data/Nat/Factorization/PrimePow.lean +++ b/Mathlib/Data/Nat/Factorization/PrimePow.lean @@ -137,6 +137,6 @@ theorem Nat.mul_divisors_filter_prime_pow {a b : ℕ} (hab : a.Coprime b) : · simp only [Nat.coprime_zero_right] at hab simp [hab, Finset.filter_singleton, not_isPrimePow_one] ext n - simp only [ha, hb, Finset.mem_union, Finset.mem_filter, Nat.mul_eq_zero, and_true_iff, Ne, + simp only [ha, hb, Finset.mem_union, Finset.mem_filter, Nat.mul_eq_zero, and_true, Ne, and_congr_left_iff, not_false_iff, Nat.mem_divisors, or_self_iff] apply hab.isPrimePow_dvd_mul diff --git a/Mathlib/Data/Nat/Lattice.lean b/Mathlib/Data/Nat/Lattice.lean index dadaa5b902c92..46be67ef9f99d 100644 --- a/Mathlib/Data/Nat/Lattice.lean +++ b/Mathlib/Data/Nat/Lattice.lean @@ -45,9 +45,9 @@ theorem _root_.Set.Infinite.Nat.sSup_eq_zero {s : Set ℕ} (h : s.Infinite) : sS theorem sInf_eq_zero {s : Set ℕ} : sInf s = 0 ↔ 0 ∈ s ∨ s = ∅ := by cases eq_empty_or_nonempty s with | inl h => subst h - simp only [or_true_iff, eq_self_iff_true, iff_true_iff, iInf, InfSet.sInf, + simp only [or_true, eq_self_iff_true, iInf, InfSet.sInf, mem_empty_iff_false, exists_false, dif_neg, not_false_iff] - | inr h => simp only [h.ne_empty, or_false_iff, Nat.sInf_def, h, Nat.find_eq_zero] + | inr h => simp only [h.ne_empty, or_false, Nat.sInf_def, h, Nat.find_eq_zero] @[simp] theorem sInf_empty : sInf ∅ = 0 := by diff --git a/Mathlib/Data/Nat/Nth.lean b/Mathlib/Data/Nat/Nth.lean index e31e9bcb4390d..96774ac4e05d1 100644 --- a/Mathlib/Data/Nat/Nth.lean +++ b/Mathlib/Data/Nat/Nth.lean @@ -237,7 +237,7 @@ theorem nth_eq_zero {n} : exacts [nth_zero_of_zero h₀, nth_of_card_le hf hle] theorem nth_eq_zero_mono (h₀ : ¬p 0) {a b : ℕ} (hab : a ≤ b) (ha : nth p a = 0) : nth p b = 0 := by - simp only [nth_eq_zero, h₀, false_and_iff, false_or_iff] at ha ⊢ + simp only [nth_eq_zero, h₀, false_and, false_or] at ha ⊢ exact ha.imp fun hf hle => hle.trans hab theorem le_nth_of_lt_nth_succ {k a : ℕ} (h : a < nth p (k + 1)) (ha : p a) : a ≤ nth p k := by diff --git a/Mathlib/Data/Nat/PartENat.lean b/Mathlib/Data/Nat/PartENat.lean index 1111f4009e1b7..cc13db102cac5 100644 --- a/Mathlib/Data/Nat/PartENat.lean +++ b/Mathlib/Data/Nat/PartENat.lean @@ -88,8 +88,8 @@ instance addCommMonoid : AddCommMonoid PartENat where add := (· + ·) zero := 0 add_comm x y := Part.ext' and_comm fun _ _ => add_comm _ _ - zero_add x := Part.ext' (true_and_iff _) fun _ _ => zero_add _ - add_zero x := Part.ext' (and_true_iff _) fun _ _ => add_zero _ + zero_add x := Part.ext' (iff_of_eq (true_and _)) fun _ _ => zero_add _ + add_zero x := Part.ext' (iff_of_eq (and_true _)) fun _ _ => add_zero _ add_assoc x y z := Part.ext' and_assoc fun _ _ => add_assoc _ _ _ nsmul := nsmulRec @@ -98,7 +98,7 @@ instance : AddCommMonoidWithOne PartENat := one := 1 natCast := some natCast_zero := rfl - natCast_succ := fun _ => Part.ext' (true_and_iff _).symm fun _ _ => rfl } + natCast_succ := fun _ => Part.ext' (iff_of_eq (true_and _)).symm fun _ _ => rfl } theorem some_eq_natCast (n : ℕ) : some n = n := rfl @@ -157,7 +157,7 @@ protected theorem casesOn {P : PartENat → Prop} : ∀ a : PartENat, P ⊤ → -- not a simp lemma as we will provide a `LinearOrderedAddCommMonoidWithTop` instance later theorem top_add (x : PartENat) : ⊤ + x = ⊤ := - Part.ext' (false_and_iff _) fun h => h.left.elim + Part.ext' (iff_of_eq (false_and _)) fun h => h.left.elim -- not a simp lemma as we will provide a `LinearOrderedAddCommMonoidWithTop` instance later theorem add_top (x : PartENat) : x + ⊤ = ⊤ := by rw [add_comm, top_add] @@ -379,7 +379,7 @@ theorem eq_top_iff_forall_le (x : PartENat) : x = ⊤ ↔ ∀ n : ℕ, (n : Part theorem pos_iff_one_le {x : PartENat} : 0 < x ↔ 1 ≤ x := PartENat.casesOn x - (by simp only [iff_true_iff, le_top, natCast_lt_top, ← @Nat.cast_zero PartENat]) + (by simp only [le_top, natCast_lt_top, ← @Nat.cast_zero PartENat]) fun n => by rw [← Nat.cast_zero, ← Nat.cast_one, PartENat.coe_lt_coe, PartENat.coe_le_coe] rfl diff --git a/Mathlib/Data/Nat/Periodic.lean b/Mathlib/Data/Nat/Periodic.lean index 5836ce4f952c9..1e2e335bec976 100644 --- a/Mathlib/Data/Nat/Periodic.lean +++ b/Mathlib/Data/Nat/Periodic.lean @@ -24,7 +24,7 @@ theorem periodic_gcd (a : ℕ) : Periodic (gcd a) a := by simp only [forall_const, gcd_add_self_right, eq_self_iff_true, Periodic] theorem periodic_coprime (a : ℕ) : Periodic (Coprime a) a := by - simp only [coprime_add_self_right, forall_const, iff_self_iff, eq_iff_iff, Periodic] + simp only [coprime_add_self_right, forall_const, eq_iff_iff, Periodic] theorem periodic_mod (a : ℕ) : Periodic (fun n => n % a) a := by simp only [forall_const, eq_self_iff_true, add_mod_right, Periodic] diff --git a/Mathlib/Data/Nat/Prime/Basic.lean b/Mathlib/Data/Nat/Prime/Basic.lean index e75aca751b05e..d0fe2d6bad6ba 100644 --- a/Mathlib/Data/Nat/Prime/Basic.lean +++ b/Mathlib/Data/Nat/Prime/Basic.lean @@ -156,7 +156,7 @@ theorem Prime.eq_one_of_pow {x n : ℕ} (h : (x ^ n).Prime) : n = 1 := theorem Prime.pow_eq_iff {p a k : ℕ} (hp : p.Prime) : a ^ k = p ↔ a = p ∧ k = 1 := by refine ⟨fun h => ?_, fun h => by rw [h.1, h.2, pow_one]⟩ rw [← h] at hp - rw [← h, hp.eq_one_of_pow, eq_self_iff_true, and_true_iff, pow_one] + rw [← h, hp.eq_one_of_pow, eq_self_iff_true, _root_.and_true, pow_one] theorem pow_minFac {n k : ℕ} (hk : k ≠ 0) : (n ^ k).minFac = n.minFac := by rcases eq_or_ne n 1 with (rfl | hn) @@ -247,7 +247,7 @@ theorem ne_one_iff_exists_prime_dvd : ∀ {n}, n ≠ 1 ↔ ∃ p : ℕ, p.Prime | n + 2 => by let a := n + 2 let ha : a ≠ 1 := Nat.succ_succ_ne_one n - simp only [true_iff_iff, Ne, not_false_iff, ha] + simp only [true_iff, Ne, not_false_iff, ha] exact ⟨a.minFac, Nat.minFac_prime ha, a.minFac_dvd⟩ theorem eq_one_iff_not_exists_prime_dvd {n : ℕ} : n = 1 ↔ ∀ p : ℕ, p.Prime → ¬p ∣ n := by diff --git a/Mathlib/Data/Nat/Prime/Defs.lean b/Mathlib/Data/Nat/Prime/Defs.lean index 8cff72a3c2aa9..a7b24cb865b8f 100644 --- a/Mathlib/Data/Nat/Prime/Defs.lean +++ b/Mathlib/Data/Nat/Prime/Defs.lean @@ -161,7 +161,7 @@ theorem Prime.not_dvd_one {p : ℕ} (pp : Prime p) : ¬p ∣ 1 := Irreducible.not_dvd_one pp theorem prime_mul_iff {a b : ℕ} : Nat.Prime (a * b) ↔ a.Prime ∧ b = 1 ∨ b.Prime ∧ a = 1 := by - simp only [iff_self_iff, irreducible_mul_iff, ← irreducible_iff_nat_prime, Nat.isUnit_iff] + simp only [irreducible_mul_iff, ← irreducible_iff_nat_prime, Nat.isUnit_iff] theorem not_prime_mul {a b : ℕ} (a1 : a ≠ 1) (b1 : b ≠ 1) : ¬Prime (a * b) := by simp [prime_mul_iff, _root_.not_or, *] @@ -337,7 +337,7 @@ theorem minFac_le_div {n : ℕ} (pos : 0 < n) (np : ¬Prime n) : minFac n ≤ n | ⟨0, h0⟩ => absurd pos <| by rw [h0, mul_zero]; decide | ⟨1, h1⟩ => by rw [mul_one] at h1 - rw [prime_def_minFac, not_and_or, ← h1, eq_self_iff_true, _root_.not_true, or_false_iff, + rw [prime_def_minFac, not_and_or, ← h1, eq_self_iff_true, _root_.not_true, _root_.or_false, not_le] at np rw [le_antisymm (le_of_lt_succ np) (succ_le_of_lt pos), minFac_one, Nat.div_one] | ⟨x + 2, hx⟩ => by diff --git a/Mathlib/Data/Nat/Squarefree.lean b/Mathlib/Data/Nat/Squarefree.lean index 8e5ebe1de87d1..d11d003343934 100644 --- a/Mathlib/Data/Nat/Squarefree.lean +++ b/Mathlib/Data/Nat/Squarefree.lean @@ -52,7 +52,7 @@ theorem _root_.Squarefree.natFactorization_le_one {n : ℕ} (p : ℕ) (hn : Squa rw [multiplicity.squarefree_iff_multiplicity_le_one] at hn by_cases hp : p.Prime · have := hn p - simp only [multiplicity_eq_factorization hp hn', Nat.isUnit_iff, hp.ne_one, or_false_iff] + simp only [multiplicity_eq_factorization hp hn', Nat.isUnit_iff, hp.ne_one, or_false] at this exact mod_cast this · rw [factorization_eq_zero_of_non_prime _ hp] diff --git a/Mathlib/Data/Nat/Totient.lean b/Mathlib/Data/Nat/Totient.lean index 5602a10425923..ec08e85c219d9 100644 --- a/Mathlib/Data/Nat/Totient.lean +++ b/Mathlib/Data/Nat/Totient.lean @@ -229,7 +229,7 @@ theorem prime_iff_card_units (p : ℕ) [Fintype (ZMod p)ˣ] : p.Prime ↔ Fintype.card (ZMod p)ˣ = p - 1 := by cases' eq_zero_or_neZero p with hp hp · subst hp - simp only [ZMod, not_prime_zero, false_iff_iff, zero_tsub] + simp only [ZMod, not_prime_zero, false_iff, zero_tsub] -- the subst created a non-defeq but subsingleton instance diamond; resolve it suffices Fintype.card ℤˣ ≠ 0 by convert this simp @@ -245,7 +245,7 @@ theorem totient_eq_one_iff : ∀ {n : ℕ}, n.totient = 1 ↔ n = 1 ∨ n = 2 | 2 => by simp | n + 3 => by have : 3 ≤ n + 3 := le_add_self - simp only [succ_succ_ne_one, false_or_iff] + simp only [succ_succ_ne_one, false_or] exact ⟨fun h => not_even_one.elim <| h ▸ totient_even this, by rintro ⟨⟩⟩ theorem dvd_two_of_totient_le_one {a : ℕ} (han : 0 < a) (ha : a.totient ≤ 1) : a ∣ 2 := by diff --git a/Mathlib/Data/PEquiv.lean b/Mathlib/Data/PEquiv.lean index b80eb899c2f5e..920084d283b41 100644 --- a/Mathlib/Data/PEquiv.lean +++ b/Mathlib/Data/PEquiv.lean @@ -390,7 +390,7 @@ instance [DecidableEq α] [DecidableEq β] : SemilatticeInf (α ≃. β) := · contrapose! h2 rw [h2] rw [← h1, hf, h2] at hg - simp only [mem_def, true_iff_iff, eq_self_iff_true] at hg + simp only [mem_def, true_iff, eq_self_iff_true] at hg rw [hg] · contrapose! h1 rw [h1] at hf h2 diff --git a/Mathlib/Data/PFunctor/Univariate/M.lean b/Mathlib/Data/PFunctor/Univariate/M.lean index f3eb8596e529a..73fe251a4dda5 100644 --- a/Mathlib/Data/PFunctor/Univariate/M.lean +++ b/Mathlib/Data/PFunctor/Univariate/M.lean @@ -91,7 +91,7 @@ theorem truncate_eq_of_agree {n : ℕ} (x : CofixA F n) (y : CofixA F (succ n)) · rfl · -- cases' h with _ _ _ _ _ h₀ h₁ cases h - simp only [truncate, Function.comp_def, true_and_iff, eq_self_iff_true, heq_iff_eq] + simp only [truncate, Function.comp_def, eq_self_iff_true, heq_iff_eq] -- Porting note: used to be `ext y` rename_i n_ih a f y h₁ suffices (fun x => truncate (y x)) = f @@ -479,8 +479,8 @@ theorem ext_aux [Inhabited (M F)] [DecidableEq F.A] {n : ℕ} (x y z : M F) (hx induction y using PFunctor.M.casesOn' simp only [iselect_nil] at hrec subst hrec - simp only [approx_mk, true_and_iff, eq_self_iff_true, heq_iff_eq, zero_eq, CofixA.intro.injEq, - heq_eq_eq, eq_iff_true_of_subsingleton, and_self] + simp only [approx_mk, eq_self_iff_true, heq_iff_eq, zero_eq, CofixA.intro.injEq, + heq_eq_eq, eq_iff_true_of_subsingleton, and_self] · cases hx cases hy induction x using PFunctor.M.casesOn' @@ -488,7 +488,7 @@ theorem ext_aux [Inhabited (M F)] [DecidableEq F.A] {n : ℕ} (x y z : M F) (hx subst z iterate 3 (have := mk_inj ‹_›; cases this) rename_i n_ih a f₃ f₂ hAgree₂ _ _ h₂ _ _ f₁ h₁ hAgree₁ clr - simp only [approx_mk, true_and_iff, eq_self_iff_true, heq_iff_eq] + simp only [approx_mk, eq_self_iff_true, heq_iff_eq] have := mk_inj h₁ cases this; clear h₁ diff --git a/Mathlib/Data/Set/Basic.lean b/Mathlib/Data/Set/Basic.lean index dba783a5ea356..5a4257688c30d 100644 --- a/Mathlib/Data/Set/Basic.lean +++ b/Mathlib/Data/Set/Basic.lean @@ -622,11 +622,11 @@ theorem union_self (a : Set α) : a ∪ a = a := @[simp] theorem union_empty (a : Set α) : a ∪ ∅ = a := - ext fun _ => or_false_iff _ + ext fun _ => iff_of_eq (or_false _) @[simp] theorem empty_union (a : Set α) : ∅ ∪ a = a := - ext fun _ => false_or_iff _ + ext fun _ => iff_of_eq (false_or _) theorem union_comm (a b : Set α) : a ∪ b = b ∪ a := ext fun _ => or_comm @@ -740,11 +740,11 @@ theorem inter_self (a : Set α) : a ∩ a = a := @[simp] theorem inter_empty (a : Set α) : a ∩ ∅ = ∅ := - ext fun _ => and_false_iff _ + ext fun _ => iff_of_eq (and_false _) @[simp] theorem empty_inter (a : Set α) : ∅ ∩ a = ∅ := - ext fun _ => false_and_iff _ + ext fun _ => iff_of_eq (false_and _) theorem inter_comm (a b : Set α) : a ∩ b = b ∩ a := ext fun _ => and_comm @@ -1134,7 +1134,7 @@ theorem sep_eq_self_iff_mem_true : { x ∈ s | p x } = s ↔ ∀ x ∈ s, p x := @[simp] theorem sep_eq_empty_iff_mem_false : { x ∈ s | p x } = ∅ ↔ ∀ x ∈ s, ¬p x := by - simp_rw [Set.ext_iff, mem_sep_iff, mem_empty_iff_false, iff_false_iff, not_and] + simp_rw [Set.ext_iff, mem_sep_iff, mem_empty_iff_false, iff_false, not_and] --Porting note (#10618): removed `simp` attribute because `simp` can prove it theorem sep_true : { x ∈ s | True } = s := @@ -1187,7 +1187,7 @@ theorem Nonempty.subset_singleton_iff (h : s.Nonempty) : s ⊆ {a} ↔ s = {a} : subset_singleton_iff_eq.trans <| or_iff_right h.ne_empty theorem ssubset_singleton_iff {s : Set α} {x : α} : s ⊂ {x} ↔ s = ∅ := by - rw [ssubset_iff_subset_ne, subset_singleton_iff_eq, or_and_right, and_not_self_iff, or_false_iff, + rw [ssubset_iff_subset_ne, subset_singleton_iff_eq, or_and_right, and_not_self_iff, or_false, and_iff_left_iff_imp] exact fun h => h ▸ (singleton_ne_empty _).symm @@ -1581,7 +1581,7 @@ theorem insert_diff_self_of_not_mem {a : α} {s : Set α} (h : a ∉ s) : insert theorem insert_diff_eq_singleton {a : α} {s : Set α} (h : a ∉ s) : insert a s \ s = {a} := by ext rw [Set.mem_diff, Set.mem_insert_iff, Set.mem_singleton_iff, or_and_right, and_not_self_iff, - or_false_iff, and_iff_left_iff_imp] + or_false, and_iff_left_iff_imp] rintro rfl exact h diff --git a/Mathlib/Data/Set/Countable.lean b/Mathlib/Data/Set/Countable.lean index 074e6b876960c..4add89c35bb25 100644 --- a/Mathlib/Data/Set/Countable.lean +++ b/Mathlib/Data/Set/Countable.lean @@ -226,7 +226,7 @@ theorem Countable.of_diff {s t : Set α} (h : (s \ t).Countable) (ht : t.Countab @[simp] theorem countable_insert {s : Set α} {a : α} : (insert a s).Countable ↔ s.Countable := by - simp only [insert_eq, countable_union, countable_singleton, true_and_iff] + simp only [insert_eq, countable_union, countable_singleton, true_and] protected theorem Countable.insert {s : Set α} (a : α) (h : s.Countable) : (insert a s).Countable := countable_insert.2 h diff --git a/Mathlib/Data/Set/Pairwise/Basic.lean b/Mathlib/Data/Set/Pairwise/Basic.lean index 64e482ceb8858..73033eaeb2298 100644 --- a/Mathlib/Data/Set/Pairwise/Basic.lean +++ b/Mathlib/Data/Set/Pairwise/Basic.lean @@ -129,8 +129,7 @@ theorem pairwise_union_of_symmetric (hr : Symmetric r) : theorem pairwise_insert : (insert a s).Pairwise r ↔ s.Pairwise r ∧ ∀ b ∈ s, a ≠ b → r a b ∧ r b a := by - simp only [insert_eq, pairwise_union, pairwise_singleton, true_and_iff, mem_singleton_iff, - forall_eq] + simp only [insert_eq, pairwise_union, pairwise_singleton, true_and, mem_singleton_iff, forall_eq] theorem pairwise_insert_of_not_mem (ha : a ∉ s) : (insert a s).Pairwise r ↔ s.Pairwise r ∧ ∀ b ∈ s, r a b ∧ r b a := diff --git a/Mathlib/Data/Set/Prod.lean b/Mathlib/Data/Set/Prod.lean index 73ce4876bc230..acc80442bafa7 100644 --- a/Mathlib/Data/Set/Prod.lean +++ b/Mathlib/Data/Set/Prod.lean @@ -72,17 +72,17 @@ theorem exists_prod_set {p : α × β → Prop} : (∃ x ∈ s ×ˢ t, p x) ↔ @[simp] theorem prod_empty : s ×ˢ (∅ : Set β) = ∅ := by ext - exact and_false_iff _ + exact iff_of_eq (and_false _) @[simp] theorem empty_prod : (∅ : Set α) ×ˢ t = ∅ := by ext - exact false_and_iff _ + exact iff_of_eq (false_and _) @[simp, mfld_simps] theorem univ_prod_univ : @univ α ×ˢ @univ β = univ := by ext - exact true_and_iff _ + exact iff_of_eq (true_and _) theorem univ_prod {t : Set β} : (univ : Set α) ×ˢ t = Prod.snd ⁻¹' t := by simp [prod_eq] @@ -326,7 +326,7 @@ theorem prod_subset_prod_iff : s ×ˢ t ⊆ s₁ ×ˢ t₁ ↔ s ⊆ s₁ ∧ t · have := image_subset (Prod.snd : α × β → β) H rwa [snd_image_prod st.1, snd_image_prod (h.mono H).fst] at this · intro H - simp only [st.1.ne_empty, st.2.ne_empty, or_false_iff] at H + simp only [st.1.ne_empty, st.2.ne_empty, or_false] at H exact prod_mono H.1 H.2 theorem prod_eq_prod_iff_of_nonempty (h : (s ×ˢ t).Nonempty) : @@ -335,7 +335,7 @@ theorem prod_eq_prod_iff_of_nonempty (h : (s ×ˢ t).Nonempty) : · intro heq have h₁ : (s₁ ×ˢ t₁ : Set _).Nonempty := by rwa [← heq] rw [prod_nonempty_iff] at h h₁ - rw [← fst_image_prod s h.2, ← fst_image_prod s₁ h₁.2, heq, eq_self_iff_true, true_and_iff, ← + rw [← fst_image_prod s h.2, ← fst_image_prod s₁ h₁.2, heq, eq_self_iff_true, true_and, ← snd_image_prod h.1 t, ← snd_image_prod h₁.1 t₁, heq] · rintro ⟨rfl, rfl⟩ rfl @@ -344,18 +344,17 @@ theorem prod_eq_prod_iff : s ×ˢ t = s₁ ×ˢ t₁ ↔ s = s₁ ∧ t = t₁ ∨ (s = ∅ ∨ t = ∅) ∧ (s₁ = ∅ ∨ t₁ = ∅) := by symm rcases eq_empty_or_nonempty (s ×ˢ t) with h | h - · simp_rw [h, @eq_comm _ ∅, prod_eq_empty_iff, prod_eq_empty_iff.mp h, true_and_iff, + · simp_rw [h, @eq_comm _ ∅, prod_eq_empty_iff, prod_eq_empty_iff.mp h, true_and, or_iff_right_iff_imp] rintro ⟨rfl, rfl⟩ exact prod_eq_empty_iff.mp h rw [prod_eq_prod_iff_of_nonempty h] rw [nonempty_iff_ne_empty, Ne, prod_eq_empty_iff] at h - simp_rw [h, false_and_iff, or_false_iff] + simp_rw [h, false_and, or_false] @[simp] theorem prod_eq_iff_eq (ht : t.Nonempty) : s ×ˢ t = s₁ ×ˢ t ↔ s = s₁ := by - simp_rw [prod_eq_prod_iff, ht.ne_empty, and_true_iff, or_iff_left_iff_imp, - or_false_iff] + simp_rw [prod_eq_prod_iff, ht.ne_empty, and_true, or_iff_left_iff_imp, or_false] rintro ⟨rfl, rfl⟩ rfl @@ -631,7 +630,7 @@ theorem pi_congr (h : s₁ = s₂) (h' : ∀ i ∈ s₁, t₁ i = t₂ i) : s₁ theorem pi_eq_empty (hs : i ∈ s) (ht : t i = ∅) : s.pi t = ∅ := by ext f - simp only [mem_empty_iff_false, not_forall, iff_false_iff, mem_pi, Classical.not_imp] + simp only [mem_empty_iff_false, not_forall, iff_false, mem_pi, Classical.not_imp] exact ⟨i, hs, by simp [ht]⟩ theorem univ_pi_eq_empty (ht : t i = ∅) : pi univ t = ∅ := diff --git a/Mathlib/Data/Set/Sigma.lean b/Mathlib/Data/Set/Sigma.lean index b10a51f01db79..7201429b16acc 100644 --- a/Mathlib/Data/Set/Sigma.lean +++ b/Mathlib/Data/Set/Sigma.lean @@ -68,15 +68,17 @@ theorem exists_sigma_iff {p : (Σi, α i) → Prop} : (∃ x ∈ s.sigma t, p x) ↔ ∃ i ∈ s, ∃ a ∈ t i, p ⟨i, a⟩ := ⟨fun ⟨⟨i, a⟩, ha, h⟩ ↦ ⟨i, ha.1, a, ha.2, h⟩, fun ⟨i, hi, a, ha, h⟩ ↦ ⟨⟨i, a⟩, ⟨hi, ha⟩, h⟩⟩ -@[simp] theorem sigma_empty : s.sigma (fun i ↦ (∅ : Set (α i))) = ∅ := ext fun _ ↦ and_false_iff _ +@[simp] theorem sigma_empty : s.sigma (fun i ↦ (∅ : Set (α i))) = ∅ := + ext fun _ ↦ iff_of_eq (and_false _) -@[simp] theorem empty_sigma : (∅ : Set ι).sigma t = ∅ := ext fun _ ↦ false_and_iff _ +@[simp] theorem empty_sigma : (∅ : Set ι).sigma t = ∅ := ext fun _ ↦ iff_of_eq (false_and _) -theorem univ_sigma_univ : (@univ ι).sigma (fun _ ↦ @univ (α i)) = univ := ext fun _ ↦ true_and_iff _ +theorem univ_sigma_univ : (@univ ι).sigma (fun _ ↦ @univ (α i)) = univ := + ext fun _ ↦ iff_of_eq (true_and _) @[simp] theorem sigma_univ : s.sigma (fun _ ↦ univ : ∀ i, Set (α i)) = Sigma.fst ⁻¹' s := - ext fun _ ↦ and_true_iff _ + ext fun _ ↦ iff_of_eq (and_true _) @[simp] theorem univ_sigma_preimage_mk (s : Set (Σ i, α i)) : (univ : Set ι).sigma (fun i ↦ Sigma.mk i ⁻¹' s) = s := diff --git a/Mathlib/Data/Set/Subsingleton.lean b/Mathlib/Data/Set/Subsingleton.lean index b31273026245f..fa736da6540c2 100644 --- a/Mathlib/Data/Set/Subsingleton.lean +++ b/Mathlib/Data/Set/Subsingleton.lean @@ -215,7 +215,7 @@ theorem Nontrivial.ne_singleton {x} (hs : s.Nontrivial) : s ≠ {x} := fun H => exact not_nontrivial_singleton hs theorem Nontrivial.not_subset_singleton {x} (hs : s.Nontrivial) : ¬s ⊆ {x} := - (not_congr subset_singleton_iff_eq).2 (not_or_of_not hs.ne_empty hs.ne_singleton) + (not_congr subset_singleton_iff_eq).2 (not_or_intro hs.ne_empty hs.ne_singleton) theorem nontrivial_univ [Nontrivial α] : (univ : Set α).Nontrivial := let ⟨x, y, hxy⟩ := exists_pair_ne α diff --git a/Mathlib/Data/Sym/Sym2.lean b/Mathlib/Data/Sym/Sym2.lean index e94c465684a68..945179cd192f3 100644 --- a/Mathlib/Data/Sym/Sym2.lean +++ b/Mathlib/Data/Sym/Sym2.lean @@ -299,8 +299,7 @@ instance : SetLike (Sym2 α) α where induction' z with x y induction' z' with x' y' have hx := h x; have hy := h y; have hx' := h x'; have hy' := h y' - simp only [mem_iff', eq_self_iff_true, or_true_iff, iff_true_iff, - true_or_iff, true_iff_iff] at hx hy hx' hy' + simp only [mem_iff', eq_self_iff_true] at hx hy hx' hy' aesop @[simp] diff --git a/Mathlib/Data/Vector3.lean b/Mathlib/Data/Vector3.lean index 4e8bcb368cc7e..9ec3dea934529 100644 --- a/Mathlib/Data/Vector3.lean +++ b/Mathlib/Data/Vector3.lean @@ -235,7 +235,7 @@ theorem vectorAllP_singleton (p : α → Prop) (x : α) : VectorAllP p (cons x [ @[simp] theorem vectorAllP_cons (p : α → Prop) (x : α) (v : Vector3 α n) : VectorAllP p (x :: v) ↔ p x ∧ VectorAllP p v := - Vector3.recOn v (and_true_iff _).symm fun _ _ _ => Iff.rfl + Vector3.recOn v (iff_of_eq (and_true _)).symm fun _ _ _ => Iff.rfl theorem vectorAllP_iff_forall (p : α → Prop) (v : Vector3 α n) : VectorAllP p v ↔ ∀ i, p (v i) := by diff --git a/Mathlib/Data/ZMod/Basic.lean b/Mathlib/Data/ZMod/Basic.lean index a1afa41447cfb..587f8afc8aa50 100644 --- a/Mathlib/Data/ZMod/Basic.lean +++ b/Mathlib/Data/ZMod/Basic.lean @@ -435,7 +435,7 @@ theorem castHom_bijective [Fintype R] (h : Fintype.card R = n) : intro hn rw [hn] at h exact (Fintype.card_eq_zero_iff.mp h).elim' 0⟩ - rw [Fintype.bijective_iff_injective_and_card, ZMod.card, h, eq_self_iff_true, and_true_iff] + rw [Fintype.bijective_iff_injective_and_card, ZMod.card, h, eq_self_iff_true, and_true] apply ZMod.castHom_injective /-- The unique ring isomorphism between `ZMod n` and a ring `R` diff --git a/Mathlib/FieldTheory/Cardinality.lean b/Mathlib/FieldTheory/Cardinality.lean index 498e3b8d326c7..8a0482682912f 100644 --- a/Mathlib/FieldTheory/Cardinality.lean +++ b/Mathlib/FieldTheory/Cardinality.lean @@ -75,5 +75,5 @@ theorem Field.nonempty_iff {α : Type u} : Nonempty (Field α) ↔ IsPrimePow # rw [Cardinal.isPrimePow_iff] cases' fintypeOrInfinite α with h h · simpa only [Cardinal.mk_fintype, Nat.cast_inj, exists_eq_left', - (Cardinal.nat_lt_aleph0 _).not_le, false_or_iff] using Fintype.nonempty_field_iff - · simpa only [← Cardinal.infinite_iff, h, true_or_iff, iff_true_iff] using Infinite.nonempty_field + (Cardinal.nat_lt_aleph0 _).not_le, false_or] using Fintype.nonempty_field_iff + · simpa only [← Cardinal.infinite_iff, h, true_or, iff_true] using Infinite.nonempty_field diff --git a/Mathlib/FieldTheory/Finite/Basic.lean b/Mathlib/FieldTheory/Finite/Basic.lean index a6f4ffff2c1aa..6b5c4835a87c6 100644 --- a/Mathlib/FieldTheory/Finite/Basic.lean +++ b/Mathlib/FieldTheory/Finite/Basic.lean @@ -293,7 +293,7 @@ theorem sum_pow_lt_card_sub_one (i : ℕ) (h : i < q - 1) : ∑ x : K, x ^ i = 0 let φ : Kˣ ↪ K := ⟨fun x ↦ x, Units.ext⟩ have : univ.map φ = univ \ {0} := by ext x - simpa only [mem_map, mem_univ, Function.Embedding.coeFn_mk, true_and_iff, mem_sdiff, + simpa only [mem_map, mem_univ, Function.Embedding.coeFn_mk, true_and, mem_sdiff, mem_singleton, φ] using isUnit_iff_ne_zero calc ∑ x : K, x ^ i = ∑ x ∈ univ \ {(0 : K)}, x ^ i := by diff --git a/Mathlib/FieldTheory/IsAlgClosed/Classification.lean b/Mathlib/FieldTheory/IsAlgClosed/Classification.lean index 1a9f0ce2972f9..eadc63373a37b 100644 --- a/Mathlib/FieldTheory/IsAlgClosed/Classification.lean +++ b/Mathlib/FieldTheory/IsAlgClosed/Classification.lean @@ -54,7 +54,7 @@ theorem cardinal_mk_le_sigma_polynomial : intro h simp? at h says simp only [Set.coe_setOf, ne_eq, Set.mem_setOf_eq, Sigma.mk.inj_iff] at h refine (Subtype.heq_iff_coe_eq ?_).1 h.2 - simp only [h.1, iff_self_iff, forall_true_iff] + simp only [h.1, forall_true_iff] /-- The cardinality of an algebraic extension is at most the maximum of the cardinality of the base ring or `ℵ₀` -/ diff --git a/Mathlib/Geometry/Euclidean/Angle/Oriented/Affine.lean b/Mathlib/Geometry/Euclidean/Angle/Oriented/Affine.lean index 1a4c7dfdf186b..40018bcc026e4 100644 --- a/Mathlib/Geometry/Euclidean/Angle/Oriented/Affine.lean +++ b/Mathlib/Geometry/Euclidean/Angle/Oriented/Affine.lean @@ -622,7 +622,7 @@ theorem _root_.Collinear.oangle_sign_of_sameRay_vsub {p₁ p₂ p₃ p₄ : P} ( have hf : ContinuousOn (fun p : P × P × P => ∡ p.1 p.2.1 p.2.2) s := by refine ContinuousAt.continuousOn fun p hp => continuousAt_oangle ?_ ?_ all_goals - simp_rw [s, Set.mem_image, Set.mem_prod, Set.mem_univ, true_and_iff, Prod.ext_iff] at hp + simp_rw [s, Set.mem_image, Set.mem_prod, Set.mem_univ, true_and, Prod.ext_iff] at hp obtain ⟨q₁, q₅, q₂⟩ := p dsimp only at hp ⊢ obtain ⟨⟨⟨q, hq⟩, v⟩, hv, rfl, rfl, rfl⟩ := hp @@ -638,7 +638,7 @@ theorem _root_.Collinear.oangle_sign_of_sameRay_vsub {p₁ p₂ p₃ p₄ : P} ( exact smul_vsub_rev_mem_vectorSpan_pair _ _ _ have hsp : ∀ p : P × P × P, p ∈ s → ∡ p.1 p.2.1 p.2.2 ≠ 0 ∧ ∡ p.1 p.2.1 p.2.2 ≠ π := by intro p hp - simp_rw [s, Set.mem_image, Set.mem_prod, Set.mem_setOf, Set.mem_univ, true_and_iff, + simp_rw [s, Set.mem_image, Set.mem_prod, Set.mem_setOf, Set.mem_univ, true_and, Prod.ext_iff] at hp obtain ⟨q₁, q₅, q₂⟩ := p dsimp only at hp ⊢ @@ -656,13 +656,13 @@ theorem _root_.Collinear.oangle_sign_of_sameRay_vsub {p₁ p₂ p₃ p₄ : P} ( rw [direction_affineSpan] exact smul_vsub_rev_mem_vectorSpan_pair _ _ _ have hp₁p₂s : (p₁, p₅, p₂) ∈ s := by - simp_rw [s, Set.mem_image, Set.mem_prod, Set.mem_setOf, Set.mem_univ, true_and_iff, + simp_rw [s, Set.mem_image, Set.mem_prod, Set.mem_setOf, Set.mem_univ, true_and, Prod.ext_iff] refine ⟨⟨⟨p₁, left_mem_affineSpan_pair ℝ _ _⟩, p₂ -ᵥ p₁⟩, ⟨SameRay.rfl, vsub_ne_zero.2 hp₁p₂.symm⟩, ?_⟩ simp have hp₃p₄s : (p₃, p₅, p₄) ∈ s := by - simp_rw [s, Set.mem_image, Set.mem_prod, Set.mem_setOf, Set.mem_univ, true_and_iff, + simp_rw [s, Set.mem_image, Set.mem_prod, Set.mem_setOf, Set.mem_univ, true_and, Prod.ext_iff] refine ⟨⟨⟨p₃, hc.mem_affineSpan_of_mem_of_ne (Set.mem_insert _ _) (Set.mem_insert_of_mem _ (Set.mem_insert _ _)) diff --git a/Mathlib/Geometry/Euclidean/Basic.lean b/Mathlib/Geometry/Euclidean/Basic.lean index 3431ffc86bbca..3ac078670754f 100644 --- a/Mathlib/Geometry/Euclidean/Basic.lean +++ b/Mathlib/Geometry/Euclidean/Basic.lean @@ -171,13 +171,13 @@ theorem eq_of_dist_eq_of_dist_eq_of_mem_of_finrank_eq_two {s : AffineSubspace exact ⟨t₁, t₂, hv⟩ rcases hv (p -ᵥ p₁) (vsub_mem_direction hps hp₁s) with ⟨t₁, t₂, hpt⟩ simp only [hpt, inner_add_right, inner_smul_right, ho, mul_zero, add_zero, - mul_eq_zero, inner_self_eq_zero, vsub_eq_zero_iff_eq, hc.symm, or_false_iff] at hop + mul_eq_zero, inner_self_eq_zero, vsub_eq_zero_iff_eq, hc.symm, or_false] at hop rw [hop, zero_smul, zero_add, ← eq_vadd_iff_vsub_eq] at hpt subst hpt have hp' : (p₂ -ᵥ p₁ : V) ≠ 0 := by simp [hp.symm] have hp₂ : dist ((1 : ℝ) • (p₂ -ᵥ p₁) +ᵥ p₁) c₁ = r₁ := by simp [hp₂c₁] rw [← hp₁c₁, dist_smul_vadd_eq_dist _ _ hp'] at hpc₁ hp₂ - simp only [one_ne_zero, false_or_iff] at hp₂ + simp only [one_ne_zero, false_or] at hp₂ rw [hp₂.symm] at hpc₁ cases' hpc₁ with hpc₁ hpc₁ <;> simp [hpc₁] diff --git a/Mathlib/Geometry/Euclidean/Circumcenter.lean b/Mathlib/Geometry/Euclidean/Circumcenter.lean index 137f90c80106a..ed19fa6831d20 100644 --- a/Mathlib/Geometry/Euclidean/Circumcenter.lean +++ b/Mathlib/Geometry/Euclidean/Circumcenter.lean @@ -316,11 +316,11 @@ theorem eq_circumradius_of_dist_eq {n : ℕ} (s : Simplex ℝ P n) {p : P} r = s.circumradius := by have h := s.circumsphere_unique_dist_eq.2 ⟨p, r⟩ simp only [hp, hr, forall_const, eq_self_iff_true, subset_sphere, Sphere.ext_iff, - Set.forall_mem_range, mem_sphere, true_and_iff] at h + Set.forall_mem_range, mem_sphere] at h -- Porting note: added the next three lines (`simp` less powerful) rw [subset_sphere (s := ⟨p, r⟩)] at h simp only [hp, hr, forall_const, eq_self_iff_true, subset_sphere, Sphere.ext_iff, - Set.forall_mem_range, mem_sphere, true_and_iff] at h + Set.forall_mem_range, mem_sphere, true_and] at h exact h.2 /-- The circumradius is non-negative. -/ @@ -866,7 +866,7 @@ theorem eq_or_eq_reflection_of_dist_eq {n : ℕ} {s : Simplex ℝ P n} {p p₁ p by_cases hp : p = s.orthogonalProjectionSpan p · rw [Simplex.orthogonalProjectionSpan] at hp rw [hp₁, hp₂, ← hp] - simp only [true_or_iff, eq_self_iff_true, smul_zero, vsub_self] + simp only [true_or, eq_self_iff_true, smul_zero, vsub_self] · have hz : ⟪p -ᵥ orthogonalProjection span_s p, p -ᵥ orthogonalProjection span_s p⟫ ≠ 0 := by simpa only [Ne, vsub_eq_zero_iff_eq, inner_self_eq_zero] using hp rw [mul_left_inj' hz, mul_self_eq_mul_self_iff] at hd₁ diff --git a/Mathlib/Geometry/Manifold/ContMDiff/Basic.lean b/Mathlib/Geometry/Manifold/ContMDiff/Basic.lean index ac0ff5b1c6d65..c8a972fc9462a 100644 --- a/Mathlib/Geometry/Manifold/ContMDiff/Basic.lean +++ b/Mathlib/Geometry/Manifold/ContMDiff/Basic.lean @@ -64,13 +64,13 @@ theorem ContMDiffWithinAt.comp {t : Set M'} {g : M' → M''} (x : M) filter_upwards [hf.1.tendsto (extChartAt_source_mem_nhds I' (f x)), inter_mem_nhdsWithin s (extChartAt_source_mem_nhds I x)] rintro x' (hfx' : f x' ∈ e'.source) ⟨hx's, hx'⟩ - simp only [e.map_source hx', true_and_iff, e.left_inv hx', st hx's, *] + simp only [e.map_source hx', true_and, e.left_inv hx', st hx's, *] refine ((hg.2.comp _ (hf.2.mono inter_subset_right) inter_subset_left).mono_of_mem (inter_mem ?_ self_mem_nhdsWithin)).congr_of_eventuallyEq ?_ ?_ · filter_upwards [A] rintro x' ⟨ht, hfx'⟩ simp only [*, mem_preimage, writtenInExtChartAt, (· ∘ ·), mem_inter_iff, e'.left_inv, - true_and_iff] + true_and] exact mem_range_self _ · filter_upwards [A] rintro x' ⟨-, hfx'⟩ diff --git a/Mathlib/Geometry/Manifold/Instances/Real.lean b/Mathlib/Geometry/Manifold/Instances/Real.lean index 32997ffec1950..5a704749a28a5 100644 --- a/Mathlib/Geometry/Manifold/Instances/Real.lean +++ b/Mathlib/Geometry/Manifold/Instances/Real.lean @@ -151,7 +151,7 @@ model for manifolds with corners -/ def modelWithCornersEuclideanQuadrant (n : ℕ) : ModelWithCorners ℝ (EuclideanSpace ℝ (Fin n)) (EuclideanQuadrant n) where toFun := Subtype.val - invFun x := ⟨fun i => max (x i) 0, fun i => by simp only [le_refl, or_true_iff, le_max_iff]⟩ + invFun x := ⟨fun i => max (x i) 0, fun i => by simp only [le_refl, or_true, le_max_iff]⟩ source := univ target := { x | ∀ i, 0 ≤ x i } map_source' x _ := x.property diff --git a/Mathlib/Geometry/Manifold/LocalInvariantProperties.lean b/Mathlib/Geometry/Manifold/LocalInvariantProperties.lean index 0fff8f9f28819..f1f298984aaa0 100644 --- a/Mathlib/Geometry/Manifold/LocalInvariantProperties.lean +++ b/Mathlib/Geometry/Manifold/LocalInvariantProperties.lean @@ -302,7 +302,7 @@ theorem liftPropWithinAt_indep_chart_target [HasGroupoid M' G'] (hf : f ∈ G'.m LiftPropWithinAt P g s x ↔ ContinuousWithinAt g s x ∧ LiftPropWithinAt P (f ∘ g) s x := by rw [liftPropWithinAt_self_target, liftPropWithinAt_iff', and_congr_right_iff] intro hg - simp_rw [(f.continuousAt xf).comp_continuousWithinAt hg, true_and_iff] + simp_rw [(f.continuousAt xf).comp_continuousWithinAt hg, true_and] exact hG.liftPropWithinAt_indep_chart_target_aux (mem_chart_source _ _) (chart_mem_maximalAtlas _ _) (mem_chart_source _ _) hf xf hg @@ -430,7 +430,7 @@ theorem liftPropOn_of_liftProp (mono : ∀ ⦃s x t⦄ ⦃f : H → H'⦄, t ⊆ theorem liftPropAt_of_mem_maximalAtlas [HasGroupoid M G] (hG : G.LocalInvariantProp G Q) (hQ : ∀ y, Q id univ y) (he : e ∈ maximalAtlas M G) (hx : x ∈ e.source) : LiftPropAt Q e x := by simp_rw [LiftPropAt, hG.liftPropWithinAt_indep_chart he hx G.id_mem_maximalAtlas (mem_univ _), - (e.continuousAt hx).continuousWithinAt, true_and_iff] + (e.continuousAt hx).continuousWithinAt, true_and] exact hG.congr' (e.eventually_right_inverse' hx) (hQ _) theorem liftPropOn_of_mem_maximalAtlas [HasGroupoid M G] (hG : G.LocalInvariantProp G Q) @@ -482,7 +482,7 @@ theorem liftPropOn_of_mem_groupoid (hG : G.LocalInvariantProp G Q) (hQ : ∀ y, theorem liftProp_id (hG : G.LocalInvariantProp G Q) (hQ : ∀ y, Q id univ y) : LiftProp Q (id : M → M) := by - simp_rw [liftProp_iff, continuous_id, true_and_iff] + simp_rw [liftProp_iff, continuous_id, true_and] exact fun x ↦ hG.congr' ((chartAt H x).eventually_right_inverse <| mem_chart_target H x) (hQ _) theorem liftPropAt_iff_comp_subtype_val (hG : LocalInvariantProp G G' P) {U : Opens M} diff --git a/Mathlib/Geometry/Manifold/VectorBundle/FiberwiseLinear.lean b/Mathlib/Geometry/Manifold/VectorBundle/FiberwiseLinear.lean index 4b107cd53d598..2a0e077c860aa 100644 --- a/Mathlib/Geometry/Manifold/VectorBundle/FiberwiseLinear.lean +++ b/Mathlib/Geometry/Manifold/VectorBundle/FiberwiseLinear.lean @@ -122,7 +122,7 @@ theorem SmoothFiberwiseLinear.locality_aux₁ (e : PartialHomeomorph (B × F) (B have hu' : ∀ p : e.source, (p : B × F).fst ∈ u p := by intro p have : (p : B × F) ∈ e.source ∩ s p := ⟨p.prop, hsp p⟩ - simpa only [hesu, mem_prod, mem_univ, and_true_iff] using this + simpa only [hesu, mem_prod, mem_univ, and_true] using this have heu : ∀ p : e.source, ∀ q : B × F, q.fst ∈ u p → q ∈ e.source := by intro p q hq have : q ∈ u p ×ˢ (univ : Set F) := ⟨hq, trivial⟩ diff --git a/Mathlib/GroupTheory/CoprodI.lean b/Mathlib/GroupTheory/CoprodI.lean index e0f94293ea1b6..67d5fdc6738bd 100644 --- a/Mathlib/GroupTheory/CoprodI.lean +++ b/Mathlib/GroupTheory/CoprodI.lean @@ -364,7 +364,7 @@ theorem rcons_inj {i} : Function.Injective (rcons : Pair M i → Word M) := by rw [← he] at h' exact h' rfl · have : m = m' ∧ w.toList = w'.toList := by - simpa [cons, rcons, dif_neg hm, dif_neg hm', true_and_iff, eq_self_iff_true, Subtype.mk_eq_mk, + simpa [cons, rcons, dif_neg hm, dif_neg hm', eq_self_iff_true, Subtype.mk_eq_mk, heq_iff_eq, ← Subtype.ext_iff_val] using he rcases this with ⟨rfl, h⟩ congr diff --git a/Mathlib/GroupTheory/DoubleCoset.lean b/Mathlib/GroupTheory/DoubleCoset.lean index 126c81594b8d8..5973451afc4e5 100644 --- a/Mathlib/GroupTheory/DoubleCoset.lean +++ b/Mathlib/GroupTheory/DoubleCoset.lean @@ -140,7 +140,7 @@ theorem disjoint_out' {H K : Subgroup G} {a b : Quotient H.1 K} : theorem union_quotToDoset (H K : Subgroup G) : ⋃ q, quotToDoset H K q = Set.univ := by ext x simp only [Set.mem_iUnion, quotToDoset, mem_doset, SetLike.mem_coe, exists_prop, Set.mem_univ, - iff_true_iff] + iff_true] use mk H K x obtain ⟨h, k, h3, h4, h5⟩ := mk_out'_eq_mul H K x refine ⟨h⁻¹, H.inv_mem h3, k⁻¹, K.inv_mem h4, ?_⟩ diff --git a/Mathlib/GroupTheory/FreeAbelianGroupFinsupp.lean b/Mathlib/GroupTheory/FreeAbelianGroupFinsupp.lean index 8cc24301aaf7c..477433e03185a 100644 --- a/Mathlib/GroupTheory/FreeAbelianGroupFinsupp.lean +++ b/Mathlib/GroupTheory/FreeAbelianGroupFinsupp.lean @@ -153,7 +153,7 @@ theorem support_zsmul (k : ℤ) (h : k ≠ 0) (a : FreeAbelianGroup X) : support (k • a) = support a := by ext x simp only [mem_support_iff, AddMonoidHom.map_zsmul] - simp only [h, zsmul_int_int, false_or_iff, Ne, mul_eq_zero] + simp only [h, zsmul_int_int, false_or, Ne, mul_eq_zero] @[simp] theorem support_nsmul (k : ℕ) (h : k ≠ 0) (a : FreeAbelianGroup X) : diff --git a/Mathlib/GroupTheory/Index.lean b/Mathlib/GroupTheory/Index.lean index 0fc048268ebf6..9085d649c8d8c 100644 --- a/Mathlib/GroupTheory/Index.lean +++ b/Mathlib/GroupTheory/Index.lean @@ -158,9 +158,9 @@ theorem index_eq_two_iff : H.index = 2 ↔ ∃ a, ∀ b, Xor' (b * a ∈ H) (b @[to_additive] theorem mul_mem_iff_of_index_two (h : H.index = 2) {a b : G} : a * b ∈ H ↔ (a ∈ H ↔ b ∈ H) := by - by_cases ha : a ∈ H; · simp only [ha, true_iff_iff, mul_mem_cancel_left ha] - by_cases hb : b ∈ H; · simp only [hb, iff_true_iff, mul_mem_cancel_right hb] - simp only [ha, hb, iff_self_iff, iff_true_iff] + by_cases ha : a ∈ H; · simp only [ha, true_iff, mul_mem_cancel_left ha] + by_cases hb : b ∈ H; · simp only [hb, iff_true, mul_mem_cancel_right hb] + simp only [ha, hb, iff_true] rcases index_eq_two_iff.1 h with ⟨c, hc⟩ refine (hc _).or.resolve_left ?_ rwa [mul_assoc, mul_mem_cancel_right ((hc _).or.resolve_right hb)] diff --git a/Mathlib/GroupTheory/Nilpotent.lean b/Mathlib/GroupTheory/Nilpotent.lean index a4b435fb7cba6..5e2d9a0bdeadd 100644 --- a/Mathlib/GroupTheory/Nilpotent.lean +++ b/Mathlib/GroupTheory/Nilpotent.lean @@ -285,7 +285,7 @@ instance lowerCentralSeries_normal (n : ℕ) : Normal (lowerCentralSeries G n) : theorem lowerCentralSeries_antitone : Antitone (lowerCentralSeries G) := by refine antitone_nat_of_succ_le fun n x hx => ?_ simp only [mem_lowerCentralSeries_succ_iff, exists_prop, mem_top, exists_true_left, - true_and_iff] at hx + true_and] at hx refine closure_induction hx ?_ (Subgroup.one_mem _) (@Subgroup.mul_mem _ _ _) (@Subgroup.inv_mem _ _ _) rintro y ⟨z, hz, a, ha⟩ diff --git a/Mathlib/GroupTheory/Perm/Basic.lean b/Mathlib/GroupTheory/Perm/Basic.lean index 4ff40a3a4a271..7392c6b050125 100644 --- a/Mathlib/GroupTheory/Perm/Basic.lean +++ b/Mathlib/GroupTheory/Perm/Basic.lean @@ -396,7 +396,7 @@ theorem ofSubtype_apply_of_not_mem (f : Perm (Subtype p)) (ha : ¬p a) : ofSubty theorem mem_iff_ofSubtype_apply_mem (f : Perm (Subtype p)) (x : α) : p x ↔ p ((ofSubtype f : α → α) x) := if h : p x then by - simpa only [h, true_iff_iff, MonoidHom.coe_mk, ofSubtype_apply_of_mem f h] using (f ⟨x, h⟩).2 + simpa only [h, true_iff, MonoidHom.coe_mk, ofSubtype_apply_of_mem f h] using (f ⟨x, h⟩).2 else by simp [h, ofSubtype_apply_of_not_mem f h] @[simp] diff --git a/Mathlib/GroupTheory/Perm/Cycle/Factors.lean b/Mathlib/GroupTheory/Perm/Cycle/Factors.lean index 48526c90ffbdf..d6af59c007b46 100644 --- a/Mathlib/GroupTheory/Perm/Cycle/Factors.lean +++ b/Mathlib/GroupTheory/Perm/Cycle/Factors.lean @@ -225,7 +225,7 @@ theorem mem_support_cycleOf_iff [DecidableEq α] [Fintype α] : simp [hx] · rw [mem_support, cycleOf_apply] split_ifs with hy - · simp only [hx, hy, iff_true_iff, Ne, not_false_iff, and_self_iff, mem_support] + · simp only [hx, hy, Ne, not_false_iff, and_self_iff, mem_support] rcases hy with ⟨k, rfl⟩ rw [← not_mem_support] simpa using hx @@ -308,7 +308,7 @@ def cycleFactorsAux [DecidableEq α] [Fintype α] (l : List α) (f : Perm α) { l : List (Perm α) // l.prod = f ∧ (∀ g ∈ l, IsCycle g) ∧ l.Pairwise Disjoint } := match l with | [] => ⟨[], by - { simp only [imp_false, List.Pairwise.nil, List.not_mem_nil, forall_const, and_true_iff, + { simp only [imp_false, List.Pairwise.nil, List.not_mem_nil, forall_const, and_true, forall_prop_of_false, Classical.not_not, not_false_iff, List.prod_nil] at * ext simp [*]}⟩ diff --git a/Mathlib/GroupTheory/Perm/Sign.lean b/Mathlib/GroupTheory/Perm/Sign.lean index f4a6c982a61b4..a7d22e67b1e55 100644 --- a/Mathlib/GroupTheory/Perm/Sign.lean +++ b/Mathlib/GroupTheory/Perm/Sign.lean @@ -159,7 +159,7 @@ def finPairsLT (n : ℕ) : Finset (Σ_ : Fin n, Fin n) := (univ : Finset (Fin n)).sigma fun a => (range a).attachFin fun _ hm => (mem_range.1 hm).trans a.2 theorem mem_finPairsLT {n : ℕ} {a : Σ_ : Fin n, Fin n} : a ∈ finPairsLT n ↔ a.2 < a.1 := by - simp only [finPairsLT, Fin.lt_iff_val_lt_val, true_and_iff, mem_attachFin, mem_range, mem_univ, + simp only [finPairsLT, Fin.lt_iff_val_lt_val, true_and, mem_attachFin, mem_range, mem_univ, mem_sigma] /-- `signAux σ` is the sign of a permutation on `Fin n`, defined as the parity of the number of @@ -255,7 +255,7 @@ private theorem signAux_swap_zero_one' (n : ℕ) : signAux (swap (0 : Fin (n + 2 rcases a₁.zero_le.eq_or_lt with (rfl | H) · exact absurd a₂.zero_le ha₁.not_le rcases a₂.zero_le.eq_or_lt with (rfl | H') - · simp only [and_true_iff, eq_self_iff_true, heq_iff_eq, mem_singleton, Sigma.mk.inj_iff] at ha₂ + · simp only [and_true, eq_self_iff_true, heq_iff_eq, mem_singleton, Sigma.mk.inj_iff] at ha₂ have : 1 < a₁ := lt_of_le_of_ne (Nat.succ_le_of_lt ha₁) (Ne.symm (by intro h; apply ha₂; simp [h])) have h01 : Equiv.swap (0 : Fin (n + 2)) 1 0 = 1 := by simp @@ -526,7 +526,7 @@ theorem prod_prodExtendRight {α : Type*} [DecidableEq α] (σ : α → Perm β) · rw [← ha'] at * refine Or.inl ⟨l.mem_cons_self a, ?_⟩ rw [prodExtendRight_apply_eq] - · refine Or.inr ⟨fun h => not_or_of_not ha' not_mem_l ((List.mem_cons).mp h), ?_⟩ + · refine Or.inr ⟨fun h => not_or_intro ha' not_mem_l ((List.mem_cons).mp h), ?_⟩ rw [prodExtendRight_apply_ne _ ha'] section congr diff --git a/Mathlib/GroupTheory/Perm/Support.lean b/Mathlib/GroupTheory/Perm/Support.lean index 63593ca5a6715..fce28868366c4 100644 --- a/Mathlib/GroupTheory/Perm/Support.lean +++ b/Mathlib/GroupTheory/Perm/Support.lean @@ -279,7 +279,7 @@ theorem coe_support_eq_set_support (f : Perm α) : (f.support : Set α) = { x | @[simp] theorem support_eq_empty_iff {σ : Perm α} : σ.support = ∅ ↔ σ = 1 := by - simp_rw [Finset.ext_iff, mem_support, Finset.not_mem_empty, iff_false_iff, not_not, + simp_rw [Finset.ext_iff, mem_support, Finset.not_mem_empty, iff_false, not_not, Equiv.Perm.ext_iff, one_apply] @[simp] @@ -390,7 +390,7 @@ theorem support_swap_iff (x y : α) : support (swap x y) = {x, y} ↔ x ≠ y := theorem support_swap_mul_swap {x y z : α} (h : List.Nodup [x, y, z]) : support (swap x y * swap y z) = {x, y, z} := by - simp only [List.not_mem_nil, and_true_iff, List.mem_cons, not_false_iff, List.nodup_cons, + simp only [List.not_mem_nil, and_true, List.mem_cons, not_false_iff, List.nodup_cons, List.mem_singleton, and_self_iff, List.nodup_nil] at h push_neg at h apply le_antisymm @@ -487,7 +487,7 @@ theorem support_extend_domain (f : α ≃ Subtype p) {g : Perm α} : rw [eq_symm_apply] exact Subtype.coe_injective ha · rw [extendDomain_apply_not_subtype _ _ pb] - simp only [not_exists, false_iff_iff, not_and, eq_self_iff_true, not_true] + simp only [not_exists, false_iff, not_and, eq_self_iff_true, not_true] rintro a _ rfl exact pb (Subtype.prop _) diff --git a/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean b/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean index d77dab25be0a0..eda2a253359f6 100644 --- a/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean +++ b/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean @@ -428,7 +428,7 @@ private theorem card_orderOf_eq_totient_aux₁ : refine Finset.sum_congr rfl fun m hm => ?_ simp only [mem_filter, mem_range, mem_properDivisors] at hm refine IH m hm.2 (hm.1.trans hd) (Finset.card_pos.2 ⟨a ^ (d / m), ?_⟩) - simp only [mem_filter, mem_univ, orderOf_pow a, ha, true_and_iff, + simp only [mem_filter, mem_univ, orderOf_pow a, ha, true_and, Nat.gcd_eq_right (div_dvd_of_dvd hm.1), Nat.div_div_self hm.1 hd0] have h2 : (∑ m ∈ d.divisors, (univ.filter fun a : α => orderOf a = m).card) = diff --git a/Mathlib/Init/Logic.lean b/Mathlib/Init/Logic.lean index 69a4350392f2e..d5eb7baedafbf 100644 --- a/Mathlib/Init/Logic.lean +++ b/Mathlib/Init/Logic.lean @@ -69,27 +69,33 @@ attribute [symm] Iff.symm variable (p) --- FIXME: remove _iff and add _eq for the lean 4 core versions +@[deprecated and_true (since := "2024-09-12")] theorem and_true_iff : p ∧ True ↔ p := iff_of_eq (and_true _) +@[deprecated true_and (since := "2024-09-12")] theorem true_and_iff : True ∧ p ↔ p := iff_of_eq (true_and _) +@[deprecated and_false (since := "2024-09-12")] theorem and_false_iff : p ∧ False ↔ False := iff_of_eq (and_false _) +@[deprecated false_and (since := "2024-09-12")] theorem false_and_iff : False ∧ p ↔ False := iff_of_eq (false_and _) - -theorem true_or_iff : True ∨ p ↔ True := iff_of_eq (true_or _) +@[deprecated or_true (since := "2024-09-12")] theorem or_true_iff : p ∨ True ↔ True := iff_of_eq (or_true _) -theorem false_or_iff : False ∨ p ↔ p := iff_of_eq (false_or _) +@[deprecated true_or (since := "2024-09-12")] +theorem true_or_iff : True ∨ p ↔ True := iff_of_eq (true_or _) +@[deprecated or_false (since := "2024-09-12")] theorem or_false_iff : p ∨ False ↔ p := iff_of_eq (or_false _) - -theorem not_or_of_not : ¬a → ¬b → ¬(a ∨ b) := fun h1 h2 ↦ not_or.2 ⟨h1, h2⟩ - +@[deprecated false_or (since := "2024-09-12")] +theorem false_or_iff : False ∨ p ↔ p := iff_of_eq (false_or _) +@[deprecated iff_true (since := "2024-09-12")] theorem iff_true_iff : (a ↔ True) ↔ a := iff_of_eq (iff_true _) +@[deprecated true_iff (since := "2024-09-12")] theorem true_iff_iff : (True ↔ a) ↔ a := iff_of_eq (true_iff _) - +@[deprecated iff_false (since := "2024-09-12")] theorem iff_false_iff : (a ↔ False) ↔ ¬a := iff_of_eq (iff_false _) - +@[deprecated false_iff (since := "2024-09-12")] theorem false_iff_iff : (False ↔ a) ↔ ¬a := iff_of_eq (false_iff _) - +@[deprecated iff_self (since := "2024-09-12")] theorem iff_self_iff (a : Prop) : (a ↔ a) ↔ True := iff_of_eq (iff_self _) +@[deprecated (since := "2024-09-12")] alias not_or_of_not := not_or_intro /- decidable -/ diff --git a/Mathlib/LinearAlgebra/Basis/Cardinality.lean b/Mathlib/LinearAlgebra/Basis/Cardinality.lean index 56f77ef131e1c..284ad5ed30aa4 100644 --- a/Mathlib/LinearAlgebra/Basis/Cardinality.lean +++ b/Mathlib/LinearAlgebra/Basis/Cardinality.lean @@ -54,7 +54,7 @@ lemma basis_finite_of_finite_spans (w : Set M) (hw : w.Finite) (s : span R w = intro x m rw [← b.linearCombination_repr x, span_image_eq_map_linearCombination, Submodule.mem_map] use b.repr x - simp only [and_true_iff, eq_self_iff_true, Finsupp.mem_supported] + simp only [and_true, eq_self_iff_true, Finsupp.mem_supported] rw [Finset.coe_subset, ← Finset.le_iff_subset] exact Finset.le_sup (f := fun x : w ↦ (b.repr ↑x).support) (Finset.mem_univ (⟨x, m⟩ : w)) -- Thus this finite subset of the basis elements spans the entire module. diff --git a/Mathlib/LinearAlgebra/FinsuppVectorSpace.lean b/Mathlib/LinearAlgebra/FinsuppVectorSpace.lean index 6c96e6b6ad6ae..a25ac55d3ce5c 100644 --- a/Mathlib/LinearAlgebra/FinsuppVectorSpace.lean +++ b/Mathlib/LinearAlgebra/FinsuppVectorSpace.lean @@ -110,8 +110,7 @@ theorem coe_basis {φ : ι → Type*} (b : ∀ i, Basis (φ i) R M) : Finsupp.single_apply_left sigma_mk_injective] · have : Sigma.mk i x ≠ Sigma.mk j y := fun h' => h <| congrArg (fun s => s.fst) h' -- Porting note: previously `this` not needed - simp only [basis_repr, single_apply, h, this, false_and_iff, if_false, LinearEquiv.map_zero, - zero_apply] + simp only [basis_repr, single_apply, h, this, if_false, LinearEquiv.map_zero, zero_apply] /-- The basis on `ι →₀ M` with basis vectors `fun i ↦ single i 1`. -/ @[simps] diff --git a/Mathlib/LinearAlgebra/FreeModule/PID.lean b/Mathlib/LinearAlgebra/FreeModule/PID.lean index ea581374ba3ee..c5c0346f9a5de 100644 --- a/Mathlib/LinearAlgebra/FreeModule/PID.lean +++ b/Mathlib/LinearAlgebra/FreeModule/PID.lean @@ -234,7 +234,7 @@ theorem Submodule.basis_of_pid_aux [Finite ι] {O : Type*} [AddCommGroup O] [Mod rw [LinearMap.mem_ker] at hx' have hc' : (c • ⟨y', y'M⟩ + ⟨x, xM⟩ : M) = 0 := by exact @Subtype.coe_injective O (· ∈ M) _ _ hc simpa only [LinearMap.map_add, LinearMap.map_zero, LinearMap.map_smul, smul_eq_mul, add_zero, - mul_eq_zero, ϕy'_ne_zero, hx', or_false_iff] using congr_arg ϕ hc' + mul_eq_zero, ϕy'_ne_zero, hx', or_false] using congr_arg ϕ hc' -- And `a • y'` is orthogonal to `N'`. have ay'_ortho_N' : ∀ (c : R), ∀ z ∈ N', c • a • y' + z = 0 → c = 0 := by intro c z zN' hc diff --git a/Mathlib/LinearAlgebra/Lagrange.lean b/Mathlib/LinearAlgebra/Lagrange.lean index 982898e6d7590..ab7d725e81277 100644 --- a/Mathlib/LinearAlgebra/Lagrange.lean +++ b/Mathlib/LinearAlgebra/Lagrange.lean @@ -145,7 +145,7 @@ theorem basisDivisor_self : basisDivisor x x = 0 := by simp only [basisDivisor, sub_self, inv_zero, map_zero, zero_mul] theorem basisDivisor_inj (hxy : basisDivisor x y = 0) : x = y := by - simp_rw [basisDivisor, mul_eq_zero, X_sub_C_ne_zero, or_false_iff, C_eq_zero, inv_eq_zero, + simp_rw [basisDivisor, mul_eq_zero, X_sub_C_ne_zero, or_false, C_eq_zero, inv_eq_zero, sub_eq_zero] at hxy exact hxy diff --git a/Mathlib/LinearAlgebra/LinearPMap.lean b/Mathlib/LinearAlgebra/LinearPMap.lean index a22821e0e63c5..bae8ae629f0f7 100644 --- a/Mathlib/LinearAlgebra/LinearPMap.lean +++ b/Mathlib/LinearAlgebra/LinearPMap.lean @@ -847,7 +847,7 @@ theorem le_graph_of_le {f g : E →ₗ.[R] F} (h : f ≤ g) : f.graph ≤ g.grap rw [mem_graph_iff] at hx ⊢ cases' hx with y hx use ⟨y, h.1 y.2⟩ - simp only [hx, Submodule.coe_mk, eq_self_iff_true, true_and_iff] + simp only [hx, Submodule.coe_mk, eq_self_iff_true, true_and] convert hx.2 using 1 refine (h.2 ?_).symm simp only [hx.1, Submodule.coe_mk] diff --git a/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean b/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean index 85a32f9416529..f29c7888f644b 100644 --- a/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean +++ b/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean @@ -136,7 +136,7 @@ theorem det_mul (M N : Matrix n n R) : det (M * N) = det M * det N := ∑ σ : Perm n, ε σ * ∏ i, M (σ i) (p i) * N (p i) i := (Eq.symm <| sum_subset (filter_subset _ _) fun f _ hbij => - det_mul_aux <| by simpa only [true_and_iff, mem_filter, mem_univ] using hbij) + det_mul_aux <| by simpa only [true_and, mem_filter, mem_univ] using hbij) _ = ∑ τ : Perm n, ∑ σ : Perm n, ε σ * ∏ i, M (σ i) (τ i) * N (τ i) i := sum_bij (fun p h ↦ Equiv.ofBijective p (mem_filter.1 h).2) (fun _ _ ↦ mem_univ _) (fun _ _ _ _ h ↦ by injection h) diff --git a/Mathlib/LinearAlgebra/Matrix/Polynomial.lean b/Mathlib/LinearAlgebra/Matrix/Polynomial.lean index 72e1f37352e62..64469ff133624 100644 --- a/Mathlib/LinearAlgebra/Matrix/Polynomial.lean +++ b/Mathlib/LinearAlgebra/Matrix/Polynomial.lean @@ -38,7 +38,7 @@ theorem natDegree_det_X_add_C_le (A B : Matrix n n α) : rw [det_apply] refine (natDegree_sum_le _ _).trans ?_ refine Multiset.max_le_of_forall_le _ _ ?_ - simp only [forall_apply_eq_imp_iff, true_and_iff, Function.comp_apply, Multiset.map_map, + simp only [forall_apply_eq_imp_iff, true_and, Function.comp_apply, Multiset.map_map, Multiset.mem_map, exists_imp, Finset.mem_univ_val] intro g calc diff --git a/Mathlib/LinearAlgebra/Matrix/Transvection.lean b/Mathlib/LinearAlgebra/Matrix/Transvection.lean index c360c3c4d9f0f..9a13c74d20d49 100644 --- a/Mathlib/LinearAlgebra/Matrix/Transvection.lean +++ b/Mathlib/LinearAlgebra/Matrix/Transvection.lean @@ -101,7 +101,7 @@ theorem updateRow_eq_transvection [Finite n] (c : R) : StdBasisMatrix.apply_of_ne] · simp only [updateRow_ne, transvection, ha, Ne.symm ha, StdBasisMatrix.apply_of_ne, add_zero, Algebra.id.smul_eq_mul, Ne, not_false_iff, DMatrix.add_apply, Pi.smul_apply, - mul_zero, false_and_iff, add_apply] + mul_zero, false_and, add_apply] variable [Fintype n] @@ -424,7 +424,7 @@ theorem listTransvecCol_mul_last_col (hM : M (inr unit) (inr unit) ≠ 0) (i : F rcases le_or_lt (n + 1) i with (hi | hi) · simp only [hi, n.le_succ.trans hi, if_true] · rw [if_neg, if_neg] - · simpa only [hni.symm, not_le, or_false_iff] using Nat.lt_succ_iff_lt_or_eq.1 hi + · simpa only [hni.symm, not_le, or_false] using Nat.lt_succ_iff_lt_or_eq.1 hi · simpa only [not_le] using hi | self => simp only [length_listTransvecCol, le_refl, List.drop_eq_nil_of_le, List.prod_nil, @@ -500,7 +500,7 @@ theorem mul_listTransvecRow_last_row (hM : M (inr unit) (inr unit) ≠ 0) (i : F · simp [hi, n.le_succ.trans hi, if_true] · rw [if_neg, if_neg] · simpa only [not_le] using hi - · simpa only [hni.symm, not_le, or_false_iff] using Nat.lt_succ_iff_lt_or_eq.1 hi + · simpa only [hni.symm, not_le, or_false] using Nat.lt_succ_iff_lt_or_eq.1 hi /-- Multiplying by all the matrices either in `listTransvecCol M` and `listTransvecRow M` kills all the coefficients in the last row but the last one. -/ diff --git a/Mathlib/LinearAlgebra/Multilinear/Basic.lean b/Mathlib/LinearAlgebra/Multilinear/Basic.lean index 417f0dde15dc6..cf3f3f112438e 100644 --- a/Mathlib/LinearAlgebra/Multilinear/Basic.lean +++ b/Mathlib/LinearAlgebra/Multilinear/Basic.lean @@ -522,7 +522,7 @@ theorem map_sum_finset_aux [DecidableEq ι] [Fintype ι] {n : ℕ} (h : (∑ i, simpa [C] using hj rw [this] simp only [B, mem_sdiff, eq_self_iff_true, not_true, not_false_iff, Finset.mem_singleton, - update_same, and_false_iff] + update_same, and_false] · simp [hi] have Beq : Function.update (fun i => ∑ j ∈ A i, g i j) i₀ (∑ j ∈ B i₀, g i₀ j) = fun i => diff --git a/Mathlib/LinearAlgebra/Prod.lean b/Mathlib/LinearAlgebra/Prod.lean index 4db6ffb5b5ed8..fb36914029e72 100644 --- a/Mathlib/LinearAlgebra/Prod.lean +++ b/Mathlib/LinearAlgebra/Prod.lean @@ -449,7 +449,7 @@ theorem ker_coprod_of_disjoint_range {M₂ : Type*} [AddCommGroup M₂] [Module rintro ⟨y, z⟩ h simp only [mem_ker, mem_prod, coprod_apply] at h ⊢ have : f y ∈ (range f) ⊓ (range g) := by - simp only [true_and_iff, mem_range, mem_inf, exists_apply_eq_apply] + simp only [true_and, mem_range, mem_inf, exists_apply_eq_apply] use -z rwa [eq_comm, map_neg, ← sub_eq_zero, sub_neg_eq_add] rw [hd.eq_bot, mem_bot] at this diff --git a/Mathlib/LinearAlgebra/Ray.lean b/Mathlib/LinearAlgebra/Ray.lean index 117d5b1c1a242..87778718498eb 100644 --- a/Mathlib/LinearAlgebra/Ray.lean +++ b/Mathlib/LinearAlgebra/Ray.lean @@ -466,7 +466,7 @@ theorem sameRay_smul_right_iff {v : M} {r : R} : SameRay R v (r • v) ↔ 0 ≤ is positive. -/ theorem sameRay_smul_right_iff_of_ne {v : M} (hv : v ≠ 0) {r : R} (hr : r ≠ 0) : SameRay R v (r • v) ↔ 0 < r := by - simp only [sameRay_smul_right_iff, hv, or_false_iff, hr.symm.le_iff_lt] + simp only [sameRay_smul_right_iff, hv, or_false, hr.symm.le_iff_lt] @[simp] theorem sameRay_smul_left_iff {v : M} {r : R} : SameRay R (r • v) v ↔ 0 ≤ r ∨ v = 0 := @@ -484,7 +484,7 @@ theorem sameRay_neg_smul_right_iff {v : M} {r : R} : SameRay R (-v) (r • v) theorem sameRay_neg_smul_right_iff_of_ne {v : M} {r : R} (hv : v ≠ 0) (hr : r ≠ 0) : SameRay R (-v) (r • v) ↔ r < 0 := by - simp only [sameRay_neg_smul_right_iff, hv, or_false_iff, hr.le_iff_lt] + simp only [sameRay_neg_smul_right_iff, hv, or_false, hr.le_iff_lt] @[simp] theorem sameRay_neg_smul_left_iff {v : M} {r : R} : SameRay R (r • v) (-v) ↔ r ≤ 0 ∨ v = 0 := diff --git a/Mathlib/LinearAlgebra/TensorProduct/Basic.lean b/Mathlib/LinearAlgebra/TensorProduct/Basic.lean index bcb3daf37d060..42c39d96c9f27 100644 --- a/Mathlib/LinearAlgebra/TensorProduct/Basic.lean +++ b/Mathlib/LinearAlgebra/TensorProduct/Basic.lean @@ -449,7 +449,7 @@ variable (R M N) /-- The simple (aka pure) elements span the tensor product. -/ theorem span_tmul_eq_top : Submodule.span R { t : M ⊗[R] N | ∃ m n, m ⊗ₜ n = t } = ⊤ := by - ext t; simp only [Submodule.mem_top, iff_true_iff] + ext t; simp only [Submodule.mem_top, iff_true] refine t.induction_on ?_ ?_ ?_ · exact Submodule.zero_mem _ · intro m n diff --git a/Mathlib/Logic/Denumerable.lean b/Mathlib/Logic/Denumerable.lean index 94274fd7d2531..cf0e64c97bd1d 100644 --- a/Mathlib/Logic/Denumerable.lean +++ b/Mathlib/Logic/Denumerable.lean @@ -294,7 +294,7 @@ private theorem right_inverse_aux : ∀ n, toFunAux (ofNat s n) = n simp only [Finset.ext_iff, mem_insert, mem_range, mem_filter] exact fun m => ⟨fun h => by - simp only [h.2, and_true_iff] + simp only [h.2, and_true] exact Or.symm (lt_or_eq_of_le ((@lt_succ_iff_le _ _ _ ⟨m, h.2⟩ _).1 h.1)), fun h => h.elim (fun h => h.symm ▸ ⟨lt_succ_self _, (ofNat s n).prop⟩) fun h => diff --git a/Mathlib/Logic/Equiv/Option.lean b/Mathlib/Logic/Equiv/Option.lean index 65c8c77633e6a..9c0c3d02003ee 100644 --- a/Mathlib/Logic/Equiv/Option.lean +++ b/Mathlib/Logic/Equiv/Option.lean @@ -130,7 +130,7 @@ theorem some_removeNone_iff {x : α} : some (removeNone e x) = e none ↔ e.symm · rw [removeNone_some _ ⟨a, h⟩] have h1 := congr_arg e.symm h rw [symm_apply_apply] at h1 - simp only [false_iff_iff, apply_eq_iff_eq, reduceCtorEq] + simp only [apply_eq_iff_eq, reduceCtorEq] simp [h1, apply_eq_iff_eq] @[simp] diff --git a/Mathlib/MeasureTheory/Constructions/BorelSpace/Order.lean b/Mathlib/MeasureTheory/Constructions/BorelSpace/Order.lean index 42987950d650c..425393be2aff9 100644 --- a/Mathlib/MeasureTheory/Constructions/BorelSpace/Order.lean +++ b/Mathlib/MeasureTheory/Constructions/BorelSpace/Order.lean @@ -281,7 +281,7 @@ theorem Dense.borel_eq_generateFrom_Ico_mem {α : Type*} [TopologicalSpace α] [ theorem borel_eq_generateFrom_Ico (α : Type*) [TopologicalSpace α] [SecondCountableTopology α] [LinearOrder α] [OrderTopology α] : borel α = .generateFrom { S : Set α | ∃ (l u : α), l < u ∧ Ico l u = S } := by - simpa only [exists_prop, mem_univ, true_and_iff] using + simpa only [exists_prop, mem_univ, true_and] using (@dense_univ α _).borel_eq_generateFrom_Ico_mem_aux (fun _ _ => mem_univ _) fun _ _ _ _ => mem_univ _ @@ -307,7 +307,7 @@ theorem Dense.borel_eq_generateFrom_Ioc_mem {α : Type*} [TopologicalSpace α] [ theorem borel_eq_generateFrom_Ioc (α : Type*) [TopologicalSpace α] [SecondCountableTopology α] [LinearOrder α] [OrderTopology α] : borel α = .generateFrom { S : Set α | ∃ l u, l < u ∧ Ioc l u = S } := by - simpa only [exists_prop, mem_univ, true_and_iff] using + simpa only [exists_prop, mem_univ, true_and] using (@dense_univ α _).borel_eq_generateFrom_Ioc_mem_aux (fun _ _ => mem_univ _) fun _ _ _ _ => mem_univ _ @@ -925,8 +925,8 @@ theorem measure_eq_measure_preimage_add_measure_tsum_Ico_zpow {α : Type*} [Meas ext x simp only [mem_singleton_iff, mem_union, mem_Ioo, mem_Ioi, mem_preimage] obtain (H | H) : f x = ∞ ∨ f x < ∞ := eq_or_lt_of_le le_top - · simp only [H, eq_self_iff_true, or_false_iff, ENNReal.zero_lt_top, not_top_lt, and_false] - · simp only [H, H.ne, and_true_iff, false_or_iff] + · simp only [H, eq_self_iff_true, or_false, ENNReal.zero_lt_top, not_top_lt, and_false] + · simp only [H, H.ne, and_true, false_or] · refine disjoint_left.2 fun x hx h'x => ?_ have : f x < ∞ := h'x.2.2 exact lt_irrefl _ (this.trans_le (le_of_eq hx.2.symm)) diff --git a/Mathlib/MeasureTheory/Constructions/BorelSpace/Real.lean b/Mathlib/MeasureTheory/Constructions/BorelSpace/Real.lean index b566424678fd2..a4d7400e4d89f 100644 --- a/Mathlib/MeasureTheory/Constructions/BorelSpace/Real.lean +++ b/Mathlib/MeasureTheory/Constructions/BorelSpace/Real.lean @@ -480,7 +480,7 @@ theorem exists_spanning_measurableSet_le {α : Type*} {m : MeasurableSpace α} { let norm_sets := fun n : ℕ => { x | f x ≤ n } have norm_sets_spanning : ⋃ n, norm_sets n = Set.univ := by ext1 x - simp only [Set.mem_iUnion, Set.mem_setOf_eq, Set.mem_univ, iff_true_iff] + simp only [Set.mem_iUnion, Set.mem_setOf_eq, Set.mem_univ, iff_true] exact exists_nat_ge (f x) let sets n := sigma_finite_sets n ∩ norm_sets n have h_meas : ∀ n, MeasurableSet (sets n) := by diff --git a/Mathlib/MeasureTheory/Constructions/Prod/Integral.lean b/Mathlib/MeasureTheory/Constructions/Prod/Integral.lean index 87e289fa69c30..25d335405cc0f 100644 --- a/Mathlib/MeasureTheory/Constructions/Prod/Integral.lean +++ b/Mathlib/MeasureTheory/Constructions/Prod/Integral.lean @@ -61,7 +61,7 @@ along one of the variables (using either the Lebesgue or Bochner integral) is me theorem measurableSet_integrable [SFinite ν] ⦃f : α → β → E⦄ (hf : StronglyMeasurable (uncurry f)) : MeasurableSet {x | Integrable (f x) ν} := by - simp_rw [Integrable, hf.of_uncurry_left.aestronglyMeasurable, true_and_iff] + simp_rw [Integrable, hf.of_uncurry_left.aestronglyMeasurable, true_and] exact measurableSet_lt (Measurable.lintegral_prod_right hf.ennnorm) measurable_const section diff --git a/Mathlib/MeasureTheory/Covering/Besicovitch.lean b/Mathlib/MeasureTheory/Covering/Besicovitch.lean index 5fc2bf194df58..466e9d9155efd 100644 --- a/Mathlib/MeasureTheory/Covering/Besicovitch.lean +++ b/Mathlib/MeasureTheory/Covering/Besicovitch.lean @@ -340,7 +340,7 @@ theorem color_lt {i : Ordinal.{u}} (hi : i < p.lastStep) {N : ℕ} have color_i : p.color i = sInf (univ \ A) := by rw [color] rw [color_i] have N_mem : N ∈ univ \ A := by - simp only [A, not_exists, true_and_iff, exists_prop, mem_iUnion, mem_singleton_iff, + simp only [A, not_exists, true_and, exists_prop, mem_iUnion, mem_singleton_iff, mem_closedBall, not_and, mem_univ, mem_diff, Subtype.exists, Subtype.coe_mk] intro j ji _ exact (IH j ji (ji.trans hi)).ne' @@ -356,7 +356,7 @@ theorem color_lt {i : Ordinal.{u}} (hi : i < p.lastStep) {N : ℕ} intro k hk rw [← Inf_eq_N] at hk have : k ∈ A := by - simpa only [true_and_iff, mem_univ, Classical.not_not, mem_diff] using + simpa only [true_and, mem_univ, Classical.not_not, mem_diff] using Nat.not_mem_of_lt_sInf hk simp only [mem_iUnion, mem_singleton_iff, exists_prop, Subtype.exists, exists_and_right, and_assoc] at this @@ -491,15 +491,15 @@ theorem exist_disjoint_covering_families {N : ℕ} {τ : ℝ} (hτ : 1 < τ) rw [color_j] apply csInf_mem refine ⟨N, ?_⟩ - simp only [A, not_exists, true_and_iff, exists_prop, mem_iUnion, mem_singleton_iff, not_and, + simp only [A, not_exists, true_and, exists_prop, mem_iUnion, mem_singleton_iff, not_and, mem_univ, mem_diff, Subtype.exists, Subtype.coe_mk] intro k hk _ exact (p.color_lt (hk.trans jy_lt) hN).ne' - simp only [A, not_exists, true_and_iff, exists_prop, mem_iUnion, mem_singleton_iff, not_and, + simp only [A, not_exists, true_and, exists_prop, mem_iUnion, mem_singleton_iff, not_and, mem_univ, mem_diff, Subtype.exists, Subtype.coe_mk] at h specialize h jx jxy contrapose! h - simpa only [jxi, jyi, and_true_iff, eq_self_iff_true, ← not_disjoint_iff_nonempty_inter] using h + simpa only [jxi, jyi, and_true, eq_self_iff_true, ← not_disjoint_iff_nonempty_inter] using h · -- show that the balls of color at most `N` cover every center. refine range_subset_iff.2 fun b => ?_ obtain ⟨a, ha⟩ : @@ -708,7 +708,7 @@ theorem exists_disjoint_closedBall_covering_ae_of_finiteMeasure_aux (μ : Measur exist_finset_disjoint_balls_large_measure μ hτ hN s' r (fun x hx => (rI x hx).1) fun x hx => (rI x hx).2.le refine ⟨t ∪ Finset.image (fun x => (x, r x)) v, Finset.subset_union_left, ⟨?_, ?_, ?_⟩, ?_⟩ - · simp only [Finset.coe_union, pairwiseDisjoint_union, ht.1, true_and_iff, Finset.coe_image] + · simp only [Finset.coe_union, pairwiseDisjoint_union, ht.1, true_and, Finset.coe_image] constructor · intro p hp q hq hpq rcases (mem_image _ _ _).1 hp with ⟨p', p'v, rfl⟩ @@ -781,8 +781,8 @@ theorem exists_disjoint_closedBall_covering_ae_of_finiteMeasure_aux (μ : Measur rw [ENNReal.div_lt_iff, one_mul] · conv_lhs => rw [← add_zero (N : ℝ≥0∞)] exact ENNReal.add_lt_add_left (ENNReal.natCast_ne_top N) zero_lt_one - · simp only [true_or_iff, add_eq_zero, Ne, not_false_iff, one_ne_zero, and_false_iff] - · simp only [ENNReal.natCast_ne_top, Ne, not_false_iff, or_true_iff] + · simp only [true_or, add_eq_zero, Ne, not_false_iff, one_ne_zero, and_false] + · simp only [ENNReal.natCast_ne_top, Ne, not_false_iff, or_true] rw [zero_mul] at C apply le_bot_iff.1 exact le_of_tendsto_of_tendsto' tendsto_const_nhds C fun n => (A n).trans (B n) @@ -921,7 +921,7 @@ theorem exists_closedBall_covering_tsum_measure_le (μ : Measure α) [SFinite μ refine ⟨t0 ∪ ⋃ i : Fin N, ((↑) : s' → α) '' S i, r, ?_, ?_, ?_, ?_, ?_⟩ -- it remains to check that they have the desired properties · exact t0_count.union (countable_iUnion fun i => (S_count i).image _) - · simp only [t0s, true_and_iff, union_subset_iff, image_subset_iff, iUnion_subset_iff] + · simp only [t0s, true_and, union_subset_iff, image_subset_iff, iUnion_subset_iff] intro i x _ exact s's x.2 · intro x hx diff --git a/Mathlib/MeasureTheory/Covering/BesicovitchVectorSpace.lean b/Mathlib/MeasureTheory/Covering/BesicovitchVectorSpace.lean index 97a0b59d00769..60ca4af5a4851 100644 --- a/Mathlib/MeasureTheory/Covering/BesicovitchVectorSpace.lean +++ b/Mathlib/MeasureTheory/Covering/BesicovitchVectorSpace.lean @@ -461,13 +461,13 @@ theorem exists_normalized {N : ℕ} {τ : ℝ} (a : SatelliteConfig E N τ) (las · simp_rw [c', Hj, hij.trans Hj, if_true] exact exists_normalized_aux1 a lastr hτ δ hδ1 hδ2 i j inej -- case `2 < ‖c j‖` - · have H'j : ‖a.c j‖ ≤ 2 ↔ False := by simpa only [not_le, iff_false_iff] using Hj + · have H'j : ‖a.c j‖ ≤ 2 ↔ False := by simpa only [not_le, iff_false] using Hj rcases le_or_lt ‖a.c i‖ 2 with (Hi | Hi) · -- case `‖c i‖ ≤ 2` simp_rw [c', Hi, if_true, H'j, if_false] exact exists_normalized_aux2 a lastc lastr hτ δ hδ1 hδ2 i j inej Hi Hj · -- case `2 < ‖c i‖` - have H'i : ‖a.c i‖ ≤ 2 ↔ False := by simpa only [not_le, iff_false_iff] using Hi + have H'i : ‖a.c i‖ ≤ 2 ↔ False := by simpa only [not_le, iff_false] using Hi simp_rw [c', H'i, if_false, H'j, if_false] exact exists_normalized_aux3 a lastc lastr hτ δ hδ1 i j inej Hi hij diff --git a/Mathlib/MeasureTheory/Covering/DensityTheorem.lean b/Mathlib/MeasureTheory/Covering/DensityTheorem.lean index 35d1a28b7705d..51f94e3d8e12d 100644 --- a/Mathlib/MeasureTheory/Covering/DensityTheorem.lean +++ b/Mathlib/MeasureTheory/Covering/DensityTheorem.lean @@ -69,7 +69,7 @@ theorem closedBall_mem_vitaliFamily_of_dist_le_mul {K : ℝ} {x y : α} {r : ℝ (rpos : 0 < r) : closedBall y r ∈ (vitaliFamily μ K).setsAt x := by let R := scalingScaleOf μ (max (4 * K + 3) 3) simp only [vitaliFamily, VitaliFamily.enlarge, Vitali.vitaliFamily, mem_union, mem_setOf_eq, - isClosed_ball, true_and_iff, (nonempty_ball.2 rpos).mono ball_subset_interior_closedBall, + isClosed_ball, true_and, (nonempty_ball.2 rpos).mono ball_subset_interior_closedBall, measurableSet_closedBall] /- The measure is doubling on scales smaller than `R`. Therefore, we treat differently small and large balls. For large balls, this follows directly from the enlargement we used in the diff --git a/Mathlib/MeasureTheory/Covering/Differentiation.lean b/Mathlib/MeasureTheory/Covering/Differentiation.lean index abe8892caee43..4af0418c40a8e 100644 --- a/Mathlib/MeasureTheory/Covering/Differentiation.lean +++ b/Mathlib/MeasureTheory/Covering/Differentiation.lean @@ -245,7 +245,7 @@ theorem ae_tendsto_div : ∀ᵐ x ∂μ, ∃ c, Tendsto (fun a => ρ a / μ a) ( intro x h1x _ apply h1x.mono fun a ha => ?_ refine (ENNReal.div_le_iff_le_mul ?_ (Or.inr (bot_le.trans_lt ha).ne')).1 ha.le - simp only [ENNReal.coe_ne_top, Ne, or_true_iff, not_false_iff] + simp only [ENNReal.coe_ne_top, Ne, or_true, not_false_iff] · simp only [and_imp, exists_prop, not_frequently, not_and, not_lt, not_le, not_eventually, mem_setOf_eq, mem_compl_iff, not_forall] intro x _ h2x @@ -324,7 +324,7 @@ theorem exists_measurable_supersets_limRatio {p q : ℝ≥0} (hpq : p < q) : (toMeasurable μ sᶜ ∪ ⋃ n, toMeasurable (ρ + μ) (w n)) ⊆ toMeasurable μ sᶜ ∪ ⋃ (m) (n), toMeasurable (ρ + μ) (u m) ∩ toMeasurable (ρ + μ) (w n) := by - simp only [inter_union_distrib_left, union_inter_distrib_right, true_and_iff, + simp only [inter_union_distrib_left, union_inter_distrib_right, true_and, subset_union_left, union_subset_iff, inter_self] refine ⟨?_, ?_, ?_⟩ · exact inter_subset_right.trans subset_union_left @@ -368,7 +368,7 @@ theorem exists_measurable_supersets_limRatio {p q : ℝ≥0} (hpq : p < q) : apply I.frequently.mono fun a ha => ?_ rw [coe_nnreal_smul_apply] refine (ENNReal.div_le_iff_le_mul ?_ (Or.inr (bot_le.trans_lt ha).ne')).1 ha.le - simp only [ENNReal.coe_ne_top, Ne, or_true_iff, not_false_iff] + simp only [ENNReal.coe_ne_top, Ne, or_true, not_false_iff] _ = p * μ (toMeasurable (ρ + μ) (u m) ∩ toMeasurable (ρ + μ) (w n)) := by simp only [coe_nnreal_smul_apply, measure_toMeasurable_add_inter_right (measurableSet_toMeasurable _ _) I] @@ -448,7 +448,7 @@ theorem measure_le_mul_of_subset_limRatioMeas_lt {p : ℝ≥0} {s : Set α} apply I.frequently.mono fun a ha => ?_ rw [coe_nnreal_smul_apply] refine (ENNReal.div_le_iff_le_mul ?_ (Or.inr (bot_le.trans_lt ha).ne')).1 ha.le - simp only [ENNReal.coe_ne_top, Ne, or_true_iff, not_false_iff] + simp only [ENNReal.coe_ne_top, Ne, or_true, not_false_iff] /-- If, for all `x` in a set `s`, one has frequently `q < ρ a / μ a`, then `q * μ s ≤ ρ s`, as proved in `measure_le_of_frequently_le`. Since `ρ a / μ a` tends almost everywhere to @@ -485,7 +485,7 @@ theorem measure_limRatioMeas_top : μ {x | v.limRatioMeas hρ x = ∞} = 0 := by intro y hy have : v.limRatioMeas hρ y = ∞ := hy.1 simp only [this, ENNReal.coe_lt_top, mem_setOf_eq] - · simp only [(zero_lt_one.trans_le hq).ne', true_or_iff, ENNReal.coe_eq_zero, Ne, + · simp only [(zero_lt_one.trans_le hq).ne', true_or, ENNReal.coe_eq_zero, Ne, not_false_iff] have B : Tendsto (fun q : ℝ≥0 => (q : ℝ≥0∞)⁻¹ * ρ s) atTop (𝓝 (∞⁻¹ * ρ s)) := by apply ENNReal.Tendsto.mul_const _ (Or.inr ρs) @@ -649,8 +649,8 @@ theorem withDensity_limRatioMeas_eq : μ.withDensity (v.limRatioMeas hρ) = ρ : ((t : ℝ≥0∞) ^ 2 * ρ s : ℝ≥0∞)) (𝓝[>] 1) (𝓝 ((1 : ℝ≥0∞) ^ 2 * ρ s)) := by refine ENNReal.Tendsto.mul ?_ ?_ tendsto_const_nhds ?_ · exact ENNReal.Tendsto.pow (ENNReal.tendsto_coe.2 nhdsWithin_le_nhds) - · simp only [one_pow, ENNReal.coe_one, true_or_iff, Ne, not_false_iff, one_ne_zero] - · simp only [one_pow, ENNReal.coe_one, Ne, or_true_iff, ENNReal.one_ne_top, not_false_iff] + · simp only [one_pow, ENNReal.coe_one, true_or, Ne, not_false_iff, one_ne_zero] + · simp only [one_pow, ENNReal.coe_one, Ne, or_true, ENNReal.one_ne_top, not_false_iff] simp only [one_pow, one_mul, ENNReal.coe_one] at this refine ge_of_tendsto this ?_ filter_upwards [self_mem_nhdsWithin] with _ ht @@ -659,7 +659,7 @@ theorem withDensity_limRatioMeas_eq : μ.withDensity (v.limRatioMeas hρ) = ρ : Tendsto (fun t : ℝ≥0 => (t : ℝ≥0∞) * μ.withDensity (v.limRatioMeas hρ) s) (𝓝[>] 1) (𝓝 ((1 : ℝ≥0∞) * μ.withDensity (v.limRatioMeas hρ) s)) := by refine ENNReal.Tendsto.mul_const (ENNReal.tendsto_coe.2 nhdsWithin_le_nhds) ?_ - simp only [ENNReal.coe_one, true_or_iff, Ne, not_false_iff, one_ne_zero] + simp only [ENNReal.coe_one, true_or, Ne, not_false_iff, one_ne_zero] simp only [one_mul, ENNReal.coe_one] at this refine ge_of_tendsto this ?_ filter_upwards [self_mem_nhdsWithin] with _ ht @@ -809,7 +809,7 @@ theorem ae_tendsto_lintegral_nnnorm_sub_div'_of_integrable {f : α → E} (hf : _ < ∞ + ∞ := haveI I : Integrable ((A.set n).indicator fun _ : α => c) μ := by simp only [integrable_indicator_iff (IsOpen.measurableSet (A.set_mem n)), - integrableOn_const, A.finite n, or_true_iff] + integrableOn_const, A.finite n, or_true] ENNReal.add_lt_add hf.2 I.2 filter_upwards [main, v.ae_eventually_measure_pos] with x hx h'x have M : diff --git a/Mathlib/MeasureTheory/Decomposition/SignedHahn.lean b/Mathlib/MeasureTheory/Decomposition/SignedHahn.lean index 3891d1175c65d..5ad19400c81ce 100644 --- a/Mathlib/MeasureTheory/Decomposition/SignedHahn.lean +++ b/Mathlib/MeasureTheory/Decomposition/SignedHahn.lean @@ -247,7 +247,7 @@ private theorem exists_subset_restrict_nonpos' (hi₁ : MeasurableSet i) (hi₂ rw [Set.mem_iUnion, exists_prop, and_iff_right_iff_imp] exact fun _ => h · convert le_of_eq s.empty.symm - ext; simp only [exists_prop, Set.mem_empty_iff_false, Set.mem_iUnion, not_and, iff_false_iff] + ext; simp only [exists_prop, Set.mem_empty_iff_false, Set.mem_iUnion, not_and, iff_false] exact fun h' => False.elim (h h') · intro; exact MeasurableSet.iUnion fun _ => restrictNonposSeq_measurableSet _ · intro a b hab diff --git a/Mathlib/MeasureTheory/Function/AEEqOfIntegral.lean b/Mathlib/MeasureTheory/Function/AEEqOfIntegral.lean index bb4121b69171b..37c95e1795cb8 100644 --- a/Mathlib/MeasureTheory/Function/AEEqOfIntegral.lean +++ b/Mathlib/MeasureTheory/Function/AEEqOfIntegral.lean @@ -180,7 +180,7 @@ theorem ae_le_of_forall_setLIntegral_le_of_sigmaFinite₀ [SigmaFinite μ] have B : (∫⁻ x in s, g x ∂μ) ≠ ∞ := (setLIntegral_lt_top_of_le_nnreal s_lt_top.ne ⟨N, fun _ h ↦ h.1.2⟩).ne have : (ε : ℝ≥0∞) * μ s ≤ 0 := ENNReal.le_of_add_le_add_left B A - simpa only [ENNReal.coe_eq_zero, nonpos_iff_eq_zero, mul_eq_zero, εpos.ne', false_or_iff] + simpa only [ENNReal.coe_eq_zero, nonpos_iff_eq_zero, mul_eq_zero, εpos.ne', false_or] obtain ⟨u, _, u_pos, u_lim⟩ : ∃ u : ℕ → ℝ≥0, StrictAnti u ∧ (∀ n, 0 < u n) ∧ Tendsto u atTop (𝓝 0) := exists_seq_strictAnti_tendsto (0 : ℝ≥0) diff --git a/Mathlib/MeasureTheory/Function/ConditionalExpectation/Basic.lean b/Mathlib/MeasureTheory/Function/ConditionalExpectation/Basic.lean index 04cb417cd938d..daab2ab684f35 100644 --- a/Mathlib/MeasureTheory/Function/ConditionalExpectation/Basic.lean +++ b/Mathlib/MeasureTheory/Function/ConditionalExpectation/Basic.lean @@ -109,7 +109,7 @@ theorem condexp_of_sigmaFinite (hm : m ≤ m0) [hμm : SigmaFinite (μ.trim hm)] else aestronglyMeasurable'_condexpL1.mk (condexpL1 hm μ f) else 0 := by rw [condexp, dif_pos hm] - simp only [hμm, Ne, true_and_iff] + simp only [hμm, Ne, true_and] by_cases hf : Integrable f μ · rw [dif_pos hf, if_pos hf] · rw [dif_neg hf, if_neg hf] diff --git a/Mathlib/MeasureTheory/Function/ConditionalExpectation/CondexpL1.lean b/Mathlib/MeasureTheory/Function/ConditionalExpectation/CondexpL1.lean index 779e0ed7f4ba6..a0b70b96c1884 100644 --- a/Mathlib/MeasureTheory/Function/ConditionalExpectation/CondexpL1.lean +++ b/Mathlib/MeasureTheory/Function/ConditionalExpectation/CondexpL1.lean @@ -176,11 +176,11 @@ theorem condexpIndL1_of_measurableSet_of_measure_ne_top (hs : MeasurableSet s) ( theorem condexpIndL1_of_measure_eq_top (hμs : μ s = ∞) (x : G) : condexpIndL1 hm μ s x = 0 := by simp only [condexpIndL1, hμs, eq_self_iff_true, not_true, Ne, dif_neg, not_false_iff, - and_false_iff] + and_false] theorem condexpIndL1_of_not_measurableSet (hs : ¬MeasurableSet s) (x : G) : condexpIndL1 hm μ s x = 0 := by - simp only [condexpIndL1, hs, dif_neg, not_false_iff, false_and_iff] + simp only [condexpIndL1, hs, dif_neg, not_false_iff, false_and] theorem condexpIndL1_add (x y : G) : condexpIndL1 hm μ s (x + y) = condexpIndL1 hm μ s x + condexpIndL1 hm μ s y := by diff --git a/Mathlib/MeasureTheory/Function/ContinuousMapDense.lean b/Mathlib/MeasureTheory/Function/ContinuousMapDense.lean index be9517c3b1db9..633e66b2ab991 100644 --- a/Mathlib/MeasureTheory/Function/ContinuousMapDense.lean +++ b/Mathlib/MeasureTheory/Function/ContinuousMapDense.lean @@ -120,7 +120,7 @@ theorem exists_continuous_eLpNorm_sub_le_of_closed [μ.OuterRegular] (hp : p ≠ have : eLpNorm (v.indicator fun _x => (1 : ℝ)) p μ < ⊤ := by refine (eLpNorm_indicator_const_le _ _).trans_lt ?_ simp only [lt_top_iff_ne_top, hμv.ne, nnnorm_one, ENNReal.coe_one, one_div, one_mul, Ne, - ENNReal.rpow_eq_top_iff, inv_lt_zero, false_and_iff, or_false_iff, not_and, not_lt, + ENNReal.rpow_eq_top_iff, inv_lt_zero, false_and, or_false, not_and, not_lt, ENNReal.toReal_nonneg, imp_true_iff] refine (eLpNorm_mono fun x => ?_).trans_lt this by_cases hx : x ∈ v diff --git a/Mathlib/MeasureTheory/Function/ConvergenceInMeasure.lean b/Mathlib/MeasureTheory/Function/ConvergenceInMeasure.lean index 00e13342ecf7e..31b0277e42b2f 100644 --- a/Mathlib/MeasureTheory/Function/ConvergenceInMeasure.lean +++ b/Mathlib/MeasureTheory/Function/ConvergenceInMeasure.lean @@ -306,7 +306,7 @@ theorem tendstoInMeasure_of_tendsto_eLpNorm_top {E} [NormedAddCommGroup E] {f : specialize hfg (ENNReal.ofReal δ / 2) (ENNReal.div_pos_iff.2 ⟨(ENNReal.ofReal_pos.2 hδ).ne.symm, ENNReal.two_ne_top⟩) refine hfg.mono fun n hn => ?_ - simp only [true_and_iff, gt_iff_lt, zero_tsub, zero_le, zero_add, Set.mem_Icc, + simp only [gt_iff_lt, zero_tsub, zero_le, zero_add, Set.mem_Icc, Pi.sub_apply] at * have : essSup (fun x : α => (‖f n x - g x‖₊ : ℝ≥0∞)) μ < ENNReal.ofReal δ := lt_of_le_of_lt hn diff --git a/Mathlib/MeasureTheory/Function/Jacobian.lean b/Mathlib/MeasureTheory/Function/Jacobian.lean index 38f7a8882415c..2aa9266c5528d 100644 --- a/Mathlib/MeasureTheory/Function/Jacobian.lean +++ b/Mathlib/MeasureTheory/Function/Jacobian.lean @@ -377,7 +377,7 @@ theorem addHaar_image_le_mul_of_det_lt (A : E →L[ℝ] E) {m : ℝ≥0} have L : Tendsto (fun a => (m : ℝ≥0∞) * (μ s + a)) (𝓝[>] 0) (𝓝 (m * (μ s + 0))) := by apply Tendsto.mono_left _ nhdsWithin_le_nhds apply ENNReal.Tendsto.const_mul (tendsto_const_nhds.add tendsto_id) - simp only [ENNReal.coe_ne_top, Ne, or_true_iff, not_false_iff] + simp only [ENNReal.coe_ne_top, Ne, or_true, not_false_iff] rw [add_zero] at L exact ge_of_tendsto L J @@ -420,10 +420,10 @@ theorem mul_le_addHaar_image_of_lt_det (A : E →L[ℝ] E) {m : ℝ≥0} -- record smallness conditions for `δ` that will be needed to apply `hδ₀` below. have L1 : ∀ᶠ δ in 𝓝 (0 : ℝ≥0), Subsingleton E ∨ δ < ‖(B.symm : E →L[ℝ] E)‖₊⁻¹ := by by_cases h : Subsingleton E - · simp only [h, true_or_iff, eventually_const] - simp only [h, false_or_iff] + · simp only [h, true_or, eventually_const] + simp only [h, false_or] apply Iio_mem_nhds - simpa only [h, false_or_iff, inv_pos] using B.subsingleton_or_nnnorm_symm_pos + simpa only [h, false_or, inv_pos] using B.subsingleton_or_nnnorm_symm_pos have L2 : ∀ᶠ δ in 𝓝 (0 : ℝ≥0), ‖(B.symm : E →L[ℝ] E)‖₊ * (‖(B.symm : E →L[ℝ] E)‖₊⁻¹ - δ)⁻¹ * δ < δ₀ := by have : @@ -448,7 +448,7 @@ theorem mul_le_addHaar_image_of_lt_det (A : E →L[ℝ] E) {m : ℝ≥0} mul_comm, ← ENNReal.coe_inv mpos.ne'] · apply Or.inl simpa only [ENNReal.coe_eq_zero, Ne] using mpos.ne' - · simp only [ENNReal.coe_ne_top, true_or_iff, Ne, not_false_iff] + · simp only [ENNReal.coe_ne_top, true_or, Ne, not_false_iff] -- as `f⁻¹` is well approximated by `B⁻¹`, the conclusion follows from `hδ₀` -- and our choice of `δ`. exact hδ₀ _ _ ((hf'.to_inv h1δ).mono_num h2δ.le) @@ -806,7 +806,7 @@ theorem addHaar_image_le_lintegral_abs_det_fderiv_aux1 (hs : MeasurableSet s) rw [dist_eq_norm] calc ‖B - A‖ ≤ (min δ δ'' : ℝ≥0) := hB - _ ≤ δ'' := by simp only [le_refl, NNReal.coe_min, min_le_iff, or_true_iff] + _ ≤ δ'' := by simp only [le_refl, NNReal.coe_min, min_le_iff, or_true] _ < δ' := half_lt_self δ'pos · intro t g htg exact h t g (htg.mono_num (min_le_left _ _)) @@ -941,7 +941,7 @@ theorem lintegral_abs_det_fderiv_le_addHaar_image_aux1 (hs : MeasurableSet s) refine ⟨min δ δ'', lt_min δpos (half_pos δ'pos), ?_, ?_⟩ · intro B hB apply I'' _ (hB.trans _) - simp only [le_refl, NNReal.coe_min, min_le_iff, or_true_iff] + simp only [le_refl, NNReal.coe_min, min_le_iff, or_true] · intro t g htg rcases eq_or_ne (μ t) ∞ with (ht | ht) · simp only [ht, εpos.ne', ENNReal.mul_top, ENNReal.coe_eq_zero, le_top, Ne, diff --git a/Mathlib/MeasureTheory/Function/L1Space.lean b/Mathlib/MeasureTheory/Function/L1Space.lean index 8951025e6bfb6..4e6447b519aa1 100644 --- a/Mathlib/MeasureTheory/Function/L1Space.lean +++ b/Mathlib/MeasureTheory/Function/L1Space.lean @@ -878,8 +878,7 @@ theorem integrable_withDensity_iff_integrable_coe_smul {f : α → ℝ≥0} (hf {g : α → E} : Integrable g (μ.withDensity fun x => f x) ↔ Integrable (fun x => (f x : ℝ) • g x) μ := by by_cases H : AEStronglyMeasurable (fun x : α => (f x : ℝ) • g x) μ - · simp only [Integrable, aestronglyMeasurable_withDensity_iff hf, HasFiniteIntegral, H, - true_and_iff] + · simp only [Integrable, aestronglyMeasurable_withDensity_iff hf, HasFiniteIntegral, H] rw [lintegral_withDensity_eq_lintegral_mul₀' hf.coe_nnreal_ennreal.aemeasurable] · rw [iff_iff_eq] congr @@ -889,7 +888,7 @@ theorem integrable_withDensity_iff_integrable_coe_smul {f : α → ℝ≥0} (hf convert H.ennnorm using 1 ext1 x simp only [nnnorm_smul, NNReal.nnnorm_eq, coe_mul] - · simp only [Integrable, aestronglyMeasurable_withDensity_iff hf, H, false_and_iff] + · simp only [Integrable, aestronglyMeasurable_withDensity_iff hf, H, false_and] theorem integrable_withDensity_iff_integrable_smul {f : α → ℝ≥0} (hf : Measurable f) {g : α → E} : Integrable g (μ.withDensity fun x => f x) ↔ Integrable (fun x => f x • g x) μ := diff --git a/Mathlib/MeasureTheory/Function/LocallyIntegrable.lean b/Mathlib/MeasureTheory/Function/LocallyIntegrable.lean index 7ef4263f03c4a..db051f36a5bc3 100644 --- a/Mathlib/MeasureTheory/Function/LocallyIntegrable.lean +++ b/Mathlib/MeasureTheory/Function/LocallyIntegrable.lean @@ -498,7 +498,7 @@ theorem MonotoneOn.integrableOn_of_measure_ne_top (hmono : MonotoneOn f s) {a b have : IsBounded (f '' s) := Metric.isBounded_of_bddAbove_of_bddBelow habove hbelow rcases isBounded_iff_forall_norm_le.mp this with ⟨C, hC⟩ have A : IntegrableOn (fun _ => C) s μ := by - simp only [hs.lt_top, integrableOn_const, or_true_iff] + simp only [hs.lt_top, integrableOn_const, or_true] exact Integrable.mono' A (aemeasurable_restrict_of_monotoneOn h's hmono).aestronglyMeasurable ((ae_restrict_iff' h's).mpr <| ae_of_all _ fun y hy => hC (f y) (mem_image_of_mem f hy)) diff --git a/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean b/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean index e7150f18e4c57..75f2bb1b23b55 100644 --- a/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean +++ b/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean @@ -404,18 +404,18 @@ theorem eLpNorm_const_lt_top_iff {p : ℝ≥0∞} {c : F} (hp_ne_zero : p ≠ 0) eLpNorm (fun _ : α => c) p μ < ∞ ↔ c = 0 ∨ μ Set.univ < ∞ := by have hp : 0 < p.toReal := ENNReal.toReal_pos hp_ne_zero hp_ne_top by_cases hμ : μ = 0 - · simp only [hμ, Measure.coe_zero, Pi.zero_apply, or_true_iff, ENNReal.zero_lt_top, + · simp only [hμ, Measure.coe_zero, Pi.zero_apply, or_true, ENNReal.zero_lt_top, eLpNorm_measure_zero] by_cases hc : c = 0 - · simp only [hc, true_or_iff, eq_self_iff_true, ENNReal.zero_lt_top, eLpNorm_zero'] + · simp only [hc, true_or, eq_self_iff_true, ENNReal.zero_lt_top, eLpNorm_zero'] rw [eLpNorm_const' c hp_ne_zero hp_ne_top] by_cases hμ_top : μ Set.univ = ∞ · simp [hc, hμ_top, hp] rw [ENNReal.mul_lt_top_iff] - simp only [true_and_iff, one_div, ENNReal.rpow_eq_zero_iff, hμ, false_or_iff, or_false_iff, + simp only [true_and, one_div, ENNReal.rpow_eq_zero_iff, hμ, false_or, or_false, ENNReal.coe_lt_top, nnnorm_eq_zero, ENNReal.coe_eq_zero, - MeasureTheory.Measure.measure_univ_eq_zero, hp, inv_lt_zero, hc, and_false_iff, false_and_iff, - inv_pos, or_self_iff, hμ_top, Ne.lt_top hμ_top, iff_true_iff] + MeasureTheory.Measure.measure_univ_eq_zero, hp, inv_lt_zero, hc, false_and, + inv_pos, or_self_iff, hμ_top, Ne.lt_top hμ_top, iff_true] exact ENNReal.rpow_lt_top_of_nonneg (inv_nonneg.mpr hp.le) hμ_top @[deprecated (since := "2024-07-27")] diff --git a/Mathlib/MeasureTheory/Function/LpSeminorm/ChebyshevMarkov.lean b/Mathlib/MeasureTheory/Function/LpSeminorm/ChebyshevMarkov.lean index 32599fde917d0..6248e4a2c5925 100644 --- a/Mathlib/MeasureTheory/Function/LpSeminorm/ChebyshevMarkov.lean +++ b/Mathlib/MeasureTheory/Function/LpSeminorm/ChebyshevMarkov.lean @@ -34,7 +34,7 @@ theorem mul_meas_ge_le_pow_eLpNorm (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ∞ have : 1 / p.toReal * p.toReal = 1 := by refine one_div_mul_cancel ?_ rw [Ne, ENNReal.toReal_eq_zero_iff] - exact not_or_of_not hp_ne_zero hp_ne_top + exact not_or_intro hp_ne_zero hp_ne_top rw [← ENNReal.rpow_one (ε * μ { x | ε ≤ (‖f x‖₊ : ℝ≥0∞) ^ p.toReal }), ← this, ENNReal.rpow_mul] gcongr exact pow_mul_meas_ge_le_eLpNorm μ hp_ne_zero hp_ne_top hf ε diff --git a/Mathlib/MeasureTheory/Function/SimpleFunc.lean b/Mathlib/MeasureTheory/Function/SimpleFunc.lean index 90386bdd03776..5615eaebce091 100644 --- a/Mathlib/MeasureTheory/Function/SimpleFunc.lean +++ b/Mathlib/MeasureTheory/Function/SimpleFunc.lean @@ -987,7 +987,7 @@ open Finset Function theorem support_eq [MeasurableSpace α] [Zero β] (f : α →ₛ β) : support f = ⋃ y ∈ f.range.filter fun y => y ≠ 0, f ⁻¹' {y} := Set.ext fun x => by - simp only [mem_support, Set.mem_preimage, mem_filter, mem_range_self, true_and_iff, exists_prop, + simp only [mem_support, Set.mem_preimage, mem_filter, mem_range_self, true_and, exists_prop, mem_iUnion, Set.mem_range, mem_singleton_iff, exists_eq_right'] variable {m : MeasurableSpace α} [Zero β] [Zero γ] {μ : Measure α} {f : α →ₛ β} diff --git a/Mathlib/MeasureTheory/Function/SimpleFuncDenseLp.lean b/Mathlib/MeasureTheory/Function/SimpleFuncDenseLp.lean index 679c905a01e48..0386ca8b2e120 100644 --- a/Mathlib/MeasureTheory/Function/SimpleFuncDenseLp.lean +++ b/Mathlib/MeasureTheory/Function/SimpleFuncDenseLp.lean @@ -299,7 +299,7 @@ theorem measure_preimage_lt_top_of_memℒp (hp_pos : p ≠ 0) (hp_ne_top : p ≠ · suffices h_empty : f ⁻¹' {y} = ∅ by rw [h_empty, measure_empty]; exact ENNReal.coe_lt_top ext1 x - rw [Set.mem_preimage, Set.mem_singleton_iff, mem_empty_iff_false, iff_false_iff] + rw [Set.mem_preimage, Set.mem_singleton_iff, mem_empty_iff_false, iff_false] refine fun hxy => hyf ?_ rw [mem_range, Set.mem_range] exact ⟨x, hxy⟩ @@ -773,7 +773,7 @@ theorem denseRange_coeSimpleFuncNonnegToLpNonneg [hp : Fact (1 ≤ p)] (hp_ne_to rw [mem_closure_iff_seq_limit] have hg_memℒp : Memℒp (g : α → G) p μ := Lp.memℒp (g : Lp G p μ) have zero_mem : (0 : G) ∈ (range (g : α → G) ∪ {0} : Set G) ∩ { y | 0 ≤ y } := by - simp only [union_singleton, mem_inter_iff, mem_insert_iff, eq_self_iff_true, true_or_iff, + simp only [union_singleton, mem_inter_iff, mem_insert_iff, eq_self_iff_true, true_or, mem_setOf_eq, le_refl, and_self_iff] have : SeparableSpace ((range (g : α → G) ∪ {0}) ∩ { y | 0 ≤ y } : Set G) := by apply IsSeparable.separableSpace diff --git a/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean b/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean index b2fb9342d5eb9..d4eafeb177727 100644 --- a/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean +++ b/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean @@ -887,7 +887,7 @@ theorem stronglyMeasurable_of_measurableSpace_le_on {α E} {m m₂ : MeasurableS exact MeasurableSet.empty ext1 y simp only [mem_inter_iff, mem_preimage, mem_singleton_iff, mem_compl_iff, - mem_empty_iff_false, iff_false_iff, not_and, not_not_mem] + mem_empty_iff_false, iff_false, not_and, not_not_mem] refine Function.mtr fun hys => ?_ rw [hg_seq_zero y hys n] exact Ne.symm hx diff --git a/Mathlib/MeasureTheory/Function/UniformIntegrable.lean b/Mathlib/MeasureTheory/Function/UniformIntegrable.lean index b55ab2cc0293d..52d0280c4dbba 100644 --- a/Mathlib/MeasureTheory/Function/UniformIntegrable.lean +++ b/Mathlib/MeasureTheory/Function/UniformIntegrable.lean @@ -213,7 +213,7 @@ theorem Memℒp.integral_indicator_norm_ge_le (hf : Memℒp f 1 μ) (hmeas : Str · assumption rw [ENNReal.tendsto_atTop_zero] at this obtain ⟨M, hM⟩ := this (ENNReal.ofReal ε) (ENNReal.ofReal_pos.2 hε) - simp only [true_and_iff, zero_tsub, zero_le, sub_zero, zero_add, coe_nnnorm, + simp only [zero_tsub, zero_le, sub_zero, zero_add, coe_nnnorm, Set.mem_Icc] at hM refine ⟨M, ?_⟩ convert hM M le_rfl diff --git a/Mathlib/MeasureTheory/Group/Measure.lean b/Mathlib/MeasureTheory/Group/Measure.lean index 68d787245397d..66eafadc4d30a 100644 --- a/Mathlib/MeasureTheory/Group/Measure.lean +++ b/Mathlib/MeasureTheory/Group/Measure.lean @@ -518,7 +518,7 @@ theorem null_iff_of_isMulLeftInvariant [Regular μ] {s : Set G} (hs : IsOpen s) μ s = 0 ↔ s = ∅ ∨ μ = 0 := by rcases eq_zero_or_neZero μ with rfl|hμ · simp - · simp only [or_false_iff, hs.measure_eq_zero_iff μ, NeZero.ne μ] + · simp only [or_false, hs.measure_eq_zero_iff μ, NeZero.ne μ] @[to_additive] theorem measure_ne_zero_iff_nonempty_of_isMulLeftInvariant [Regular μ] (hμ : μ ≠ 0) {s : Set G} diff --git a/Mathlib/MeasureTheory/Group/Prod.lean b/Mathlib/MeasureTheory/Group/Prod.lean index 775c06472f6f6..38b8691d5bc0d 100644 --- a/Mathlib/MeasureTheory/Group/Prod.lean +++ b/Mathlib/MeasureTheory/Group/Prod.lean @@ -243,7 +243,7 @@ theorem absolutelyContinuous_of_isMulLeftInvariant [IsMulLeftInvariant ν] (hν refine AbsolutelyContinuous.mk fun s sm hνs => ?_ have h1 := measure_mul_lintegral_eq μ ν sm 1 measurable_one simp_rw [Pi.one_apply, lintegral_one, mul_one, (measure_mul_right_null ν _).mpr hνs, - lintegral_zero, mul_eq_zero (M₀ := ℝ≥0∞), measure_univ_eq_zero.not.mpr hν, or_false_iff] at h1 + lintegral_zero, mul_eq_zero (M₀ := ℝ≥0∞), measure_univ_eq_zero.not.mpr hν, or_false] at h1 exact h1 section SigmaFinite diff --git a/Mathlib/MeasureTheory/Integral/Bochner.lean b/Mathlib/MeasureTheory/Integral/Bochner.lean index e26b72d02951b..43d62ecfb9b43 100644 --- a/Mathlib/MeasureTheory/Integral/Bochner.lean +++ b/Mathlib/MeasureTheory/Integral/Bochner.lean @@ -1044,7 +1044,7 @@ theorem integral_eq_lintegral_of_nonneg_ae {f : α → ℝ} (hf : 0 ≤ᵐ[μ] f · exact measurable_ofReal.comp_aemeasurable hfm.aemeasurable.neg rw [h_min, zero_toReal, _root_.sub_zero] · rw [integral_undef hfi] - simp_rw [Integrable, hfm, hasFiniteIntegral_iff_norm, lt_top_iff_ne_top, Ne, true_and_iff, + simp_rw [Integrable, hfm, hasFiniteIntegral_iff_norm, lt_top_iff_ne_top, Ne, true_and, Classical.not_not] at hfi have : ∫⁻ a : α, ENNReal.ofReal (f a) ∂μ = ∫⁻ a, ENNReal.ofReal ‖f a‖ ∂μ := by refine lintegral_congr_ae (hf.mono fun a h => ?_) @@ -1123,7 +1123,7 @@ theorem integral_nonpos {f : α → ℝ} (hf : f ≤ 0) : ∫ a, f a ∂μ ≤ 0 theorem integral_eq_zero_iff_of_nonneg_ae {f : α → ℝ} (hf : 0 ≤ᵐ[μ] f) (hfi : Integrable f μ) : ∫ x, f x ∂μ = 0 ↔ f =ᵐ[μ] 0 := by simp_rw [integral_eq_lintegral_of_nonneg_ae hf hfi.1, ENNReal.toReal_eq_zero_iff, - ← ENNReal.not_lt_top, ← hasFiniteIntegral_iff_ofReal hf, hfi.2, not_true_eq_false, or_false_iff] + ← ENNReal.not_lt_top, ← hasFiniteIntegral_iff_ofReal hf, hfi.2, not_true_eq_false, or_false] -- Porting note: split into parts, to make `rw` and `simp` work rw [lintegral_eq_zero_iff'] · rw [← hf.le_iff_eq, Filter.EventuallyEq, Filter.EventuallyLE] diff --git a/Mathlib/MeasureTheory/Integral/FundThmCalculus.lean b/Mathlib/MeasureTheory/Integral/FundThmCalculus.lean index 45c64597e7b0c..4a55c25968c90 100644 --- a/Mathlib/MeasureTheory/Integral/FundThmCalculus.lean +++ b/Mathlib/MeasureTheory/Integral/FundThmCalculus.lean @@ -1002,7 +1002,7 @@ theorem sub_le_integral_of_hasDeriv_right_of_le_Ico (hab : a ≤ b) _ ≤ ∫ w in t..u, (G' w).toReal := by rw [intervalIntegral.integral_of_le hu.1.le, ← integral_Icc_eq_integral_Ioc] apply setIntegral_mono_ae_restrict - · simp only [integrableOn_const, Real.volume_Icc, ENNReal.ofReal_lt_top, or_true_iff] + · simp only [integrableOn_const, Real.volume_Icc, ENNReal.ofReal_lt_top, or_true] · exact IntegrableOn.mono_set G'int I · have C1 : ∀ᵐ x : ℝ ∂volume.restrict (Icc t u), G' x < ∞ := ae_mono (Measure.restrict_mono I le_rfl) G'lt_top @@ -1252,10 +1252,10 @@ theorem intervalIntegrable_deriv_of_nonneg (hcont : ContinuousOn g (uIcc a b)) (hpos : ∀ x ∈ Ioo (min a b) (max a b), 0 ≤ g' x) : IntervalIntegrable g' volume a b := by rcases le_total a b with hab | hab · simp only [uIcc_of_le, min_eq_left, max_eq_right, hab, IntervalIntegrable, hab, - Ioc_eq_empty_of_le, integrableOn_empty, and_true_iff] at hcont hderiv hpos ⊢ + Ioc_eq_empty_of_le, integrableOn_empty, and_true] at hcont hderiv hpos ⊢ exact integrableOn_deriv_of_nonneg hcont hderiv hpos · simp only [uIcc_of_ge, min_eq_right, max_eq_left, hab, IntervalIntegrable, Ioc_eq_empty_of_le, - integrableOn_empty, true_and_iff] at hcont hderiv hpos ⊢ + integrableOn_empty, true_and] at hcont hderiv hpos ⊢ exact integrableOn_deriv_of_nonneg hcont hderiv hpos /-! diff --git a/Mathlib/MeasureTheory/Integral/IntegrableOn.lean b/Mathlib/MeasureTheory/Integral/IntegrableOn.lean index 20ef2efcbbe01..ae1b813c974c1 100644 --- a/Mathlib/MeasureTheory/Integral/IntegrableOn.lean +++ b/Mathlib/MeasureTheory/Integral/IntegrableOn.lean @@ -622,7 +622,7 @@ theorem integrableOn_Icc_iff_integrableOn_Ioc' (ha : μ {a} ≠ ∞) : IntegrableOn f (Icc a b) μ ↔ IntegrableOn f (Ioc a b) μ := by by_cases hab : a ≤ b · rw [← Ioc_union_left hab, integrableOn_union, - eq_true (integrableOn_singleton_iff.mpr <| Or.inr ha.lt_top), and_true_iff] + eq_true (integrableOn_singleton_iff.mpr <| Or.inr ha.lt_top), and_true] · rw [Icc_eq_empty hab, Ioc_eq_empty] contrapose! hab exact hab.le @@ -631,7 +631,7 @@ theorem integrableOn_Icc_iff_integrableOn_Ico' (hb : μ {b} ≠ ∞) : IntegrableOn f (Icc a b) μ ↔ IntegrableOn f (Ico a b) μ := by by_cases hab : a ≤ b · rw [← Ico_union_right hab, integrableOn_union, - eq_true (integrableOn_singleton_iff.mpr <| Or.inr hb.lt_top), and_true_iff] + eq_true (integrableOn_singleton_iff.mpr <| Or.inr hb.lt_top), and_true] · rw [Icc_eq_empty hab, Ico_eq_empty] contrapose! hab exact hab.le @@ -640,14 +640,14 @@ theorem integrableOn_Ico_iff_integrableOn_Ioo' (ha : μ {a} ≠ ∞) : IntegrableOn f (Ico a b) μ ↔ IntegrableOn f (Ioo a b) μ := by by_cases hab : a < b · rw [← Ioo_union_left hab, integrableOn_union, - eq_true (integrableOn_singleton_iff.mpr <| Or.inr ha.lt_top), and_true_iff] + eq_true (integrableOn_singleton_iff.mpr <| Or.inr ha.lt_top), and_true] · rw [Ioo_eq_empty hab, Ico_eq_empty hab] theorem integrableOn_Ioc_iff_integrableOn_Ioo' (hb : μ {b} ≠ ∞) : IntegrableOn f (Ioc a b) μ ↔ IntegrableOn f (Ioo a b) μ := by by_cases hab : a < b · rw [← Ioo_union_right hab, integrableOn_union, - eq_true (integrableOn_singleton_iff.mpr <| Or.inr hb.lt_top), and_true_iff] + eq_true (integrableOn_singleton_iff.mpr <| Or.inr hb.lt_top), and_true] · rw [Ioo_eq_empty hab, Ioc_eq_empty hab] theorem integrableOn_Icc_iff_integrableOn_Ioo' (ha : μ {a} ≠ ∞) (hb : μ {b} ≠ ∞) : @@ -657,12 +657,12 @@ theorem integrableOn_Icc_iff_integrableOn_Ioo' (ha : μ {a} ≠ ∞) (hb : μ {b theorem integrableOn_Ici_iff_integrableOn_Ioi' (hb : μ {b} ≠ ∞) : IntegrableOn f (Ici b) μ ↔ IntegrableOn f (Ioi b) μ := by rw [← Ioi_union_left, integrableOn_union, - eq_true (integrableOn_singleton_iff.mpr <| Or.inr hb.lt_top), and_true_iff] + eq_true (integrableOn_singleton_iff.mpr <| Or.inr hb.lt_top), and_true] theorem integrableOn_Iic_iff_integrableOn_Iio' (hb : μ {b} ≠ ∞) : IntegrableOn f (Iic b) μ ↔ IntegrableOn f (Iio b) μ := by rw [← Iio_union_right, integrableOn_union, - eq_true (integrableOn_singleton_iff.mpr <| Or.inr hb.lt_top), and_true_iff] + eq_true (integrableOn_singleton_iff.mpr <| Or.inr hb.lt_top), and_true] variable [NoAtoms μ] diff --git a/Mathlib/MeasureTheory/Integral/IntervalIntegral.lean b/Mathlib/MeasureTheory/Integral/IntervalIntegral.lean index 0acaf9efc4688..fc342e5424d56 100644 --- a/Mathlib/MeasureTheory/Integral/IntervalIntegral.lean +++ b/Mathlib/MeasureTheory/Integral/IntervalIntegral.lean @@ -878,7 +878,7 @@ theorem integral_const_of_cdf [CompleteSpace E] [IsFiniteMeasure μ] (c : E) : ∫ _ in a..b, c ∂μ = ((μ (Iic b)).toReal - (μ (Iic a)).toReal) • c := by simp only [sub_smul, ← setIntegral_const] refine (integral_Iic_sub_Iic ?_ ?_).symm <;> - simp only [integrableOn_const, measure_lt_top, or_true_iff] + simp only [integrableOn_const, measure_lt_top, or_true] theorem integral_eq_integral_of_support_subset {a b} (h : support f ⊆ Ioc a b) : ∫ x in a..b, f x ∂μ = ∫ x, f x ∂μ := by @@ -936,9 +936,9 @@ theorem integral_pos_iff_support_of_nonneg_ae' (hf : 0 ≤ᵐ[μ.restrict (Ι a (0 < ∫ x in a..b, f x ∂μ) ↔ a < b ∧ 0 < μ (support f ∩ Ioc a b) := by cases' lt_or_le a b with hab hba · rw [uIoc_of_le hab.le] at hf - simp only [hab, true_and_iff, integral_of_le hab.le, + simp only [hab, true_and, integral_of_le hab.le, setIntegral_pos_iff_support_of_nonneg_ae hf hfi.1] - · suffices (∫ x in a..b, f x ∂μ) ≤ 0 by simp only [this.not_lt, hba.not_lt, false_and_iff] + · suffices (∫ x in a..b, f x ∂μ) ≤ 0 by simp only [this.not_lt, hba.not_lt, false_and] rw [integral_of_ge hba, neg_nonpos] rw [uIoc_comm, uIoc_of_le hba] at hf exact integral_nonneg_of_ae hf diff --git a/Mathlib/MeasureTheory/Integral/Layercake.lean b/Mathlib/MeasureTheory/Integral/Layercake.lean index e0771f0991c48..e1602854efd97 100644 --- a/Mathlib/MeasureTheory/Integral/Layercake.lean +++ b/Mathlib/MeasureTheory/Integral/Layercake.lean @@ -141,7 +141,7 @@ theorem lintegral_comp_eq_lintegral_meas_le_mul_of_measurable_of_sigmaFinite · simp only [h_copy, h h', indicator_of_not_mem, not_false_iff, mem_Ici, not_le, mul_zero] · have : s ∉ Ioi (0 : ℝ) := h' simp only [this, h', indicator_of_not_mem, not_false_iff, mul_zero, - zero_mul, mem_Ioc, false_and_iff] + zero_mul, mem_Ioc, false_and] simp_rw [aux₁] rw [lintegral_const_mul'] swap diff --git a/Mathlib/MeasureTheory/Integral/Lebesgue.lean b/Mathlib/MeasureTheory/Integral/Lebesgue.lean index f41a9e5db44de..cc0f07957494b 100644 --- a/Mathlib/MeasureTheory/Integral/Lebesgue.lean +++ b/Mathlib/MeasureTheory/Integral/Lebesgue.lean @@ -358,7 +358,7 @@ theorem lintegral_iSup {f : ℕ → α → ℝ≥0∞} (hf : ∀ n, Measurable ( have eq : ∀ p, rs.map c ⁻¹' {p} = ⋃ n, rs.map c ⁻¹' {p} ∩ { a | p ≤ f n a } := by intro p rw [← inter_iUnion]; nth_rw 1 [← inter_univ (map c rs ⁻¹' {p})] - refine Set.ext fun x => and_congr_right fun hx => true_iff_iff.2 ?_ + refine Set.ext fun x => and_congr_right fun hx => (iff_of_eq (true_iff _)).2 ?_ by_cases p_eq : p = 0 · simp [p_eq] simp only [coe_map, mem_preimage, Function.comp_apply, mem_singleton_iff] at hx @@ -1967,8 +1967,8 @@ theorem SimpleFunc.exists_lt_lintegral_simpleFunc_of_lt_lintegral {m : Measurabl simp only [hc, ENNReal.coe_zero, zero_mul, not_lt_zero] at hL have : L / c < μ s := by rwa [ENNReal.div_lt_iff, mul_comm] - · simp only [c_ne_zero, Ne, ENNReal.coe_eq_zero, not_false_iff, true_or_iff] - · simp only [Ne, coe_ne_top, not_false_iff, true_or_iff] + · simp only [c_ne_zero, Ne, ENNReal.coe_eq_zero, not_false_iff, true_or] + · simp only [Ne, coe_ne_top, not_false_iff, true_or] obtain ⟨t, ht, ts, mlt, t_top⟩ : ∃ t : Set α, MeasurableSet t ∧ t ⊆ s ∧ L / ↑c < μ t ∧ μ t < ∞ := Measure.exists_subset_measure_lt_top hs this @@ -1982,8 +1982,8 @@ theorem SimpleFunc.exists_lt_lintegral_simpleFunc_of_lt_lintegral {m : Measurabl piecewise_eq_indicator, ENNReal.coe_indicator, Function.const_apply, lintegral_indicator, lintegral_const, Measure.restrict_apply', univ_inter] rwa [mul_comm, ← ENNReal.div_lt_iff] - · simp only [c_ne_zero, Ne, ENNReal.coe_eq_zero, not_false_iff, true_or_iff] - · simp only [Ne, coe_ne_top, not_false_iff, true_or_iff] + · simp only [c_ne_zero, Ne, ENNReal.coe_eq_zero, not_false_iff, true_or] + · simp only [Ne, coe_ne_top, not_false_iff, true_or] · replace hL : L < ∫⁻ x, f₁ x ∂μ + ∫⁻ x, f₂ x ∂μ := by rwa [← lintegral_add_left f₁.measurable.coe_nnreal_ennreal] by_cases hf₁ : ∫⁻ x, f₁ x ∂μ = 0 diff --git a/Mathlib/MeasureTheory/Integral/SetIntegral.lean b/Mathlib/MeasureTheory/Integral/SetIntegral.lean index f7e029875e697..eebb4d8f3b6bf 100644 --- a/Mathlib/MeasureTheory/Integral/SetIntegral.lean +++ b/Mathlib/MeasureTheory/Integral/SetIntegral.lean @@ -378,7 +378,7 @@ theorem setIntegral_eq_of_subset_of_ae_diff_eq_zero_aux (hts : s ⊆ t) intro h'x by_cases xs : x ∈ s · simp only [xs, hts xs] - · simp only [xs, iff_false_iff] + · simp only [xs, iff_false] intro xt exact h'x (hx ⟨xt, xs⟩) _ = ∫ x in s ∩ k, f x ∂μ + ∫ x in s \ k, f x ∂μ := by diff --git a/Mathlib/MeasureTheory/Integral/SetToL1.lean b/Mathlib/MeasureTheory/Integral/SetToL1.lean index 02f0849ff37db..02dbf1cdd368b 100644 --- a/Mathlib/MeasureTheory/Integral/SetToL1.lean +++ b/Mathlib/MeasureTheory/Integral/SetToL1.lean @@ -116,14 +116,14 @@ theorem of_smul_measure (c : ℝ≥0∞) (hc_ne_top : c ≠ ∞) (hT : FinMeasAd FinMeasAdditive μ T := by refine of_eq_top_imp_eq_top (fun s _ hμs => ?_) hT rw [Measure.smul_apply, smul_eq_mul, ENNReal.mul_eq_top] at hμs - simp only [hc_ne_top, or_false_iff, Ne, false_and_iff] at hμs + simp only [hc_ne_top, or_false, Ne, false_and] at hμs exact hμs.2 theorem smul_measure (c : ℝ≥0∞) (hc_ne_zero : c ≠ 0) (hT : FinMeasAdditive μ T) : FinMeasAdditive (c • μ) T := by refine of_eq_top_imp_eq_top (fun s _ hμs => ?_) hT rw [Measure.smul_apply, smul_eq_mul, ENNReal.mul_eq_top] - simp only [hc_ne_zero, true_and_iff, Ne, not_false_iff] + simp only [hc_ne_zero, true_and, Ne, not_false_iff] exact Or.inl hμs theorem smul_measure_iff (c : ℝ≥0∞) (hc_ne_zero : c ≠ 0) (hc_ne_top : c ≠ ∞) : @@ -231,8 +231,8 @@ theorem of_smul_measure (c : ℝ≥0∞) (hc_ne_top : c ≠ ∞) (hT : Dominated DominatedFinMeasAdditive μ T (c.toReal * C) := by have h : ∀ s, MeasurableSet s → c • μ s = ∞ → μ s = ∞ := by intro s _ hcμs - simp only [hc_ne_top, Algebra.id.smul_eq_mul, ENNReal.mul_eq_top, or_false_iff, Ne, - false_and_iff] at hcμs + simp only [hc_ne_top, Algebra.id.smul_eq_mul, ENNReal.mul_eq_top, or_false, Ne, + false_and] at hcμs exact hcμs.2 refine ⟨hT.1.of_eq_top_imp_eq_top (μ := c • μ) h, fun s hs hμs => ?_⟩ have hcμs : c • μ s ≠ ∞ := mt (h s hs) hμs.ne @@ -1502,7 +1502,7 @@ theorem setToFun_top_smul_measure (hT : DominatedFinMeasAdditive (∞ • μ) T setToFun (∞ • μ) T hT f = 0 := by refine setToFun_measure_zero' hT fun s _ hμs => ?_ rw [lt_top_iff_ne_top] at hμs - simp only [true_and_iff, Measure.smul_apply, ENNReal.mul_eq_top, eq_self_iff_true, + simp only [true_and, Measure.smul_apply, ENNReal.mul_eq_top, eq_self_iff_true, top_ne_zero, Ne, not_false_iff, not_or, Classical.not_not, smul_eq_mul] at hμs simp only [hμs.right, Measure.smul_apply, mul_zero, smul_eq_mul] diff --git a/Mathlib/MeasureTheory/Integral/VitaliCaratheodory.lean b/Mathlib/MeasureTheory/Integral/VitaliCaratheodory.lean index 15a4b3437ddd3..2aa0a2e287af9 100644 --- a/Mathlib/MeasureTheory/Integral/VitaliCaratheodory.lean +++ b/Mathlib/MeasureTheory/Integral/VitaliCaratheodory.lean @@ -110,11 +110,11 @@ theorem SimpleFunc.exists_le_lowerSemicontinuous_lintegral_ge (f : α →ₛ ℝ · simp only [lintegral_const, zero_mul, zero_le, ENNReal.coe_zero] have ne_top : μ s ≠ ⊤ := by classical - simpa [f, hs, hc, lt_top_iff_ne_top, true_and_iff, SimpleFunc.coe_const, + simpa [f, hs, hc, lt_top_iff_ne_top, SimpleFunc.coe_const, Function.const_apply, lintegral_const, ENNReal.coe_indicator, Set.univ_inter, ENNReal.coe_ne_top, MeasurableSet.univ, ENNReal.mul_eq_top, SimpleFunc.const_zero, - or_false_iff, lintegral_indicator, ENNReal.coe_eq_zero, Ne, not_false_iff, - SimpleFunc.coe_zero, Set.piecewise_eq_indicator, SimpleFunc.coe_piecewise, false_and_iff, + lintegral_indicator, ENNReal.coe_eq_zero, Ne, not_false_iff, + SimpleFunc.coe_zero, Set.piecewise_eq_indicator, SimpleFunc.coe_piecewise, restrict_apply] using h have : μ s < μ s + ε / c := by have : (0 : ℝ≥0∞) < ε / c := ENNReal.div_pos_iff.2 ⟨ε0, ENNReal.coe_ne_top⟩ @@ -326,12 +326,12 @@ theorem SimpleFunc.exists_upperSemicontinuous_le_lintegral_le (f : α →ₛ ℝ Set.piecewise_eq_indicator, ENNReal.coe_zero, SimpleFunc.coe_piecewise, zero_le] have μs_lt_top : μ s < ∞ := by classical - simpa only [hs, hc, lt_top_iff_ne_top, true_and_iff, SimpleFunc.coe_const, or_false_iff, + simpa only [hs, hc, lt_top_iff_ne_top, true_and, SimpleFunc.coe_const, or_false, lintegral_const, ENNReal.coe_indicator, Set.univ_inter, ENNReal.coe_ne_top, Measure.restrict_apply MeasurableSet.univ, ENNReal.mul_eq_top, SimpleFunc.const_zero, Function.const_apply, lintegral_indicator, ENNReal.coe_eq_zero, Ne, not_false_iff, SimpleFunc.coe_zero, Set.piecewise_eq_indicator, SimpleFunc.coe_piecewise, - false_and_iff] using int_f + false_and] using int_f have : (0 : ℝ≥0∞) < ε / c := ENNReal.div_pos_iff.2 ⟨ε0, ENNReal.coe_ne_top⟩ obtain ⟨F, Fs, F_closed, μF⟩ : ∃ (F : _), F ⊆ s ∧ IsClosed F ∧ μ s < μ F + ε / c := hs.exists_isClosed_lt_add μs_lt_top.ne this.ne' diff --git a/Mathlib/MeasureTheory/Measure/AEMeasurable.lean b/Mathlib/MeasureTheory/Measure/AEMeasurable.lean index f9ad11c6fbc66..7ea4ac9be4b94 100644 --- a/Mathlib/MeasureTheory/Measure/AEMeasurable.lean +++ b/Mathlib/MeasureTheory/Measure/AEMeasurable.lean @@ -197,7 +197,7 @@ theorem exists_ae_eq_range_subset (H : AEMeasurable f μ) {t : Set β} (ht : ∀ simp only [g, hx, piecewise_eq_of_not_mem, not_false_iff] contrapose! hx apply subset_toMeasurable - simp only [hx, mem_compl_iff, mem_setOf_eq, false_and_iff, not_false_iff] + simp only [hx, mem_compl_iff, mem_setOf_eq, false_and, not_false_iff] theorem exists_measurable_nonneg {β} [Preorder β] [Zero β] {mβ : MeasurableSpace β} {f : α → β} (hf : AEMeasurable f μ) (f_nn : ∀ᵐ t ∂μ, 0 ≤ f t) : ∃ g, Measurable g ∧ 0 ≤ g ∧ f =ᵐ[μ] g := by diff --git a/Mathlib/MeasureTheory/Measure/AddContent.lean b/Mathlib/MeasureTheory/Measure/AddContent.lean index 362cd5b634339..b2e51584933b0 100644 --- a/Mathlib/MeasureTheory/Measure/AddContent.lean +++ b/Mathlib/MeasureTheory/Measure/AddContent.lean @@ -91,7 +91,7 @@ lemma addContent_union' (hs : s ∈ C) (ht : t ∈ C) (hst : s ∪ t ∈ C) (h_d rotate_left · simp only [coe_pair, Set.insert_subset_iff, hs, ht, Set.singleton_subset_iff, and_self_iff] · simp only [coe_pair, Set.pairwiseDisjoint_insert, pairwiseDisjoint_singleton, - mem_singleton_iff, Ne, id, forall_eq, true_and_iff] + mem_singleton_iff, Ne, id, forall_eq, true_and] exact fun _ => h_dis · simp only [coe_pair, sUnion_insert, sUnion_singleton] exact hst diff --git a/Mathlib/MeasureTheory/Measure/Haar/Basic.lean b/Mathlib/MeasureTheory/Measure/Haar/Basic.lean index b8bf558a35d4e..c09a144500014 100644 --- a/Mathlib/MeasureTheory/Measure/Haar/Basic.lean +++ b/Mathlib/MeasureTheory/Measure/Haar/Basic.lean @@ -199,7 +199,7 @@ theorem index_union_le (K₁ K₂ : Compacts G) {V : Set G} (hV : (interior V).N apply Nat.sInf_le; refine ⟨_, ?_, rfl⟩; rw [mem_setOf_eq] apply union_subset <;> refine Subset.trans (by assumption) ?_ <;> apply biUnion_subset_biUnion_left <;> intro g hg <;> simp only [mem_def] at hg <;> - simp only [mem_def, Multiset.mem_union, Finset.union_val, hg, or_true_iff, true_or_iff] + simp only [mem_def, Multiset.mem_union, Finset.union_val, hg, or_true, true_or] @[to_additive addIndex_union_eq] theorem index_union_eq (K₁ K₂ : Compacts G) {V : Set G} (hV : (interior V).Nonempty) @@ -216,7 +216,7 @@ theorem index_union_eq (K₁ K₂ : Compacts G) {V : Set G} (hV : (interior V).N intro g hg; rcases hK hg with ⟨_, ⟨g₀, rfl⟩, _, ⟨h1g₀, rfl⟩, h2g₀⟩ simp only [mem_preimage] at h2g₀ simp only [mem_iUnion]; use g₀; constructor; swap - · simp only [Finset.mem_filter, h1g₀, true_and_iff]; use g + · simp only [Finset.mem_filter, h1g₀, true_and]; use g simp only [hg, h2g₀, mem_inter_iff, mem_preimage, and_self_iff] exact h2g₀ refine @@ -432,7 +432,7 @@ theorem chaar_sup_eq {K₀ : PositiveCompacts G} mem_of_subset_of_mem _ (chaar_mem_clPrehaar K₀ ⟨⟨V⁻¹, (h2V₁.inter h2V₂).preimage continuous_inv⟩, by - simp only [V, mem_inv, inv_one, h3V₁, h3V₂, mem_inter_iff, true_and_iff]⟩) + simp only [V, mem_inv, inv_one, h3V₁, h3V₂, mem_inter_iff, true_and]⟩) unfold clPrehaar; rw [IsClosed.closure_subset_iff] · rintro _ ⟨U, ⟨h1U, h2U, h3U⟩, rfl⟩ simp only [eval, mem_preimage, sub_eq_zero, mem_singleton_iff]; rw [eq_comm] diff --git a/Mathlib/MeasureTheory/Measure/Lebesgue/EqHaar.lean b/Mathlib/MeasureTheory/Measure/Lebesgue/EqHaar.lean index fb8dff12af177..264bf06f72a92 100644 --- a/Mathlib/MeasureTheory/Measure/Lebesgue/EqHaar.lean +++ b/Mathlib/MeasureTheory/Measure/Lebesgue/EqHaar.lean @@ -512,8 +512,8 @@ theorem addHaar_singleton_add_smul_div_singleton_add_smul {r : ℝ} (hr : r ≠ (μ s * (μ t)⁻¹) := by rw [ENNReal.mul_inv] · ring - · simp only [pow_pos (abs_pos.mpr hr), ENNReal.ofReal_eq_zero, not_le, Ne, true_or_iff] - · simp only [ENNReal.ofReal_ne_top, true_or_iff, Ne, not_false_iff] + · simp only [pow_pos (abs_pos.mpr hr), ENNReal.ofReal_eq_zero, not_le, Ne, true_or] + · simp only [ENNReal.ofReal_ne_top, true_or, Ne, not_false_iff] _ = μ s / μ t := by rw [ENNReal.mul_inv_cancel, one_mul, div_eq_mul_inv] · simp only [pow_pos (abs_pos.mpr hr), ENNReal.ofReal_eq_zero, not_le, Ne] @@ -624,7 +624,7 @@ theorem tendsto_addHaar_inter_smul_zero_of_density_zero_aux1 (s : Set E) (x : E) (𝓝[>] 0) (𝓝 (0 * (μ (closedBall x 1) / μ ({x} + u)))) := by apply ENNReal.Tendsto.mul A _ B (Or.inr ENNReal.zero_ne_top) simp only [ne_eq, not_true, singleton_add, image_add_left, measure_preimage_add, false_or, - ENNReal.div_eq_top, h'u, false_or_iff, not_and, and_false_iff] + ENNReal.div_eq_top, h'u, not_and, and_false] intro aux exact (measure_closedBall_lt_top.ne aux).elim -- Porting note: it used to be enough to pass `measure_closedBall_lt_top.ne` to `simp` @@ -780,11 +780,10 @@ theorem tendsto_addHaar_inter_smul_one_of_density_one_aux (s : Set E) (hs : Meas rintro r (rpos : 0 < r) refine I ({x} + r • t) s ?_ ?_ hs · simp only [h't, abs_of_nonneg rpos.le, pow_pos rpos, addHaar_smul, image_add_left, - ENNReal.ofReal_eq_zero, not_le, or_false_iff, Ne, measure_preimage_add, abs_pow, + ENNReal.ofReal_eq_zero, not_le, or_false, Ne, measure_preimage_add, abs_pow, singleton_add, mul_eq_zero] · simp [h''t, ENNReal.ofReal_ne_top, addHaar_smul, image_add_left, ENNReal.mul_eq_top, - Ne, not_false_iff, measure_preimage_add, singleton_add, and_false_iff, false_and_iff, - or_self_iff] + Ne, not_false_iff, measure_preimage_add, singleton_add, or_self_iff] /-- Consider a point `x` at which a set `s` has density one, with respect to closed balls (i.e., a Lebesgue density point of `s`). Then `s` has also density one at `x` with respect to any diff --git a/Mathlib/MeasureTheory/Measure/MeasureSpace.lean b/Mathlib/MeasureTheory/Measure/MeasureSpace.lean index e14f83eb3ffb8..84929ceccafa0 100644 --- a/Mathlib/MeasureTheory/Measure/MeasureSpace.lean +++ b/Mathlib/MeasureTheory/Measure/MeasureSpace.lean @@ -1842,7 +1842,7 @@ theorem tendsto_measure_Iic_atTop [SemilatticeSup α] [(atTop : Filter α).IsCou obtain ⟨xs, hxs_mono, hxs_tendsto⟩ := exists_seq_monotone_tendsto_atTop_atTop α have h_univ : (univ : Set α) = ⋃ n, Iic (xs n) := by ext1 x - simp only [mem_univ, mem_iUnion, mem_Iic, true_iff_iff] + simp only [mem_univ, mem_iUnion, mem_Iic, true_iff] obtain ⟨n, hn⟩ := tendsto_atTop_atTop.mp hxs_tendsto x exact ⟨n, hn n le_rfl⟩ rw [h_univ, measure_iUnion_eq_iSup, iSup_eq_iSup_subseq_of_monotone h_mono hxs_tendsto] diff --git a/Mathlib/MeasureTheory/Measure/Regular.lean b/Mathlib/MeasureTheory/Measure/Regular.lean index cb00bad6d11e4..a38a269f631cc 100644 --- a/Mathlib/MeasureTheory/Measure/Regular.lean +++ b/Mathlib/MeasureTheory/Measure/Regular.lean @@ -574,7 +574,7 @@ theorem weaklyRegular_of_finite [BorelSpace α] (μ : Measure α) [IsFiniteMeasu refine ⟨Uᶜ, compl_subset_compl.2 hsU, Fᶜ, compl_subset_compl.2 hFs, hUo.isClosed_compl, hFc.isOpen_compl, ?_⟩ - simp only [measure_compl_le_add_iff, *, hUo.measurableSet, hFc.measurableSet, true_and_iff] + simp only [measure_compl_le_add_iff, *, hUo.measurableSet, hFc.measurableSet, true_and] -- check for disjoint unions · intro s hsd hsm H ε ε0 have ε0' : ε / 2 ≠ 0 := (ENNReal.half_pos ε0).ne' @@ -667,7 +667,7 @@ lemma innerRegularWRT_isClosed_isOpen [R1Space α] [OpensMeasurableSpace α] [h theorem exists_compact_not_null [InnerRegular μ] : (∃ K, IsCompact K ∧ μ K ≠ 0) ↔ μ ≠ 0 := by simp_rw [Ne, ← measure_univ_eq_zero, MeasurableSet.univ.measure_eq_iSup_isCompact, - ENNReal.iSup_eq_zero, not_forall, exists_prop, subset_univ, true_and_iff] + ENNReal.iSup_eq_zero, not_forall, exists_prop, subset_univ, true_and] /-- If `μ` is inner regular, then any measurable set can be approximated by a compact subset. See also `MeasurableSet.exists_isCompact_lt_add_of_ne_top`. -/ @@ -993,7 +993,7 @@ theorem _root_.IsOpen.measure_eq_iSup_isCompact ⦃U : Set α⦄ (hU : IsOpen U) theorem exists_compact_not_null [Regular μ] : (∃ K, IsCompact K ∧ μ K ≠ 0) ↔ μ ≠ 0 := by simp_rw [Ne, ← measure_univ_eq_zero, isOpen_univ.measure_eq_iSup_isCompact, - ENNReal.iSup_eq_zero, not_forall, exists_prop, subset_univ, true_and_iff] + ENNReal.iSup_eq_zero, not_forall, exists_prop, subset_univ, true_and] /-- If `μ` is a regular measure, then any measurable set of finite measure can be approximated by a compact subset. See also `MeasurableSet.exists_isCompact_lt_add` and diff --git a/Mathlib/MeasureTheory/Measure/Stieltjes.lean b/Mathlib/MeasureTheory/Measure/Stieltjes.lean index 150f72aaf7a86..3d6f90733bc29 100644 --- a/Mathlib/MeasureTheory/Measure/Stieltjes.lean +++ b/Mathlib/MeasureTheory/Measure/Stieltjes.lean @@ -211,7 +211,7 @@ theorem length_subadditive_Icc_Ioo {a b : ℝ} {c d : ℕ → ℝ} (ss : Icc a b ⟨s, _, hf, hs⟩ have e : ⋃ i ∈ (hf.toFinset : Set ℕ), Ioo (c i) (d i) = ⋃ i ∈ s, Ioo (c i) (d i) := by simp only [Set.ext_iff, exists_prop, Finset.set_biUnion_coe, mem_iUnion, forall_const, - iff_self_iff, Finite.mem_toFinset] + Finite.mem_toFinset] rw [ENNReal.tsum_eq_iSup_sum] refine le_trans ?_ (le_iSup _ hf.toFinset) exact this hf.toFinset _ (by simpa only [e] ) @@ -312,7 +312,7 @@ theorem measurableSet_Ioi {c : ℝ} : MeasurableSet[f.outer.caratheodory] (Ioi c sub_add_sub_cancel, le_refl, max_eq_right] · simp only [hbc, le_refl, Ioc_eq_empty, Ioc_inter_Ioi, min_eq_left, Ioc_diff_Ioi, f.length_empty, - zero_add, or_true_iff, le_sup_iff, f.length_Ioc, not_lt] + zero_add, or_true, le_sup_iff, f.length_Ioc, not_lt] · simp only [hac, hbc, Ioc_inter_Ioi, Ioc_diff_Ioi, f.length_Ioc, min_eq_right, _root_.sup_eq_max, le_refl, Ioc_eq_empty, add_zero, max_eq_left, f.length_empty, not_lt] diff --git a/Mathlib/MeasureTheory/Measure/WithDensity.lean b/Mathlib/MeasureTheory/Measure/WithDensity.lean index 2822135c290aa..4af7da3eaff23 100644 --- a/Mathlib/MeasureTheory/Measure/WithDensity.lean +++ b/Mathlib/MeasureTheory/Measure/WithDensity.lean @@ -235,14 +235,14 @@ theorem withDensity_apply_eq_zero' {f : α → ℝ≥0∞} {s : Set α} (hf : AE simp only [Pi.zero_apply, mem_setOf_eq, Filter.mem_mk] at A convert A using 2 ext x - simp only [and_comm, exists_prop, mem_inter_iff, iff_self_iff, mem_setOf_eq, + simp only [and_comm, exists_prop, mem_inter_iff, mem_setOf_eq, mem_compl_iff, not_forall] · intro hs let t := toMeasurable μ ({ x | f x ≠ 0 } ∩ s) have A : s ⊆ t ∪ { x | f x = 0 } := by intro x hx rcases eq_or_ne (f x) 0 with (fx | fx) - · simp only [fx, mem_union, mem_setOf_eq, eq_self_iff_true, or_true_iff] + · simp only [fx, mem_union, mem_setOf_eq, eq_self_iff_true, or_true] · left apply subset_toMeasurable _ _ exact ⟨fx, hx⟩ @@ -271,7 +271,7 @@ theorem ae_withDensity_iff' {p : α → Prop} {f : α → ℝ≥0∞} (hf : AEMe rw [ae_iff, ae_iff, withDensity_apply_eq_zero' hf, iff_iff_eq] congr ext x - simp only [exists_prop, mem_inter_iff, iff_self_iff, mem_setOf_eq, not_forall] + simp only [exists_prop, mem_inter_iff, mem_setOf_eq, not_forall] theorem ae_withDensity_iff {p : α → Prop} {f : α → ℝ≥0∞} (hf : Measurable f) : (∀ᵐ x ∂μ.withDensity f, p x) ↔ ∀ᵐ x ∂μ, f x ≠ 0 → p x := diff --git a/Mathlib/MeasureTheory/PiSystem.lean b/Mathlib/MeasureTheory/PiSystem.lean index eb36ea8a28489..aed108095fc15 100644 --- a/Mathlib/MeasureTheory/PiSystem.lean +++ b/Mathlib/MeasureTheory/PiSystem.lean @@ -260,8 +260,7 @@ theorem mem_generatePiSystem_iUnion_elim {α β} {g : β → Set (Set α)} (h_pi rw [← forall_and] constructor <;> intro h1 b <;> by_cases hbs : b ∈ T_s <;> by_cases hbt : b ∈ T_t' <;> specialize h1 b <;> - simp only [hbs, hbt, if_true, if_false, true_imp_iff, and_self_iff, false_imp_iff, - and_true_iff, true_and_iff] at h1 ⊢ + simp only [hbs, hbt, if_true, if_false, true_imp_iff, and_self_iff, false_imp_iff] at h1 ⊢ all_goals exact h1 intro b h_b split_ifs with hbs hbt hbt @@ -338,11 +337,10 @@ theorem piiUnionInter_singleton (π : ι → Set (Set α)) (i : ι) : exact Or.inl (hfπ i hi) · have ht_empty : t = ∅ := by ext1 x - simp only [Finset.not_mem_empty, iff_false_iff] + simp only [Finset.not_mem_empty, iff_false] exact fun hx => hi (hti x hx ▸ hx) -- Porting note: `Finset.not_mem_empty` required - simp [ht_empty, Finset.not_mem_empty, iInter_false, iInter_univ, Set.mem_singleton univ, - or_true_iff] + simp [ht_empty, Finset.not_mem_empty, iInter_false, iInter_univ, Set.mem_singleton univ] · cases' h with hs hs · refine ⟨{i}, ?_, fun _ => s, ⟨fun x hx => ?_, ?_⟩⟩ · rw [Finset.coe_singleton] @@ -351,7 +349,7 @@ theorem piiUnionInter_singleton (π : ι → Set (Set α)) (i : ι) : · simp only [Finset.mem_singleton, iInter_iInter_eq_left] · refine ⟨∅, ?_⟩ simpa only [Finset.coe_empty, subset_singleton_iff, mem_empty_iff_false, IsEmpty.forall_iff, - imp_true_iff, Finset.not_mem_empty, iInter_false, iInter_univ, true_and_iff, + imp_true_iff, Finset.not_mem_empty, iInter_false, iInter_univ, true_and, exists_const] using hs theorem piiUnionInter_singleton_left (s : ι → Set α) (S : Set ι) : diff --git a/Mathlib/MeasureTheory/SetSemiring.lean b/Mathlib/MeasureTheory/SetSemiring.lean index 5da29e8340326..0727d10d8f8f5 100644 --- a/Mathlib/MeasureTheory/SetSemiring.lean +++ b/Mathlib/MeasureTheory/SetSemiring.lean @@ -72,7 +72,7 @@ lemma empty_not_mem_diffFinset (hC : IsSetSemiring C) (hs : s ∈ C) (ht : t ∈ ∅ ∉ hC.diffFinset hs ht := by classical simp only [diffFinset, mem_sdiff, Finset.mem_singleton, eq_self_iff_true, not_true, - and_false_iff, not_false_iff] + and_false, not_false_iff] lemma diffFinset_subset (hC : IsSetSemiring C) (hs : s ∈ C) (ht : t ∈ C) : ↑(hC.diffFinset hs ht) ⊆ C := by @@ -209,7 +209,7 @@ lemma empty_not_mem_diffFinset₀ (hC : IsSetSemiring C) (hs : s ∈ C) (hI : ∅ ∉ hC.diffFinset₀ hs hI := by classical simp only [diffFinset₀, mem_sdiff, Finset.mem_singleton, eq_self_iff_true, not_true, - and_false_iff, not_false_iff] + and_false, not_false_iff] lemma diffFinset₀_subset (hC : IsSetSemiring C) (hs : s ∈ C) (hI : ↑I ⊆ C) : ↑(hC.diffFinset₀ hs hI) ⊆ C := by diff --git a/Mathlib/ModelTheory/Encoding.lean b/Mathlib/ModelTheory/Encoding.lean index 98775b13439f3..fab8363491989 100644 --- a/Mathlib/ModelTheory/Encoding.lean +++ b/Mathlib/ModelTheory/Encoding.lean @@ -146,7 +146,7 @@ instance [Encodable α] [Encodable (Σi, L.Functions i)] : Encodable (L.Term α) instance [h1 : Countable α] [h2 : Countable (Σl, L.Functions l)] : Countable (L.Term α) := by refine mk_le_aleph0_iff.1 (card_le.trans (max_le_iff.2 ?_)) - simp only [le_refl, mk_sum, add_le_aleph0, lift_le_aleph0, true_and_iff] + simp only [le_refl, mk_sum, add_le_aleph0, lift_le_aleph0, true_and] exact ⟨Cardinal.mk_le_aleph0, Cardinal.mk_le_aleph0⟩ instance small [Small.{u} α] : Small.{u} (L.Term α) := diff --git a/Mathlib/ModelTheory/Satisfiability.lean b/Mathlib/ModelTheory/Satisfiability.lean index 848d9f1321768..8be017397efda 100644 --- a/Mathlib/ModelTheory/Satisfiability.lean +++ b/Mathlib/ModelTheory/Satisfiability.lean @@ -397,12 +397,12 @@ namespace IsComplete theorem models_not_iff (h : T.IsComplete) (φ : L.Sentence) : T ⊨ᵇ φ.not ↔ ¬T ⊨ᵇ φ := by cases' h.2 φ with hφ hφn - · simp only [hφ, not_true, iff_false_iff] + · simp only [hφ, not_true, iff_false] rw [models_sentence_iff, not_forall] refine ⟨h.1.some, ?_⟩ simp only [Sentence.realize_not, Classical.not_not] exact models_sentence_iff.1 hφ _ - · simp only [hφn, true_iff_iff] + · simp only [hφn, true_iff] intro hφ rw [models_sentence_iff] at * exact hφn h.1.some (hφ _) diff --git a/Mathlib/ModelTheory/Semantics.lean b/Mathlib/ModelTheory/Semantics.lean index 64044e8cbeccd..0827f5b0b226a 100644 --- a/Mathlib/ModelTheory/Semantics.lean +++ b/Mathlib/ModelTheory/Semantics.lean @@ -946,7 +946,7 @@ theorem Sentence.realize_cardGe (n) : M ⊨ Sentence.cardGe L n ↔ ↑n ≤ #M BoundedFormula.realize_exs] simp_rw [BoundedFormula.realize_foldr_inf] simp only [Function.comp_apply, List.mem_map, Prod.exists, Ne, List.mem_product, - List.mem_finRange, forall_exists_index, and_imp, List.mem_filter, true_and_iff] + List.mem_finRange, forall_exists_index, and_imp, List.mem_filter, true_and] refine ⟨?_, fun xs => ⟨xs.some, ?_⟩⟩ · rintro ⟨xs, h⟩ refine ⟨⟨xs, fun i j ij => ?_⟩⟩ diff --git a/Mathlib/ModelTheory/Substructures.lean b/Mathlib/ModelTheory/Substructures.lean index 6d782ceaf8e00..0fc0993999295 100644 --- a/Mathlib/ModelTheory/Substructures.lean +++ b/Mathlib/ModelTheory/Substructures.lean @@ -739,7 +739,7 @@ theorem closure_withConstants_eq : refine closure_eq_of_le ((A.subset_union_right).trans subset_closure) ?_ rw [← (L.lhomWithConstants A).substructureReduct.le_iff_le] simp only [subset_closure, reduct_withConstants, closure_le, LHom.coe_substructureReduct, - Set.union_subset_iff, and_true_iff] + Set.union_subset_iff, and_true] exact subset_closure_withConstants end Substructure @@ -926,7 +926,7 @@ namespace Equiv theorem toHom_range (f : M ≃[L] N) : f.toHom.range = ⊤ := by ext n - simp only [Hom.mem_range, coe_toHom, Substructure.mem_top, iff_true_iff] + simp only [Hom.mem_range, coe_toHom, Substructure.mem_top, iff_true] exact ⟨f.symm n, apply_symm_apply _ _⟩ end Equiv diff --git a/Mathlib/ModelTheory/Syntax.lean b/Mathlib/ModelTheory/Syntax.lean index b75f1a92257f2..c302df830a56f 100644 --- a/Mathlib/ModelTheory/Syntax.lean +++ b/Mathlib/ModelTheory/Syntax.lean @@ -417,7 +417,7 @@ theorem castLE_castLE {k m n} (km : k ≤ m) (mn : m ≤ n) (φ : L.BoundedFormu | equal => simp | rel => intros - simp only [castLE, eq_self_iff_true, heq_iff_eq, true_and_iff] + simp only [castLE, eq_self_iff_true, heq_iff_eq] rw [← Function.comp.assoc, Term.relabel_comp_relabel] simp | imp _ _ ih1 ih2 => simp [ih1, ih2] diff --git a/Mathlib/NumberTheory/ArithmeticFunction.lean b/Mathlib/NumberTheory/ArithmeticFunction.lean index 442363c3a7429..9e6e2e67edaef 100644 --- a/Mathlib/NumberTheory/ArithmeticFunction.lean +++ b/Mathlib/NumberTheory/ArithmeticFunction.lean @@ -625,7 +625,7 @@ theorem mul [CommSemiring R] {f g : ArithmeticFunction R} (hf : f.IsMultiplicati constructor · ring rw [Nat.mul_eq_zero] at * - apply not_or_of_not ha hb + apply not_or_intro ha hb · simp only [Set.InjOn, mem_coe, mem_divisorsAntidiagonal, Ne, mem_product, Prod.mk.inj_iff] rintro ⟨⟨a1, a2⟩, ⟨b1, b2⟩⟩ ⟨⟨rfl, ha⟩, ⟨rfl, hb⟩⟩ ⟨⟨c1, c2⟩, ⟨d1, d2⟩⟩ hcd h simp only [Prod.mk.inj_iff] at h diff --git a/Mathlib/NumberTheory/Bernoulli.lean b/Mathlib/NumberTheory/Bernoulli.lean index 167675460f60a..bda018aa1e2ea 100644 --- a/Mathlib/NumberTheory/Bernoulli.lean +++ b/Mathlib/NumberTheory/Bernoulli.lean @@ -253,7 +253,7 @@ theorem bernoulliPowerSeries_mul_exp_sub_one : bernoulliPowerSeries A * (exp A - simp only [bernoulliPowerSeries, coeff_mul, coeff_X, sum_antidiagonal_succ', one_div, coeff_mk, coeff_one, coeff_exp, LinearMap.map_sub, factorial, if_pos, cast_succ, cast_one, cast_mul, sub_zero, RingHom.map_one, add_eq_zero, if_false, _root_.inv_one, zero_add, one_ne_zero, - mul_zero, and_false_iff, sub_self, ← RingHom.map_mul, ← map_sum] + mul_zero, and_false, sub_self, ← RingHom.map_mul, ← map_sum] cases' n with n · simp rw [if_neg n.succ_succ_ne_one] diff --git a/Mathlib/NumberTheory/Divisors.lean b/Mathlib/NumberTheory/Divisors.lean index 78221e801b0c2..0c98cefad6ecd 100644 --- a/Mathlib/NumberTheory/Divisors.lean +++ b/Mathlib/NumberTheory/Divisors.lean @@ -82,7 +82,7 @@ theorem cons_self_properDivisors (h : n ≠ 0) : @[simp] theorem mem_divisors {m : ℕ} : n ∈ divisors m ↔ n ∣ m ∧ m ≠ 0 := by rcases eq_or_ne m 0 with (rfl | hm); · simp [divisors] - simp only [hm, Ne, not_false_iff, and_true_iff, ← filter_dvd_eq_divisors hm, mem_filter, + simp only [hm, Ne, not_false_iff, and_true, ← filter_dvd_eq_divisors hm, mem_filter, mem_range, and_iff_right_iff_imp, Nat.lt_succ_iff] exact le_of_dvd hm.bot_lt @@ -109,7 +109,7 @@ theorem mem_divisorsAntidiagonal {x : ℕ × ℕ} : · rw [Nat.lt_add_one_iff, Nat.lt_add_one_iff] rw [mul_eq_zero, not_or] at h simp only [succ_le_of_lt (Nat.pos_of_ne_zero h.1), succ_le_of_lt (Nat.pos_of_ne_zero h.2), - true_and_iff] + true_and] exact ⟨Nat.le_mul_of_pos_right _ (Nat.pos_of_ne_zero h.2), Nat.le_mul_of_pos_left _ (Nat.pos_of_ne_zero h.1)⟩ @@ -130,7 +130,7 @@ lemma right_ne_zero_of_mem_divisorsAntidiagonal {p : ℕ × ℕ} (hp : p ∈ n.d theorem divisor_le {m : ℕ} : n ∈ divisors m → n ≤ m := by cases' m with m · simp - · simp only [mem_divisors, Nat.succ_ne_zero m, and_true_iff, Ne, not_false_iff] + · simp only [mem_divisors, Nat.succ_ne_zero m, and_true, Ne, not_false_iff] exact Nat.le_of_dvd (Nat.succ_pos m) theorem divisors_subset_of_dvd {m : ℕ} (hzero : n ≠ 0) (h : m ∣ n) : divisors m ⊆ divisors n := diff --git a/Mathlib/NumberTheory/LegendreSymbol/GaussEisensteinLemmas.lean b/Mathlib/NumberTheory/LegendreSymbol/GaussEisensteinLemmas.lean index cd03518fe7e21..392736e26084d 100644 --- a/Mathlib/NumberTheory/LegendreSymbol/GaussEisensteinLemmas.lean +++ b/Mathlib/NumberTheory/LegendreSymbol/GaussEisensteinLemmas.lean @@ -205,7 +205,7 @@ theorem sum_mul_div_add_sum_mul_div_eq_mul (p q : ℕ) [hp : Fact p.Prime] (hq0 simpa [hq0] using congr_arg ((↑) : ℕ → ZMod p) (le_antisymm hpq hqp) apply_fun ZMod.val at this rw [val_cast_of_lt hxp, val_zero] at this - simp only [this, nonpos_iff_eq_zero, mem_Ico, one_ne_zero, false_and_iff, mem_product] at hx + simp only [this, nonpos_iff_eq_zero, mem_Ico, one_ne_zero, false_and, mem_product] at hx have hunion : (((Ico 1 (p / 2).succ ×ˢ Ico 1 (q / 2).succ).filter fun x : ℕ × ℕ => x.2 * p ≤ x.1 * q) ∪ (Ico 1 (p / 2).succ ×ˢ Ico 1 (q / 2).succ).filter fun x : ℕ × ℕ => x.1 * q ≤ x.2 * p) = diff --git a/Mathlib/NumberTheory/LegendreSymbol/QuadraticChar/Basic.lean b/Mathlib/NumberTheory/LegendreSymbol/QuadraticChar/Basic.lean index 4c18eb5f5ff10..9ba104ca9690a 100644 --- a/Mathlib/NumberTheory/LegendreSymbol/QuadraticChar/Basic.lean +++ b/Mathlib/NumberTheory/LegendreSymbol/QuadraticChar/Basic.lean @@ -64,7 +64,7 @@ theorem quadraticCharFun_eq_zero_iff {a : F} : quadraticCharFun F a = 0 ↔ a = simp only [quadraticCharFun] by_cases ha : a = 0 · simp only [ha, if_true] - · simp only [ha, if_false, iff_false_iff] + · simp only [ha, if_false] split_ifs <;> simp only [neg_eq_zero, one_ne_zero, not_false_iff] @[simp] @@ -247,7 +247,7 @@ theorem quadraticChar_card_sqrts (hF : ringChar F ≠ 2) (a : F) : ext1 -- Porting note(https://github.com/leanprover-community/mathlib4/issues/5026): -- added (Set.mem_toFinset), Set.mem_setOf - simp only [(Set.mem_toFinset), Set.mem_setOf, not_mem_empty, iff_false_iff] + simp only [(Set.mem_toFinset), Set.mem_setOf, not_mem_empty, iff_false] rw [isSquare_iff_exists_sq] at h exact fun h' ↦ h ⟨_, h'.symm⟩ @@ -284,7 +284,7 @@ theorem quadraticChar_neg_one [DecidableEq F] (hF : ringChar F ≠ 2) : theorem FiniteField.isSquare_neg_one_iff : IsSquare (-1 : F) ↔ Fintype.card F % 4 ≠ 3 := by classical -- suggested by the linter (instead of `[DecidableEq F]`) by_cases hF : ringChar F = 2 - · simp only [FiniteField.isSquare_of_char_two hF, Ne, true_iff_iff] + · simp only [FiniteField.isSquare_of_char_two hF, Ne, true_iff] exact fun hf ↦ one_ne_zero <| (Nat.odd_of_mod_four_eq_three hf).symm.trans <| FiniteField.even_card_of_char_two hF diff --git a/Mathlib/NumberTheory/LegendreSymbol/QuadraticChar/GaussSum.lean b/Mathlib/NumberTheory/LegendreSymbol/QuadraticChar/GaussSum.lean index 15a271ae8ab8f..b10a0d651815b 100644 --- a/Mathlib/NumberTheory/LegendreSymbol/QuadraticChar/GaussSum.lean +++ b/Mathlib/NumberTheory/LegendreSymbol/QuadraticChar/GaussSum.lean @@ -41,7 +41,7 @@ theorem FiniteField.isSquare_two_iff : classical by_cases hF : ringChar F = 2 · have h := FiniteField.even_card_of_char_two hF - simp only [FiniteField.isSquare_of_char_two hF, true_iff_iff] + simp only [FiniteField.isSquare_of_char_two hF, true_iff] omega · have h := FiniteField.odd_card_of_char_ne_two hF rw [← quadraticChar_one_iff_isSquare (Ring.two_ne_zero hF), quadraticChar_two hF, @@ -60,7 +60,7 @@ theorem FiniteField.isSquare_neg_two_iff : classical by_cases hF : ringChar F = 2 · have h := FiniteField.even_card_of_char_two hF - simp only [FiniteField.isSquare_of_char_two hF, true_iff_iff] + simp only [FiniteField.isSquare_of_char_two hF, true_iff] omega · have h := FiniteField.odd_card_of_char_ne_two hF rw [← quadraticChar_one_iff_isSquare (neg_ne_zero.mpr (Ring.two_ne_zero hF)), @@ -102,7 +102,7 @@ theorem FiniteField.isSquare_odd_prime_iff (hF : ringChar F ≠ 2) {p : ℕ} [Fa classical by_cases hFp : ringChar F = p · rw [show (p : F) = 0 by rw [← hFp]; exact ringChar.Nat.cast_ringChar] - simp only [isSquare_zero, Ne, true_iff_iff, map_mul] + simp only [isSquare_zero, Ne, true_iff, map_mul] obtain ⟨n, _, hc⟩ := FiniteField.card F (ringChar F) have hchar : ringChar F = ringChar (ZMod p) := by rw [hFp]; exact (ringChar_zmod_n p).symm conv => enter [1, 1, 2]; rw [hc, Nat.cast_pow, map_pow, hchar, map_ringChar] diff --git a/Mathlib/NumberTheory/MulChar/Basic.lean b/Mathlib/NumberTheory/MulChar/Basic.lean index 7c8c94626c010..e4ed0932f9876 100644 --- a/Mathlib/NumberTheory/MulChar/Basic.lean +++ b/Mathlib/NumberTheory/MulChar/Basic.lean @@ -159,14 +159,14 @@ noncomputable def ofUnitHom (f : Rˣ →* R'ˣ) : MulChar R R' where classical intro x y by_cases hx : IsUnit x - · simp only [hx, IsUnit.mul_iff, true_and_iff, dif_pos] + · simp only [hx, IsUnit.mul_iff, true_and, dif_pos] by_cases hy : IsUnit y · simp only [hy, dif_pos] have hm : (IsUnit.mul_iff.mpr ⟨hx, hy⟩).unit = hx.unit * hy.unit := Units.eq_iff.mp rfl rw [hm, map_mul] norm_cast · simp only [hy, not_false_iff, dif_neg, mul_zero] - · simp only [hx, IsUnit.mul_iff, false_and_iff, not_false_iff, dif_neg, zero_mul] + · simp only [hx, IsUnit.mul_iff, false_and, not_false_iff, dif_neg, zero_mul] map_nonunit' := by intro a ha simp only [ha, not_false_iff, dif_neg] @@ -577,7 +577,7 @@ theorem sum_one_eq_card_units [DecidableEq R] : · exact map_nonunit _ h · congr ext a - simp only [Finset.mem_filter, Finset.mem_univ, true_and_iff, Finset.mem_map, + simp only [Finset.mem_filter, Finset.mem_univ, true_and, Finset.mem_map, Function.Embedding.coeFn_mk, exists_true_left, IsUnit] end sum diff --git a/Mathlib/NumberTheory/Multiplicity.lean b/Mathlib/NumberTheory/Multiplicity.lean index dd1a44061a628..43de4a819a527 100644 --- a/Mathlib/NumberTheory/Multiplicity.lean +++ b/Mathlib/NumberTheory/Multiplicity.lean @@ -389,7 +389,7 @@ theorem pow_add_pow (hxy : p ∣ x + y) (hx : ¬p ∣ x) {n : ℕ} (hn : Odd n) iterate 3 rw [padicValNat_def, PartENat.natCast_get] · exact multiplicity.Nat.pow_add_pow hp.out hp1 hxy hx hn · exact Odd.pos hn - · simp only [add_pos_iff, Nat.succ_pos', or_true_iff] + · simp only [add_pos_iff, Nat.succ_pos', or_true] · exact Nat.lt_add_left _ (pow_pos y.succ_pos _) end padicValNat diff --git a/Mathlib/NumberTheory/Padics/PadicIntegers.lean b/Mathlib/NumberTheory/Padics/PadicIntegers.lean index 5e22e36d43bd8..02820e33e2483 100644 --- a/Mathlib/NumberTheory/Padics/PadicIntegers.lean +++ b/Mathlib/NumberTheory/Padics/PadicIntegers.lean @@ -494,7 +494,7 @@ theorem norm_le_pow_iff_mem_span_pow (x : ℤ_[p]) (n : ℕ) : ‖x‖ ≤ (p : ℝ) ^ (-n : ℤ) ↔ x ∈ (Ideal.span {(p : ℤ_[p]) ^ n} : Ideal ℤ_[p]) := by by_cases hx : x = 0 · subst hx - simp only [norm_zero, zpow_neg, zpow_natCast, inv_nonneg, iff_true_iff, Submodule.zero_mem] + simp only [norm_zero, zpow_neg, zpow_natCast, inv_nonneg, iff_true, Submodule.zero_mem] exact mod_cast Nat.zero_le _ rw [norm_le_pow_iff_le_valuation x hx, mem_span_pow_iff_le_valuation x hx] diff --git a/Mathlib/NumberTheory/Padics/PadicNorm.lean b/Mathlib/NumberTheory/Padics/PadicNorm.lean index 164a53abcb173..9a46b58826307 100644 --- a/Mathlib/NumberTheory/Padics/PadicNorm.lean +++ b/Mathlib/NumberTheory/Padics/PadicNorm.lean @@ -262,7 +262,7 @@ theorem int_eq_one_iff (m : ℤ) : padicNorm p m = 1 ↔ ¬(p : ℤ) ∣ m := by theorem int_lt_one_iff (m : ℤ) : padicNorm p m < 1 ↔ (p : ℤ) ∣ m := by rw [← not_iff_not, ← int_eq_one_iff, eq_iff_le_not_lt] - simp only [padicNorm.of_int, true_and_iff] + simp only [padicNorm.of_int, true_and] theorem of_nat (m : ℕ) : padicNorm p m ≤ 1 := padicNorm.of_int (m : ℤ) diff --git a/Mathlib/NumberTheory/Pell.lean b/Mathlib/NumberTheory/Pell.lean index 4ffe309128a95..8c063efdbcb77 100644 --- a/Mathlib/NumberTheory/Pell.lean +++ b/Mathlib/NumberTheory/Pell.lean @@ -305,8 +305,8 @@ theorem exists_pos_variant (h₀ : 0 < d) (a : Solution₁ d) : (lt_or_gt_of_ne (a.x_ne_zero h₀.le)).elim ((le_total 0 a.y).elim (fun hy hx => ⟨-a⁻¹, ?_, ?_, ?_⟩) fun hy hx => ⟨-a, ?_, ?_, ?_⟩) ((le_total 0 a.y).elim (fun hy hx => ⟨a, hx, hy, ?_⟩) fun hy hx => ⟨a⁻¹, hx, ?_, ?_⟩) <;> - simp only [neg_neg, inv_inv, neg_inv, Set.mem_insert_iff, Set.mem_singleton_iff, true_or_iff, - eq_self_iff_true, x_neg, x_inv, y_neg, y_inv, neg_pos, neg_nonneg, or_true_iff] <;> + simp only [neg_neg, inv_inv, neg_inv, Set.mem_insert_iff, Set.mem_singleton_iff, true_or, + eq_self_iff_true, x_neg, x_inv, y_neg, y_inv, neg_pos, neg_nonneg, or_true] <;> assumption end Solution₁ diff --git a/Mathlib/NumberTheory/Zsqrtd/Basic.lean b/Mathlib/NumberTheory/Zsqrtd/Basic.lean index acbc027998af2..c73a74bb7a6b6 100644 --- a/Mathlib/NumberTheory/Zsqrtd/Basic.lean +++ b/Mathlib/NumberTheory/Zsqrtd/Basic.lean @@ -893,8 +893,8 @@ theorem norm_eq_zero {d : ℤ} (h_nonsquare : ∀ n : ℤ, d ≠ n * n) (a : ℤ · push_neg at h suffices a.re * a.re = 0 by rw [eq_zero_of_mul_self_eq_zero this] at ha ⊢ - simpa only [true_and_iff, or_self_right, zero_re, zero_im, eq_self_iff_true, zero_eq_mul, - mul_zero, mul_eq_zero, h.ne, false_or_iff, or_self_iff] using ha + simpa only [true_and, or_self_right, zero_re, zero_im, eq_self_iff_true, zero_eq_mul, + mul_zero, mul_eq_zero, h.ne, false_or, or_self_iff] using ha apply _root_.le_antisymm _ (mul_self_nonneg _) rw [ha, mul_assoc] exact mul_nonpos_of_nonpos_of_nonneg h.le (mul_self_nonneg _) diff --git a/Mathlib/Order/Bounds/Basic.lean b/Mathlib/Order/Bounds/Basic.lean index 534374839904e..86984514ba3e7 100644 --- a/Mathlib/Order/Bounds/Basic.lean +++ b/Mathlib/Order/Bounds/Basic.lean @@ -759,7 +759,7 @@ theorem nonempty_of_not_bddBelow [Nonempty α] (h : ¬BddBelow s) : s.Nonempty : @[simp] theorem bddAbove_insert [IsDirected α (· ≤ ·)] {s : Set α} {a : α} : BddAbove (insert a s) ↔ BddAbove s := by - simp only [insert_eq, bddAbove_union, bddAbove_singleton, true_and_iff] + simp only [insert_eq, bddAbove_union, bddAbove_singleton, true_and] protected theorem BddAbove.insert [IsDirected α (· ≤ ·)] {s : Set α} (a : α) : BddAbove s → BddAbove (insert a s) := @@ -769,7 +769,7 @@ protected theorem BddAbove.insert [IsDirected α (· ≤ ·)] {s : Set α} (a : @[simp] theorem bddBelow_insert [IsDirected α (· ≥ ·)] {s : Set α} {a : α} : BddBelow (insert a s) ↔ BddBelow s := by - simp only [insert_eq, bddBelow_union, bddBelow_singleton, true_and_iff] + simp only [insert_eq, bddBelow_union, bddBelow_singleton, true_and] protected theorem BddBelow.insert [IsDirected α (· ≥ ·)] {s : Set α} (a : α) : BddBelow s → BddBelow (insert a s) := diff --git a/Mathlib/Order/CompactlyGenerated/Basic.lean b/Mathlib/Order/CompactlyGenerated/Basic.lean index 5a2890bd05342..1d5c2b6ad6f39 100644 --- a/Mathlib/Order/CompactlyGenerated/Basic.lean +++ b/Mathlib/Order/CompactlyGenerated/Basic.lean @@ -129,7 +129,7 @@ theorem isCompactElement_iff_le_of_directed_sSup_le (k : α) : apply sSup_le_sSup intro x hx use {x} - simpa only [and_true_iff, id, Finset.coe_singleton, eq_self_iff_true, + simpa only [and_true, id, Finset.coe_singleton, eq_self_iff_true, Finset.sup_singleton, Set.singleton_subset_iff] have Sne : S.Nonempty := by suffices ⊥ ∈ S from Set.nonempty_of_mem this diff --git a/Mathlib/Order/Compare.lean b/Mathlib/Order/Compare.lean index 4bd82d405b971..adbc0a9916738 100644 --- a/Mathlib/Order/Compare.lean +++ b/Mathlib/Order/Compare.lean @@ -31,12 +31,12 @@ def cmpLE {α} [LE α] [@DecidableRel α (· ≤ ·)] (x y : α) : Ordering := theorem cmpLE_swap {α} [LE α] [IsTotal α (· ≤ ·)] [@DecidableRel α (· ≤ ·)] (x y : α) : (cmpLE x y).swap = cmpLE y x := by by_cases xy : x ≤ y <;> by_cases yx : y ≤ x <;> simp [cmpLE, *, Ordering.swap] - cases not_or_of_not xy yx (total_of _ _ _) + cases not_or_intro xy yx (total_of _ _ _) theorem cmpLE_eq_cmp {α} [Preorder α] [IsTotal α (· ≤ ·)] [@DecidableRel α (· ≤ ·)] [@DecidableRel α (· < ·)] (x y : α) : cmpLE x y = cmp x y := by by_cases xy : x ≤ y <;> by_cases yx : y ≤ x <;> simp [cmpLE, lt_iff_le_not_le, *, cmp, cmpUsing] - cases not_or_of_not xy yx (total_of _ _ _) + cases not_or_intro xy yx (total_of _ _ _) namespace Ordering diff --git a/Mathlib/Order/Cover.lean b/Mathlib/Order/Cover.lean index 945aa4514d439..17ba059d015c4 100644 --- a/Mathlib/Order/Cover.lean +++ b/Mathlib/Order/Cover.lean @@ -80,7 +80,7 @@ theorem wcovBy_congr_right (hab : AntisymmRel (· ≤ ·) a b) : c ⩿ a ↔ c /-- If `a ≤ b`, then `b` does not cover `a` iff there's an element in between. -/ theorem not_wcovBy_iff (h : a ≤ b) : ¬a ⩿ b ↔ ∃ c, a < c ∧ c < b := by - simp_rw [WCovBy, h, true_and_iff, not_forall, exists_prop, not_not] + simp_rw [WCovBy, h, true_and, not_forall, exists_prop, not_not] instance WCovBy.isRefl : IsRefl α (· ⩿ ·) := ⟨WCovBy.refl⟩ @@ -208,7 +208,7 @@ theorem CovBy.lt (h : a ⋖ b) : a < b := /-- If `a < b`, then `b` does not cover `a` iff there's an element in between. -/ theorem not_covBy_iff (h : a < b) : ¬a ⋖ b ↔ ∃ c, a < c ∧ c < b := by - simp_rw [CovBy, h, true_and_iff, not_forall, exists_prop, not_not] + simp_rw [CovBy, h, true_and, not_forall, exists_prop, not_not] alias ⟨exists_lt_lt_of_not_covBy, _⟩ := not_covBy_iff diff --git a/Mathlib/Order/Defs.lean b/Mathlib/Order/Defs.lean index 9a0471732ac3c..59abe8bdb2289 100644 --- a/Mathlib/Order/Defs.lean +++ b/Mathlib/Order/Defs.lean @@ -515,14 +515,14 @@ lemma compare_gt_iff_gt : compare a b = .gt ↔ a > b := by split_ifs <;> simp only [*, lt_irrefl, not_lt_of_gt, reduceCtorEq] case _ h₁ h₂ => have h : b < a := lt_trichotomy a b |>.resolve_left h₁ |>.resolve_left h₂ - exact true_iff_iff.2 h + rwa [true_iff] lemma compare_eq_iff_eq : compare a b = .eq ↔ a = b := by rw [LinearOrder.compare_eq_compareOfLessAndEq, compareOfLessAndEq] split_ifs <;> try simp only [reduceCtorEq] - case _ h => exact false_iff_iff.2 <| ne_iff_lt_or_gt.2 <| .inl h - case _ _ h => exact true_iff_iff.2 h - case _ _ h => exact false_iff_iff.2 h + case _ h => rw [false_iff]; exact ne_iff_lt_or_gt.2 <| .inl h + case _ _ h => rwa [true_iff] + case _ _ h => rwa [false_iff] lemma compare_le_iff_le : compare a b ≠ .gt ↔ a ≤ b := by cases h : compare a b <;> simp diff --git a/Mathlib/Order/Directed.lean b/Mathlib/Order/Directed.lean index b69b6cfa1122f..39d27f9b220b1 100644 --- a/Mathlib/Order/Directed.lean +++ b/Mathlib/Order/Directed.lean @@ -108,7 +108,7 @@ theorem Directed.extend_bot [Preorder α] [OrderBot α] {e : ι → β} {f : ι simp [Function.extend_apply' _ _ _ hb] rcases hf i j with ⟨k, hi, hj⟩ use e k - simp only [he.extend_apply, *, true_and_iff] + simp only [he.extend_apply, *, true_and] /-- A set stable by infimum is `≥`-directed. -/ theorem directedOn_of_inf_mem [SemilatticeInf α] {S : Set α} diff --git a/Mathlib/Order/Filter/AtTopBot.lean b/Mathlib/Order/Filter/AtTopBot.lean index d52d2c11a9b5d..eda822fe73f47 100644 --- a/Mathlib/Order/Filter/AtTopBot.lean +++ b/Mathlib/Order/Filter/AtTopBot.lean @@ -258,7 +258,7 @@ variable [Nonempty α] lemma atTop_neBot : NeBot (atTop : Filter α) := atTop_basis.neBot_iff.2 fun _ => nonempty_Ici @[simp] lemma mem_atTop_sets {s : Set α} : s ∈ (atTop : Filter α) ↔ ∃ a : α, ∀ b ≥ a, b ∈ s := - atTop_basis.mem_iff.trans <| exists_congr fun _ => true_and_iff _ + atTop_basis.mem_iff.trans <| exists_congr fun _ => iff_of_eq (true_and _) @[simp] lemma eventually_atTop : (∀ᶠ x in atTop, p x) ↔ ∃ a, ∀ b ≥ a, p b := mem_atTop_sets diff --git a/Mathlib/Order/Filter/Bases.lean b/Mathlib/Order/Filter/Bases.lean index 4891e31ed0e66..dfe2683487ad8 100644 --- a/Mathlib/Order/Filter/Bases.lean +++ b/Mathlib/Order/Filter/Bases.lean @@ -543,7 +543,7 @@ theorem hasBasis_iSup {ι : Sort*} {ι' : ι → Type*} {l : ι → Filter α} { theorem HasBasis.sup_principal (hl : l.HasBasis p s) (t : Set α) : (l ⊔ 𝓟 t).HasBasis p fun i => s i ∪ t := ⟨fun u => by - simp only [(hl.sup' (hasBasis_principal t)).mem_iff, PProd.exists, exists_prop, and_true_iff, + simp only [(hl.sup' (hasBasis_principal t)).mem_iff, PProd.exists, exists_prop, and_true, Unique.exists_iff]⟩ theorem HasBasis.sup_pure (hl : l.HasBasis p s) (x : α) : diff --git a/Mathlib/Order/Filter/Basic.lean b/Mathlib/Order/Filter/Basic.lean index 1cfc122eac281..d73d46a740880 100644 --- a/Mathlib/Order/Filter/Basic.lean +++ b/Mathlib/Order/Filter/Basic.lean @@ -523,7 +523,7 @@ theorem mem_sSup {x : Set α} {s : Set (Filter α)} : x ∈ sSup s ↔ ∀ f ∈ @[simp] theorem mem_iSup {x : Set α} {f : ι → Filter α} : x ∈ iSup f ↔ ∀ i, x ∈ f i := by - simp only [← Filter.mem_sets, iSup_sets_eq, iff_self_iff, mem_iInter] + simp only [← Filter.mem_sets, iSup_sets_eq, mem_iInter] @[simp] theorem iSup_neBot {f : ι → Filter α} : (⨆ i, f i).NeBot ↔ ∃ i, (f i).NeBot := by @@ -575,7 +575,7 @@ theorem mem_iInf' {ι} {s : ι → Filter α} {U : Set α} : exacts [hV ⟨i,_⟩, univ_mem] · exact dif_neg hi · simp only [iInter_dite, biInter_eq_iInter, dif_pos (Subtype.coe_prop _), Subtype.coe_eta, - iInter_univ, inter_univ, eq_self_iff_true, true_and_iff] + iInter_univ, inter_univ, eq_self_iff_true, true_and] theorem exists_iInter_of_mem_iInf {ι : Type*} {α : Type*} {f : ι → Filter α} {s} (hs : s ∈ ⨅ i, f i) : ∃ t : ι → Set α, (∀ i, t i ∈ f i) ∧ s = ⋂ i, t i := @@ -595,7 +595,7 @@ theorem Iic_principal (s : Set α) : Iic (𝓟 s) = { l | s ∈ l } := Set.ext fun _ => le_principal_iff theorem principal_mono {s t : Set α} : 𝓟 s ≤ 𝓟 t ↔ s ⊆ t := by - simp only [le_principal_iff, iff_self_iff, mem_principal] + simp only [le_principal_iff, mem_principal] @[gcongr] alias ⟨_, _root_.GCongr.filter_principal_mono⟩ := principal_mono @@ -1623,7 +1623,7 @@ theorem set_eventuallyLE_iff_mem_inf_principal {s t : Set α} {l : Filter α} : theorem set_eventuallyLE_iff_inf_principal_le {s t : Set α} {l : Filter α} : s ≤ᶠ[l] t ↔ l ⊓ 𝓟 s ≤ l ⊓ 𝓟 t := set_eventuallyLE_iff_mem_inf_principal.trans <| by - simp only [le_inf_iff, inf_le_left, true_and_iff, le_principal_iff] + simp only [le_inf_iff, inf_le_left, true_and, le_principal_iff] theorem set_eventuallyEq_iff_inf_principal {s t : Set α} {l : Filter α} : s =ᶠ[l] t ↔ l ⊓ 𝓟 s = l ⊓ 𝓟 t := by @@ -2456,7 +2456,7 @@ theorem mem_seq_def {f : Filter (α → β)} {g : Filter α} {s : Set β} : theorem mem_seq_iff {f : Filter (α → β)} {g : Filter α} {s : Set β} : s ∈ f.seq g ↔ ∃ u ∈ f, ∃ t ∈ g, Set.seq u t ⊆ s := by - simp only [mem_seq_def, seq_subset, exists_prop, iff_self_iff] + simp only [mem_seq_def, seq_subset, exists_prop] theorem mem_map_seq_iff {f : Filter α} {g : Filter β} {m : α → β → γ} {s : Set γ} : s ∈ (f.map m).seq g ↔ ∃ t u, t ∈ g ∧ u ∈ f ∧ ∀ x ∈ u, ∀ y ∈ t, m x y ∈ s := @@ -2746,7 +2746,7 @@ theorem map_eq_of_inverse {f : Filter α} {g : Filter β} {φ : α → β} (ψ : theorem tendsto_inf {f : α → β} {x : Filter α} {y₁ y₂ : Filter β} : Tendsto f x (y₁ ⊓ y₂) ↔ Tendsto f x y₁ ∧ Tendsto f x y₂ := by - simp only [Tendsto, le_inf_iff, iff_self_iff] + simp only [Tendsto, le_inf_iff] theorem tendsto_inf_left {f : α → β} {x₁ x₂ : Filter α} {y : Filter β} (h : Tendsto f x₁ y) : Tendsto f (x₁ ⊓ x₂) y := @@ -2763,7 +2763,7 @@ theorem Tendsto.inf {f : α → β} {x₁ x₂ : Filter α} {y₁ y₂ : Filter @[simp] theorem tendsto_iInf {f : α → β} {x : Filter α} {y : ι → Filter β} : Tendsto f x (⨅ i, y i) ↔ ∀ i, Tendsto f x (y i) := by - simp only [Tendsto, iff_self_iff, le_iInf_iff] + simp only [Tendsto, le_iInf_iff] theorem tendsto_iInf' {f : α → β} {x : ι → Filter α} {y : Filter β} (i : ι) (hi : Tendsto f (x i) y) : Tendsto f (⨅ i, x i) y := diff --git a/Mathlib/Order/Filter/Lift.lean b/Mathlib/Order/Filter/Lift.lean index 9c69971f95604..d292896c54051 100644 --- a/Mathlib/Order/Filter/Lift.lean +++ b/Mathlib/Order/Filter/Lift.lean @@ -362,7 +362,7 @@ theorem prod_same_eq : f ×ˢ f = f.lift' fun t : Set α => t ×ˢ t := theorem tendsto_prod_self_iff {f : α × α → β} {x : Filter α} {y : Filter β} : Filter.Tendsto f (x ×ˢ x) y ↔ ∀ W ∈ y, ∃ U ∈ x, ∀ x x' : α, x ∈ U → x' ∈ U → f (x, x') ∈ W := by - simp only [tendsto_def, mem_prod_same_iff, prod_sub_preimage_iff, exists_prop, iff_self_iff] + simp only [tendsto_def, mem_prod_same_iff, prod_sub_preimage_iff, exists_prop] variable {α₁ : Type*} {α₂ : Type*} {β₁ : Type*} {β₂ : Type*} diff --git a/Mathlib/Order/Filter/Pointwise.lean b/Mathlib/Order/Filter/Pointwise.lean index 3b54059256e8c..a1c46ad3a9f3f 100644 --- a/Mathlib/Order/Filter/Pointwise.lean +++ b/Mathlib/Order/Filter/Pointwise.lean @@ -715,7 +715,7 @@ theorem isUnit_pure (a : α) : IsUnit (pure a : Filter α) := @[simp] theorem isUnit_iff_singleton : IsUnit f ↔ ∃ a, f = pure a := by - simp only [isUnit_iff, Group.isUnit, and_true_iff] + simp only [isUnit_iff, Group.isUnit, and_true] @[to_additive] theorem map_inv' : f⁻¹.map m = (f.map m)⁻¹ := diff --git a/Mathlib/Order/Filter/Prod.lean b/Mathlib/Order/Filter/Prod.lean index ff1c3c33a3247..4cf3233718cf7 100644 --- a/Mathlib/Order/Filter/Prod.lean +++ b/Mathlib/Order/Filter/Prod.lean @@ -416,7 +416,7 @@ theorem frequently_prod_and {p : α → Prop} {q : β → Prop} : theorem tendsto_prod_iff {f : α × β → γ} {x : Filter α} {y : Filter β} {z : Filter γ} : Tendsto f (x ×ˢ y) z ↔ ∀ W ∈ z, ∃ U ∈ x, ∃ V ∈ y, ∀ x y, x ∈ U → y ∈ V → f (x, y) ∈ W := by - simp only [tendsto_def, mem_prod_iff, prod_sub_preimage_iff, exists_prop, iff_self_iff] + simp only [tendsto_def, mem_prod_iff, prod_sub_preimage_iff, exists_prop] theorem tendsto_prod_iff' {g' : Filter γ} {s : α → β × γ} : Tendsto s f (g ×ˢ g') ↔ Tendsto (fun n => (s n).1) f g ∧ Tendsto (fun n => (s n).2) f g' := by diff --git a/Mathlib/Order/Fin/Tuple.lean b/Mathlib/Order/Fin/Tuple.lean index 428d79cc5cd09..1f92421e88291 100644 --- a/Mathlib/Order/Fin/Tuple.lean +++ b/Mathlib/Order/Fin/Tuple.lean @@ -39,12 +39,12 @@ lemma insertNth_mem_Icc {i : Fin (n + 1)} {x : α i} {p : ∀ j, α (i.succAbove lemma preimage_insertNth_Icc_of_mem {i : Fin (n + 1)} {x : α i} {q₁ q₂ : ∀ j, α j} (hx : x ∈ Icc (q₁ i) (q₂ i)) : i.insertNth x ⁻¹' Icc q₁ q₂ = Icc (fun j ↦ q₁ (i.succAbove j)) fun j ↦ q₂ (i.succAbove j) := - Set.ext fun p ↦ by simp only [mem_preimage, insertNth_mem_Icc, hx, true_and_iff] + Set.ext fun p ↦ by simp only [mem_preimage, insertNth_mem_Icc, hx, true_and] lemma preimage_insertNth_Icc_of_not_mem {i : Fin (n + 1)} {x : α i} {q₁ q₂ : ∀ j, α j} (hx : x ∉ Icc (q₁ i) (q₂ i)) : i.insertNth x ⁻¹' Icc q₁ q₂ = ∅ := Set.ext fun p ↦ by - simp only [mem_preimage, insertNth_mem_Icc, hx, false_and_iff, mem_empty_iff_false] + simp only [mem_preimage, insertNth_mem_Icc, hx, false_and, mem_empty_iff_false] end Fin diff --git a/Mathlib/Order/InitialSeg.lean b/Mathlib/Order/InitialSeg.lean index 678affe70ff2c..2ae4fcd5010cc 100644 --- a/Mathlib/Order/InitialSeg.lean +++ b/Mathlib/Order/InitialSeg.lean @@ -161,7 +161,7 @@ theorem eq_or_principal [IsWellOrder β s] (f : r ≼i s) : ⟨IH _, fun ⟨a, e⟩ => by rw [← e] exact (trichotomous _ _).resolve_right - (not_or_of_not (hn a) fun hl => not_exists.2 hn (f.init hl))⟩⟩ + (not_or_intro (hn a) fun hl => not_exists.2 hn (f.init hl))⟩⟩ /-- Restrict the codomain of an initial segment -/ def codRestrict (p : Set β) (f : r ≼i s) (H : ∀ a, f a ∈ p) : r ≼i Subrel s p := diff --git a/Mathlib/Order/Interval/Finset/Basic.lean b/Mathlib/Order/Interval/Finset/Basic.lean index e0b90ecd39119..8f54103e53fcc 100644 --- a/Mathlib/Order/Interval/Finset/Basic.lean +++ b/Mathlib/Order/Interval/Finset/Basic.lean @@ -113,19 +113,19 @@ theorem Ioo_eq_empty_of_le (h : b ≤ a) : Ioo a b = ∅ := -- porting note (#10618): simp can prove this -- @[simp] -theorem left_mem_Icc : a ∈ Icc a b ↔ a ≤ b := by simp only [mem_Icc, true_and_iff, le_rfl] +theorem left_mem_Icc : a ∈ Icc a b ↔ a ≤ b := by simp only [mem_Icc, true_and, le_rfl] -- porting note (#10618): simp can prove this -- @[simp] -theorem left_mem_Ico : a ∈ Ico a b ↔ a < b := by simp only [mem_Ico, true_and_iff, le_refl] +theorem left_mem_Ico : a ∈ Ico a b ↔ a < b := by simp only [mem_Ico, true_and, le_refl] -- porting note (#10618): simp can prove this -- @[simp] -theorem right_mem_Icc : b ∈ Icc a b ↔ a ≤ b := by simp only [mem_Icc, and_true_iff, le_rfl] +theorem right_mem_Icc : b ∈ Icc a b ↔ a ≤ b := by simp only [mem_Icc, and_true, le_rfl] -- porting note (#10618): simp can prove this -- @[simp] -theorem right_mem_Ioc : b ∈ Ioc a b ↔ a < b := by simp only [mem_Ioc, and_true_iff, le_rfl] +theorem right_mem_Ioc : b ∈ Ioc a b ↔ a < b := by simp only [mem_Ioc, and_true, le_rfl] -- porting note (#10618): simp can prove this -- @[simp] diff --git a/Mathlib/Order/Interval/Set/UnorderedInterval.lean b/Mathlib/Order/Interval/Set/UnorderedInterval.lean index a0556c3932ff6..8c14071911f3d 100644 --- a/Mathlib/Order/Interval/Set/UnorderedInterval.lean +++ b/Mathlib/Order/Interval/Set/UnorderedInterval.lean @@ -294,8 +294,8 @@ lemma uIoc_injective_right (a : α) : Injective fun b => Ι b a := by rw [Set.ext_iff] at h obtain ha | ha := le_or_lt b a · have hb := (h b).not - simp only [ha, left_mem_uIoc, not_lt, true_iff_iff, not_mem_uIoc, ← not_le, - and_true_iff, not_true, false_and_iff, not_false_iff, true_iff_iff, or_false_iff] at hb + simp only [ha, left_mem_uIoc, not_lt, true_iff, not_mem_uIoc, ← not_le, + and_true, not_true, false_and, not_false_iff, or_false] at hb refine hb.eq_of_not_lt fun hc => ?_ simpa [ha, and_iff_right hc, ← @not_le _ _ _ a, iff_not_self, -not_le] using h c · refine diff --git a/Mathlib/Order/LiminfLimsup.lean b/Mathlib/Order/LiminfLimsup.lean index 07583f615a980..62ee5119d70d2 100644 --- a/Mathlib/Order/LiminfLimsup.lean +++ b/Mathlib/Order/LiminfLimsup.lean @@ -1151,12 +1151,12 @@ theorem cofinite.bliminf_set_eq : bliminf s cofinite p = { x | { n | p n ∧ x /-- In other words, `limsup cofinite s` is the set of elements lying inside the family `s` infinitely often. -/ theorem cofinite.limsup_set_eq : limsup s cofinite = { x | { n | x ∈ s n }.Infinite } := by - simp only [← cofinite.blimsup_true s, cofinite.blimsup_set_eq, true_and_iff] + simp only [← cofinite.blimsup_true s, cofinite.blimsup_set_eq, true_and] /-- In other words, `liminf cofinite s` is the set of elements lying outside the family `s` finitely often. -/ theorem cofinite.liminf_set_eq : liminf s cofinite = { x | { n | x ∉ s n }.Finite } := by - simp only [← cofinite.bliminf_true s, cofinite.bliminf_set_eq, true_and_iff] + simp only [← cofinite.bliminf_true s, cofinite.bliminf_set_eq, true_and] theorem exists_forall_mem_of_hasBasis_mem_blimsup {l : Filter β} {b : ι → Set β} {q : ι → Prop} (hl : l.HasBasis q b) {u : β → Set α} {p : β → Prop} {x : α} (hx : x ∈ blimsup u l p) : diff --git a/Mathlib/Order/ModularLattice.lean b/Mathlib/Order/ModularLattice.lean index 76b1789fce3e3..6dfc3e2db07ba 100644 --- a/Mathlib/Order/ModularLattice.lean +++ b/Mathlib/Order/ModularLattice.lean @@ -262,7 +262,7 @@ def infIccOrderIsoIccSup (a b : α) : Set.Icc (a ⊓ b) a ≃o Set.Icc b (a ⊔ change a ⊓ ↑x ⊔ b = ↑x rw [inf_comm, inf_sup_assoc_of_le _ x.prop.1, inf_eq_left.2 x.prop.2]) map_rel_iff' := @fun x y => by - simp only [Subtype.mk_le_mk, Equiv.coe_fn_mk, and_true_iff, le_sup_right] + simp only [Subtype.mk_le_mk, Equiv.coe_fn_mk, le_sup_right] rw [← Subtype.coe_le_coe] refine ⟨fun h => ?_, fun h => sup_le_sup_right h _⟩ rw [← sup_eq_right.2 x.prop.1, inf_sup_assoc_of_le _ x.prop.2, sup_comm, ← diff --git a/Mathlib/Order/RelClasses.lean b/Mathlib/Order/RelClasses.lean index dcfbcc1773326..e9910df2bafde 100644 --- a/Mathlib/Order/RelClasses.lean +++ b/Mathlib/Order/RelClasses.lean @@ -170,7 +170,7 @@ abbrev partialOrderOfSO (r) [IsStrictOrder α r] : PartialOrder α where | _, _, Or.inl rfl => rfl | _, Or.inr h₁, Or.inr h₂ => (asymm h₁ h₂).elim lt_iff_le_not_le x y := - ⟨fun h => ⟨Or.inr h, not_or_of_not (fun e => by rw [e] at h; exact irrefl _ h) (asymm h)⟩, + ⟨fun h => ⟨Or.inr h, not_or_intro (fun e => by rw [e] at h; exact irrefl _ h) (asymm h)⟩, fun ⟨h₁, h₂⟩ => h₁.resolve_left fun e => h₂ <| e ▸ Or.inl rfl⟩ /-- Construct a linear order from an `IsStrictTotalOrder` relation. diff --git a/Mathlib/Order/SupIndep.lean b/Mathlib/Order/SupIndep.lean index b973420ec6a20..cb3c3c12cac6a 100644 --- a/Mathlib/Order/SupIndep.lean +++ b/Mathlib/Order/SupIndep.lean @@ -131,7 +131,7 @@ theorem supIndep_pair [DecidableEq ι] {i j : ι} (hij : i ≠ j) : have : ({i, k} : Finset ι).erase k = {i} := by ext rw [mem_erase, mem_insert, mem_singleton, mem_singleton, and_or_left, Ne, - not_and_self_iff, or_false_iff, and_iff_right_of_imp] + not_and_self_iff, or_false, and_iff_right_of_imp] rintro rfl exact hij rw [this, Finset.sup_singleton]⟩ @@ -174,7 +174,7 @@ theorem supIndep_attach : (s.attach.SupIndep fun a => f a) ↔ s.SupIndep f := b convert h (filter_subset (fun (i : { x // x ∈ s }) => (i : ι) ∈ t) _) (mem_attach _ ⟨i, ‹_›⟩) fun hi => hit <| by simpa using hi using 1 refine eq_of_forall_ge_iff ?_ - simp only [Finset.sup_le_iff, mem_filter, mem_attach, true_and_iff, Function.comp_apply, + simp only [Finset.sup_le_iff, mem_filter, mem_attach, true_and, Function.comp_apply, Subtype.forall, Subtype.coe_mk] exact fun a => forall_congr' fun j => ⟨fun h _ => h, fun h hj => h (ht hj) hj⟩ diff --git a/Mathlib/Order/WellFoundedSet.lean b/Mathlib/Order/WellFoundedSet.lean index 5e939f8f9b787..af6576042206c 100644 --- a/Mathlib/Order/WellFoundedSet.lean +++ b/Mathlib/Order/WellFoundedSet.lean @@ -167,10 +167,10 @@ theorem wellFoundedOn_iff_no_descending_seq : · rintro ⟨⟨f, hf⟩⟩ have H : ∀ n, f n ∈ s := fun n => (hf.2 n.lt_succ_self).2.2 refine ⟨⟨f, ?_⟩, H⟩ - simpa only [H, and_true_iff] using @hf + simpa only [H, and_true] using @hf · rintro ⟨⟨f, hf⟩, hfs : ∀ n, f n ∈ s⟩ refine ⟨⟨f, ?_⟩⟩ - simpa only [hfs, and_true_iff] using @hf + simpa only [hfs, and_true] using @hf theorem WellFoundedOn.union (hs : s.WellFoundedOn r) (ht : t.WellFoundedOn r) : (s ∪ t).WellFoundedOn r := by @@ -310,7 +310,7 @@ theorem Subsingleton.partiallyWellOrderedOn (hs : s.Subsingleton) : PartiallyWel theorem partiallyWellOrderedOn_insert : PartiallyWellOrderedOn (insert a s) r ↔ PartiallyWellOrderedOn s r := by simp only [← singleton_union, partiallyWellOrderedOn_union, - partiallyWellOrderedOn_singleton, true_and_iff] + partiallyWellOrderedOn_singleton, true_and] protected theorem PartiallyWellOrderedOn.insert (h : PartiallyWellOrderedOn s r) (a : α) : PartiallyWellOrderedOn (insert a s) r := @@ -431,7 +431,7 @@ protected theorem Subsingleton.isPWO (hs : s.Subsingleton) : IsPWO s := hs.finit @[simp] theorem isPWO_insert {a} : IsPWO (insert a s) ↔ IsPWO s := by - simp only [← singleton_union, isPWO_union, isPWO_singleton, true_and_iff] + simp only [← singleton_union, isPWO_union, isPWO_singleton, true_and] protected theorem IsPWO.insert (h : IsPWO s) (a : α) : IsPWO (insert a s) := isPWO_insert.2 h @@ -444,7 +444,7 @@ protected theorem Subsingleton.isWF (hs : s.Subsingleton) : IsWF s := hs.isPWO.i @[simp] theorem isWF_insert {a} : IsWF (insert a s) ↔ IsWF s := by - simp only [← singleton_union, isWF_union, isWF_singleton, true_and_iff] + simp only [← singleton_union, isWF_union, isWF_singleton, true_and] protected theorem IsWF.insert (h : IsWF s) (a : α) : IsWF (insert a s) := isWF_insert.2 h @@ -468,7 +468,7 @@ protected theorem Subsingleton.wellFoundedOn (hs : s.Subsingleton) : s.WellFound @[simp] theorem wellFoundedOn_insert : WellFoundedOn (insert a s) r ↔ WellFoundedOn s r := by - simp only [← singleton_union, wellFoundedOn_union, wellFoundedOn_singleton, true_and_iff] + simp only [← singleton_union, wellFoundedOn_union, wellFoundedOn_singleton, true_and] protected theorem WellFoundedOn.insert (h : WellFoundedOn s r) (a : α) : WellFoundedOn (insert a s) r := diff --git a/Mathlib/Probability/Distributions/Uniform.lean b/Mathlib/Probability/Distributions/Uniform.lean index 67f4746b80fcf..c639f53703f5b 100644 --- a/Mathlib/Probability/Distributions/Uniform.lean +++ b/Mathlib/Probability/Distributions/Uniform.lean @@ -346,7 +346,7 @@ theorem mem_support_ofMultiset_iff (a : α) : a ∈ (ofMultiset s hs).support theorem ofMultiset_apply_of_not_mem {a : α} (ha : a ∉ s) : ofMultiset s hs a = 0 := by simpa only [ofMultiset_apply, ENNReal.div_eq_zero_iff, Nat.cast_eq_zero, Multiset.count_eq_zero, - ENNReal.natCast_ne_top, or_false_iff] using ha + ENNReal.natCast_ne_top, or_false] using ha section Measure diff --git a/Mathlib/Probability/Kernel/Composition.lean b/Mathlib/Probability/Kernel/Composition.lean index 97547030b3738..43797429a7831 100644 --- a/Mathlib/Probability/Kernel/Composition.lean +++ b/Mathlib/Probability/Kernel/Composition.lean @@ -320,9 +320,9 @@ theorem compProd_restrict {s : Set β} {t : Set γ} (hs : MeasurableSet s) (ht : classical rw [Set.indicator_apply] split_ifs with h - · simp only [h, true_and_iff] + · simp only [h, true_and] rfl - · simp only [h, false_and_iff, and_false_iff, Set.setOf_false, measure_empty] + · simp only [h, false_and, and_false, Set.setOf_false, measure_empty] simp_rw [this] rw [lintegral_indicator _ hs] diff --git a/Mathlib/Probability/Kernel/Disintegration/CDFToKernel.lean b/Mathlib/Probability/Kernel/Disintegration/CDFToKernel.lean index 18005da5094fd..48889e94338a1 100644 --- a/Mathlib/Probability/Kernel/Disintegration/CDFToKernel.lean +++ b/Mathlib/Probability/Kernel/Disintegration/CDFToKernel.lean @@ -621,7 +621,7 @@ lemma lintegral_toKernel_mem [IsFiniteKernel κ] (hf : IsCondKernelCDF f κ ν) simp only [mem_setOf_eq] at ht₁ ht₂ have h_prod_eq_snd : ∀ a ∈ t₁, {x : ℝ | (a, x) ∈ t₁ ×ˢ t₂} = t₂ := by intro a ha - simp only [ha, prod_mk_mem_set_prod_eq, true_and_iff, setOf_mem_eq] + simp only [ha, prod_mk_mem_set_prod_eq, true_and, setOf_mem_eq] rw [← lintegral_add_compl _ ht₁] have h_eq1 : ∫⁻ x in t₁, hf.toKernel f (a, x) {y : ℝ | (x, y) ∈ t₁ ×ˢ t₂} ∂(ν a) = ∫⁻ x in t₁, hf.toKernel f (a, x) t₂ ∂(ν a) := by @@ -635,7 +635,7 @@ lemma lintegral_toKernel_mem [IsFiniteKernel κ] (hf : IsCondKernelCDF f κ ν) simp only [lintegral_const, zero_mul] intro a hat₁ rw [mem_compl_iff] at hat₁ - simp only [hat₁, prod_mk_mem_set_prod_eq, false_and_iff, setOf_false, measure_empty] + simp only [hat₁, prod_mk_mem_set_prod_eq, false_and, setOf_false, measure_empty] rw [h_eq1, h_eq2, add_zero] exact setLIntegral_toKernel_prod hf a ht₁ ht₂ · intro t ht ht_eq @@ -667,7 +667,7 @@ lemma lintegral_toKernel_mem [IsFiniteKernel κ] (hf : IsCondKernelCDF f κ ν) have h_disj := hf_disj hij rw [Function.onFun, disjoint_iff_inter_eq_empty] at h_disj ⊢ ext1 x - simp only [mem_inter_iff, mem_setOf_eq, mem_empty_iff_false, iff_false_iff] + simp only [mem_inter_iff, mem_setOf_eq, mem_empty_iff_false, iff_false] intro h_mem_both suffices (a, x) ∈ ∅ by rwa [mem_empty_iff_false] at this rwa [← h_disj, mem_inter_iff] diff --git a/Mathlib/Probability/Kernel/IntegralCompProd.lean b/Mathlib/Probability/Kernel/IntegralCompProd.lean index a6cb2cd870a84..a8ad7e64ece6c 100644 --- a/Mathlib/Probability/Kernel/IntegralCompProd.lean +++ b/Mathlib/Probability/Kernel/IntegralCompProd.lean @@ -116,7 +116,7 @@ theorem integrable_compProd_iff ⦃f : β × γ → E⦄ (hf : AEStronglyMeasura (∀ᵐ x ∂κ a, Integrable (fun y => f (x, y)) (η (a, x))) ∧ Integrable (fun x => ∫ y, ‖f (x, y)‖ ∂η (a, x)) (κ a) := by simp only [Integrable, hasFiniteIntegral_compProd_iff' hf, hf.norm.integral_kernel_compProd, - hf, hf.compProd_mk_left, eventually_and, true_and_iff] + hf, hf.compProd_mk_left, eventually_and, true_and] theorem _root_.MeasureTheory.Integrable.compProd_mk_left_ae ⦃f : β × γ → E⦄ (hf : Integrable f ((κ ⊗ₖ η) a)) : ∀ᵐ x ∂κ a, Integrable (fun y => f (x, y)) (η (a, x)) := diff --git a/Mathlib/Probability/Kernel/MeasurableIntegral.lean b/Mathlib/Probability/Kernel/MeasurableIntegral.lean index 312093ef4376d..9d5915b707937 100644 --- a/Mathlib/Probability/Kernel/MeasurableIntegral.lean +++ b/Mathlib/Probability/Kernel/MeasurableIntegral.lean @@ -64,7 +64,7 @@ theorem measurable_kernel_prod_mk_left_of_finite {t : Set (α × β)} (ht : Meas have h_eq_sdiff : ∀ a, Prod.mk a ⁻¹' t'ᶜ = Set.univ \ Prod.mk a ⁻¹' t' := by intro a ext1 b - simp only [mem_compl_iff, mem_preimage, mem_diff, mem_univ, true_and_iff] + simp only [mem_compl_iff, mem_preimage, mem_diff, mem_univ, true_and] simp_rw [h_eq_sdiff] have : (fun a => κ a (Set.univ \ Prod.mk a ⁻¹' t')) = fun a => @@ -236,7 +236,7 @@ variable {E : Type*} [NormedAddCommGroup E] [IsSFiniteKernel κ] [IsSFiniteKerne theorem measurableSet_kernel_integrable ⦃f : α → β → E⦄ (hf : StronglyMeasurable (uncurry f)) : MeasurableSet {x | Integrable (f x) (κ x)} := by - simp_rw [Integrable, hf.of_uncurry_left.aestronglyMeasurable, true_and_iff] + simp_rw [Integrable, hf.of_uncurry_left.aestronglyMeasurable, true_and] exact measurableSet_lt (Measurable.lintegral_kernel_prod_right hf.ennnorm) measurable_const end ProbabilityTheory diff --git a/Mathlib/Probability/Martingale/Convergence.lean b/Mathlib/Probability/Martingale/Convergence.lean index 831ab2d4dae51..5cbec74af7d1e 100644 --- a/Mathlib/Probability/Martingale/Convergence.lean +++ b/Mathlib/Probability/Martingale/Convergence.lean @@ -173,8 +173,8 @@ theorem Submartingale.upcrossings_ae_lt_top' [IsFiniteMeasure μ] (hf : Submarti exact norm_nonneg _ · simp only [Ne, ENNReal.coe_ne_top, not_false_iff] · simp only [hab, Ne, ENNReal.ofReal_eq_zero, sub_nonpos, not_le] - · simp only [hab, Ne, ENNReal.ofReal_eq_zero, sub_nonpos, not_le, true_or_iff] - · simp only [Ne, ENNReal.ofReal_ne_top, not_false_iff, true_or_iff] + · simp only [hab, Ne, ENNReal.ofReal_eq_zero, sub_nonpos, not_le, true_or] + · simp only [Ne, ENNReal.ofReal_ne_top, not_false_iff, true_or] theorem Submartingale.upcrossings_ae_lt_top [IsFiniteMeasure μ] (hf : Submartingale f ℱ μ) (hbdd : ∀ n, eLpNorm (f n) 1 μ ≤ R) : ∀ᵐ ω ∂μ, ∀ a b : ℚ, a < b → upcrossings a b f ω < ∞ := by diff --git a/Mathlib/Probability/Martingale/OptionalSampling.lean b/Mathlib/Probability/Martingale/OptionalSampling.lean index 3132d43cb03f9..535963b759aba 100644 --- a/Mathlib/Probability/Martingale/OptionalSampling.lean +++ b/Mathlib/Probability/Martingale/OptionalSampling.lean @@ -66,7 +66,7 @@ theorem condexp_stopping_time_ae_eq_restrict_eq_const_of_le_const (h : Martingal rw [Set.inter_comm _ t, IsStoppingTime.measurableSet_inter_eq_iff] · suffices {x : Ω | τ x = i} = ∅ by simp [this]; norm_cast ext1 x - simp only [Set.mem_setOf_eq, Set.mem_empty_iff_false, iff_false_iff] + simp only [Set.mem_setOf_eq, Set.mem_empty_iff_false, iff_false] rintro rfl exact hin (hτ_le x) @@ -88,7 +88,7 @@ theorem stoppedValue_ae_eq_condexp_of_le_const_of_countable_range (h : Martingal stoppedValue f τ =ᵐ[μ] μ[f n|hτ.measurableSpace] := by have : Set.univ = ⋃ i ∈ Set.range τ, {x | τ x = i} := by ext1 x - simp only [Set.mem_univ, Set.mem_range, true_and_iff, Set.iUnion_exists, Set.iUnion_iUnion_eq', + simp only [Set.mem_univ, Set.mem_range, Set.iUnion_exists, Set.iUnion_iUnion_eq', Set.mem_iUnion, Set.mem_setOf_eq, exists_apply_eq_apply'] nth_rw 1 [← @Measure.restrict_univ Ω _ μ] rw [this, ae_eq_restrict_biUnion_iff _ h_countable_range] diff --git a/Mathlib/Probability/Martingale/Upcrossing.lean b/Mathlib/Probability/Martingale/Upcrossing.lean index 0d3e293008a55..a1694af83142d 100644 --- a/Mathlib/Probability/Martingale/Upcrossing.lean +++ b/Mathlib/Probability/Martingale/Upcrossing.lean @@ -449,7 +449,7 @@ theorem crossing_eq_crossing_of_lowerCrossingTime_lt {M : ℕ} (hNM : N ≤ M) lt_of_le_of_lt upperCrossingTime_le_lowerCrossingTime h induction' n with k ih · simp only [upperCrossingTime_zero, bot_eq_zero', eq_self_iff_true, - lowerCrossingTime_zero, true_and_iff, eq_comm] + lowerCrossingTime_zero, true_and, eq_comm] refine hitting_eq_hitting_of_exists hNM ?_ rw [lowerCrossingTime, hitting_lt_iff] at h · obtain ⟨j, hj₁, hj₂⟩ := h diff --git a/Mathlib/Probability/Moments.lean b/Mathlib/Probability/Moments.lean index 9f39ce4a332a9..3ad21345e302f 100644 --- a/Mathlib/Probability/Moments.lean +++ b/Mathlib/Probability/Moments.lean @@ -166,7 +166,7 @@ theorem mgf_pos' (hμ : μ ≠ 0) (h_int_X : Integrable (fun ω => exp (t * X ω rw [this, setIntegral_pos_iff_support_of_nonneg_ae _ _] · have h_eq_univ : (Function.support fun x : Ω => exp (t * X x)) = Set.univ := by ext1 x - simp only [Function.mem_support, Set.mem_univ, iff_true_iff] + simp only [Function.mem_support, Set.mem_univ, iff_true] exact (exp_pos _).ne' rw [h_eq_univ, Set.inter_univ _] refine Ne.bot_lt ?_ diff --git a/Mathlib/Probability/Process/HittingTime.lean b/Mathlib/Probability/Process/HittingTime.lean index 28eb2728601cf..161d19118807f 100644 --- a/Mathlib/Probability/Process/HittingTime.lean +++ b/Mathlib/Probability/Process/HittingTime.lean @@ -164,7 +164,7 @@ theorem hitting_le_iff_of_lt [IsWellOrder ι (· < ·)] {m : ι} (i : ι) (hi : · rw [hitting_le_iff_of_exists h_exists] · simp_rw [hitting, if_neg h_exists] push_neg at h_exists - simp only [not_le.mpr hi, Set.mem_Icc, false_iff_iff, not_exists, not_and, and_imp] + simp only [not_le.mpr hi, Set.mem_Icc, false_iff, not_exists, not_and, and_imp] exact fun k hkn hki => h_exists k ⟨hkn, hki.trans hi.le⟩ theorem hitting_lt_iff [IsWellOrder ι (· < ·)] {m : ι} (i : ι) (hi : i ≤ m) : @@ -251,7 +251,7 @@ theorem isStoppingTime_hitting_isStoppingTime [ConditionallyCompleteLinearOrder have h₂ : ⋃ i > n, {x | τ x = i} ∩ {x | hitting u s i N x ≤ n} = ∅ := by ext x simp only [gt_iff_lt, Set.mem_iUnion, Set.mem_inter_iff, Set.mem_setOf_eq, exists_prop, - Set.mem_empty_iff_false, iff_false_iff, not_exists, not_and, not_le] + Set.mem_empty_iff_false, iff_false, not_exists, not_and, not_le] rintro m hm rfl exact lt_of_lt_of_le hm (le_hitting (hτbdd _) _) rw [h₁, h₂, Set.union_empty] @@ -283,7 +283,7 @@ theorem hitting_bot_le_iff {i n : ι} {ω : Ω} (hx : ∃ j, j ≤ n ∧ u j ω cases' lt_or_le i n with hi hi · rw [hitting_le_iff_of_lt _ hi] simp - · simp only [(hitting_le ω).trans hi, true_iff_iff] + · simp only [(hitting_le ω).trans hi, true_iff] obtain ⟨j, hj₁, hj₂⟩ := hx exact ⟨j, hj₁.trans hi, hj₂⟩ diff --git a/Mathlib/Probability/Process/Stopping.lean b/Mathlib/Probability/Process/Stopping.lean index 9d48460885ec5..0b6419fc37ae6 100644 --- a/Mathlib/Probability/Process/Stopping.lean +++ b/Mathlib/Probability/Process/Stopping.lean @@ -69,7 +69,7 @@ theorem IsStoppingTime.measurableSet_lt_of_pred [PredOrder ι] (hτ : IsStopping by_cases hi_min : IsMin i · suffices {ω : Ω | τ ω < i} = ∅ by rw [this]; exact @MeasurableSet.empty _ (f i) ext1 ω - simp only [Set.mem_setOf_eq, Set.mem_empty_iff_false, iff_false_iff] + simp only [Set.mem_setOf_eq, Set.mem_empty_iff_false, iff_false] rw [isMin_iff_forall_not_lt] at hi_min exact hi_min (τ ω) have : {ω : Ω | τ ω < i} = τ ⁻¹' Set.Iic (pred i) := by ext; simp [Iic_pred_of_not_isMin hi_min] @@ -152,7 +152,7 @@ theorem IsStoppingTime.measurableSet_lt_of_isLUB (hτ : IsStoppingTime f τ) (i by_cases hi_min : IsMin i · suffices {ω | τ ω < i} = ∅ by rw [this]; exact @MeasurableSet.empty _ (f i) ext1 ω - simp only [Set.mem_setOf_eq, Set.mem_empty_iff_false, iff_false_iff] + simp only [Set.mem_setOf_eq, Set.mem_empty_iff_false, iff_false] exact isMin_iff_forall_not_lt.mp hi_min (τ ω) obtain ⟨seq, -, -, h_tendsto, h_bound⟩ : ∃ seq : ℕ → ι, Monotone seq ∧ (∀ j, seq j ≤ i) ∧ Tendsto seq atTop (𝓝 i) ∧ ∀ j, seq j < i := @@ -261,7 +261,7 @@ theorem add_const_nat {f : Filtration ℕ m} {τ : Ω → ℕ} (hτ : IsStopping · rw [not_le] at hij convert @MeasurableSet.empty _ (f.1 j) ext ω - simp only [Set.mem_empty_iff_false, iff_false_iff, Set.mem_setOf] + simp only [Set.mem_empty_iff_false, iff_false, Set.mem_setOf] omega -- generalize to certain countable type? @@ -567,18 +567,17 @@ theorem measurableSet_inter_le [TopologicalSpace ι] [SecondCountableTopology ι s ∩ {ω | τ ω ≤ i} ∩ {ω | min (τ ω) (π ω) ≤ i} ∩ {ω | min (τ ω) i ≤ min (min (τ ω) (π ω)) i} := by ext1 ω - simp only [min_le_iff, Set.mem_inter_iff, Set.mem_setOf_eq, le_min_iff, le_refl, true_and_iff, - and_true_iff, true_or_iff, or_true_iff] + simp only [min_le_iff, Set.mem_inter_iff, Set.mem_setOf_eq, le_min_iff, le_refl, true_and, + true_or] by_cases hτi : τ ω ≤ i - · simp only [hτi, true_or_iff, and_true_iff, and_congr_right_iff] + · simp only [hτi, true_or, and_true, and_congr_right_iff] intro constructor <;> intro h · exact Or.inl h · cases' h with h h · exact h · exact hτi.trans h - simp only [hτi, false_or_iff, and_false_iff, false_and_iff, iff_false_iff, not_and, not_le, - and_imp] + simp only [hτi, false_or, and_false, false_and, iff_false, not_and, not_le, and_imp] refine fun _ hτ_le_π => lt_of_lt_of_le ?_ hτ_le_π rw [← not_le] exact hτi @@ -617,10 +616,10 @@ theorem measurableSet_le_stopping_time [TopologicalSpace ι] [SecondCountableTop intro j have : {ω | τ ω ≤ π ω} ∩ {ω | τ ω ≤ j} = {ω | min (τ ω) j ≤ min (π ω) j} ∩ {ω | τ ω ≤ j} := by ext1 ω - simp only [Set.mem_inter_iff, Set.mem_setOf_eq, min_le_iff, le_min_iff, le_refl, and_true_iff, + simp only [Set.mem_inter_iff, Set.mem_setOf_eq, min_le_iff, le_min_iff, le_refl, and_congr_left_iff] intro h - simp only [h, or_self_iff, and_true_iff] + simp only [h, or_self_iff, and_true] rw [Iff.comm, or_iff_left_iff_imp] exact h.trans rw [this] diff --git a/Mathlib/Probability/StrongLaw.lean b/Mathlib/Probability/StrongLaw.lean index ec1e988d50e44..24a5607e58036 100644 --- a/Mathlib/Probability/StrongLaw.lean +++ b/Mathlib/Probability/StrongLaw.lean @@ -107,11 +107,11 @@ theorem truncation_eq_of_nonneg {f : α → ℝ} {A : ℝ} (h : ∀ x, 0 ≤ f x truncation f A = indicator (Set.Ioc 0 A) id ∘ f := by ext x rcases (h x).lt_or_eq with (hx | hx) - · simp only [truncation, indicator, hx, Set.mem_Ioc, id, Function.comp_apply, true_and_iff] + · simp only [truncation, indicator, hx, Set.mem_Ioc, id, Function.comp_apply] by_cases h'x : f x ≤ A · have : -A < f x := by linarith [h x] - simp only [this, true_and_iff] - · simp only [h'x, and_false_iff] + simp only [this, true_and] + · simp only [h'x, and_false] · simp only [truncation, indicator, hx, id, Function.comp_apply, ite_self] theorem truncation_nonneg {f : α → ℝ} (A : ℝ) {x : α} (h : 0 ≤ f x) : 0 ≤ truncation f A x := @@ -546,7 +546,7 @@ theorem strong_law_aux5 : · have : -(n : ℝ) < X n ω := by apply lt_of_lt_of_le _ (hnonneg n ω) simpa only [Right.neg_neg_iff, Nat.cast_pos] using npos - simp only [this, true_and_iff, not_le] at h + simp only [this, true_and, not_le] at h exact (hn h).elim filter_upwards [B] with ω hω convert isLittleO_sum_range_of_tendsto_zero hω using 1 diff --git a/Mathlib/Probability/Variance.lean b/Mathlib/Probability/Variance.lean index d2036c459f53a..9b0a9bcb2fe5c 100644 --- a/Mathlib/Probability/Variance.lean +++ b/Mathlib/Probability/Variance.lean @@ -228,7 +228,7 @@ theorem evariance_def' [@IsProbabilityMeasure Ω _ ℙ] {X : Ω → ℝ} (hX : A rw [Memℒp, not_and] at hℒ specialize hℒ hX simp only [eLpNorm_eq_lintegral_rpow_nnnorm two_ne_zero ENNReal.two_ne_top, not_lt, top_le_iff, - ENNReal.toReal_ofNat, one_div, ENNReal.rpow_eq_top_iff, inv_lt_zero, inv_pos, and_true_iff, + ENNReal.toReal_ofNat, one_div, ENNReal.rpow_eq_top_iff, inv_lt_zero, inv_pos, and_true, or_iff_not_imp_left, not_and_or, zero_lt_two] at hℒ exact mod_cast hℒ fun _ => zero_le_two @@ -255,7 +255,7 @@ theorem meas_ge_le_variance_div_sq [@IsFiniteMeasure Ω _ ℙ] {X : Ω → ℝ} (hc : 0 < c) : ℙ {ω | c ≤ |X ω - 𝔼[X]|} ≤ ENNReal.ofReal (Var[X] / c ^ 2) := by rw [ENNReal.ofReal_div_of_pos (sq_pos_of_ne_zero hc.ne.symm), hX.ofReal_variance_eq] convert @meas_ge_le_evariance_div_sq _ _ _ hX.1 c.toNNReal (by simp [hc]) using 1 - · simp only [Real.coe_toNNReal', max_le_iff, abs_nonneg, and_true_iff] + · simp only [Real.coe_toNNReal', max_le_iff, abs_nonneg, and_true] · rw [ENNReal.ofReal_pow hc.le] rfl diff --git a/Mathlib/RingTheory/Algebraic.lean b/Mathlib/RingTheory/Algebraic.lean index 0326b17fa9001..2f80a41c0e3cb 100644 --- a/Mathlib/RingTheory/Algebraic.lean +++ b/Mathlib/RingTheory/Algebraic.lean @@ -79,7 +79,7 @@ theorem Subalgebra.isAlgebraic_iff (S : Subalgebra R A) : /-- An algebra is algebraic if and only if it is algebraic as a subalgebra. -/ theorem Algebra.isAlgebraic_iff : Algebra.IsAlgebraic R A ↔ (⊤ : Subalgebra R A).IsAlgebraic := by delta Subalgebra.IsAlgebraic - simp only [Algebra.isAlgebraic_def, Algebra.mem_top, forall_prop_of_true, iff_self_iff] + simp only [Algebra.isAlgebraic_def, Algebra.mem_top, forall_prop_of_true] theorem isAlgebraic_iff_not_injective {x : A} : IsAlgebraic R x ↔ ¬Function.Injective (Polynomial.aeval x : R[X] →ₐ[R] A) := by diff --git a/Mathlib/RingTheory/Coprime/Ideal.lean b/Mathlib/RingTheory/Coprime/Ideal.lean index f9d40e059c801..3f7176c9ae4c7 100644 --- a/Mathlib/RingTheory/Coprime/Ideal.lean +++ b/Mathlib/RingTheory/Coprime/Ideal.lean @@ -32,7 +32,7 @@ theorem iSup_iInf_eq_top_iff_pairwise {t : Finset ι} (h : t.Nonempty) (I : ι haveI : DecidableEq ι := Classical.decEq ι rw [eq_top_iff_one, Submodule.mem_iSup_finset_iff_exists_sum] refine h.cons_induction ?_ ?_ <;> clear t h - · simp only [Finset.sum_singleton, Finset.coe_singleton, Set.pairwise_singleton, iff_true_iff] + · simp only [Finset.sum_singleton, Finset.coe_singleton, Set.pairwise_singleton, iff_true] refine fun a => ⟨fun i => if h : i = a then ⟨1, ?_⟩ else 0, ?_⟩ · simp [h] · simp only [dif_pos, Submodule.coe_mk, eq_self_iff_true] diff --git a/Mathlib/RingTheory/DedekindDomain/FiniteAdeleRing.lean b/Mathlib/RingTheory/DedekindDomain/FiniteAdeleRing.lean index 8591f904f4b2b..c3d41814f077b 100644 --- a/Mathlib/RingTheory/DedekindDomain/FiniteAdeleRing.lean +++ b/Mathlib/RingTheory/DedekindDomain/FiniteAdeleRing.lean @@ -215,7 +215,7 @@ theorem zero : (0 : K_hat R K).IsFiniteAdele := by rw [IsFiniteAdele, Filter.eventually_cofinite] have h_empty : {v : HeightOneSpectrum R | ¬(0 : v.adicCompletion K) ∈ v.adicCompletionIntegers K} = ∅ := by - ext v; rw [mem_empty_iff_false, iff_false_iff]; intro hv + ext v; rw [mem_empty_iff_false, iff_false]; intro hv rw [mem_setOf] at hv; apply hv; rw [mem_adicCompletionIntegers] have h_zero : (Valued.v (0 : v.adicCompletion K) : WithZero (Multiplicative ℤ)) = 0 := Valued.v.map_zero' @@ -259,7 +259,7 @@ theorem one : (1 : K_hat R K).IsFiniteAdele := by rw [IsFiniteAdele, Filter.eventually_cofinite] have h_empty : {v : HeightOneSpectrum R | ¬(1 : v.adicCompletion K) ∈ v.adicCompletionIntegers K} = ∅ := by - ext v; rw [mem_empty_iff_false, iff_false_iff]; intro hv + ext v; rw [mem_empty_iff_false, iff_false]; intro hv rw [mem_setOf] at hv; apply hv; rw [mem_adicCompletionIntegers] exact le_of_eq Valued.v.map_one' -- Porting note: was `exact`, but `OfNat` got in the way. diff --git a/Mathlib/RingTheory/DedekindDomain/Ideal.lean b/Mathlib/RingTheory/DedekindDomain/Ideal.lean index 67e3680bfcb02..798b6e6abf371 100644 --- a/Mathlib/RingTheory/DedekindDomain/Ideal.lean +++ b/Mathlib/RingTheory/DedekindDomain/Ideal.lean @@ -630,7 +630,7 @@ instance Ideal.uniqueFactorizationMonoid : UniqueFactorizationMonoid (Ideal A) : rintro ⟨⟨x, x_mem, x_not_mem⟩, ⟨y, y_mem, y_not_mem⟩⟩ exact ⟨x * y, Ideal.mul_mem_mul x_mem y_mem, - mt this.isPrime.mem_or_mem (not_or_of_not x_not_mem y_not_mem)⟩⟩, Prime.irreducible⟩ } + mt this.isPrime.mem_or_mem (not_or_intro x_not_mem y_not_mem)⟩⟩, Prime.irreducible⟩ } instance Ideal.normalizationMonoid : NormalizationMonoid (Ideal A) := normalizationMonoidOfUniqueUnits diff --git a/Mathlib/RingTheory/DedekindDomain/IntegralClosure.lean b/Mathlib/RingTheory/DedekindDomain/IntegralClosure.lean index 8f37a34bfe33b..f8aba734476d2 100644 --- a/Mathlib/RingTheory/DedekindDomain/IntegralClosure.lean +++ b/Mathlib/RingTheory/DedekindDomain/IntegralClosure.lean @@ -156,7 +156,7 @@ theorem FiniteDimensional.exists_is_basis_integral : · intro x; simp only [mul_inv_cancel_left₀ hy'] · rintro ⟨x', hx'⟩ simp only [Algebra.smul_def, Finset.mem_image, exists_prop, Finset.mem_univ, - true_and_iff] at his' + true_and] at his' simp only [Basis.map_apply, LinearEquiv.coe_mk] exact his' _ ⟨_, rfl⟩ diff --git a/Mathlib/RingTheory/DiscreteValuationRing/Basic.lean b/Mathlib/RingTheory/DiscreteValuationRing/Basic.lean index d461a65055d40..214213dcd11a6 100644 --- a/Mathlib/RingTheory/DiscreteValuationRing/Basic.lean +++ b/Mathlib/RingTheory/DiscreteValuationRing/Basic.lean @@ -193,7 +193,7 @@ theorem toUniqueFactorizationMonoid (hR : HasUnitMulPowIrreducibleFactorization intro a b h by_cases ha : a = 0 · rw [ha] - simp only [true_or_iff, dvd_zero] + simp only [true_or, dvd_zero] obtain ⟨m, u, rfl⟩ := spec.2 ha rw [mul_assoc, mul_left_comm, Units.dvd_mul_left] at h rw [Units.dvd_mul_right] @@ -222,8 +222,7 @@ theorem of_ufd_of_unique_irreducible [UniqueFactorizationMonoid R] (h₁ : ∃ p congr 1 symm rw [Multiset.eq_replicate] - simp only [true_and_iff, and_imp, Multiset.card_map, eq_self_iff_true, Multiset.mem_map, - exists_imp] + simp only [true_and, and_imp, Multiset.card_map, eq_self_iff_true, Multiset.mem_map, exists_imp] rintro _ q hq rfl rw [Associates.mk_eq_mk_iff_associated] apply h₂ (hfx.1 _ hq) hp @@ -311,8 +310,7 @@ theorem associated_pow_irreducible {x : R} (hx : x ≠ 0) {ϖ : R} (hirr : Irred rw [← H, ← Associates.prod_mk, Associates.mk_pow, ← Multiset.prod_replicate] congr 1 rw [Multiset.eq_replicate] - simp only [true_and_iff, and_imp, Multiset.card_map, eq_self_iff_true, Multiset.mem_map, - exists_imp] + simp only [true_and, and_imp, Multiset.card_map, eq_self_iff_true, Multiset.mem_map, exists_imp] rintro _ _ _ rfl rw [Associates.mk_eq_mk_iff_associated] refine associated_of_irreducible _ ?_ hirr diff --git a/Mathlib/RingTheory/HahnSeries/Multiplication.lean b/Mathlib/RingTheory/HahnSeries/Multiplication.lean index 981c0ad06b1fa..f3c1ef6a86bc5 100644 --- a/Mathlib/RingTheory/HahnSeries/Multiplication.lean +++ b/Mathlib/RingTheory/HahnSeries/Multiplication.lean @@ -253,7 +253,7 @@ theorem single_smul_coeff_add [MulZeroClass R] [SMulWithZero R V] {r : R} {x : H rw [sum_congr _ fun _ _ => rfl, sum_empty] ext ⟨a1, a2⟩ simp only [not_mem_empty, not_and, Set.mem_singleton_iff, Classical.not_not, - mem_vaddAntidiagonal, Set.mem_setOf_eq, iff_false_iff] + mem_vaddAntidiagonal, Set.mem_setOf_eq, iff_false] rintro rfl h2 h1 rw [IsCancelVAdd.left_cancel a1 a2 a h1] at h2 exact h2 hx @@ -387,7 +387,7 @@ theorem mul_single_coeff_add [NonUnitalNonAssocSemiring R] {r : R} {x : HahnSeri rw [sum_congr _ fun _ _ => rfl, sum_empty] ext ⟨a1, a2⟩ simp only [not_mem_empty, not_and, Set.mem_singleton_iff, Classical.not_not, - mem_addAntidiagonal, Set.mem_setOf_eq, iff_false_iff] + mem_addAntidiagonal, Set.mem_setOf_eq, iff_false] rintro h2 rfl h1 rw [← add_right_cancel h1] at hx exact h2 hx @@ -754,7 +754,7 @@ instance [Nontrivial Γ] [Nontrivial R] : Nontrivial (Subalgebra R (HahnSeries rw [Ne, SetLike.ext_iff, not_forall] obtain ⟨a, ha⟩ := exists_ne (0 : Γ) refine ⟨single a 1, ?_⟩ - simp only [Algebra.mem_bot, not_exists, Set.mem_range, iff_true_iff, Algebra.mem_top] + simp only [Algebra.mem_bot, not_exists, Set.mem_range, iff_true, Algebra.mem_top] intro x rw [HahnSeries.ext_iff, Function.funext_iff, not_forall] refine ⟨a, ?_⟩ diff --git a/Mathlib/RingTheory/HahnSeries/Summable.lean b/Mathlib/RingTheory/HahnSeries/Summable.lean index 517632b481a87..871e001eae93b 100644 --- a/Mathlib/RingTheory/HahnSeries/Summable.lean +++ b/Mathlib/RingTheory/HahnSeries/Summable.lean @@ -418,8 +418,8 @@ def powers (x : HahnSeries Γ R) (hx : 0 < x.orderTop) : SummableFamily Γ R ℕ · exact Set.mem_union_right _ (Set.mem_singleton 0) · obtain ⟨i, hi, j, hj, rfl⟩ := support_mul_subset_add_support hn refine Set.mem_union_left _ ⟨n, Set.mem_iUnion.2 ⟨⟨j, i⟩, Set.mem_iUnion.2 ⟨?_, hi⟩⟩, rfl⟩ - simp only [and_true_iff, Set.mem_iUnion, mem_addAntidiagonal, mem_coe, eq_self_iff_true, - Ne, mem_support, Set.mem_setOf_eq] + simp only [Set.mem_iUnion, mem_addAntidiagonal, mem_coe, eq_self_iff_true, Ne, mem_support, + Set.mem_setOf_eq] exact ⟨hj, ⟨n, hi⟩, add_comm j i⟩ variable {x : HahnSeries Γ R} (hx : 0 < x.orderTop) diff --git a/Mathlib/RingTheory/Ideal/AssociatedPrime.lean b/Mathlib/RingTheory/Ideal/AssociatedPrime.lean index e37e8ce6c85bf..b4ddec4e36127 100644 --- a/Mathlib/RingTheory/Ideal/AssociatedPrime.lean +++ b/Mathlib/RingTheory/Ideal/AssociatedPrime.lean @@ -104,7 +104,7 @@ theorem LinearEquiv.AssociatedPrimes.eq (l : M ≃ₗ[R] M') : (associatedPrimes.subset_of_injective l.symm l.symm.injective) theorem associatedPrimes.eq_empty_of_subsingleton [Subsingleton M] : associatedPrimes R M = ∅ := by - ext; simp only [Set.mem_empty_iff_false, iff_false_iff] + ext; simp only [Set.mem_empty_iff_false, iff_false] apply not_isAssociatedPrime_of_subsingleton variable (R M) diff --git a/Mathlib/RingTheory/Ideal/Operations.lean b/Mathlib/RingTheory/Ideal/Operations.lean index 839647966fb3a..e24dba67ac093 100644 --- a/Mathlib/RingTheory/Ideal/Operations.lean +++ b/Mathlib/RingTheory/Ideal/Operations.lean @@ -1017,7 +1017,7 @@ theorem subset_union_prime' {R : Type u} [CommRing R] {s : Finset ι} {f : ι rw [Finset.card_eq_zero] at hn subst hn rw [Finset.coe_empty, Set.biUnion_empty, Set.union_empty, subset_union] at h - simpa only [exists_prop, Finset.not_mem_empty, false_and_iff, exists_false, or_false_iff] + simpa only [exists_prop, Finset.not_mem_empty, false_and, exists_false, or_false] classical replace hn : ∃ (i : ι) (t : Finset ι), i ∉ t ∧ insert i t = s ∧ t.card = n := Finset.card_eq_succ.1 hn diff --git a/Mathlib/RingTheory/Jacobson.lean b/Mathlib/RingTheory/Jacobson.lean index f9ca45cdd5b0b..29cd5339dde32 100644 --- a/Mathlib/RingTheory/Jacobson.lean +++ b/Mathlib/RingTheory/Jacobson.lean @@ -127,7 +127,7 @@ theorem isJacobson_of_isIntegral [Algebra R S] [Algebra.IsIntegral R S] (hR : Is ((isJacobson_iff_prime_eq.1 hR) (comap (algebraMap R S) P) (comap_isPrime _ _)), comap_jacobson] refine sInf_le_sInf fun J hJ => ?_ - simp only [true_and_iff, Set.mem_image, bot_le, Set.mem_setOf_eq] + simp only [true_and, Set.mem_image, bot_le, Set.mem_setOf_eq] have : J.IsMaximal := by simpa using hJ exact exists_ideal_over_maximal_of_isIntegral J (comap_bot_le_of_injective _ algebraMap_quotient_injective) diff --git a/Mathlib/RingTheory/LocalRing/Basic.lean b/Mathlib/RingTheory/LocalRing/Basic.lean index 8e057c3e33536..a7bb4b615fa7b 100644 --- a/Mathlib/RingTheory/LocalRing/Basic.lean +++ b/Mathlib/RingTheory/LocalRing/Basic.lean @@ -63,7 +63,7 @@ theorem isUnit_or_isUnit_of_isUnit_add {a b : R} (h : IsUnit (a + b)) : IsUnit a apply Or.imp _ _ (isUnit_or_isUnit_of_add_one hu) <;> exact isUnit_of_mul_isUnit_right theorem nonunits_add {a b : R} (ha : a ∈ nonunits R) (hb : b ∈ nonunits R) : a + b ∈ nonunits R := - fun H => not_or_of_not ha hb (isUnit_or_isUnit_of_isUnit_add H) + fun H => not_or_intro ha hb (isUnit_or_isUnit_of_isUnit_add H) end LocalRing diff --git a/Mathlib/RingTheory/Localization/AtPrime.lean b/Mathlib/RingTheory/Localization/AtPrime.lean index 6ff9e226beb8a..a52c0df699d51 100644 --- a/Mathlib/RingTheory/Localization/AtPrime.lean +++ b/Mathlib/RingTheory/Localization/AtPrime.lean @@ -89,7 +89,7 @@ theorem AtPrime.localRing [IsLocalization.AtPrime S P] : LocalRing S := obtain ⟨t, ht⟩ := IsLocalization.eq.1 hxyz simp only [mul_one, one_mul, Submonoid.coe_mul, Subtype.coe_mk] at ht suffices (t : R) * (sx * sy * sz) ∈ P from - not_or_of_not (mt hp.mem_or_mem <| not_or_of_not sx.2 sy.2) sz.2 + not_or_intro (mt hp.mem_or_mem <| not_or_intro sx.2 sy.2) sz.2 (hp.mem_or_mem <| (hp.mem_or_mem this).resolve_left t.2) rw [← ht] exact diff --git a/Mathlib/RingTheory/Multiplicity.lean b/Mathlib/RingTheory/Multiplicity.lean index effafa7779249..457a6e35fcef7 100644 --- a/Mathlib/RingTheory/Multiplicity.lean +++ b/Mathlib/RingTheory/Multiplicity.lean @@ -506,7 +506,7 @@ protected theorem mul' {p a b : α} (hp : Prime p) (h : (multiplicity p (a * b)) 1) ∣ a * b := fun h => - not_or_of_not (is_greatest' _ (lt_succ_self _)) (is_greatest' _ (lt_succ_self _)) + not_or_intro (is_greatest' _ (lt_succ_self _)) (is_greatest' _ (lt_succ_self _)) (_root_.succ_dvd_or_succ_dvd_of_succ_sum_dvd_mul hp hdiva hdivb h) rw [← PartENat.natCast_inj, PartENat.natCast_get, eq_coe_iff]; exact ⟨hdiv, hsucc⟩ diff --git a/Mathlib/RingTheory/MvPolynomial/Ideal.lean b/Mathlib/RingTheory/MvPolynomial/Ideal.lean index 25562372ca086..de669fd310c46 100644 --- a/Mathlib/RingTheory/MvPolynomial/Ideal.lean +++ b/Mathlib/RingTheory/MvPolynomial/Ideal.lean @@ -37,7 +37,7 @@ theorem mem_ideal_span_monomial_image_iff_dvd {x : MvPolynomial σ R} {s : Set ( x ∈ Ideal.span ((fun s => monomial s (1 : R)) '' s) ↔ ∀ xi ∈ x.support, ∃ si ∈ s, monomial si 1 ∣ monomial xi (x.coeff xi) := by refine mem_ideal_span_monomial_image.trans (forall₂_congr fun xi hxi => ?_) - simp_rw [monomial_dvd_monomial, one_dvd, and_true_iff, mem_support_iff.mp hxi, false_or_iff] + simp_rw [monomial_dvd_monomial, one_dvd, and_true, mem_support_iff.mp hxi, false_or] /-- `x` is in a monomial ideal generated by variables `X` iff every element of its support has a component in `s`. -/ diff --git a/Mathlib/RingTheory/MvPolynomial/Tower.lean b/Mathlib/RingTheory/MvPolynomial/Tower.lean index 95f6fb384851e..77228738a64aa 100644 --- a/Mathlib/RingTheory/MvPolynomial/Tower.lean +++ b/Mathlib/RingTheory/MvPolynomial/Tower.lean @@ -52,7 +52,7 @@ theorem aeval_algebraMap_apply (x : σ → A) (p : MvPolynomial σ R) : theorem aeval_algebraMap_eq_zero_iff [NoZeroSMulDivisors A B] [Nontrivial B] (x : σ → A) (p : MvPolynomial σ R) : aeval (algebraMap A B ∘ x) p = 0 ↔ aeval x p = 0 := by rw [aeval_algebraMap_apply, Algebra.algebraMap_eq_smul_one, smul_eq_zero, - iff_false_intro (one_ne_zero' B), or_false_iff] + iff_false_intro (one_ne_zero' B), or_false] theorem aeval_algebraMap_eq_zero_iff_of_injective {x : σ → A} {p : MvPolynomial σ R} (h : Function.Injective (algebraMap A B)) : diff --git a/Mathlib/RingTheory/MvPowerSeries/Basic.lean b/Mathlib/RingTheory/MvPowerSeries/Basic.lean index 8e2b9f0fb7e79..3e649d881d0c8 100644 --- a/Mathlib/RingTheory/MvPowerSeries/Basic.lean +++ b/Mathlib/RingTheory/MvPowerSeries/Basic.lean @@ -730,7 +730,7 @@ instance [Nonempty σ] [Nontrivial R] : Nontrivial (Subalgebra R (MvPowerSeries rw [Ne, SetLike.ext_iff, not_forall] inhabit σ refine ⟨X default, ?_⟩ - simp only [Algebra.mem_bot, not_exists, Set.mem_range, iff_true_iff, Algebra.mem_top] + simp only [Algebra.mem_bot, not_exists, Set.mem_range, iff_true, Algebra.mem_top] intro x rw [MvPowerSeries.ext_iff, not_forall] refine ⟨Finsupp.single default 1, ?_⟩ diff --git a/Mathlib/RingTheory/Polynomial/Basic.lean b/Mathlib/RingTheory/Polynomial/Basic.lean index 6e18f0a0e8ef5..68388199b41c9 100644 --- a/Mathlib/RingTheory/Polynomial/Basic.lean +++ b/Mathlib/RingTheory/Polynomial/Basic.lean @@ -757,8 +757,7 @@ theorem isPrime_map_C_iff_isPrime (P : Ideal R) : · rw [← not_le] intro hnj exact (add_lt_add_of_lt_of_le hmi hnj).ne hij.2.symm - · simp only [eq_self_iff_true, not_true, false_or_iff, add_right_inj, - not_and_self_iff] at hij + · simp only [eq_self_iff_true, not_true, false_or, add_right_inj, not_and_self_iff] at hij · rw [mul_comm] apply P.mul_mem_left exact Classical.not_not.1 (Nat.find_min hf hi) diff --git a/Mathlib/RingTheory/Polynomial/Cyclotomic/Eval.lean b/Mathlib/RingTheory/Polynomial/Cyclotomic/Eval.lean index af7686e449eba..4a1ad15664ee8 100644 --- a/Mathlib/RingTheory/Polynomial/Cyclotomic/Eval.lean +++ b/Mathlib/RingTheory/Polynomial/Cyclotomic/Eval.lean @@ -147,7 +147,7 @@ theorem eval_one_cyclotomic_not_prime_pow {R : Type*} [Ring R] {n : ℕ} rw [← Finset.prod_sdiff <| show {n} ⊆ _ from _] at this swap · simp only [singleton_subset_iff, mem_sdiff, mem_erase, Ne, mem_divisors, dvd_refl, - true_and_iff, mem_image, mem_range, exists_prop, not_exists, not_and] + true_and, mem_image, mem_range, exists_prop, not_exists, not_and] exact ⟨⟨hn.ne', hn'.ne'⟩, fun t _ => h hp _⟩ rw [← Int.natAbs_ofNat p, Int.natAbs_dvd_natAbs] at hpe obtain ⟨t, ht⟩ := hpe @@ -203,7 +203,7 @@ theorem sub_one_pow_totient_lt_cyclotomic_eval {n : ℕ} {q : ℝ} (hn' : 2 ≤ Units.val_le_val, ← NNReal.coe_le_coe, Complex.abs.nonneg, hq'.le, Units.val_mk0, Real.coe_toNNReal', coe_nnnorm, Complex.norm_eq_abs, max_le_iff, tsub_le_iff_right] intro x hx - simpa only [and_true_iff, tsub_le_iff_right] using hfor x hx + simpa only [and_true, tsub_le_iff_right] using hfor x hx · simp only [Subtype.coe_mk, Finset.mem_attach, exists_true_left, Subtype.exists, ← NNReal.coe_lt_coe, ← Units.val_lt_val, Units.val_mk0 _, coe_nnnorm] simpa [hq'.le, Real.coe_toNNReal', max_eq_left, sub_nonneg] using hex diff --git a/Mathlib/RingTheory/Polynomial/Dickson.lean b/Mathlib/RingTheory/Polynomial/Dickson.lean index 6f3bfc0a478bf..fdb6b3b533b34 100644 --- a/Mathlib/RingTheory/Polynomial/Dickson.lean +++ b/Mathlib/RingTheory/Polynomial/Dickson.lean @@ -231,7 +231,7 @@ theorem dickson_one_one_zmod_p (p : ℕ) [Fact p.Prime] : dickson 1 (1 : ZMod p) mem_roots hφ, IsRoot, eval_add, eval_sub, eval_pow, eval_mul, eval_X, eval_C, eval_one, Multiset.mem_singleton] by_cases hy : y = 0 - · simp only [hy, eq_self_iff_true, or_true_iff] + · simp only [hy, eq_self_iff_true, or_true] apply or_congr _ Iff.rfl rw [← mul_left_inj' hy, eq_comm, ← sub_eq_zero, add_mul, inv_mul_cancel₀ hy] apply eq_iff_eq_cancel_right.mpr @@ -241,9 +241,9 @@ theorem dickson_one_one_zmod_p (p : ℕ) [Fact p.Prime] : dickson 1 (1 : ZMod p) intro x simp only [exists_prop, Set.mem_iUnion, Set.bind_def, Ne, Set.mem_setOf_eq] by_cases hx : x = 0 - · simp only [hx, and_true_iff, eq_self_iff_true, inv_zero, or_true_iff] + · simp only [hx, and_true, eq_self_iff_true, inv_zero, or_true] exact ⟨_, 1, rfl, one_ne_zero⟩ - · simp only [hx, or_false_iff, exists_eq_right] + · simp only [hx, or_false, exists_eq_right] exact ⟨_, rfl, hx⟩ theorem dickson_one_one_charP (p : ℕ) [Fact p.Prime] [CharP R p] : dickson 1 (1 : R) p = X ^ p := by diff --git a/Mathlib/RingTheory/Polynomial/Selmer.lean b/Mathlib/RingTheory/Polynomial/Selmer.lean index 765741980303a..c6f4ed8e30afd 100644 --- a/Mathlib/RingTheory/Polynomial/Selmer.lean +++ b/Mathlib/RingTheory/Polynomial/Selmer.lean @@ -34,7 +34,7 @@ theorem X_pow_sub_X_sub_one_irreducible_aux (z : ℂ) : ¬(z ^ n = z + 1 ∧ z ^ rw [← Nat.mod_add_div n 3, pow_add, pow_mul, h3, one_pow, mul_one] have : n % 3 < 3 := Nat.mod_lt n zero_lt_three interval_cases n % 3 <;> - simp only [this, pow_zero, pow_one, eq_self_iff_true, or_true_iff, true_or_iff] + simp only [this, pow_zero, pow_one, eq_self_iff_true, or_true, true_or] have z_ne_zero : z ≠ 0 := fun h => zero_ne_one ((zero_pow three_ne_zero).symm.trans (show (0 : ℂ) ^ 3 = 1 from h ▸ h3)) rcases key with (key | key | key) diff --git a/Mathlib/RingTheory/Polynomial/Tower.lean b/Mathlib/RingTheory/Polynomial/Tower.lean index 2c16e8138d977..5e49599ddbda9 100644 --- a/Mathlib/RingTheory/Polynomial/Tower.lean +++ b/Mathlib/RingTheory/Polynomial/Tower.lean @@ -56,7 +56,7 @@ theorem aeval_algebraMap_apply (x : A) (p : R[X]) : theorem aeval_algebraMap_eq_zero_iff [NoZeroSMulDivisors A B] [Nontrivial B] (x : A) (p : R[X]) : aeval (algebraMap A B x) p = 0 ↔ aeval x p = 0 := by rw [aeval_algebraMap_apply, Algebra.algebraMap_eq_smul_one, smul_eq_zero, - iff_false_intro (one_ne_zero' B), or_false_iff] + iff_false_intro (one_ne_zero' B), or_false] variable {B} diff --git a/Mathlib/RingTheory/RootsOfUnity/Basic.lean b/Mathlib/RingTheory/RootsOfUnity/Basic.lean index 04f0e45847dbc..fbfd625bc62c1 100644 --- a/Mathlib/RingTheory/RootsOfUnity/Basic.lean +++ b/Mathlib/RingTheory/RootsOfUnity/Basic.lean @@ -442,7 +442,7 @@ then there is a `b`-th primitive root of unity in `R`. -/ theorem pow {n : ℕ} {a b : ℕ} (hn : 0 < n) (h : IsPrimitiveRoot ζ n) (hprod : n = a * b) : IsPrimitiveRoot (ζ ^ a) b := by subst n - simp only [iff_def, ← pow_mul, h.pow_eq_one, eq_self_iff_true, true_and_iff] + simp only [iff_def, ← pow_mul, h.pow_eq_one, eq_self_iff_true, true_and] intro l hl -- Porting note: was `by rintro rfl; simpa only [Nat.not_lt_zero, zero_mul] using hn` have ha0 : a ≠ 0 := left_ne_zero_of_mul hn.ne' @@ -868,14 +868,14 @@ theorem card_primitiveRoots {ζ : R} {k : ℕ} (h : IsPrimitiveRoot ζ k) : · simp [h0] symm refine Finset.card_bij (fun i _ ↦ ζ ^ i) ?_ ?_ ?_ - · simp only [true_and_iff, and_imp, mem_filter, mem_range, mem_univ] + · simp only [and_imp, mem_filter, mem_range, mem_univ] rintro i - hi rw [mem_primitiveRoots (Nat.pos_of_ne_zero h0)] exact h.pow_of_coprime i hi.symm - · simp only [true_and_iff, and_imp, mem_filter, mem_range, mem_univ] + · simp only [and_imp, mem_filter, mem_range, mem_univ] rintro i hi - j hj - H exact h.pow_inj hi hj H - · simp only [exists_prop, true_and_iff, mem_filter, mem_range, mem_univ] + · simp only [exists_prop, mem_filter, mem_range, mem_univ] intro ξ hξ rw [mem_primitiveRoots (Nat.pos_of_ne_zero h0), h.isPrimitiveRoot_iff (Nat.pos_of_ne_zero h0)] at hξ @@ -898,7 +898,7 @@ theorem nthRoots_one_eq_biUnion_primitiveRoots' {ζ : R} {n : ℕ+} (h : IsPrimi · intro x simp only [nthRootsFinset, ← Multiset.toFinset_eq (nthRoots_one_nodup h), exists_prop, Finset.mem_biUnion, Finset.mem_filter, Finset.mem_range, mem_nthRoots, Finset.mem_mk, - Nat.mem_divisors, and_true_iff, Ne, PNat.ne_zero, PNat.pos, not_false_iff] + Nat.mem_divisors, and_true, Ne, PNat.ne_zero, PNat.pos, not_false_iff] rintro ⟨a, ⟨d, hd⟩, ha⟩ have hazero : 0 < a := by contrapose! hd with ha0 diff --git a/Mathlib/RingTheory/UniqueFactorizationDomain.lean b/Mathlib/RingTheory/UniqueFactorizationDomain.lean index df5977afdc34d..619ab4c415bcb 100644 --- a/Mathlib/RingTheory/UniqueFactorizationDomain.lean +++ b/Mathlib/RingTheory/UniqueFactorizationDomain.lean @@ -1216,9 +1216,9 @@ theorem prod_mono : ∀ {a b : FactorSet α}, a ≤ b → a.prod ≤ b.prod theorem FactorSet.prod_eq_zero_iff [Nontrivial α] (p : FactorSet α) : p.prod = 0 ↔ p = ⊤ := by unfold FactorSet at p induction p -- TODO: `induction_eliminator` doesn't work with `abbrev` - · simp only [iff_self_iff, eq_self_iff_true, Associates.prod_top] + · simp only [eq_self_iff_true, Associates.prod_top] · rw [prod_coe, Multiset.prod_eq_zero_iff, Multiset.mem_map, eq_false WithTop.coe_ne_top, - iff_false_iff, not_exists] + iff_false, not_exists] exact fun a => not_and_of_not_right _ a.prop.ne_zero section count @@ -1871,7 +1871,7 @@ noncomputable def fintypeSubtypeDvd {M : Type*} [CancelCommMonoidWithZero M] (((normalizedFactors y).powerset.toFinset ×ˢ (Finset.univ : Finset Mˣ)).image fun s => (s.snd : M) * s.fst.prod) fun x => ?_ - simp only [exists_prop, Finset.mem_image, Finset.mem_product, Finset.mem_univ, and_true_iff, + simp only [exists_prop, Finset.mem_image, Finset.mem_product, Finset.mem_univ, and_true, Multiset.mem_toFinset, Multiset.mem_powerset, exists_eq_right, Multiset.mem_map] constructor · rintro ⟨s, hs, rfl⟩ diff --git a/Mathlib/RingTheory/WittVector/MulCoeff.lean b/Mathlib/RingTheory/WittVector/MulCoeff.lean index 52cb2c116b25b..fa74982d2d6fa 100644 --- a/Mathlib/RingTheory/WittVector/MulCoeff.lean +++ b/Mathlib/RingTheory/WittVector/MulCoeff.lean @@ -118,11 +118,10 @@ theorem mul_polyOfInterest_aux1 (n : ℕ) : congr 1 have hsupp : (Finsupp.single i (p ^ (n - i))).support = {i} := by rw [Finsupp.support_eq_singleton] - simp only [and_true_iff, Finsupp.single_eq_same, eq_self_iff_true, Ne] + simp only [and_true, Finsupp.single_eq_same, eq_self_iff_true, Ne] exact pow_ne_zero _ hp.out.ne_zero simp only [bind₁_monomial, hsupp, Int.cast_natCast, prod_singleton, eq_intCast, - Finsupp.single_eq_same, C_pow, mul_eq_mul_left_iff, true_or_iff, eq_self_iff_true, - Int.cast_pow] + Finsupp.single_eq_same, C_pow, mul_eq_mul_left_iff, eq_self_iff_true, Int.cast_pow] · simp only [map_mul, bind₁_X_right] theorem mul_polyOfInterest_aux2 (n : ℕ) : diff --git a/Mathlib/RingTheory/WittVector/StructurePolynomial.lean b/Mathlib/RingTheory/WittVector/StructurePolynomial.lean index ad4f137a05748..5a89447cb88ff 100644 --- a/Mathlib/RingTheory/WittVector/StructurePolynomial.lean +++ b/Mathlib/RingTheory/WittVector/StructurePolynomial.lean @@ -370,7 +370,7 @@ theorem wittStructureRat_vars [Fintype idx] (Φ : MvPolynomial idx ℚ) (n : ℕ (wittStructureRat p Φ n).vars ⊆ Finset.univ ×ˢ Finset.range (n + 1) := by rw [wittStructureRat] intro x hx - simp only [Finset.mem_product, true_and_iff, Finset.mem_univ, Finset.mem_range] + simp only [Finset.mem_product, true_and, Finset.mem_univ, Finset.mem_range] obtain ⟨k, hk, hx'⟩ := mem_vars_bind₁ _ _ hx obtain ⟨i, -, hx''⟩ := mem_vars_bind₁ _ _ hx' obtain ⟨j, hj, rfl⟩ := mem_vars_rename _ _ hx'' diff --git a/Mathlib/RingTheory/WittVector/WittPolynomial.lean b/Mathlib/RingTheory/WittVector/WittPolynomial.lean index aab99256bb5a2..ea728fea38133 100644 --- a/Mathlib/RingTheory/WittVector/WittPolynomial.lean +++ b/Mathlib/RingTheory/WittVector/WittPolynomial.lean @@ -230,7 +230,7 @@ theorem xInTermsOfW_vars_aux (n : ℕ) : rw [xInTermsOfW_eq, mul_comm, vars_C_mul _ (Invertible.ne_zero _), vars_sub_of_disjoint, vars_X, range_succ, insert_eq] on_goal 1 => - simp only [true_and_iff, true_or_iff, eq_self_iff_true, mem_union, mem_singleton] + simp only [true_and, true_or, eq_self_iff_true, mem_union, mem_singleton] intro i rw [mem_union, mem_union] apply Or.imp id diff --git a/Mathlib/SetTheory/Cardinal/Basic.lean b/Mathlib/SetTheory/Cardinal/Basic.lean index 37627111901d9..e27e28d53e996 100644 --- a/Mathlib/SetTheory/Cardinal/Basic.lean +++ b/Mathlib/SetTheory/Cardinal/Basic.lean @@ -1425,7 +1425,7 @@ theorem nsmul_lt_aleph0_iff {n : ℕ} {a : Cardinal} : n • a < ℵ₀ ↔ n = cases n with | zero => simpa using nat_lt_aleph0 0 | succ n => - simp only [Nat.succ_ne_zero, false_or_iff] + simp only [Nat.succ_ne_zero, false_or] induction' n with n ih · simp rw [succ_nsmul, add_lt_aleph0_iff, ih, and_self_iff] diff --git a/Mathlib/SetTheory/Cardinal/Divisibility.lean b/Mathlib/SetTheory/Cardinal/Divisibility.lean index 658f7943b383d..4dac7f8531de2 100644 --- a/Mathlib/SetTheory/Cardinal/Divisibility.lean +++ b/Mathlib/SetTheory/Cardinal/Divisibility.lean @@ -134,7 +134,7 @@ theorem is_prime_iff {a : Cardinal} : Prime a ↔ ℵ₀ ≤ a ∨ ∃ p : ℕ, theorem isPrimePow_iff {a : Cardinal} : IsPrimePow a ↔ ℵ₀ ≤ a ∨ ∃ n : ℕ, a = n ∧ IsPrimePow n := by by_cases h : ℵ₀ ≤ a · simp [h, (prime_of_aleph0_le h).isPrimePow] - simp only [h, Nat.cast_inj, exists_eq_left', false_or_iff, isPrimePow_nat_iff] + simp only [h, Nat.cast_inj, exists_eq_left', false_or, isPrimePow_nat_iff] lift a to ℕ using not_le.mp h rw [isPrimePow_def] refine diff --git a/Mathlib/SetTheory/Ordinal/Arithmetic.lean b/Mathlib/SetTheory/Ordinal/Arithmetic.lean index ab6c9728bd05f..d7673f6897010 100644 --- a/Mathlib/SetTheory/Ordinal/Arithmetic.lean +++ b/Mathlib/SetTheory/Ordinal/Arithmetic.lean @@ -173,7 +173,7 @@ theorem pred_eq_iff_not_succ' {o} : pred o = o ↔ ∀ a, o ≠ succ a := by simpa using pred_eq_iff_not_succ theorem pred_lt_iff_is_succ {o} : pred o < o ↔ ∃ a, o = succ a := - Iff.trans (by simp only [le_antisymm_iff, pred_le_self, true_and_iff, not_le]) + Iff.trans (by simp only [le_antisymm_iff, pred_le_self, true_and, not_le]) (iff_not_comm.1 pred_eq_iff_not_succ).symm @[simp] @@ -572,15 +572,15 @@ instance monoid : Monoid Ordinal.{u} where Quotient.sound ⟨⟨punitProd _, @fun a b => by rcases a with ⟨⟨⟨⟩⟩, a⟩; rcases b with ⟨⟨⟨⟩⟩, b⟩ - simp only [Prod.lex_def, EmptyRelation, false_or_iff] - simp only [eq_self_iff_true, true_and_iff] + simp only [Prod.lex_def, EmptyRelation, false_or] + simp only [eq_self_iff_true, true_and] rfl⟩⟩ one_mul a := inductionOn a fun α r _ => Quotient.sound ⟨⟨prodPUnit _, @fun a b => by rcases a with ⟨a, ⟨⟨⟩⟩⟩; rcases b with ⟨b, ⟨⟨⟩⟩⟩ - simp only [Prod.lex_def, EmptyRelation, and_false_iff, or_false_iff] + simp only [Prod.lex_def, EmptyRelation, and_false, or_false] rfl⟩⟩ @[simp] @@ -625,8 +625,7 @@ instance leftDistribClass : LeftDistribClass Ordinal.{u} := simp only [Prod.lex_def, Sum.lex_inl_inl, Sum.Lex.sep, Sum.lex_inr_inl, Sum.lex_inr_inr, sumProdDistrib_apply_left, sumProdDistrib_apply_right, reduceCtorEq] <;> -- Porting note: `Sum.inr.inj_iff` is required. - simp only [Sum.inl.inj_iff, Sum.inr.inj_iff, - true_or_iff, false_and_iff, false_or_iff]⟩⟩⟩ + simp only [Sum.inl.inj_iff, Sum.inr.inj_iff, true_or, false_and, false_or]⟩⟩⟩ theorem mul_succ (a b : Ordinal) : a * succ b = a * b + a := mul_add_one a b @@ -686,11 +685,11 @@ private theorem mul_le_of_limit_aux {α β r s} [IsWellOrder α r] [IsWellOrder intro h by_cases e₁ : b = b₁ <;> by_cases e₂ : b = b₂ · substs b₁ b₂ - simpa only [subrel_val, Prod.lex_def, @irrefl _ s _ b, true_and_iff, false_or_iff, + simpa only [subrel_val, Prod.lex_def, @irrefl _ s _ b, true_and, false_or, eq_self_iff_true, dif_pos, Sum.lex_inr_inr] using h · subst b₁ simp only [subrel_val, Prod.lex_def, e₂, Prod.lex_def, dif_pos, subrel_val, eq_self_iff_true, - or_false_iff, dif_neg, not_false_iff, Sum.lex_inr_inl, false_and_iff] at h ⊢ + or_false, dif_neg, not_false_iff, Sum.lex_inr_inl, false_and] at h ⊢ cases' h₂ with _ _ _ _ h₂_h h₂_h <;> [exact asymm h h₂_h; exact e₂ rfl] · simp [e₂, dif_neg e₁, show b₂ ≠ b₁ from e₂ ▸ e₁] · simpa only [dif_neg e₁, dif_neg e₂, Prod.lex_def, subrel_val, Subtype.mk_eq_mk, @@ -794,8 +793,7 @@ theorem le_div {a b c : Ordinal} (c0 : c ≠ 0) : a ≤ b / c ↔ c * a ≤ b := | H₂ _ _ => rw [succ_le_iff, lt_div c0] | H₃ _ h₁ h₂ => revert h₁ h₂ - simp (config := { contextual := true }) only [mul_le_of_limit, limit_le, iff_self_iff, - forall_true_iff] + simp (config := { contextual := true }) only [mul_le_of_limit, limit_le, forall_true_iff] theorem div_lt {a b c : Ordinal} (b0 : b ≠ 0) : a / b < c ↔ a < b * c := lt_iff_lt_of_le_iff_le <| le_div b0 diff --git a/Mathlib/SetTheory/Ordinal/Notation.lean b/Mathlib/SetTheory/Ordinal/Notation.lean index caa7c0132c5b6..f6b982ade0d82 100644 --- a/Mathlib/SetTheory/Ordinal/Notation.lean +++ b/Mathlib/SetTheory/Ordinal/Notation.lean @@ -1017,7 +1017,7 @@ theorem fundamentalSequence_has_prop (o) : FundamentalSequenceProp o (fundamenta simp only [repr, iha, ihb, opow_lt_opow_iff_right one_lt_omega, add_lt_add_iff_left, add_zero, eq_self_iff_true, lt_add_iff_pos_right, lt_def, mul_one, Nat.cast_zero, Nat.cast_succ, Nat.succPNat_coe, opow_succ, opow_zero, mul_add_one, PNat.one_coe, succ_zero, - true_and_iff, _root_.zero_add, zero_def] + _root_.zero_add, zero_def] · decide · exact ⟨rfl, inferInstance⟩ · have := opow_pos (repr a') omega_pos diff --git a/Mathlib/SetTheory/Ordinal/Principal.lean b/Mathlib/SetTheory/Ordinal/Principal.lean index 8d13d1a443605..0f967f92bc0c3 100644 --- a/Mathlib/SetTheory/Ordinal/Principal.lean +++ b/Mathlib/SetTheory/Ordinal/Principal.lean @@ -200,7 +200,7 @@ theorem principal_add_iff_zero_or_omega_opow {o : Ordinal} : rcases eq_or_ne o 0 with (rfl | ho) · simp only [principal_zero, Or.inl] · rw [principal_add_iff_add_left_eq_self] - simp only [ho, false_or_iff] + simp only [ho, false_or] refine ⟨fun H => ⟨_, ((lt_or_eq_of_le (opow_log_le_self _ ho)).resolve_left fun h => ?_)⟩, fun ⟨b, e⟩ => e.symm ▸ fun a => add_omega_opow⟩ diff --git a/Mathlib/SetTheory/ZFC/Basic.lean b/Mathlib/SetTheory/ZFC/Basic.lean index a556cdab50f44..bede4eaab6232 100644 --- a/Mathlib/SetTheory/ZFC/Basic.lean +++ b/Mathlib/SetTheory/ZFC/Basic.lean @@ -1120,7 +1120,7 @@ theorem pair_injective : Function.Injective2 pair := fun x x' y y' H => by rw [mem_singleton.mp m] have he : x = y → y = y' := by rintro rfl - cases' (ae {x, y'}).2 (by simp only [eq_self_iff_true, or_true_iff]) with xy'x xy'xx + cases' (ae {x, y'}).2 (by simp only [eq_self_iff_true, or_true]) with xy'x xy'xx · rw [eq_comm, ← mem_singleton, ← xy'x, mem_pair] exact Or.inr rfl · simpa [eq_comm] using (ZFSet.ext_iff.1 xy'xx y').1 (by simp) @@ -1282,7 +1282,7 @@ theorem not_empty_hom (x : ZFSet.{u}) : ¬(∅ : Class.{u}) x := @[simp] theorem mem_univ {A : Class.{u}} : A ∈ univ.{u} ↔ ∃ x : ZFSet.{u}, ↑x = A := - exists_congr fun _ => and_true_iff _ + exists_congr fun _ => iff_of_eq (and_true _) @[simp] theorem mem_univ_hom (x : ZFSet.{u}) : univ.{u} x := @@ -1333,7 +1333,7 @@ def congToClass (x : Set Class.{u}) : Class.{u} := @[simp] theorem congToClass_empty : congToClass ∅ = ∅ := by ext z - simp only [congToClass, not_empty_hom, iff_false_iff] + simp only [congToClass, not_empty_hom, iff_false] exact Set.not_mem_empty z /-- Convert a class into a conglomerate (a collection of classes) -/ @@ -1391,7 +1391,7 @@ theorem coe_sep (p : Class.{u}) (x : ZFSet.{u}) : @[simp, norm_cast] theorem coe_empty : ↑(∅ : ZFSet.{u}) = (∅ : Class.{u}) := - ext fun y => iff_false_iff.2 <| ZFSet.not_mem_empty y + ext fun y => iff_false _ ▸ ZFSet.not_mem_empty y @[simp, norm_cast] theorem coe_insert (x y : ZFSet.{u}) : ↑(insert x y) = @insert ZFSet.{u} Class.{u} _ x y := diff --git a/Mathlib/Tactic/CC/Lemmas.lean b/Mathlib/Tactic/CC/Lemmas.lean index 0f66873939d73..12348c1e7b4cb 100644 --- a/Mathlib/Tactic/CC/Lemmas.lean +++ b/Mathlib/Tactic/CC/Lemmas.lean @@ -4,47 +4,45 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura -/ -import Mathlib.Init.Logic - /-! Lemmas use by the congruence closure module -/ namespace Mathlib.Tactic.CC theorem iff_eq_of_eq_true_left {a b : Prop} (h : a = True) : (a ↔ b) = b := - h.symm ▸ propext true_iff_iff + h.symm ▸ true_iff _ theorem iff_eq_of_eq_true_right {a b : Prop} (h : b = True) : (a ↔ b) = a := - h.symm ▸ propext iff_true_iff + h.symm ▸ iff_true _ theorem iff_eq_true_of_eq {a b : Prop} (h : a = b) : (a ↔ b) = True := - h ▸ propext (iff_self_iff _) + h ▸ iff_self _ theorem and_eq_of_eq_true_left {a b : Prop} (h : a = True) : (a ∧ b) = b := - h.symm ▸ propext (true_and_iff _) + h.symm ▸ true_and _ theorem and_eq_of_eq_true_right {a b : Prop} (h : b = True) : (a ∧ b) = a := - h.symm ▸ propext (and_true_iff _) + h.symm ▸ and_true _ theorem and_eq_of_eq_false_left {a b : Prop} (h : a = False) : (a ∧ b) = False := - h.symm ▸ propext (false_and_iff _) + h.symm ▸ false_and _ theorem and_eq_of_eq_false_right {a b : Prop} (h : b = False) : (a ∧ b) = False := - h.symm ▸ propext (and_false_iff _) + h.symm ▸ and_false _ theorem and_eq_of_eq {a b : Prop} (h : a = b) : (a ∧ b) = a := h ▸ propext and_self_iff theorem or_eq_of_eq_true_left {a b : Prop} (h : a = True) : (a ∨ b) = True := - h.symm ▸ propext (true_or_iff _) + h.symm ▸ true_or _ theorem or_eq_of_eq_true_right {a b : Prop} (h : b = True) : (a ∨ b) = True := - h.symm ▸ propext (or_true_iff _) + h.symm ▸ or_true _ theorem or_eq_of_eq_false_left {a b : Prop} (h : a = False) : (a ∨ b) = b := - h.symm ▸ propext (false_or_iff _) + h.symm ▸ false_or _ theorem or_eq_of_eq_false_right {a b : Prop} (h : b = False) : (a ∨ b) = a := - h.symm ▸ propext (or_false_iff _) + h.symm ▸ or_false _ theorem or_eq_of_eq {a b : Prop} (h : a = b) : (a ∨ b) = a := h ▸ propext or_self_iff diff --git a/Mathlib/Testing/SlimCheck/Functions.lean b/Mathlib/Testing/SlimCheck/Functions.lean index ddc576d88d833..12325734b9e55 100644 --- a/Mathlib/Testing/SlimCheck/Functions.lean +++ b/Mathlib/Testing/SlimCheck/Functions.lean @@ -177,8 +177,7 @@ def applyFinsupp (tf : TotalFunction α β) : α →₀ β where · intro h use (A.dlookup a).getD (0 : β) rw [← List.dlookup_dedupKeys] at h ⊢ - simp only [h, ← List.mem_dlookup_iff A.nodupKeys_dedupKeys, and_true_iff, not_false_iff, - Option.mem_def] + simp only [h, ← List.mem_dlookup_iff A.nodupKeys_dedupKeys, not_false_iff, Option.mem_def] cases haA : List.dlookup a A.dedupKeys · simp [haA] at h · simp @@ -324,12 +323,12 @@ theorem applyId_mem_iff [DecidableEq α] {xs ys : List α} (h₀ : List.Nodup xs dsimp [List.dlookup] at h₃; split_ifs at h₃ with h · rw [Option.some_inj] at h₃ subst x'; subst val - simp only [List.mem_cons, true_or_iff, eq_self_iff_true] + simp only [List.mem_cons, true_or, eq_self_iff_true] · cases' h₀ with _ _ h₀ h₅ cases' h₂ with _ _ h₂ h₄ have h₆ := Nat.succ.inj h₁ specialize xs_ih h₅ h₃ h₄ h₆ - simp only [Ne.symm h, xs_ih, List.mem_cons, false_or_iff] + simp only [Ne.symm h, xs_ih, List.mem_cons] suffices val ∈ ys by tauto erw [← Option.mem_def, List.mem_dlookup_iff] at h₃ · simp only [Prod.toSigma, List.mem_map, heq_iff_eq, Prod.exists] at h₃ @@ -454,7 +453,7 @@ protected theorem injective [DecidableEq α] (f : InjectiveFunction α) : Inject induction xs with | nil => simp only [List.zip_nil_right, List.map_nil] | cons xs_hd xs_tl xs_ih => - simp only [true_and_iff, Prod.toSigma, eq_self_iff_true, Sigma.eta, List.zip_cons_cons, + simp only [Prod.toSigma, eq_self_iff_true, Sigma.eta, List.zip_cons_cons, List.map, List.cons_inj_right] exact xs_ih revert hperm hnodup diff --git a/Mathlib/Topology/Algebra/Field.lean b/Mathlib/Topology/Algebra/Field.lean index 6a6d37b9907fc..7fc53eb7891e7 100644 --- a/Mathlib/Topology/Algebra/Field.lean +++ b/Mathlib/Topology/Algebra/Field.lean @@ -146,6 +146,6 @@ theorem IsPreconnected.eq_of_sq_eq [Field 𝕜] [HasContinuousInv₀ 𝕜] [Cont rcases hS.eq_or_eq_neg_of_sq_eq hf hg @hsq @hg_ne with (h | h) · exact h hx · rw [h _, Pi.neg_apply, neg_eq_iff_add_eq_zero, ← two_mul, mul_eq_zero, - iff_false_iff.2 (hg_ne _)] at hy' ⊢ <;> assumption + (iff_of_eq (iff_false _)).2 (hg_ne _)] at hy' ⊢ <;> assumption end Preconnected diff --git a/Mathlib/Topology/Algebra/Valued/NormedValued.lean b/Mathlib/Topology/Algebra/Valued/NormedValued.lean index 7ce1d9f1fee9d..dbe91b1484df4 100644 --- a/Mathlib/Topology/Algebra/Valued/NormedValued.lean +++ b/Mathlib/Topology/Algebra/Valued/NormedValued.lean @@ -97,7 +97,7 @@ def toNormedField : NormedField L := haveI : Nonempty { ε : ℝ // ε > 0 } := nonempty_Ioi_subtype ext U rw [hasBasis_iff.mp (Valued.hasBasis_uniformity L Γ₀), iInf_subtype', mem_iInf_of_directed] - · simp only [true_and_iff, mem_principal, Subtype.exists, gt_iff_lt, exists_prop] + · simp only [true_and, mem_principal, Subtype.exists, gt_iff_lt, exists_prop] refine ⟨fun ⟨ε, hε⟩ => ?_, fun ⟨r, hr_pos, hr⟩ => ?_⟩ · set δ : ℝ≥0 := hv.hom ε with hδ have hδ_pos : 0 < δ := by diff --git a/Mathlib/Topology/Algebra/Valued/ValuedField.lean b/Mathlib/Topology/Algebra/Valued/ValuedField.lean index 2d7836874a275..c9a159cdcd849 100644 --- a/Mathlib/Topology/Algebra/Valued/ValuedField.lean +++ b/Mathlib/Topology/Algebra/Valued/ValuedField.lean @@ -313,9 +313,9 @@ theorem closure_coe_completion_v_lt {γ : Γ₀ˣ} : suffices γ₀ ≠ 0 → (x ∈ closure ((↑) '' { x : K | v x < (γ : Γ₀) }) ↔ γ₀ < (γ : Γ₀)) by rcases eq_or_ne γ₀ 0 with h | h · simp only [h, (Valuation.zero_iff _).mp h, mem_setOf_eq, Valuation.map_zero, Units.zero_lt, - iff_true_iff] + iff_true] apply subset_closure - exact ⟨0, by simp only [mem_setOf_eq, Valuation.map_zero, Units.zero_lt, true_and_iff]; rfl⟩ + exact ⟨0, by simp only [mem_setOf_eq, Valuation.map_zero, Units.zero_lt, true_and]; rfl⟩ · exact this h intro h have hγ₀ : extension ⁻¹' {γ₀} ∈ 𝓝 x := diff --git a/Mathlib/Topology/Basic.lean b/Mathlib/Topology/Basic.lean index ce40739e493b3..388fd5adcf054 100644 --- a/Mathlib/Topology/Basic.lean +++ b/Mathlib/Topology/Basic.lean @@ -846,7 +846,7 @@ theorem tendsto_nhds {f : α → X} {l : Filter α} : theorem tendsto_atTop_nhds [Nonempty α] [SemilatticeSup α] {f : α → X} : Tendsto f atTop (𝓝 x) ↔ ∀ U : Set X, x ∈ U → IsOpen U → ∃ N, ∀ n, N ≤ n → f n ∈ U := (atTop_basis.tendsto_iff (nhds_basis_opens x)).trans <| by - simp only [and_imp, exists_prop, true_and_iff, mem_Ici] + simp only [and_imp, exists_prop, true_and, mem_Ici] theorem tendsto_const_nhds {f : Filter α} : Tendsto (fun _ : α => x) f (𝓝 x) := tendsto_nhds.mpr fun _ _ ha => univ_mem' fun _ => ha diff --git a/Mathlib/Topology/Bornology/Constructions.lean b/Mathlib/Topology/Bornology/Constructions.lean index e53a16415f67c..c38d27051cf5b 100644 --- a/Mathlib/Topology/Bornology/Constructions.lean +++ b/Mathlib/Topology/Bornology/Constructions.lean @@ -77,7 +77,7 @@ theorem isBounded_prod_of_nonempty (hne : Set.Nonempty (s ×ˢ t)) : theorem isBounded_prod : IsBounded (s ×ˢ t) ↔ s = ∅ ∨ t = ∅ ∨ IsBounded s ∧ IsBounded t := by rcases s.eq_empty_or_nonempty with (rfl | hs); · simp rcases t.eq_empty_or_nonempty with (rfl | ht); · simp - simp only [hs.ne_empty, ht.ne_empty, isBounded_prod_of_nonempty (hs.prod ht), false_or_iff] + simp only [hs.ne_empty, ht.ne_empty, isBounded_prod_of_nonempty (hs.prod ht), false_or] theorem isBounded_prod_self : IsBounded (s ×ˢ s) ↔ IsBounded s := by rcases s.eq_empty_or_nonempty with (rfl | hs); · simp @@ -109,7 +109,7 @@ theorem isBounded_pi_of_nonempty (hne : (pi univ S).Nonempty) : theorem isBounded_pi : IsBounded (pi univ S) ↔ (∃ i, S i = ∅) ∨ ∀ i, IsBounded (S i) := by by_cases hne : ∃ i, S i = ∅ · simp [hne, univ_pi_eq_empty_iff.2 hne] - · simp only [hne, false_or_iff] + · simp only [hne, false_or] simp only [not_exists, ← Ne.eq_def, ← nonempty_iff_ne_empty, ← univ_pi_nonempty_iff] at hne exact isBounded_pi_of_nonempty hne diff --git a/Mathlib/Topology/Category/Compactum.lean b/Mathlib/Topology/Category/Compactum.lean index 7cd821aa6d9d6..1bb55e129ab03 100644 --- a/Mathlib/Topology/Category/Compactum.lean +++ b/Mathlib/Topology/Category/Compactum.lean @@ -211,7 +211,7 @@ private theorem cl_cl {X : Compactum} (A : Set X) : cl (cl A) ⊆ cl A := by have claim1 : ∀ (B) (_ : B ∈ C0) (C) (_ : C ∈ C0), B ∩ C ∈ C0 := by rintro B ⟨Q, hQ, rfl⟩ C ⟨R, hR, rfl⟩ use Q ∩ R - simp only [and_true_iff, eq_self_iff_true, Set.preimage_inter] + simp only [and_true, eq_self_iff_true, Set.preimage_inter] exact inter_sets _ hQ hR -- All sets in C0 are nonempty. have claim2 : ∀ B ∈ C0, Set.Nonempty B := by diff --git a/Mathlib/Topology/Category/TopCat/Limits/Konig.lean b/Mathlib/Topology/Category/TopCat/Limits/Konig.lean index a8ca567c55146..0d6f7501547f2 100644 --- a/Mathlib/Topology/Category/TopCat/Limits/Konig.lean +++ b/Mathlib/Topology/Category/TopCat/Limits/Konig.lean @@ -133,8 +133,8 @@ theorem nonempty_limitCone_of_compact_t2_cofiltered_system (F : J ⥤ TopCat.{ma intro X Y f let G : FiniteDiagram J := ⟨{X, Y}, - {⟨X, Y, by simp only [true_or_iff, eq_self_iff_true, Finset.mem_insert], by - simp only [eq_self_iff_true, or_true_iff, Finset.mem_insert, Finset.mem_singleton], f⟩}⟩ + {⟨X, Y, by simp only [true_or, eq_self_iff_true, Finset.mem_insert], by + simp only [eq_self_iff_true, or_true, Finset.mem_insert, Finset.mem_singleton], f⟩}⟩ exact hu _ ⟨G, rfl⟩ (Finset.mem_singleton_self _) end TopologicalKonig diff --git a/Mathlib/Topology/Compactness/Compact.lean b/Mathlib/Topology/Compactness/Compact.lean index a381d9824966d..b83613bdfbf38 100644 --- a/Mathlib/Topology/Compactness/Compact.lean +++ b/Mathlib/Topology/Compactness/Compact.lean @@ -229,11 +229,11 @@ theorem IsCompact.elim_directed_family_closed {ι : Type v} [hι : Nonempty ι] hs.elim_directed_cover (compl ∘ t) (fun i => (htc i).isOpen_compl) (by simpa only [subset_def, not_forall, eq_empty_iff_forall_not_mem, mem_iUnion, exists_prop, - mem_inter_iff, not_and, iff_self_iff, mem_iInter, mem_compl_iff] using hst) + mem_inter_iff, not_and, mem_iInter, mem_compl_iff] using hst) (hdt.mono_comp _ fun _ _ => compl_subset_compl.mpr) ⟨t, by simpa only [subset_def, not_forall, eq_empty_iff_forall_not_mem, mem_iUnion, exists_prop, - mem_inter_iff, not_and, iff_self_iff, mem_iInter, mem_compl_iff] using ht⟩ + mem_inter_iff, not_and, mem_iInter, mem_compl_iff] using ht⟩ -- Porting note (#11215): TODO: reformulate using `Disjoint` /-- For every family of closed sets whose intersection avoids a compact set, diff --git a/Mathlib/Topology/Connected/Basic.lean b/Mathlib/Topology/Connected/Basic.lean index 8d8d2a542f23d..54b61ed7b660c 100644 --- a/Mathlib/Topology/Connected/Basic.lean +++ b/Mathlib/Topology/Connected/Basic.lean @@ -506,7 +506,7 @@ theorem isConnected_connectedComponent {x : α} : IsConnected (connectedComponen theorem isConnected_connectedComponentIn_iff {x : α} {F : Set α} : IsConnected (connectedComponentIn F x) ↔ x ∈ F := by simp_rw [← connectedComponentIn_nonempty_iff, IsConnected, isPreconnected_connectedComponentIn, - and_true_iff] + and_true] theorem IsPreconnected.subset_connectedComponent {x : α} {s : Set α} (H1 : IsPreconnected s) (H2 : x ∈ s) : s ⊆ connectedComponent x := fun _z hz => mem_sUnion_of_mem hz ⟨H1, H2⟩ diff --git a/Mathlib/Topology/Constructions.lean b/Mathlib/Topology/Constructions.lean index db254ae559843..9785d355d034f 100644 --- a/Mathlib/Topology/Constructions.lean +++ b/Mathlib/Topology/Constructions.lean @@ -728,7 +728,7 @@ theorem isOpen_prod_iff' {s : Set X} {t : Set Y} : rw [← snd_image_prod st.1 t] exact isOpenMap_snd _ H · intro H - simp only [st.1.ne_empty, st.2.ne_empty, not_false_iff, or_false_iff] at H + simp only [st.1.ne_empty, st.2.ne_empty, not_false_iff, or_false] at H exact H.1.prod H.2 theorem quotientMap_fst [Nonempty Y] : QuotientMap (Prod.fst : X × Y → X) := diff --git a/Mathlib/Topology/ContinuousOn.lean b/Mathlib/Topology/ContinuousOn.lean index c60aab4b1e2ea..69e02cd2a9319 100644 --- a/Mathlib/Topology/ContinuousOn.lean +++ b/Mathlib/Topology/ContinuousOn.lean @@ -685,8 +685,7 @@ theorem continuousWithinAt_singleton {f : α → β} {x : α} : ContinuousWithin @[simp] theorem continuousWithinAt_insert_self {f : α → β} {x : α} {s : Set α} : ContinuousWithinAt f (insert x s) x ↔ ContinuousWithinAt f s x := by - simp only [← singleton_union, continuousWithinAt_union, continuousWithinAt_singleton, - true_and_iff] + simp only [← singleton_union, continuousWithinAt_union, continuousWithinAt_singleton, true_and] alias ⟨_, ContinuousWithinAt.insert_self⟩ := continuousWithinAt_insert_self diff --git a/Mathlib/Topology/EMetricSpace/Defs.lean b/Mathlib/Topology/EMetricSpace/Defs.lean index 2d530cb9ad131..5afe753e4baa3 100644 --- a/Mathlib/Topology/EMetricSpace/Defs.lean +++ b/Mathlib/Topology/EMetricSpace/Defs.lean @@ -447,7 +447,7 @@ theorem tendsto_nhdsWithin_nhds {a b} : Tendsto f (𝓝[s] a) (𝓝 b) ↔ ∀ ε > 0, ∃ δ > 0, ∀ {x : α}, x ∈ s → edist x a < δ → edist (f x) b < ε := by rw [← nhdsWithin_univ b, tendsto_nhdsWithin_nhdsWithin] - simp only [mem_univ, true_and_iff] + simp only [mem_univ, true_and] theorem tendsto_nhds_nhds {a b} : Tendsto f (𝓝 a) (𝓝 b) ↔ ∀ ε > 0, ∃ δ > 0, ∀ ⦃x⦄, edist x a < δ → edist (f x) b < ε := @@ -491,7 +491,7 @@ theorem tendsto_nhds {f : Filter β} {u : β → α} {a : α} : theorem tendsto_atTop [Nonempty β] [SemilatticeSup β] {u : β → α} {a : α} : Tendsto u atTop (𝓝 a) ↔ ∀ ε > 0, ∃ N, ∀ n ≥ N, edist (u n) a < ε := (atTop_basis.tendsto_iff nhds_basis_eball).trans <| by - simp only [exists_prop, true_and_iff, mem_Ici, mem_ball] + simp only [exists_prop, true_and, mem_Ici, mem_ball] section Compact diff --git a/Mathlib/Topology/EMetricSpace/Diam.lean b/Mathlib/Topology/EMetricSpace/Diam.lean index 3c3206848162a..7fb1092cbfdc0 100644 --- a/Mathlib/Topology/EMetricSpace/Diam.lean +++ b/Mathlib/Topology/EMetricSpace/Diam.lean @@ -74,7 +74,7 @@ theorem diam_iUnion_mem_option {ι : Type*} (o : Option ι) (s : ι → Set α) theorem diam_insert : diam (insert x s) = max (⨆ y ∈ s, edist x y) (diam s) := eq_of_forall_ge_iff fun d => by simp only [diam_le_iff, forall_mem_insert, edist_self, edist_comm x, max_le_iff, iSup_le_iff, - zero_le, true_and_iff, forall_and, and_self_iff, ← and_assoc] + zero_le, true_and, forall_and, and_self_iff, ← and_assoc] theorem diam_pair : diam ({x, y} : Set α) = edist x y := by simp only [iSup_singleton, diam_insert, diam_singleton, ENNReal.max_zero_right] diff --git a/Mathlib/Topology/FiberBundle/Basic.lean b/Mathlib/Topology/FiberBundle/Basic.lean index bb2f8386ba3a3..39a1e8d2edc2b 100644 --- a/Mathlib/Topology/FiberBundle/Basic.lean +++ b/Mathlib/Topology/FiberBundle/Basic.lean @@ -439,7 +439,7 @@ def trivChange (i j : ι) : PartialHomeomorph (B × F) (B × F) where exacts [hx.1, ⟨⟨hx.1, hx.2⟩, hx.1⟩] right_inv' := by rintro ⟨x, v⟩ hx - simp only [prod_mk_mem_set_prod_eq, mem_inter_iff, and_true_iff, mem_univ] at hx + simp only [prod_mk_mem_set_prod_eq, mem_inter_iff, and_true, mem_univ] at hx dsimp only rw [Z.coordChange_comp, Z.coordChange_self] · exact hx.2 @@ -469,9 +469,9 @@ def localTrivAsPartialEquiv (i : ι) : PartialEquiv Z.TotalSpace (B × F) where invFun p := ⟨p.1, Z.coordChange i (Z.indexAt p.1) p.1 p.2⟩ toFun p := ⟨p.1, Z.coordChange (Z.indexAt p.1) i p.1 p.2⟩ map_source' p hp := by - simpa only [Set.mem_preimage, and_true_iff, Set.mem_univ, Set.prod_mk_mem_set_prod_eq] using hp + simpa only [Set.mem_preimage, and_true, Set.mem_univ, Set.prod_mk_mem_set_prod_eq] using hp map_target' p hp := by - simpa only [Set.mem_preimage, and_true_iff, Set.mem_univ, Set.mem_prod] using hp + simpa only [Set.mem_preimage, and_true, Set.mem_univ, Set.mem_prod] using hp left_inv' := by rintro ⟨x, v⟩ hx replace hx : x ∈ Z.baseSet i := hx @@ -479,7 +479,7 @@ def localTrivAsPartialEquiv (i : ι) : PartialEquiv Z.TotalSpace (B × F) where rw [Z.coordChange_comp, Z.coordChange_self] <;> apply_rules [mem_baseSet_at, mem_inter] right_inv' := by rintro ⟨x, v⟩ hx - simp only [prod_mk_mem_set_prod_eq, and_true_iff, mem_univ] at hx + simp only [prod_mk_mem_set_prod_eq, and_true, mem_univ] at hx dsimp only rw [Z.coordChange_comp, Z.coordChange_self] exacts [hx, ⟨⟨hx, Z.mem_baseSet_at _⟩, hx⟩] @@ -493,7 +493,7 @@ theorem mem_localTrivAsPartialEquiv_source (p : Z.TotalSpace) : theorem mem_localTrivAsPartialEquiv_target (p : B × F) : p ∈ (Z.localTrivAsPartialEquiv i).target ↔ p.1 ∈ Z.baseSet i := by erw [mem_prod] - simp only [and_true_iff, mem_univ] + simp only [and_true, mem_univ] theorem localTrivAsPartialEquiv_apply (p : Z.TotalSpace) : (Z.localTrivAsPartialEquiv i) p = ⟨p.1, Z.coordChange (Z.indexAt p.1) i p.1 p.2⟩ := @@ -508,9 +508,9 @@ theorem localTrivAsPartialEquiv_trans (i j : ι) : simp only [mem_localTrivAsPartialEquiv_target, mfld_simps] rfl · rintro ⟨x, v⟩ hx - simp only [trivChange, localTrivAsPartialEquiv, PartialEquiv.symm, true_and_iff, + simp only [trivChange, localTrivAsPartialEquiv, PartialEquiv.symm, Prod.mk.inj_iff, prod_mk_mem_set_prod_eq, PartialEquiv.trans_source, mem_inter_iff, - and_true_iff, mem_preimage, proj, mem_univ, eq_self_iff_true, (· ∘ ·), + mem_preimage, proj, mem_univ, eq_self_iff_true, (· ∘ ·), PartialEquiv.coe_trans, TotalSpace.proj] at hx ⊢ simp only [Z.coordChange_comp, hx, mem_inter_iff, and_self_iff, mem_baseSet_at] diff --git a/Mathlib/Topology/FiberBundle/Constructions.lean b/Mathlib/Topology/FiberBundle/Constructions.lean index 222d8ca96b186..8a7ea9e221f3c 100644 --- a/Mathlib/Topology/FiberBundle/Constructions.lean +++ b/Mathlib/Topology/FiberBundle/Constructions.lean @@ -144,7 +144,7 @@ theorem Prod.continuous_to_fun : ContinuousOn (Prod.toFun' e₁ e₂) · rw [e₁.source_eq, e₂.source_eq] exact mapsTo_preimage _ _ rintro ⟨b, v₁, v₂⟩ ⟨hb₁, _⟩ - simp only [f₃, Prod.toFun', Prod.mk.inj_iff, Function.comp_apply, and_true_iff] + simp only [f₃, Prod.toFun', Prod.mk.inj_iff, Function.comp_apply, and_true] rw [e₁.coe_fst] rw [e₁.source_eq, mem_preimage] exact hb₁ @@ -309,7 +309,7 @@ noncomputable def Trivialization.pullback (e : Trivialization F (π F E)) (f : K target := (f ⁻¹' e.baseSet) ×ˢ univ map_source' x h := by simp_rw [e.source_eq, mem_preimage, Pullback.lift_proj] at h - simp_rw [prod_mk_mem_set_prod_eq, mem_univ, and_true_iff, mem_preimage, h] + simp_rw [prod_mk_mem_set_prod_eq, mem_univ, and_true, mem_preimage, h] map_target' y h := by rw [mem_prod, mem_preimage] at h simp_rw [e.source_eq, mem_preimage, Pullback.lift_proj, h.1] @@ -317,7 +317,7 @@ noncomputable def Trivialization.pullback (e : Trivialization F (π F E)) (f : K simp_rw [mem_preimage, e.mem_source, Pullback.lift_proj] at h simp_rw [Pullback.lift, e.symm_apply_apply_mk h] right_inv' x h := by - simp_rw [mem_prod, mem_preimage, mem_univ, and_true_iff] at h + simp_rw [mem_prod, mem_preimage, mem_univ, and_true] at h simp_rw [Pullback.lift_mk, e.apply_mk_symm h] open_source := by simp_rw [e.source_eq, ← preimage_comp] diff --git a/Mathlib/Topology/FiberBundle/Trivialization.lean b/Mathlib/Topology/FiberBundle/Trivialization.lean index 72f78dd20048f..87a1362341a73 100644 --- a/Mathlib/Topology/FiberBundle/Trivialization.lean +++ b/Mathlib/Topology/FiberBundle/Trivialization.lean @@ -168,7 +168,7 @@ theorem preimage_symm_proj_inter (s : Set B) : e.toPartialEquiv.symm ⁻¹' (proj ⁻¹' s) ∩ e.baseSet ×ˢ univ = (s ∩ e.baseSet) ×ˢ univ := by ext ⟨x, y⟩ suffices x ∈ e.baseSet → (proj (e.toPartialEquiv.symm (x, y)) ∈ s ↔ x ∈ s) by - simpa only [prod_mk_mem_set_prod_eq, mem_inter_iff, and_true_iff, mem_univ, and_congr_left_iff] + simpa only [prod_mk_mem_set_prod_eq, mem_inter_iff, and_true, mem_univ, and_congr_left_iff] intro h rw [e.proj_symm_apply' h] diff --git a/Mathlib/Topology/Instances/ENNReal.lean b/Mathlib/Topology/Instances/ENNReal.lean index 631eda52d1417..b01cfd6e7865d 100644 --- a/Mathlib/Topology/Instances/ENNReal.lean +++ b/Mathlib/Topology/Instances/ENNReal.lean @@ -390,11 +390,11 @@ protected theorem continuous_pow (n : ℕ) : Continuous fun a : ℝ≥0∞ => a simp_rw [pow_add, pow_one, continuous_iff_continuousAt] intro x refine ENNReal.Tendsto.mul (IH.tendsto _) ?_ tendsto_id ?_ <;> by_cases H : x = 0 - · simp only [H, zero_ne_top, Ne, or_true_iff, not_false_iff] + · simp only [H, zero_ne_top, Ne, or_true, not_false_iff] · exact Or.inl fun h => H (pow_eq_zero h) - · simp only [H, pow_eq_top_iff, zero_ne_top, false_or_iff, eq_self_iff_true, not_true, Ne, - not_false_iff, false_and_iff] - · simp only [H, true_or_iff, Ne, not_false_iff] + · simp only [H, pow_eq_top_iff, zero_ne_top, false_or, eq_self_iff_true, not_true, Ne, + not_false_iff, false_and] + · simp only [H, true_or, Ne, not_false_iff] theorem continuousOn_sub : ContinuousOn (fun p : ℝ≥0∞ × ℝ≥0∞ => p.fst - p.snd) { p : ℝ≥0∞ × ℝ≥0∞ | p ≠ ⟨∞, ∞⟩ } := by @@ -406,7 +406,7 @@ theorem continuousOn_sub : theorem continuous_sub_left {a : ℝ≥0∞} (a_ne_top : a ≠ ∞) : Continuous (a - ·) := by change Continuous (Function.uncurry Sub.sub ∘ (a, ·)) refine continuousOn_sub.comp_continuous (Continuous.Prod.mk a) fun x => ?_ - simp only [a_ne_top, Ne, mem_setOf_eq, Prod.mk.inj_iff, false_and_iff, not_false_iff] + simp only [a_ne_top, Ne, mem_setOf_eq, Prod.mk.inj_iff, false_and, not_false_iff] theorem continuous_nnreal_sub {a : ℝ≥0} : Continuous fun x : ℝ≥0∞ => (a : ℝ≥0∞) - x := continuous_sub_left coe_ne_top @@ -423,7 +423,7 @@ theorem continuous_sub_right (a : ℝ≥0∞) : Continuous fun x : ℝ≥0∞ => · rw [show (fun x => x - a) = (fun p : ℝ≥0∞ × ℝ≥0∞ => p.fst - p.snd) ∘ fun x => ⟨x, a⟩ by rfl] apply ContinuousOn.comp_continuous continuousOn_sub (continuous_id'.prod_mk continuous_const) intro x - simp only [a_infty, Ne, mem_setOf_eq, Prod.mk.inj_iff, and_false_iff, not_false_iff] + simp only [a_infty, Ne, mem_setOf_eq, Prod.mk.inj_iff, and_false, not_false_iff] protected theorem Tendsto.pow {f : Filter α} {m : α → ℝ≥0∞} {a : ℝ≥0∞} {n : ℕ} (hm : Tendsto m f (𝓝 a)) : Tendsto (fun x => m x ^ n) f (𝓝 (a ^ n)) := @@ -805,7 +805,7 @@ theorem tsum_const_eq_top_of_ne_zero {α : Type*} [Infinite α] {c : ℝ≥0∞} ∑' _ : α, c = ∞ := by have A : Tendsto (fun n : ℕ => (n : ℝ≥0∞) * c) atTop (𝓝 (∞ * c)) := by apply ENNReal.Tendsto.mul_const tendsto_nat_nhds_top - simp only [true_or_iff, top_ne_zero, Ne, not_false_iff] + simp only [true_or, top_ne_zero, Ne, not_false_iff] have B : ∀ n : ℕ, (n : ℝ≥0∞) * c ≤ ∑' _ : α, c := fun n => by rcases Infinite.exists_subset_card_eq α n with ⟨s, hs⟩ simpa [hs] using @ENNReal.sum_le_tsum α (fun _ => c) s @@ -1236,7 +1236,7 @@ open EMetric theorem tendsto_iff_edist_tendsto_0 {l : Filter β} {f : β → α} {y : α} : Tendsto f l (𝓝 y) ↔ Tendsto (fun x => edist (f x) y) l (𝓝 0) := by simp only [EMetric.nhds_basis_eball.tendsto_right_iff, EMetric.mem_ball, - @tendsto_order ℝ≥0∞ β _ _, forall_prop_of_false ENNReal.not_lt_zero, forall_const, true_and_iff] + @tendsto_order ℝ≥0∞ β _ _, forall_prop_of_false ENNReal.not_lt_zero, forall_const, true_and] /-- Yet another metric characterization of Cauchy sequences on integers. This one is often the most efficient. -/ diff --git a/Mathlib/Topology/Maps/Basic.lean b/Mathlib/Topology/Maps/Basic.lean index e50f6343087cb..de7a4ffef4499 100644 --- a/Mathlib/Topology/Maps/Basic.lean +++ b/Mathlib/Topology/Maps/Basic.lean @@ -521,7 +521,7 @@ theorem openEmbedding_iff_embedding_open : theorem openEmbedding_of_continuous_injective_open (h₁ : Continuous f) (h₂ : Injective f) (h₃ : IsOpenMap f) : OpenEmbedding f := by - simp only [openEmbedding_iff_embedding_open, embedding_iff, inducing_iff_nhds, *, and_true_iff] + simp only [openEmbedding_iff_embedding_open, embedding_iff, inducing_iff_nhds, *, and_true] exact fun x => le_antisymm (h₁.tendsto _).le_comap (@comap_map _ _ (𝓝 x) _ h₂ ▸ comap_mono (h₃.nhds_le _)) diff --git a/Mathlib/Topology/MetricSpace/Completion.lean b/Mathlib/Topology/MetricSpace/Completion.lean index 8d667f11c999c..1ffdfa1fba1b7 100644 --- a/Mathlib/Topology/MetricSpace/Completion.lean +++ b/Mathlib/Topology/MetricSpace/Completion.lean @@ -103,7 +103,7 @@ protected theorem mem_uniformity_dist (s : Set (Completion α × Completion α)) · have Z := hε (not_le.1 h) simp only [Set.mem_setOf_eq] at Z exact Or.inr Z - simp only [not_le.mpr hxy, false_or_iff, not_le] at this + simp only [not_le.mpr hxy, false_or, not_le] at this exact ts this · /- Start from a set `s` containing an ε-neighborhood of the diagonal in `Completion α`. To show that it is an entourage, we use the fact that `dist` is uniformly continuous on diff --git a/Mathlib/Topology/MetricSpace/GromovHausdorff.lean b/Mathlib/Topology/MetricSpace/GromovHausdorff.lean index acc49b78007c1..8b58e2adf2ada 100644 --- a/Mathlib/Topology/MetricSpace/GromovHausdorff.lean +++ b/Mathlib/Topology/MetricSpace/GromovHausdorff.lean @@ -282,7 +282,7 @@ theorem hausdorffDist_optimal {X : Type u} [MetricSpace X] [CompactSpace X] [Non let F : (X ⊕ Y) × (X ⊕ Y) → ℝ := fun p => dist (f p.1) (f p.2) -- check that the induced "distance" is a candidate have Fgood : F ∈ candidates X Y := by - simp only [F, candidates, forall_const, and_true_iff, add_comm, eq_self_iff_true, + simp only [F, candidates, forall_const, add_comm, eq_self_iff_true, dist_eq_zero, and_self_iff, Set.mem_setOf_eq] repeat' constructor · exact fun x y => @@ -398,7 +398,7 @@ instance : MetricSpace GHSpace where · exact ⟨0, by rintro b ⟨⟨u, v⟩, -, rfl⟩; exact hausdorffDist_nonneg⟩ · simp only [mem_image, mem_prod, mem_setOf_eq, Prod.exists] exists y, y - simpa only [and_self_iff, hausdorffDist_self_zero, eq_self_iff_true, and_true_iff] + simpa only [and_self_iff, hausdorffDist_self_zero, eq_self_iff_true, and_true] · apply le_csInf · exact Set.Nonempty.image _ <| Set.Nonempty.prod ⟨y, hy⟩ ⟨y, hy⟩ · rintro b ⟨⟨u, v⟩, -, rfl⟩; exact hausdorffDist_nonneg diff --git a/Mathlib/Topology/MetricSpace/PiNat.lean b/Mathlib/Topology/MetricSpace/PiNat.lean index b001d3c706ab2..8c9627c5ed747 100644 --- a/Mathlib/Topology/MetricSpace/PiNat.lean +++ b/Mathlib/Topology/MetricSpace/PiNat.lean @@ -786,7 +786,7 @@ theorem min_dist_le_dist_pi (x y : ∀ i, F i) (i : ι) : theorem dist_le_dist_pi_of_dist_lt {x y : ∀ i, F i} {i : ι} (h : dist x y < (1 / 2) ^ encode i) : dist (x i) (y i) ≤ dist x y := by - simpa only [not_le.2 h, false_or_iff] using min_le_iff.1 (min_dist_le_dist_pi x y i) + simpa only [not_le.2 h, false_or] using min_le_iff.1 (min_dist_le_dist_pi x y i) open Topology Filter NNReal diff --git a/Mathlib/Topology/MetricSpace/Pseudo/Defs.lean b/Mathlib/Topology/MetricSpace/Pseudo/Defs.lean index cfaf85d8f63ea..89df0d0b39ba3 100644 --- a/Mathlib/Topology/MetricSpace/Pseudo/Defs.lean +++ b/Mathlib/Topology/MetricSpace/Pseudo/Defs.lean @@ -766,7 +766,7 @@ theorem tendsto_nhdsWithin_nhds [PseudoMetricSpace β] {f : α → β} {a b} : Tendsto f (𝓝[s] a) (𝓝 b) ↔ ∀ ε > 0, ∃ δ > 0, ∀ {x : α}, x ∈ s → dist x a < δ → dist (f x) b < ε := by rw [← nhdsWithin_univ b, tendsto_nhdsWithin_nhdsWithin] - simp only [mem_univ, true_and_iff] + simp only [mem_univ, true_and] theorem tendsto_nhds_nhds [PseudoMetricSpace β] {f : α → β} {a b} : Tendsto f (𝓝 a) (𝓝 b) ↔ ∀ ε > 0, ∃ δ > 0, ∀ {x : α}, dist x a < δ → dist (f x) b < ε := diff --git a/Mathlib/Topology/Metrizable/Uniformity.lean b/Mathlib/Topology/Metrizable/Uniformity.lean index 51973d61b6486..efdf16b745cec 100644 --- a/Mathlib/Topology/Metrizable/Uniformity.lean +++ b/Mathlib/Topology/Metrizable/Uniformity.lean @@ -206,7 +206,7 @@ protected theorem UniformSpace.metrizable_uniformity (X : Type*) [UniformSpace X split_ifs with h · rw [← not_forall] at h simp [h, pow_eq_zero_iff'] - · simpa only [not_exists, Classical.not_not, eq_self_iff_true, true_iff_iff] using h + · simpa only [not_exists, Classical.not_not, eq_self_iff_true, true_iff] using h have hd_symm : ∀ x y, d x y = d y x := by intro x y simp only [d, @SymmetricRel.mk_mem_comm _ _ (hU_symm _) x y] diff --git a/Mathlib/Topology/Order/IntermediateValue.lean b/Mathlib/Topology/Order/IntermediateValue.lean index 06c7281cea152..9b86197c50a9e 100644 --- a/Mathlib/Topology/Order/IntermediateValue.lean +++ b/Mathlib/Topology/Order/IntermediateValue.lean @@ -443,7 +443,7 @@ theorem setOf_isPreconnected_eq_of_ordered : (range Ici ∪ range Ioi ∪ range Iic ∪ range Iio ∪ {univ, ∅}) := by refine Subset.antisymm setOf_isPreconnected_subset_of_ordered ?_ simp only [subset_def, forall_mem_range, uncurry, or_imp, forall_and, mem_union, - mem_setOf_eq, insert_eq, mem_singleton_iff, forall_eq, forall_true_iff, and_true_iff, + mem_setOf_eq, insert_eq, mem_singleton_iff, forall_eq, forall_true_iff, and_true, isPreconnected_Icc, isPreconnected_Ico, isPreconnected_Ioc, isPreconnected_Ioo, isPreconnected_Ioi, isPreconnected_Iio, isPreconnected_Ici, isPreconnected_Iic, isPreconnected_univ, isPreconnected_empty] diff --git a/Mathlib/Topology/PartialHomeomorph.lean b/Mathlib/Topology/PartialHomeomorph.lean index 75f77fa794bf8..98e7d2e0d5f74 100644 --- a/Mathlib/Topology/PartialHomeomorph.lean +++ b/Mathlib/Topology/PartialHomeomorph.lean @@ -370,7 +370,7 @@ theorem preimage_eventuallyEq_target_inter_preimage_inter {e : PartialHomeomorph filter_upwards [e.open_source.mem_nhds hxe, mem_nhdsWithin_iff_eventually.mp (hf.preimage_mem_nhdsWithin ht)] intro y hy hyu - simp_rw [mem_inter_iff, mem_preimage, mem_inter_iff, e.mapsTo hy, true_and_iff, iff_self_and, + simp_rw [mem_inter_iff, mem_preimage, mem_inter_iff, e.mapsTo hy, true_and, iff_self_and, e.left_inv hy, iff_true_intro hyu] theorem isOpen_inter_preimage {s : Set Y} (hs : IsOpen s) : IsOpen (e.source ∩ e ⁻¹' s) := diff --git a/Mathlib/Topology/Sober.lean b/Mathlib/Topology/Sober.lean index 05385bd2df372..a4717443959a5 100644 --- a/Mathlib/Topology/Sober.lean +++ b/Mathlib/Topology/Sober.lean @@ -91,7 +91,7 @@ end IsGenericPoint theorem isGenericPoint_iff_forall_closed (hS : IsClosed S) (hxS : x ∈ S) : IsGenericPoint x S ↔ ∀ Z : Set α, IsClosed Z → x ∈ Z → S ⊆ Z := by have : closure {x} ⊆ S := closure_minimal (singleton_subset_iff.2 hxS) hS - simp_rw [IsGenericPoint, subset_antisymm_iff, this, true_and_iff, closure, subset_sInter_iff, + simp_rw [IsGenericPoint, subset_antisymm_iff, this, true_and, closure, subset_sInter_iff, mem_setOf_eq, and_imp, singleton_subset_iff] end genericPoint diff --git a/Mathlib/Topology/StoneCech.lean b/Mathlib/Topology/StoneCech.lean index ba2b4232abec5..34a2c255120e0 100644 --- a/Mathlib/Topology/StoneCech.lean +++ b/Mathlib/Topology/StoneCech.lean @@ -100,7 +100,7 @@ instance Ultrafilter.t2Space : T2Space (Ultrafilter α) := instance : TotallyDisconnectedSpace (Ultrafilter α) := by rw [totallyDisconnectedSpace_iff_connectedComponent_singleton] intro A - simp only [Set.eq_singleton_iff_unique_mem, mem_connectedComponent, true_and_iff] + simp only [Set.eq_singleton_iff_unique_mem, mem_connectedComponent, true_and] intro B hB rw [← Ultrafilter.coe_le_coe] intro s hs diff --git a/Mathlib/Topology/UniformSpace/Cauchy.lean b/Mathlib/Topology/UniformSpace/Cauchy.lean index b7f78b1b17c0f..7907f4159bea3 100644 --- a/Mathlib/Topology/UniformSpace/Cauchy.lean +++ b/Mathlib/Topology/UniformSpace/Cauchy.lean @@ -282,7 +282,7 @@ theorem Filter.HasBasis.cauchySeq_iff {γ} [Nonempty β] [SemilatticeSup β] {u CauchySeq u ↔ ∀ i, p i → ∃ N, ∀ m, N ≤ m → ∀ n, N ≤ n → (u m, u n) ∈ s i := by rw [cauchySeq_iff_tendsto, ← prod_atTop_atTop_eq] refine (atTop_basis.prod_self.tendsto_iff h).trans ?_ - simp only [exists_prop, true_and_iff, MapsTo, preimage, subset_def, Prod.forall, mem_prod_eq, + simp only [exists_prop, true_and, MapsTo, preimage, subset_def, Prod.forall, mem_prod_eq, mem_setOf_eq, mem_Ici, and_imp, Prod.map, @forall_swap (_ ≤ _) β] theorem Filter.HasBasis.cauchySeq_iff' {γ} [Nonempty β] [SemilatticeSup β] {u : β → α} diff --git a/Mathlib/Topology/UniformSpace/UniformConvergence.lean b/Mathlib/Topology/UniformSpace/UniformConvergence.lean index a5dbd62933646..a363dd73b8729 100644 --- a/Mathlib/Topology/UniformSpace/UniformConvergence.lean +++ b/Mathlib/Topology/UniformSpace/UniformConvergence.lean @@ -481,7 +481,7 @@ theorem UniformCauchySeqOn.prod' {β' : Type*} [UniformSpace β'] {F' : ι → a Cauchy sequence. -/ theorem UniformCauchySeqOn.cauchy_map [hp : NeBot p] (hf : UniformCauchySeqOn F p s) (hx : x ∈ s) : Cauchy (map (fun i => F i x) p) := by - simp only [cauchy_map_iff, hp, true_and_iff] + simp only [cauchy_map_iff, hp, true_and] intro u hu rw [mem_map] filter_upwards [hf u hu] with p hp using hp x hx