diff --git a/Mathlib/AlgebraicTopology/SimplicialSet/Basic.lean b/Mathlib/AlgebraicTopology/SimplicialSet/Basic.lean index 32b22e7b91ce50..814de7455bce4a 100644 --- a/Mathlib/AlgebraicTopology/SimplicialSet/Basic.lean +++ b/Mathlib/AlgebraicTopology/SimplicialSet/Basic.lean @@ -221,8 +221,7 @@ def const (n : ℕ) (i k : Fin (n+3)) (m : SimplexCategoryᵒᵖ) : Λ[n+2, i].o This edge only exists if `{i, a, b}` has cardinality less than `n`. -/ @[simps] -def edge (n : ℕ) (i a b : Fin (n+1)) (hab : a ≤ b) (H : Finset.card {i, a, b} ≤ n) : - Λ[n, i] _[1] := by +def edge (n : ℕ) (i a b : Fin (n+1)) (hab : a ≤ b) (H : #{i, a, b} ≤ n) : Λ[n, i] _[1] := by refine ⟨standardSimplex.edge n a b hab, ?range⟩ case range => suffices ∃ x, ¬i = x ∧ ¬a = x ∧ ¬b = x by diff --git a/Mathlib/Analysis/BoxIntegral/Basic.lean b/Mathlib/Analysis/BoxIntegral/Basic.lean index 77776a456f6e39..44c6028cdd79de 100644 --- a/Mathlib/Analysis/BoxIntegral/Basic.lean +++ b/Mathlib/Analysis/BoxIntegral/Basic.lean @@ -512,8 +512,8 @@ theorem dist_integralSum_sum_integral_le_of_memBaseSet_of_iUnion_eq (h : Integra -- Let us prove that the distance is less than or equal to `ε + δ` for all positive `δ`. refine le_of_forall_pos_le_add fun δ δ0 => ?_ -- First we choose some constants. - set δ' : ℝ := δ / (π₀.boxes.card + 1) - have H0 : 0 < (π₀.boxes.card + 1 : ℝ) := Nat.cast_add_one_pos _ + set δ' : ℝ := δ / (#π₀.boxes + 1) + have H0 : 0 < (#π₀.boxes + 1 : ℝ) := Nat.cast_add_one_pos _ have δ'0 : 0 < δ' := div_pos δ0 H0 set C := max π₀.distortion π₀.compl.distortion /- Next we choose a tagged partition of each `J ∈ π₀` such that the integral sum of `f` over this @@ -665,7 +665,7 @@ theorem integrable_of_bounded_and_ae_continuousWithinAt [CompleteSpace E] {I : B let t₂ (J : Box ι) : ℝⁿ := (π₂.infPrepartition π₁.toPrepartition).tag J let B := (π₁.toPrepartition ⊓ π₂.toPrepartition).boxes classical - let B' := B.filter (fun J ↦ J.toSet ⊆ U) + let B' := {J ∈ B | J.toSet ⊆ U} have hB' : B' ⊆ B := B.filter_subset (fun J ↦ J.toSet ⊆ U) have μJ_ne_top : ∀ J ∈ B, μ J ≠ ⊤ := fun J hJ ↦ lt_top_iff_ne_top.1 <| lt_of_le_of_lt (μ.mono (Prepartition.le_of_mem' _ J hJ)) μI @@ -795,7 +795,7 @@ theorem HasIntegral.of_bRiemann_eq_false_of_forall_isLittleO (hl : l.bRiemann = /- For the boxes such that `π.tag J ∈ s`, we use the fact that at most `2 ^ #ι` boxes have the same tag. -/ specialize hlH hsne - have : ∀ J ∈ π.boxes.filter fun J => π.tag J ∈ s, + have : ∀ J ∈ {J ∈ π.boxes | π.tag J ∈ s}, dist (vol J (f <| π.tag J)) (g J) ≤ εs (π.tag J) := fun J hJ ↦ by rw [Finset.mem_filter] at hJ; cases' hJ with hJ hJs refine Hδ₁ c _ ⟨π.tag_mem_Icc _, hJs⟩ _ (hεs0 _) _ (π.le_of_mem' _ hJ) ?_ @@ -815,9 +815,9 @@ theorem HasIntegral.of_bRiemann_eq_false_of_forall_isLittleO (hl : l.bRiemann = In this case the estimate is straightforward. -/ -- Porting note: avoided strange elaboration issues by rewriting using `calc` calc - dist (∑ J ∈ π.boxes.filter (¬tag π · ∈ s), vol J (f (tag π J))) - (∑ J ∈ π.boxes.filter (¬tag π · ∈ s), g J) - ≤ ∑ J ∈ π.boxes.filter (¬tag π · ∈ s), ε' * B J := dist_sum_sum_le_of_le _ fun J hJ ↦ by + dist (∑ J ∈ π.boxes with ¬tag π J ∈ s, vol J (f (tag π J))) + (∑ J ∈ π.boxes with ¬tag π J ∈ s, g J) + ≤ ∑ J ∈ π.boxes with ¬tag π J ∈ s, ε' * B J := dist_sum_sum_le_of_le _ fun J hJ ↦ by rw [Finset.mem_filter] at hJ; cases' hJ with hJ hJs refine Hδ₂ c _ ⟨π.tag_mem_Icc _, hJs⟩ _ ε'0 _ (π.le_of_mem' _ hJ) ?_ (fun hH => hπδ.2 hH J hJ) fun hD => (Finset.le_sup hJ).trans (hπδ.3 hD) diff --git a/Mathlib/Analysis/BoxIntegral/Partition/Basic.lean b/Mathlib/Analysis/BoxIntegral/Partition/Basic.lean index c8af644f555c90..1686d1eac27e6f 100644 --- a/Mathlib/Analysis/BoxIntegral/Partition/Basic.lean +++ b/Mathlib/Analysis/BoxIntegral/Partition/Basic.lean @@ -177,7 +177,7 @@ theorem injOn_setOf_mem_Icc_setOf_lower_eq (x : ι → ℝ) : /-- The set of boxes of a prepartition that contain `x` in their closures has cardinality at most `2 ^ Fintype.card ι`. -/ theorem card_filter_mem_Icc_le [Fintype ι] (x : ι → ℝ) : - (π.boxes.filter fun J : Box ι => x ∈ Box.Icc J).card ≤ 2 ^ Fintype.card ι := by + #{J ∈ π.boxes | x ∈ Box.Icc J} ≤ 2 ^ Fintype.card ι := by rw [← Fintype.card_set] refine Finset.card_le_card_of_injOn (fun J : Box ι => { i | J.lower i = x i }) (fun _ _ => Finset.mem_univ _) ?_ @@ -517,7 +517,7 @@ instance : SemilatticeInf (Prepartition I) := /-- The prepartition with boxes `{J ∈ π | p J}`. -/ @[simps] def filter (π : Prepartition I) (p : Box ι → Prop) : Prepartition I where - boxes := π.boxes.filter p + boxes := {J ∈ π.boxes | p J} le_of_mem' _ hJ := π.le_of_mem (mem_filter.1 hJ).1 pairwiseDisjoint _ h₁ _ h₂ := π.disjoint_coe_of_mem (mem_filter.1 h₁).1 (mem_filter.1 h₂).1 diff --git a/Mathlib/Analysis/BoxIntegral/Partition/Tagged.lean b/Mathlib/Analysis/BoxIntegral/Partition/Tagged.lean index 967500dc84169c..9710dd56108b37 100644 --- a/Mathlib/Analysis/BoxIntegral/Partition/Tagged.lean +++ b/Mathlib/Analysis/BoxIntegral/Partition/Tagged.lean @@ -25,10 +25,8 @@ rectangular box, box partition noncomputable section +open Finset Function ENNReal NNReal Set open scoped Classical -open ENNReal NNReal - -open Set Function namespace BoxIntegral @@ -204,10 +202,9 @@ theorem isHenstock_biUnionTagged {π : Prepartition I} {πi : ∀ J, TaggedPrepa /-- In a Henstock prepartition, there are at most `2 ^ Fintype.card ι` boxes with a given tag. -/ theorem IsHenstock.card_filter_tag_eq_le [Fintype ι] (h : π.IsHenstock) (x : ι → ℝ) : - (π.boxes.filter fun J => π.tag J = x).card ≤ 2 ^ Fintype.card ι := + #{J ∈ π.boxes | π.tag J = x} ≤ 2 ^ Fintype.card ι := calc - (π.boxes.filter fun J => π.tag J = x).card ≤ - (π.boxes.filter fun J : Box ι => x ∈ Box.Icc J).card := by + #{J ∈ π.boxes | π.tag J = x} ≤ #{J ∈ π.boxes | x ∈ Box.Icc J} := by refine Finset.card_le_card fun J hJ => ?_ rw [Finset.mem_filter] at hJ ⊢; rcases hJ with ⟨hJ, rfl⟩ exact ⟨hJ, h J hJ⟩ diff --git a/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean b/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean index 80735fc0538cf1..4947d3460259ab 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean @@ -619,7 +619,7 @@ theorem changeOrigin_toFormalMultilinearSeries [DecidableEq ι] : rw [card_compl, Fintype.card_fin, Finset.card_singleton, Nat.add_sub_cancel_left]⟩) · use fun _ _ ↦ (singleton_injective <| compl_injective <| Subtype.ext_iff.mp ·) intro ⟨s, hs⟩ - have h : sᶜ.card = 1 := by rw [card_compl, hs, Fintype.card_fin, Nat.add_sub_cancel] + have h : #sᶜ = 1 := by rw [card_compl, hs, Fintype.card_fin, Nat.add_sub_cancel] obtain ⟨a, ha⟩ := card_eq_one.mp h exact ⟨a, Subtype.ext (compl_eq_comm.mp ha)⟩ rw [Function.comp_apply, Subtype.coe_mk, compl_singleton, piecewise_erase_univ, diff --git a/Mathlib/Analysis/Convex/Birkhoff.lean b/Mathlib/Analysis/Convex/Birkhoff.lean index a754f5db135c90..702bc496c5a0d4 100644 --- a/Mathlib/Analysis/Convex/Birkhoff.lean +++ b/Mathlib/Analysis/Convex/Birkhoff.lean @@ -46,13 +46,13 @@ private lemma exists_perm_eq_zero_implies_eq_zero [Nonempty n] {s : R} (hs : 0 < (hM : ∃ M' ∈ doublyStochastic R n, M = s • M') : ∃ σ : Equiv.Perm n, ∀ i j, M i j = 0 → σ.permMatrix R i j = 0 := by rw [exists_mem_doublyStochastic_eq_smul_iff hs.le] at hM - let f (i : n) : Finset n := univ.filter (M i · ≠ 0) - have hf (A : Finset n) : A.card ≤ (A.biUnion f).card := by + let f (i : n) : Finset n := {j | M i j ≠ 0} + have hf (A : Finset n) : #A ≤ #(A.biUnion f) := by have (i) : ∑ j ∈ f i, M i j = s := by simp [sum_subset (filter_subset _ _), hM.2.1] - have h₁ : ∑ i ∈ A, ∑ j ∈ f i, M i j = A.card * s := by simp [this] - have h₂ : ∑ i, ∑ j ∈ A.biUnion f, M i j = (A.biUnion f).card * s := by + have h₁ : ∑ i ∈ A, ∑ j ∈ f i, M i j = #A * s := by simp [this] + have h₂ : ∑ i, ∑ j ∈ A.biUnion f, M i j = #(A.biUnion f) * s := by simp [sum_comm (t := A.biUnion f), hM.2.2, mul_comm s] - suffices A.card * s ≤ (A.biUnion f).card * s by exact_mod_cast le_of_mul_le_mul_right this hs + suffices #A * s ≤ #(A.biUnion f) * s by exact_mod_cast le_of_mul_le_mul_right this hs rw [← h₁, ← h₂] trans ∑ i ∈ A, ∑ j ∈ A.biUnion f, M i j · refine sum_le_sum fun i hi => ?_ @@ -86,7 +86,7 @@ private lemma doublyStochastic_sum_perm_aux (M : Matrix n n R) ∃ w : Equiv.Perm n → R, (∀ σ, 0 ≤ w σ) ∧ ∑ σ, w σ • σ.permMatrix R = M := by rcases isEmpty_or_nonempty n case inl => exact ⟨1, by simp, Subsingleton.elim _ _⟩ - set d : ℕ := (Finset.univ.filter fun i : n × n => M i.1 i.2 ≠ 0).card with ← hd + set d : ℕ := #{i : n × n | M i.1 i.2 ≠ 0} with ← hd clear_value d induction d using Nat.strongRecOn generalizing M s case ind d ih => @@ -114,7 +114,7 @@ private lemma doublyStochastic_sum_perm_aux (M : Matrix n n R) split case isTrue h => exact (hi' i' (by simp)).trans_eq (by rw [h]) case isFalse h => exact hM.1 _ _ - have hd' : (univ.filter fun i : n × n => N i.1 i.2 ≠ 0).card < d := by + have hd' : #{i : n × n | N i.1 i.2 ≠ 0} < d := by rw [← hd] refine card_lt_card ?_ rw [ssubset_iff_of_subset (monotone_filter_right _ _)] diff --git a/Mathlib/Analysis/Convex/Caratheodory.lean b/Mathlib/Analysis/Convex/Caratheodory.lean index 356e64e002984b..59b4f09c6a3ec3 100644 --- a/Mathlib/Analysis/Convex/Caratheodory.lean +++ b/Mathlib/Analysis/Convex/Caratheodory.lean @@ -55,7 +55,7 @@ theorem mem_convexHull_erase [DecidableEq E] {t : Finset E} (h : ¬AffineIndepen obtain ⟨g, gcombo, gsum, gpos⟩ := exists_nontrivial_relation_sum_zero_of_not_affine_ind h replace gpos := exists_pos_of_sum_zero_of_exists_nonzero g gsum gpos clear h - let s := @Finset.filter _ (fun z => 0 < g z) (fun _ => LinearOrder.decidableLT _ _) t + let s := {z ∈ t | 0 < g z} obtain ⟨i₀, mem, w⟩ : ∃ i₀ ∈ s, ∀ i ∈ s, f i₀ / g i₀ ≤ f i / g i := by apply s.exists_min_image fun z => f z / g z obtain ⟨x, hx, hgx⟩ : ∃ x ∈ t, 0 < g x := gpos @@ -117,13 +117,13 @@ theorem minCardFinsetOfMemConvexHull_nonempty : (minCardFinsetOfMemConvexHull hx exact ⟨x, mem_minCardFinsetOfMemConvexHull hx⟩ theorem minCardFinsetOfMemConvexHull_card_le_card {t : Finset E} (ht₁ : ↑t ⊆ s) - (ht₂ : x ∈ convexHull 𝕜 (t : Set E)) : (minCardFinsetOfMemConvexHull hx).card ≤ t.card := + (ht₂ : x ∈ convexHull 𝕜 (t : Set E)) : #(minCardFinsetOfMemConvexHull hx) ≤ #t := Function.argminOn_le _ _ _ (by exact ⟨ht₁, ht₂⟩) theorem affineIndependent_minCardFinsetOfMemConvexHull : AffineIndependent 𝕜 ((↑) : minCardFinsetOfMemConvexHull hx → E) := by - let k := (minCardFinsetOfMemConvexHull hx).card - 1 - have hk : (minCardFinsetOfMemConvexHull hx).card = k + 1 := + let k := #(minCardFinsetOfMemConvexHull hx) - 1 + have hk : #(minCardFinsetOfMemConvexHull hx) = k + 1 := (Nat.succ_pred_eq_of_pos (Finset.card_pos.mpr (minCardFinsetOfMemConvexHull_nonempty hx))).symm classical by_contra h @@ -163,7 +163,7 @@ theorem eq_pos_convex_span_of_mem_convexHull {x : E} (hx : x ∈ convexHull 𝕜 obtain ⟨t, ht₁, ht₂, ht₃⟩ := hx simp only [t.convexHull_eq, exists_prop, Set.mem_setOf_eq] at ht₃ obtain ⟨w, hw₁, hw₂, hw₃⟩ := ht₃ - let t' := t.filter fun i => w i ≠ 0 + let t' := {i ∈ t | w i ≠ 0} refine ⟨t', t'.fintypeCoeSort, ((↑) : t' → E), w ∘ ((↑) : t' → E), ?_, ?_, ?_, ?_, ?_⟩ · rw [Subtype.range_coe_subtype] exact Subset.trans (Finset.filter_subset _ t) ht₁ diff --git a/Mathlib/Analysis/Convex/Combination.lean b/Mathlib/Analysis/Convex/Combination.lean index 87bc0806851213..6fe599b4f51f72 100644 --- a/Mathlib/Analysis/Convex/Combination.lean +++ b/Mathlib/Analysis/Convex/Combination.lean @@ -124,8 +124,7 @@ theorem Finset.centerMass_subset {t' : Finset ι} (ht : t ⊆ t') (h : ∀ i ∈ intro i hit' hit rw [h i hit' hit, zero_smul, smul_zero] -theorem Finset.centerMass_filter_ne_zero : - (t.filter fun i => w i ≠ 0).centerMass w z = t.centerMass w z := +theorem Finset.centerMass_filter_ne_zero : {i ∈ t | 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, Ne, Classical.not_not] using hit' @@ -264,7 +263,7 @@ theorem Finset.centroid_mem_convexHull (s : Finset E) (hs : s.Nonempty) : rw [s.centroid_eq_centerMass hs] apply s.centerMass_id_mem_convexHull · simp only [inv_nonneg, imp_true_iff, Nat.cast_nonneg, Finset.centroidWeights_apply] - · have hs_card : (s.card : R) ≠ 0 := by simp [Finset.nonempty_iff_ne_empty.mp hs] + · have hs_card : (#s : R) ≠ 0 := by simp [Finset.nonempty_iff_ne_empty.mp hs] simp only [hs_card, Finset.sum_const, nsmul_eq_mul, mul_inv_cancel₀, Ne, not_false_iff, Finset.centroidWeights_apply, zero_lt_one] diff --git a/Mathlib/Analysis/Convex/Jensen.lean b/Mathlib/Analysis/Convex/Jensen.lean index 93f195e0143930..90692868f122b9 100644 --- a/Mathlib/Analysis/Convex/Jensen.lean +++ b/Mathlib/Analysis/Convex/Jensen.lean @@ -271,8 +271,8 @@ lemma ConvexOn.exists_ge_of_centerMass {t : Finset ι} (h : ConvexOn 𝕜 s f) ∃ i ∈ t, f (t.centerMass w p) ≤ f (p i) := by set y := t.centerMass w p -- TODO: can `rsuffices` be used to write the `exact` first, then the proof of this obtain? - obtain ⟨i, hi, hfi⟩ : ∃ i ∈ t.filter fun i => w i ≠ 0, w i • f y ≤ w i • (f ∘ p) i := by - have hw' : (0 : 𝕜) < ∑ i ∈ filter (fun i => w i ≠ 0) t, w i := by rwa [sum_filter_ne_zero] + obtain ⟨i, hi, hfi⟩ : ∃ i ∈ {i ∈ t | w i ≠ 0}, w i • f y ≤ w i • (f ∘ p) i := by + have hw' : (0 : 𝕜) < ∑ i ∈ t with w i ≠ 0, w i := by rwa [sum_filter_ne_zero] refine exists_le_of_sum_le (nonempty_of_sum_ne_zero hw'.ne') ?_ rw [← sum_smul, ← smul_le_smul_iff_of_pos_left (inv_pos.2 hw'), inv_smul_smul₀ hw'.ne', ← centerMass, centerMass_filter_ne_zero] diff --git a/Mathlib/Analysis/Convex/Radon.lean b/Mathlib/Analysis/Convex/Radon.lean index f297649b6f470e..2a65ed2e8986ab 100644 --- a/Mathlib/Analysis/Convex/Radon.lean +++ b/Mathlib/Analysis/Convex/Radon.lean @@ -41,8 +41,8 @@ theorem radon_partition {f : ι → E} (h : ¬ AffineIndependent 𝕜 f) : rw [affineIndependent_iff] at h push_neg at h obtain ⟨s, w, h_wsum, h_vsum, nonzero_w_index, h1, h2⟩ := h - let I : Finset ι := s.filter fun i ↦ 0 ≤ w i - let J : Finset ι := s.filter fun i ↦ w i < 0 + let I : Finset ι := {i ∈ s | 0 ≤ w i} + let J : Finset ι := {i ∈ s | w i < 0} let p : E := centerMass I w f -- point of intersection have hJI : ∑ j ∈ J, w j + ∑ i ∈ I, w i = 0 := by simpa only [h_wsum, not_lt] using sum_filter_add_sum_filter_not s (fun i ↦ w i < 0) w @@ -66,8 +66,8 @@ open Module /-- Corner case for `helly_theorem'`. -/ private lemma helly_theorem_corner {F : ι → Set E} {s : Finset ι} - (h_card_small : s.card ≤ finrank 𝕜 E + 1) - (h_inter : ∀ I ⊆ s, I.card ≤ finrank 𝕜 E + 1 → (⋂ i ∈ I, F i).Nonempty) : + (h_card_small : #s ≤ finrank 𝕜 E + 1) + (h_inter : ∀ I ⊆ s, #I ≤ finrank 𝕜 E + 1 → (⋂ i ∈ I, F i).Nonempty) : (⋂ i ∈ s, F i).Nonempty := h_inter s (by simp) h_card_small variable [FiniteDimensional 𝕜 E] @@ -78,12 +78,12 @@ If `F` is a finite family of convex sets in a vector space of finite dimension ` `k ≤ d + 1` sets of `F` intersect nontrivially, then all sets of `F` intersect nontrivially. -/ theorem helly_theorem' {F : ι → Set E} {s : Finset ι} (h_convex : ∀ i ∈ s, Convex 𝕜 (F i)) - (h_inter : ∀ I ⊆ s, I.card ≤ finrank 𝕜 E + 1 → (⋂ i ∈ I, F i).Nonempty) : + (h_inter : ∀ I ⊆ s, #I ≤ finrank 𝕜 E + 1 → (⋂ i ∈ I, F i).Nonempty) : (⋂ i ∈ s, F i).Nonempty := by classical - obtain h_card | h_card := lt_or_le s.card (finrank 𝕜 E + 1) + obtain h_card | h_card := lt_or_le #s (finrank 𝕜 E + 1) · exact helly_theorem_corner (le_of_lt h_card) h_inter - generalize hn : s.card = n + generalize hn : #s = n rw [hn] at h_card induction' n, h_card using Nat.le_induction with k h_card hk generalizing ι · exact helly_theorem_corner (le_of_eq hn) h_inter @@ -137,9 +137,9 @@ theorem helly_theorem' {F : ι → Set E} {s : Finset ι} If `F` is a family of `n` convex sets in a vector space of finite dimension `d`, with `n ≥ d + 1`, and any `d + 1` sets of `F` intersect nontrivially, then all sets of `F` intersect nontrivially. -/ theorem helly_theorem {F : ι → Set E} {s : Finset ι} - (h_card : finrank 𝕜 E + 1 ≤ s.card) + (h_card : finrank 𝕜 E + 1 ≤ #s) (h_convex : ∀ i ∈ s, Convex 𝕜 (F i)) - (h_inter : ∀ I ⊆ s, I.card = finrank 𝕜 E + 1 → (⋂ i ∈ I, F i).Nonempty) : + (h_inter : ∀ I ⊆ s, #I = finrank 𝕜 E + 1 → (⋂ i ∈ I, F i).Nonempty) : (⋂ i ∈ s, F i).Nonempty := by apply helly_theorem' h_convex intro I hI_ss hI_card @@ -153,7 +153,7 @@ If `F` is a finite set of convex sets in a vector space of finite dimension `d`, sets from `F` intersect nontrivially, then all sets from `F` intersect nontrivially. -/ theorem helly_theorem_set' {F : Finset (Set E)} (h_convex : ∀ X ∈ F, Convex 𝕜 X) - (h_inter : ∀ G : Finset (Set E), G ⊆ F → G.card ≤ finrank 𝕜 E + 1 → (⋂₀ G : Set E).Nonempty) : + (h_inter : ∀ G : Finset (Set E), G ⊆ F → #G ≤ finrank 𝕜 E + 1 → (⋂₀ G : Set E).Nonempty) : (⋂₀ (F : Set (Set E))).Nonempty := by classical -- for DecidableEq, required for the family version rw [show ⋂₀ F = ⋂ X ∈ F, (X : Set E) by ext; simp] @@ -168,9 +168,9 @@ If `F` is a finite set of convex sets in a vector space of finite dimension `d`, and any `d + 1` sets from `F` intersect nontrivially, then all sets from `F` intersect nontrivially. -/ theorem helly_theorem_set {F : Finset (Set E)} - (h_card : finrank 𝕜 E + 1 ≤ F.card) + (h_card : finrank 𝕜 E + 1 ≤ #F) (h_convex : ∀ X ∈ F, Convex 𝕜 X) - (h_inter : ∀ G : Finset (Set E), G ⊆ F → G.card = finrank 𝕜 E + 1 → (⋂₀ G : Set E).Nonempty) : + (h_inter : ∀ G : Finset (Set E), G ⊆ F → #G = finrank 𝕜 E + 1 → (⋂₀ G : Set E).Nonempty) : (⋂₀ (F : Set (Set E))).Nonempty := by apply helly_theorem_set' h_convex intro I hI_ss hI_card @@ -185,7 +185,7 @@ If `F` is a family of compact convex sets in a vector space of finite dimension `k ≤ d + 1` sets of `F` intersect nontrivially, then all sets of `F` intersect nontrivially. -/ theorem helly_theorem_compact' [TopologicalSpace E] [T2Space E] {F : ι → Set E} (h_convex : ∀ i, Convex 𝕜 (F i)) (h_compact : ∀ i, IsCompact (F i)) - (h_inter : ∀ I : Finset ι, I.card ≤ finrank 𝕜 E + 1 → (⋂ i ∈ I, F i).Nonempty) : + (h_inter : ∀ I : Finset ι, #I ≤ finrank 𝕜 E + 1 → (⋂ i ∈ I, F i).Nonempty) : (⋂ i, F i).Nonempty := by classical /- If `ι` is empty the statement is trivial. -/ @@ -214,11 +214,11 @@ then all sets of `F` intersect nontrivially. -/ theorem helly_theorem_compact [TopologicalSpace E] [T2Space E] {F : ι → Set E} (h_card : finrank 𝕜 E + 1 ≤ PartENat.card ι) (h_convex : ∀ i, Convex 𝕜 (F i)) (h_compact : ∀ i, IsCompact (F i)) - (h_inter : ∀ I : Finset ι, I.card = finrank 𝕜 E + 1 → (⋂ i ∈ I, F i).Nonempty) : + (h_inter : ∀ I : Finset ι, #I = finrank 𝕜 E + 1 → (⋂ i ∈ I, F i).Nonempty) : (⋂ i, F i).Nonempty := by apply helly_theorem_compact' h_convex h_compact intro I hI_card - have hJ : ∃ J : Finset ι, I ⊆ J ∧ J.card = finrank 𝕜 E + 1 := by + have hJ : ∃ J : Finset ι, I ⊆ J ∧ #J = finrank 𝕜 E + 1 := by by_cases h : Infinite ι · exact Infinite.exists_superset_card_eq _ _ hI_card · have : Finite ι := Finite.of_not_infinite h @@ -236,7 +236,7 @@ If `F` is a set of compact convex sets in a vector space of finite dimension `d` `k ≤ d + 1` sets from `F` intersect nontrivially, then all sets from `F` intersect nontrivially. -/ theorem helly_theorem_set_compact' [TopologicalSpace E] [T2Space E] {F : Set (Set E)} (h_convex : ∀ X ∈ F, Convex 𝕜 X) (h_compact : ∀ X ∈ F, IsCompact X) - (h_inter : ∀ G : Finset (Set E), (G : Set (Set E)) ⊆ F → G.card ≤ finrank 𝕜 E + 1 → + (h_inter : ∀ G : Finset (Set E), (G : Set (Set E)) ⊆ F → #G ≤ finrank 𝕜 E + 1 → (⋂₀ G : Set E).Nonempty) : (⋂₀ (F : Set (Set E))).Nonempty := by classical -- for DecidableEq, required for the family version @@ -259,7 +259,7 @@ then all sets from `F` intersect nontrivially. -/ theorem helly_theorem_set_compact [TopologicalSpace E] [T2Space E] {F : Set (Set E)} (h_card : finrank 𝕜 E + 1 ≤ F.encard) (h_convex : ∀ X ∈ F, Convex 𝕜 X) (h_compact : ∀ X ∈ F, IsCompact X) - (h_inter : ∀ G : Finset (Set E), (G : Set (Set E)) ⊆ F → G.card = finrank 𝕜 E + 1 → + (h_inter : ∀ G : Finset (Set E), (G : Set (Set E)) ⊆ F → #G = finrank 𝕜 E + 1 → (⋂₀ G : Set E).Nonempty) : (⋂₀ (F : Set (Set E))).Nonempty := by apply helly_theorem_set_compact' h_convex h_compact diff --git a/Mathlib/Analysis/MeanInequalities.lean b/Mathlib/Analysis/MeanInequalities.lean index ebc7bb9056c0fb..d75d457d35cf70 100644 --- a/Mathlib/Analysis/MeanInequalities.lean +++ b/Mathlib/Analysis/MeanInequalities.lean @@ -494,7 +494,7 @@ theorem inner_le_Lp_mul_Lq_hasSum {f g : ι → ℝ≥0} {A B : ℝ≥0} {p q : sum of the `p`-th powers of `f i`. Version for sums over finite sets, with `ℝ≥0`-valued functions. -/ theorem rpow_sum_le_const_mul_sum_rpow (f : ι → ℝ≥0) {p : ℝ} (hp : 1 ≤ p) : - (∑ i ∈ s, f i) ^ p ≤ (card s : ℝ≥0) ^ (p - 1) * ∑ i ∈ s, f i ^ p := by + (∑ i ∈ s, f i) ^ p ≤ (#s : ℝ≥0) ^ (p - 1) * ∑ i ∈ s, f i ^ p := by cases' eq_or_lt_of_le hp with hp hp · simp [← hp] let q : ℝ := p / (p - 1) @@ -620,7 +620,7 @@ theorem inner_le_Lp_mul_Lq (hpq : IsConjExponent p q) : /-- For `1 ≤ p`, the `p`-th power of the sum of `f i` is bounded above by a constant times the sum of the `p`-th powers of `f i`. Version for sums over finite sets, with `ℝ`-valued functions. -/ theorem rpow_sum_le_const_mul_sum_rpow (hp : 1 ≤ p) : - (∑ i ∈ s, |f i|) ^ p ≤ (card s : ℝ) ^ (p - 1) * ∑ i ∈ s, |f i| ^ p := by + (∑ i ∈ s, |f i|) ^ p ≤ (#s : ℝ) ^ (p - 1) * ∑ i ∈ s, |f i| ^ p := by have := NNReal.coe_le_coe.2 (NNReal.rpow_sum_le_const_mul_sum_rpow s (fun i => ⟨_, abs_nonneg (f i)⟩) hp) @@ -722,7 +722,7 @@ theorem inner_le_Lp_mul_Lq_hasSum_of_nonneg (hpq : p.IsConjExponent q) {A B : sum of the `p`-th powers of `f i`. Version for sums over finite sets, with nonnegative `ℝ`-valued functions. -/ theorem rpow_sum_le_const_mul_sum_rpow_of_nonneg (hp : 1 ≤ p) (hf : ∀ i ∈ s, 0 ≤ f i) : - (∑ i ∈ s, f i) ^ p ≤ (card s : ℝ) ^ (p - 1) * ∑ i ∈ s, f i ^ p := by + (∑ i ∈ s, f i) ^ p ≤ (#s : ℝ) ^ (p - 1) * ∑ i ∈ s, f i ^ p := by convert rpow_sum_le_const_mul_sum_rpow s f hp using 2 <;> apply sum_congr rfl <;> intro i hi <;> simp only [abs_of_nonneg, hf i hi] diff --git a/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean b/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean index 7f88338734d621..02f606b3771ea6 100644 --- a/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean +++ b/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean @@ -307,7 +307,7 @@ theorem coe_mkContinuous (C : ℝ) (H : ∀ m, ‖f m‖ ≤ C * ∏ i, ‖m i the other coordinates, then the resulting restricted function satisfies an inequality `‖f.restr v‖ ≤ C * ‖z‖^(n-k) * Π ‖v i‖` if the original function satisfies `‖f v‖ ≤ C * Π ‖v i‖`. -/ theorem restr_norm_le {k n : ℕ} (f : (MultilinearMap 𝕜 (fun _ : Fin n => G) G' : _)) - (s : Finset (Fin n)) (hk : s.card = k) (z : G) {C : ℝ} (H : ∀ m, ‖f m‖ ≤ C * ∏ i, ‖m i‖) + (s : Finset (Fin n)) (hk : #s = k) (z : G) {C : ℝ} (H : ∀ m, ‖f m‖ ≤ C * ∏ i, ‖m i‖) (v : Fin k → G) : ‖f.restr s hk z v‖ ≤ C * ‖z‖ ^ (n - k) * ∏ i, ‖v i‖ := by rw [mul_right_comm, mul_assoc] convert H _ using 2 @@ -732,12 +732,12 @@ namespace ContinuousMultilinearMap these variables, and fixing the other ones equal to a given value `z`. It is denoted by `f.restr s hk z`, where `hk` is a proof that the cardinality of `s` is `k`. The implicit identification between `Fin k` and `s` that we use is the canonical (increasing) bijection. -/ -def restr {k n : ℕ} (f : (G[×n]→L[𝕜] G' : _)) (s : Finset (Fin n)) (hk : s.card = k) (z : G) : +def restr {k n : ℕ} (f : (G[×n]→L[𝕜] G' : _)) (s : Finset (Fin n)) (hk : #s = k) (z : G) : G[×k]→L[𝕜] G' := (f.toMultilinearMap.restr s hk z).mkContinuous (‖f‖ * ‖z‖ ^ (n - k)) fun _ => MultilinearMap.restr_norm_le _ _ _ _ f.le_opNorm _ -theorem norm_restr {k n : ℕ} (f : G[×n]→L[𝕜] G') (s : Finset (Fin n)) (hk : s.card = k) (z : G) : +theorem norm_restr {k n : ℕ} (f : G[×n]→L[𝕜] G') (s : Finset (Fin n)) (hk : #s = k) (z : G) : ‖f.restr s hk z‖ ≤ ‖f‖ * ‖z‖ ^ (n - k) := by apply MultilinearMap.mkContinuous_norm_le exact mul_nonneg (norm_nonneg _) (pow_nonneg (norm_nonneg _) _) diff --git a/Mathlib/Analysis/NormedSpace/Multilinear/Curry.lean b/Mathlib/Analysis/NormedSpace/Multilinear/Curry.lean index 2ce89faee16da4..3e81916160893e 100644 --- a/Mathlib/Analysis/NormedSpace/Multilinear/Curry.lean +++ b/Mathlib/Analysis/NormedSpace/Multilinear/Curry.lean @@ -560,7 +560,7 @@ variable (𝕜 G G') {k l : ℕ} {s : Finset (Fin n)} `l`, then the space of continuous multilinear maps `G [×n]→L[𝕜] G'` of `n` variables is isomorphic to the space of continuous multilinear maps `G [×k]→L[𝕜] G [×l]→L[𝕜] G'` of `k` variables taking values in the space of continuous multilinear maps of `l` variables. -/ -def curryFinFinset {k l n : ℕ} {s : Finset (Fin n)} (hk : s.card = k) (hl : sᶜ.card = l) : +def curryFinFinset {k l n : ℕ} {s : Finset (Fin n)} (hk : #s = k) (hl : #sᶜ = l) : (G[×n]→L[𝕜] G') ≃ₗᵢ[𝕜] G[×k]→L[𝕜] G[×l]→L[𝕜] G' := (domDomCongrₗᵢ 𝕜 G G' (finSumEquivOfFinset hk hl).symm).trans (currySumEquiv 𝕜 (Fin k) (Fin l) G G') @@ -568,31 +568,31 @@ def curryFinFinset {k l n : ℕ} {s : Finset (Fin n)} (hk : s.card = k) (hl : s variable {𝕜 G G'} @[simp] -theorem curryFinFinset_apply (hk : s.card = k) (hl : sᶜ.card = l) (f : G[×n]→L[𝕜] G') +theorem curryFinFinset_apply (hk : #s = k) (hl : #sᶜ = l) (f : G[×n]→L[𝕜] G') (mk : Fin k → G) (ml : Fin l → G) : curryFinFinset 𝕜 G G' hk hl f mk ml = f fun i => Sum.elim mk ml ((finSumEquivOfFinset hk hl).symm i) := rfl @[simp] -theorem curryFinFinset_symm_apply (hk : s.card = k) (hl : sᶜ.card = l) +theorem curryFinFinset_symm_apply (hk : #s = k) (hl : #sᶜ = l) (f : G[×k]→L[𝕜] G[×l]→L[𝕜] G') (m : Fin n → G) : (curryFinFinset 𝕜 G G' hk hl).symm f m = f (fun i => m <| finSumEquivOfFinset hk hl (Sum.inl i)) fun i => m <| finSumEquivOfFinset hk hl (Sum.inr i) := rfl -theorem curryFinFinset_symm_apply_piecewise_const (hk : s.card = k) (hl : sᶜ.card = l) +theorem curryFinFinset_symm_apply_piecewise_const (hk : #s = k) (hl : #sᶜ = l) (f : G[×k]→L[𝕜] G[×l]→L[𝕜] G') (x y : G) : (curryFinFinset 𝕜 G G' hk hl).symm f (s.piecewise (fun _ => x) fun _ => y) = f (fun _ => x) fun _ => y := MultilinearMap.curryFinFinset_symm_apply_piecewise_const hk hl _ x y @[simp] -theorem curryFinFinset_symm_apply_const (hk : s.card = k) (hl : sᶜ.card = l) +theorem curryFinFinset_symm_apply_const (hk : #s = k) (hl : #sᶜ = l) (f : G[×k]→L[𝕜] G[×l]→L[𝕜] G') (x : G) : ((curryFinFinset 𝕜 G G' hk hl).symm f fun _ => x) = f (fun _ => x) fun _ => x := rfl -theorem curryFinFinset_apply_const (hk : s.card = k) (hl : sᶜ.card = l) (f : G[×n]→L[𝕜] G') +theorem curryFinFinset_apply_const (hk : #s = k) (hl : #sᶜ = l) (f : G[×n]→L[𝕜] G') (x y : G) : (curryFinFinset 𝕜 G G' hk hl f (fun _ => x) fun _ => y) = f (s.piecewise (fun _ => x) fun _ => y) := by refine (curryFinFinset_symm_apply_piecewise_const hk hl _ _ _).symm.trans ?_ diff --git a/Mathlib/Analysis/SpecificLimits/FloorPow.lean b/Mathlib/Analysis/SpecificLimits/FloorPow.lean index 0d410a3b11645b..6d020f15e36d59 100644 --- a/Mathlib/Analysis/SpecificLimits/FloorPow.lean +++ b/Mathlib/Analysis/SpecificLimits/FloorPow.lean @@ -216,7 +216,7 @@ theorem tendsto_div_of_monotone_of_tendsto_div_floor_pow (u : ℕ → ℝ) (l : /-- The sum of `1/(c^i)^2` above a threshold `j` is comparable to `1/j^2`, up to a multiplicative constant. -/ theorem sum_div_pow_sq_le_div_sq (N : ℕ) {j : ℝ} (hj : 0 < j) {c : ℝ} (hc : 1 < c) : - (∑ i ∈ (range N).filter (j < c ^ ·), (1 : ℝ) / (c ^ i) ^ 2) ≤ c ^ 3 * (c - 1)⁻¹ / j ^ 2 := by + (∑ i ∈ range N with j < c ^ i, (1 : ℝ) / (c ^ i) ^ 2) ≤ c ^ 3 * (c - 1)⁻¹ / j ^ 2 := by have cpos : 0 < c := zero_lt_one.trans hc have A : (0 : ℝ) < c⁻¹ ^ 2 := sq_pos_of_pos (inv_pos.2 cpos) have B : c ^ 2 * ((1 : ℝ) - c⁻¹ ^ 2)⁻¹ ≤ c ^ 3 * (c - 1)⁻¹ := by @@ -229,7 +229,7 @@ theorem sum_div_pow_sq_le_div_sq (N : ℕ) {j : ℝ} (hj : 0 < j) {c : ℝ} (hc simpa using pow_right_mono₀ hc.le one_le_two have C : c⁻¹ ^ 2 < 1 := pow_lt_one₀ (inv_nonneg.2 cpos.le) (inv_lt_one_of_one_lt₀ hc) two_ne_zero calc - (∑ i ∈ (range N).filter (j < c ^ ·), (1 : ℝ) / (c ^ i) ^ 2) ≤ + (∑ i ∈ range N with j < c ^ i, (1 : ℝ) / (c ^ i) ^ 2) ≤ ∑ 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 @@ -276,17 +276,16 @@ theorem mul_pow_le_nat_floor_pow {c : ℝ} (hc : 1 < c) (i : ℕ) : (1 - c⁻¹) /-- The sum of `1/⌊c^i⌋₊^2` above a threshold `j` is comparable to `1/j^2`, up to a multiplicative constant. -/ theorem sum_div_nat_floor_pow_sq_le_div_sq (N : ℕ) {j : ℝ} (hj : 0 < j) {c : ℝ} (hc : 1 < c) : - (∑ i ∈ (range N).filter (j < ⌊c ^ ·⌋₊), (1 : ℝ) / (⌊c ^ i⌋₊ : ℝ) ^ 2) ≤ + (∑ i ∈ range N with j < ⌊c ^ i⌋₊, (1 : ℝ) / (⌊c ^ i⌋₊ : ℝ) ^ 2) ≤ c ^ 5 * (c - 1)⁻¹ ^ 3 / j ^ 2 := by have cpos : 0 < c := zero_lt_one.trans hc have A : 0 < 1 - c⁻¹ := sub_pos.2 (inv_lt_one_of_one_lt₀ hc) calc - (∑ i ∈ (range N).filter (j < ⌊c ^ ·⌋₊), (1 : ℝ) / (⌊c ^ i⌋₊ : ℝ) ^ 2) ≤ - ∑ i ∈ (range N).filter (j < c ^ ·), (1 : ℝ) / (⌊c ^ i⌋₊ : ℝ) ^ 2 := by - apply sum_le_sum_of_subset_of_nonneg - · exact monotone_filter_right _ fun k hk ↦ hk.trans_le <| Nat.floor_le (by positivity) - · intros; positivity - _ ≤ ∑ i ∈ (range N).filter (j < c ^ ·), (1 - c⁻¹)⁻¹ ^ 2 * ((1 : ℝ) / (c ^ i) ^ 2) := by + (∑ i ∈ range N with j < ⌊c ^ i⌋₊, (1 : ℝ) / (⌊c ^ i⌋₊ : ℝ) ^ 2) ≤ + ∑ i ∈ range N with j < c ^ i, (1 : ℝ) / (⌊c ^ i⌋₊ : ℝ) ^ 2 := by + gcongr + exact fun k hk ↦ hk.trans_le <| Nat.floor_le (by positivity) + _ ≤ ∑ i ∈ range N with j < c ^ i, (1 - c⁻¹)⁻¹ ^ 2 * ((1 : ℝ) / (c ^ i) ^ 2) := by refine sum_le_sum fun i _hi => ?_ rw [mul_div_assoc', mul_one, div_le_div_iff]; rotate_left · apply sq_pos_of_pos diff --git a/Mathlib/Computability/AkraBazzi/AkraBazzi.lean b/Mathlib/Computability/AkraBazzi/AkraBazzi.lean index 5f7d1c28e96239..9cfbcd713b62a1 100644 --- a/Mathlib/Computability/AkraBazzi/AkraBazzi.lean +++ b/Mathlib/Computability/AkraBazzi/AkraBazzi.lean @@ -644,9 +644,9 @@ lemma eventually_atTop_sumTransform_le : _ ≤ u := by exact_mod_cast hu'.1 _ ≤ n ^ (p a b) * (∑ _u ∈ Finset.Ico (r i n) n, c₂ * g n / (r i n) ^ ((p a b) + 1)) := by gcongr with u hu; rw [Finset.mem_Ico] at hu; exact hu.1 - _ ≤ n ^ (p a b) * (Finset.Ico (r i n) n).card • (c₂ * g n / (r i n) ^ ((p a b) + 1)) := by + _ ≤ n ^ p a b * #(Ico (r i n) n) • (c₂ * g n / r i n ^ (p a b + 1)) := by gcongr; exact Finset.sum_le_card_nsmul _ _ _ (fun x _ => by rfl) - _ = n ^ (p a b) * (Finset.Ico (r i n) n).card * (c₂ * g n / (r i n) ^ ((p a b) + 1)) := by + _ = n ^ p a b * #(Ico (r i n) n) * (c₂ * g n / r i n ^ (p a b + 1)) := by rw [nsmul_eq_mul, mul_assoc] _ = n ^ (p a b) * (n - r i n) * (c₂ * g n / (r i n) ^ ((p a b) + 1)) := by congr; rw [Nat.card_Ico, Nat.cast_sub (le_of_lt <| hr_lt_n i)] @@ -682,9 +682,9 @@ lemma eventually_atTop_sumTransform_le : _ ≤ u := by exact hu.1 exact rpow_le_rpow_of_exponent_nonpos (by positivity) (by exact_mod_cast (le_of_lt hu.2)) (le_of_lt hp) - _ ≤ n ^ (p a b) * (Finset.Ico (r i n) n).card • (c₂ * g n / n ^ ((p a b) + 1)) := by + _ ≤ n ^ p a b * #(Ico (r i n) n) • (c₂ * g n / n ^ (p a b + 1)) := by gcongr; exact Finset.sum_le_card_nsmul _ _ _ (fun x _ => by rfl) - _ = n ^ (p a b) * (Finset.Ico (r i n) n).card * (c₂ * g n / n ^ ((p a b) + 1)) := by + _ = n ^ p a b * #(Ico (r i n) n) * (c₂ * g n / n ^ (p a b + 1)) := by rw [nsmul_eq_mul, mul_assoc] _ = n ^ (p a b) * (n - r i n) * (c₂ * g n / n ^ ((p a b) + 1)) := by congr; rw [Nat.card_Ico, Nat.cast_sub (le_of_lt <| hr_lt_n i)] @@ -729,9 +729,9 @@ lemma eventually_atTop_sumTransform_ge : positivity · rw [Finset.mem_Ico] at hu exact le_of_lt hu.2 - _ ≥ n ^ (p a b) * (Finset.Ico (r i n) n).card • (c₂ * g n / n ^ ((p a b) + 1)) := by + _ ≥ n ^ p a b * #(Ico (r i n) n) • (c₂ * g n / n ^ (p a b + 1)) := by gcongr; exact Finset.card_nsmul_le_sum _ _ _ (fun x _ => by rfl) - _ = n ^ (p a b) * (Finset.Ico (r i n) n).card * (c₂ * g n / n ^ ((p a b) + 1)) := by + _ = n ^ p a b * #(Ico (r i n) n) * (c₂ * g n / n ^ (p a b + 1)) := by rw [nsmul_eq_mul, mul_assoc] _ = n ^ (p a b) * (n - r i n) * (c₂ * g n / n ^ ((p a b) + 1)) := by congr; rw [Nat.card_Ico, Nat.cast_sub (le_of_lt <| hr_lt_n i)] @@ -764,12 +764,12 @@ lemma eventually_atTop_sumTransform_ge : · rw [Finset.mem_Ico] at hu exact rpow_le_rpow_of_exponent_nonpos (by positivity) (by exact_mod_cast hu.1) (le_of_lt hp) - _ ≥ n ^ (p a b) * (Finset.Ico (r i n) n).card • (c₂ * g n / (r i n) ^ ((p a b) + 1)) := by + _ ≥ n ^ p a b * #(Ico (r i n) n) • (c₂ * g n / r i n ^ (p a b + 1)) := by gcongr; exact Finset.card_nsmul_le_sum _ _ _ (fun x _ => by rfl) - _ = n ^ (p a b) * (Finset.Ico (r i n) n).card * (c₂ * g n / (r i n) ^ ((p a b) + 1)) := by + _ = n ^ p a b * #(Ico (r i n) n) * (c₂ * g n / r i n ^ (p a b + 1)) := by rw [nsmul_eq_mul, mul_assoc] - _ ≥ n ^ (p a b) * (Finset.Ico (r i n) n).card * (c₂ * g n / (c₁ * n) ^ ((p a b) + 1)) := by - gcongr n^(p a b) * (Finset.Ico (r i n) n).card * (c₂ * g n / ?_) + _ ≥ n ^ p a b * #(Ico (r i n) n) * (c₂ * g n / (c₁ * n) ^ (p a b + 1)) := by + gcongr n ^ p a b * #(Ico (r i n) n) * (c₂ * g n / ?_) exact rpow_le_rpow_of_exponent_nonpos (by positivity) (hn₁ i) (le_of_lt hp) _ = n ^ (p a b) * (n - r i n) * (c₂ * g n / (c₁ * n) ^ ((p a b) + 1)) := by congr; rw [Nat.card_Ico, Nat.cast_sub (le_of_lt <| hr_lt_n i)] diff --git a/Mathlib/FieldTheory/Finite/Basic.lean b/Mathlib/FieldTheory/Finite/Basic.lean index 715387ba0669b4..f688f79dda1655 100644 --- a/Mathlib/FieldTheory/Finite/Basic.lean +++ b/Mathlib/FieldTheory/Finite/Basic.lean @@ -61,10 +61,10 @@ open Polynomial /-- The cardinality of a field is at most `n` times the cardinality of the image of a degree `n` polynomial -/ theorem card_image_polynomial_eval [DecidableEq R] [Fintype R] {p : R[X]} (hp : 0 < p.degree) : - Fintype.card R ≤ natDegree p * (univ.image fun x => eval x p).card := + Fintype.card R ≤ natDegree p * #(univ.image fun x => eval x p) := Finset.card_le_mul_card_image _ _ (fun a _ => calc - _ = (p - C a).roots.toFinset.card := + _ = #(p - C a).roots.toFinset := congr_arg card (by simp [Finset.ext_iff, ← mem_roots_sub_C hp]) _ ≤ Multiset.card (p - C a).roots := Multiset.toFinset_card_le _ _ ≤ _ := card_roots_sub_C' hp) @@ -80,17 +80,17 @@ theorem exists_root_sum_quadratic [Fintype R] {f g : R[X]} (hf2 : degree f = 2) rcases this with ⟨x, ⟨a, _, ha⟩, ⟨b, _, hb⟩⟩ exact ⟨a, b, by rw [ha, ← hb, eval_neg, neg_add_cancel]⟩ fun hd : Disjoint _ _ => - lt_irrefl (2 * ((univ.image fun x : R => eval x f) ∪ univ.image fun x : R => eval x (-g)).card) <| - calc 2 * ((univ.image fun x : R => eval x f) ∪ univ.image fun x : R => eval x (-g)).card + lt_irrefl (2 * #((univ.image fun x : R => eval x f) ∪ univ.image fun x : R => eval x (-g))) <| + calc 2 * #((univ.image fun x : R => eval x f) ∪ univ.image fun x : R => eval x (-g)) ≤ 2 * Fintype.card R := Nat.mul_le_mul_left _ (Finset.card_le_univ _) _ = Fintype.card R + Fintype.card R := two_mul _ - _ < natDegree f * (univ.image fun x : R => eval x f).card + - natDegree (-g) * (univ.image fun x : R => eval x (-g)).card := + _ < natDegree f * #(univ.image fun x : R => eval x f) + + natDegree (-g) * #(univ.image fun x : R => eval x (-g)) := (add_lt_add_of_lt_of_le (lt_of_le_of_ne (card_image_polynomial_eval (by rw [hf2]; decide)) (mt (congr_arg (· % 2)) (by simp [natDegree_eq_of_degree_eq_some hf2, hR]))) (card_image_polynomial_eval (by rw [degree_neg, hg2]; decide))) - _ = 2 * ((univ.image fun x : R => eval x f) ∪ univ.image fun x : R => eval x (-g)).card := by + _ = 2 * #((univ.image fun x : R => eval x f) ∪ univ.image fun x : R => eval x (-g)) := by rw [card_union_of_disjoint hd] simp [natDegree_eq_of_degree_eq_some hf2, natDegree_eq_of_degree_eq_some hg2, mul_add] diff --git a/Mathlib/Geometry/Euclidean/Circumcenter.lean b/Mathlib/Geometry/Euclidean/Circumcenter.lean index ab715d6d83899e..bd9d26e0df1047 100644 --- a/Mathlib/Geometry/Euclidean/Circumcenter.lean +++ b/Mathlib/Geometry/Euclidean/Circumcenter.lean @@ -453,7 +453,7 @@ theorem orthogonalProjection_eq_circumcenter_of_dist_eq {n : ℕ} (s : Simplex /-- The orthogonal projection of the circumcenter onto a face is the circumcenter of that face. -/ theorem orthogonalProjection_circumcenter {n : ℕ} (s : Simplex ℝ P n) {fs : Finset (Fin (n + 1))} - {m : ℕ} (h : fs.card = m + 1) : + {m : ℕ} (h : #fs = m + 1) : ↑((s.face h).orthogonalProjectionSpan s.circumcenter) = (s.face h).circumcenter := haveI hr : ∃ r, ∀ i, dist ((s.face h).points i) s.circumcenter = r := by use s.circumradius @@ -564,7 +564,7 @@ theorem point_eq_affineCombination_of_pointsWithCircumcenter {n : ℕ} (s : Simp terms of `pointsWithCircumcenter`. -/ def centroidWeightsWithCircumcenter {n : ℕ} (fs : Finset (Fin (n + 1))) : PointsWithCircumcenterIndex n → ℝ - | pointIndex i => if i ∈ fs then (card fs : ℝ)⁻¹ else 0 + | pointIndex i => if i ∈ fs then (#fs : ℝ)⁻¹ else 0 | circumcenterIndex => 0 /-- `centroidWeightsWithCircumcenter` sums to 1, if the `Finset` is nonempty. -/ @@ -584,7 +584,7 @@ theorem centroid_eq_affineCombination_of_pointsWithCircumcenter {n : ℕ} (s : S simp_rw [centroid_def, affineCombination_apply, weightedVSubOfPoint_apply, sum_pointsWithCircumcenter, centroidWeightsWithCircumcenter, pointsWithCircumcenter_point, zero_smul, add_zero, centroidWeights, - ← sum_indicator_subset_of_eq_zero (Function.const (Fin (n + 1)) (card fs : ℝ)⁻¹) + ← sum_indicator_subset_of_eq_zero (Function.const (Fin (n + 1)) (#fs : ℝ)⁻¹) (fun i wi => wi • (s.points i -ᵥ Classical.choice AddTorsor.nonempty)) fs.subset_univ fun _ => zero_smul ℝ _, Set.indicator_apply] @@ -638,7 +638,7 @@ theorem reflection_circumcenter_eq_affineCombination_of_pointsWithCircumcenter { reflection (affineSpan ℝ (s.points '' {i₁, i₂})) s.circumcenter = (univ : Finset (PointsWithCircumcenterIndex n)).affineCombination ℝ s.pointsWithCircumcenter (reflectionCircumcenterWeightsWithCircumcenter i₁ i₂) := by - have hc : card ({i₁, i₂} : Finset (Fin (n + 1))) = 2 := by simp [h] + have hc : #{i₁, i₂} = 2 := by simp [h] -- Making the next line a separate definition helps the elaborator: set W : AffineSubspace ℝ P := affineSpan ℝ (s.points '' {i₁, i₂}) have h_faces : diff --git a/Mathlib/Geometry/Euclidean/MongePoint.lean b/Mathlib/Geometry/Euclidean/MongePoint.lean index d829d76dfb4ef8..f08ceefb2ff482 100644 --- a/Mathlib/Geometry/Euclidean/MongePoint.lean +++ b/Mathlib/Geometry/Euclidean/MongePoint.lean @@ -168,7 +168,7 @@ theorem mongePointVSubFaceCentroidWeightsWithCircumcenter_eq_sub {n : ℕ} {i₁ cases' i with i · rw [Pi.sub_apply, mongePointWeightsWithCircumcenter, centroidWeightsWithCircumcenter, mongePointVSubFaceCentroidWeightsWithCircumcenter] - have hu : card ({i₁, i₂}ᶜ : Finset (Fin (n + 3))) = n + 1 := by + have hu : #{i₁, i₂}ᶜ = n + 1 := by simp [card_compl, Fintype.card_fin, h] rw [hu] by_cases hi : i = i₁ ∨ i = i₂ <;> simp [compl_eq_univ_sdiff, hi] @@ -306,7 +306,7 @@ theorem eq_mongePoint_of_forall_mem_mongePlane {n : ℕ} {s : Simplex ℝ P (n + rw [hu, ← vectorSpan_image_eq_span_vsub_set_left_ne ℝ _ (Set.mem_univ _), Set.image_univ] at hi have hv : p -ᵥ s.mongePoint ∈ vectorSpan ℝ (Set.range s.points) := by let s₁ : Finset (Fin (n + 3)) := univ.erase i₁ - obtain ⟨i₂, h₂⟩ := card_pos.1 (show 0 < card s₁ by simp [s₁, card_erase_of_mem]) + obtain ⟨i₂, h₂⟩ := card_pos.1 (show 0 < #s₁ by simp [s₁, card_erase_of_mem]) have h₁₂ : i₁ ≠ i₂ := (ne_of_mem_erase h₂).symm exact (Submodule.mem_inf.1 (h' i₂ h₁₂)).2 exact Submodule.disjoint_def.1 (vectorSpan ℝ (Set.range s.points)).orthogonal_disjoint _ hv hi @@ -359,7 +359,7 @@ theorem finrank_direction_altitude {n : ℕ} (s : Simplex ℝ P (n + 1)) (i : Fi rw [direction_altitude] have h := Submodule.finrank_add_inf_finrank_orthogonal (vectorSpan_mono ℝ (Set.image_subset_range s.points ↑(univ.erase i))) - have hc : card (univ.erase i) = n + 1 := by rw [card_erase_of_mem (mem_univ _)]; simp + have hc : #(univ.erase i) = n + 1 := by rw [card_erase_of_mem (mem_univ _)]; simp refine add_left_cancel (_root_.trans h ?_) classical rw [s.independent.finrank_vectorSpan (Fintype.card_fin _), ← Finset.coe_image, diff --git a/Mathlib/GroupTheory/Index.lean b/Mathlib/GroupTheory/Index.lean index 237170a95e0dc9..66ce7103de92bc 100644 --- a/Mathlib/GroupTheory/Index.lean +++ b/Mathlib/GroupTheory/Index.lean @@ -598,7 +598,7 @@ variable {G M F : Type*} [Group G] [Fintype G] [Monoid M] [DecidableEq M] @[to_additive] lemma card_fiber_eq_of_mem_range (f : F) {x y : M} (hx : x ∈ Set.range f) (hy : y ∈ Set.range f) : - (univ.filter <| fun g => f g = x).card = (univ.filter <| fun g => f g = y).card := by + #{g | f g = x} = #{g | f g = y} := by rcases hx with ⟨x, rfl⟩ rcases hy with ⟨y, rfl⟩ rcases mul_left_surjective x y with ⟨y, rfl⟩ diff --git a/Mathlib/GroupTheory/Perm/Cycle/Basic.lean b/Mathlib/GroupTheory/Perm/Cycle/Basic.lean index 0fc24ba0a1006a..6cf77b57dbfc2f 100644 --- a/Mathlib/GroupTheory/Perm/Cycle/Basic.lean +++ b/Mathlib/GroupTheory/Perm/Cycle/Basic.lean @@ -310,7 +310,7 @@ protected theorem IsSwap.isCycle : IsSwap f → IsCycle f := by variable [Fintype α] -theorem IsCycle.two_le_card_support (h : IsCycle f) : 2 ≤ f.support.card := +theorem IsCycle.two_le_card_support (h : IsCycle f) : 2 ≤ #f.support := two_le_card_support_of_ne_one h.ne_one /-- The subgroup generated by a cycle is in bijection with its support -/ @@ -351,7 +351,7 @@ theorem IsCycle.zpowersEquivSupport_symm_apply {σ : Perm α} (hσ : IsCycle σ) ⟨σ ^ n, n, rfl⟩ := (Equiv.symm_apply_eq _).2 hσ.zpowersEquivSupport_apply -protected theorem IsCycle.orderOf (hf : IsCycle f) : orderOf f = f.support.card := by +protected theorem IsCycle.orderOf (hf : IsCycle f) : orderOf f = #f.support := by rw [← Fintype.card_zpowers, ← Fintype.card_coe] convert Fintype.card_congr (IsCycle.zpowersEquivSupport hf) @@ -438,12 +438,12 @@ theorem IsCycle.swap_mul {α : Type*} [DecidableEq α] {f : Perm α} (hf : IsCyc isCycle_swap_mul_aux₂ (i - 1) hy hi⟩ -theorem IsCycle.sign {f : Perm α} (hf : IsCycle f) : sign f = -(-1) ^ f.support.card := +theorem IsCycle.sign {f : Perm α} (hf : IsCycle f) : sign f = -(-1) ^ #f.support := let ⟨x, hx⟩ := hf calc Perm.sign f = Perm.sign (swap x (f x) * (swap x (f x) * f)) := by {rw [← mul_assoc, mul_def, mul_def, swap_swap, trans_refl]} - _ = -(-1) ^ f.support.card := + _ = -(-1) ^ #f.support := if h1 : f (f x) = x then by have h : swap x (f x) * f = 1 := by simp only [mul_def, one_def] @@ -452,14 +452,13 @@ theorem IsCycle.sign {f : Perm α} (hf : IsCycle f) : sign f = -(-1) ^ f.support hf.eq_swap_of_apply_apply_eq_self hx.1 h1, card_support_swap hx.1.symm] rfl else by - have h : card (support (swap x (f x) * f)) + 1 = card (support f) := by + have h : #(swap x (f x) * f).support + 1 = #f.support := by rw [← insert_erase (mem_support.2 hx.1), support_swap_mul_eq _ _ h1, card_insert_of_not_mem (not_mem_erase _ _), sdiff_singleton_eq_erase] - have : card (support (swap x (f x) * f)) < card (support f) := - card_support_swap_mul hx.1 + have : #(swap x (f x) * f).support < #f.support := card_support_swap_mul hx.1 rw [sign_mul, sign_swap hx.1.symm, (hf.swap_mul hx.1 h1).sign, ← h] simp only [mul_neg, neg_mul, one_mul, neg_neg, pow_add, pow_one, mul_one] -termination_by f.support.card +termination_by #f.support theorem IsCycle.of_pow {n : ℕ} (h1 : IsCycle (f ^ n)) (h2 : f.support ⊆ (f ^ n).support) : IsCycle f := by @@ -472,7 +471,7 @@ theorem IsCycle.of_pow {n : ℕ} (h1 : IsCycle (f ^ n)) (h2 : f.support ⊆ (f ^ exact ⟨n * i, by rwa [zpow_mul]⟩ -- The lemma `support_zpow_le` is relevant. It means that `h2` is equivalent to --- `σ.support = (σ ^ n).support`, as well as to `σ.support.card ≤ (σ ^ n).support.card`. +-- `σ.support = (σ ^ n).support`, as well as to `#σ.support ≤ #(σ ^ n).support`. theorem IsCycle.of_zpow {n : ℤ} (h1 : IsCycle (f ^ n)) (h2 : f.support ⊆ (f ^ n).support) : IsCycle f := by cases n @@ -645,7 +644,7 @@ section Conjugation variable [Fintype α] [DecidableEq α] {σ τ : Perm α} -theorem IsCycle.isConj (hσ : IsCycle σ) (hτ : IsCycle τ) (h : σ.support.card = τ.support.card) : +theorem IsCycle.isConj (hσ : IsCycle σ) (hτ : IsCycle τ) (h : #σ.support = #τ.support) : IsConj σ τ := by refine isConj_of_support_equiv @@ -665,7 +664,7 @@ theorem IsCycle.isConj (hσ : IsCycle σ) (hτ : IsCycle τ) (h : σ.support.car rfl theorem IsCycle.isConj_iff (hσ : IsCycle σ) (hτ : IsCycle τ) : - IsConj σ τ ↔ σ.support.card = τ.support.card where + IsConj σ τ ↔ #σ.support = #τ.support where mp h := by obtain ⟨π, rfl⟩ := (_root_.isConj_iff).1 h refine Finset.card_bij (fun a _ => π a) (fun _ ha => ?_) (fun _ _ _ _ ab => π.injective ab) @@ -770,16 +769,15 @@ protected theorem IsCycleOn.subtypePerm (hf : f.IsCycleOn s) : -- TODO: Theory of order of an element under an action theorem IsCycleOn.pow_apply_eq {s : Finset α} (hf : f.IsCycleOn s) (ha : a ∈ s) {n : ℕ} : - (f ^ n) a = a ↔ s.card ∣ n := by + (f ^ n) a = a ↔ #s ∣ n := by obtain rfl | hs := Finset.eq_singleton_or_nontrivial ha · rw [coe_singleton, isCycleOn_singleton] at hf simpa using IsFixedPt.iterate hf n classical - have h : ∀ x ∈ s.attach, ¬f ↑x = ↑x := fun x _ => hf.apply_ne hs x.2 + have h (x : s) : ¬f x = x := hf.apply_ne hs x.2 have := (hf.isCycle_subtypePerm hs).orderOf - simp only [coe_sort_coe, support_subtype_perm, ne_eq, decide_not, Bool.not_eq_true', - decide_eq_false_iff_not, mem_attach, forall_true_left, Subtype.forall, filter_true_of_mem h, - card_attach] at this + simp only [coe_sort_coe, support_subtype_perm, ne_eq, h, not_false_eq_true, univ_eq_attach, + mem_attach, imp_self, implies_true, filter_true_of_mem, card_attach] at this rw [← this, orderOf_dvd_iff_pow_eq_one, (hf.isCycle_subtypePerm hs).pow_eq_one_iff' (ne_of_apply_ne ((↑) : s → α) <| hf.apply_ne hs (⟨a, ha⟩ : s).2)] @@ -789,32 +787,32 @@ theorem IsCycleOn.pow_apply_eq {s : Finset α} (hf : f.IsCycleOn s) (ha : a ∈ simp theorem IsCycleOn.zpow_apply_eq {s : Finset α} (hf : f.IsCycleOn s) (ha : a ∈ s) : - ∀ {n : ℤ}, (f ^ n) a = a ↔ (s.card : ℤ) ∣ n + ∀ {n : ℤ}, (f ^ n) a = a ↔ (#s : ℤ) ∣ n | Int.ofNat _ => (hf.pow_apply_eq ha).trans Int.natCast_dvd_natCast.symm | Int.negSucc n => by rw [zpow_negSucc, ← inv_pow] exact (hf.inv.pow_apply_eq ha).trans (dvd_neg.trans Int.natCast_dvd_natCast).symm theorem IsCycleOn.pow_apply_eq_pow_apply {s : Finset α} (hf : f.IsCycleOn s) (ha : a ∈ s) - {m n : ℕ} : (f ^ m) a = (f ^ n) a ↔ m ≡ n [MOD s.card] := by + {m n : ℕ} : (f ^ m) a = (f ^ n) a ↔ m ≡ n [MOD #s] := by rw [Nat.modEq_iff_dvd, ← hf.zpow_apply_eq ha] simp [sub_eq_neg_add, zpow_add, eq_inv_iff_eq, eq_comm] theorem IsCycleOn.zpow_apply_eq_zpow_apply {s : Finset α} (hf : f.IsCycleOn s) (ha : a ∈ s) - {m n : ℤ} : (f ^ m) a = (f ^ n) a ↔ m ≡ n [ZMOD s.card] := by + {m n : ℤ} : (f ^ m) a = (f ^ n) a ↔ m ≡ n [ZMOD #s] := by rw [Int.modEq_iff_dvd, ← hf.zpow_apply_eq ha] simp [sub_eq_neg_add, zpow_add, eq_inv_iff_eq, eq_comm] theorem IsCycleOn.pow_card_apply {s : Finset α} (hf : f.IsCycleOn s) (ha : a ∈ s) : - (f ^ s.card) a = a := + (f ^ #s) a = a := (hf.pow_apply_eq ha).2 dvd_rfl theorem IsCycleOn.exists_pow_eq {s : Finset α} (hf : f.IsCycleOn s) (ha : a ∈ s) (hb : b ∈ s) : - ∃ n < s.card, (f ^ n) a = b := by + ∃ n < #s, (f ^ n) a = b := by classical obtain ⟨n, rfl⟩ := hf.2 ha hb - obtain ⟨k, hk⟩ := (Int.mod_modEq n s.card).symm.dvd - refine ⟨n.natMod s.card, Int.natMod_lt (Nonempty.card_pos ⟨a, ha⟩).ne', ?_⟩ + obtain ⟨k, hk⟩ := (Int.mod_modEq n #s).symm.dvd + refine ⟨n.natMod #s, Int.natMod_lt (Nonempty.card_pos ⟨a, ha⟩).ne', ?_⟩ rw [← zpow_natCast, Int.natMod, Int.toNat_of_nonneg (Int.emod_nonneg _ <| Nat.cast_ne_zero.2 (Nonempty.card_pos ⟨a, ha⟩).ne'), sub_eq_iff_eq_add'.1 hk, zpow_add, zpow_mul] @@ -931,7 +929,7 @@ namespace Finset variable {f : Perm α} {s : Finset α} theorem product_self_eq_disjiUnion_perm_aux (hf : f.IsCycleOn s) : - (range s.card : Set ℕ).PairwiseDisjoint fun k => + (range #s : Set ℕ).PairwiseDisjoint fun k => s.map ⟨fun i => (i, (f ^ k) i), fun _ _ => congr_arg Prod.fst⟩ := by obtain hs | _ := (s : Set α).subsingleton_or_nontrivial · refine Set.Subsingleton.pairwise ?_ _ @@ -959,7 +957,7 @@ The diagonals are given by the cycle `f`. -/ theorem product_self_eq_disjiUnion_perm (hf : f.IsCycleOn s) : s ×ˢ s = - (range s.card).disjiUnion + (range #s).disjiUnion (fun k => s.map ⟨fun i => (i, (f ^ k) i), fun _ _ => congr_arg Prod.fst⟩) (product_self_eq_disjiUnion_perm_aux hf) := by ext ⟨a, b⟩ @@ -978,12 +976,12 @@ namespace Finset variable [Semiring α] [AddCommMonoid β] [Module α β] {s : Finset ι} {σ : Perm ι} theorem sum_smul_sum_eq_sum_perm (hσ : σ.IsCycleOn s) (f : ι → α) (g : ι → β) : - (∑ i ∈ s, f i) • ∑ i ∈ s, g i = ∑ k ∈ range s.card, ∑ i ∈ s, f i • g ((σ ^ k) i) := by + (∑ i ∈ s, f i) • ∑ i ∈ s, g i = ∑ k ∈ range #s, ∑ i ∈ s, f i • g ((σ ^ k) i) := by rw [sum_smul_sum, ← sum_product'] simp_rw [product_self_eq_disjiUnion_perm hσ, sum_disjiUnion, sum_map, Embedding.coeFn_mk] theorem sum_mul_sum_eq_sum_perm (hσ : σ.IsCycleOn s) (f g : ι → α) : - ((∑ i ∈ s, f i) * ∑ i ∈ s, g i) = ∑ k ∈ range s.card, ∑ i ∈ s, f i * g ((σ ^ k) i) := + ((∑ i ∈ s, f i) * ∑ i ∈ s, g i) = ∑ k ∈ range #s, ∑ i ∈ s, f i * g ((σ ^ k) i) := sum_smul_sum_eq_sum_perm hσ f g end Finset @@ -1094,10 +1092,10 @@ theorem zpow_eq_ofSubtype_subtypePerm_iff theorem cycle_zpow_mem_support_iff {g : Perm α} (hg : g.IsCycle) {n : ℤ} {x : α} (hx : g x ≠ x) : - (g ^ n) x = x ↔ n % g.support.card = 0 := by - set q := n / g.support.card - set r := n % g.support.card - have div_euc : r + g.support.card * q = n ∧ 0 ≤ r ∧ r < g.support.card := by + (g ^ n) x = x ↔ n % #g.support = 0 := by + set q := n / #g.support + set r := n % #g.support + have div_euc : r + #g.support * q = n ∧ 0 ≤ r ∧ r < #g.support := by rw [← Int.ediv_emod_unique _] · exact ⟨rfl, rfl⟩ simp only [Int.natCast_pos] diff --git a/Mathlib/GroupTheory/Perm/Cycle/Factors.lean b/Mathlib/GroupTheory/Perm/Cycle/Factors.lean index e2fc3224e97be7..aa5260cd2a4c38 100644 --- a/Mathlib/GroupTheory/Perm/Cycle/Factors.lean +++ b/Mathlib/GroupTheory/Perm/Cycle/Factors.lean @@ -166,7 +166,7 @@ theorem isCycle_cycleOf (f : Perm α) [DecidableRel f.SameCycle] (hx : f x ≠ x @[simp] theorem two_le_card_support_cycleOf_iff [DecidableEq α] [Fintype α] : - 2 ≤ card (cycleOf f x).support ↔ f x ≠ x := by + 2 ≤ #(cycleOf f x).support ↔ f x ≠ x := by refine ⟨fun h => ?_, fun h => by simpa using (isCycle_cycleOf _ h).two_le_card_support⟩ contrapose! h rw [← cycleOf_eq_one_iff] at h @@ -179,7 +179,7 @@ theorem two_le_card_support_cycleOf_iff [DecidableEq α] [Fintype α] : @[deprecated support_cycleOf_nonempty (since := "2024-06-16")] theorem card_support_cycleOf_pos_iff [DecidableEq α] [Fintype α] : - 0 < card (cycleOf f x).support ↔ f x ≠ x := by + 0 < #(cycleOf f x).support ↔ f x ≠ x := by rw [card_pos, support_cycleOf_nonempty] theorem pow_mod_orderOf_cycleOf_apply (f : Perm α) [DecidableRel f.SameCycle] (n : ℕ) (x : α) : @@ -241,7 +241,7 @@ theorem SameCycle.mem_support_iff {f} [DecidableEq α] [Fintype α] (h : SameCyc support_cycleOf_le f y (mem_support_cycleOf_iff.mpr ⟨h.symm, hy⟩)⟩ theorem pow_mod_card_support_cycleOf_self_apply [DecidableEq α] [Fintype α] - (f : Perm α) (n : ℕ) (x : α) : (f ^ (n % (f.cycleOf x).support.card)) x = (f ^ n) x := by + (f : Perm α) (n : ℕ) (x : α) : (f ^ (n % #(f.cycleOf x).support)) x = (f ^ n) x := by by_cases hx : f x = x · rw [pow_apply_eq_self_of_apply_eq_self hx, pow_apply_eq_self_of_apply_eq_self hx] · rw [← cycleOf_pow_apply_self, ← cycleOf_pow_apply_self f, ← (isCycle_cycleOf f hx).orderOf, @@ -268,17 +268,17 @@ theorem isCycleOn_support_cycleOf [DecidableEq α] [Fintype α] (f : Perm α) (x exact ha.1.symm.trans hb.1⟩ theorem SameCycle.exists_pow_eq_of_mem_support {f} [DecidableEq α] [Fintype α] (h : SameCycle f x y) - (hx : x ∈ f.support) : ∃ i < (f.cycleOf x).support.card, (f ^ i) x = y := by + (hx : x ∈ f.support) : ∃ i < #(f.cycleOf x).support, (f ^ i) x = y := by rw [mem_support] at hx exact Equiv.Perm.IsCycleOn.exists_pow_eq (b := y) (f.isCycleOn_support_cycleOf x) (by rw [mem_support_cycleOf_iff' hx]) (by rwa [mem_support_cycleOf_iff' hx]) theorem SameCycle.exists_pow_eq [DecidableEq α] [Fintype α] (f : Perm α) (h : SameCycle f x y) : - ∃ i : ℕ, 0 < i ∧ i ≤ (f.cycleOf x).support.card + 1 ∧ (f ^ i) x = y := by + ∃ i : ℕ, 0 < i ∧ i ≤ #(f.cycleOf x).support + 1 ∧ (f ^ i) x = y := by by_cases hx : x ∈ f.support · obtain ⟨k, hk, hk'⟩ := h.exists_pow_eq_of_mem_support hx cases' k with k - · refine ⟨(f.cycleOf x).support.card, ?_, self_le_add_right _ _, ?_⟩ + · refine ⟨#(f.cycleOf x).support, ?_, self_le_add_right _ _, ?_⟩ · refine zero_lt_one.trans (one_lt_card_support_of_ne_one ?_) simpa using hx · simp only [pow_zero, coe_one, id_eq] at hk' @@ -293,8 +293,7 @@ theorem SameCycle.exists_pow_eq [DecidableEq α] [Fintype α] (f : Perm α) (h : theorem zpow_eq_zpow_on_iff [DecidableEq α] [Fintype α] (g : Perm α) {m n : ℤ} {x : α} (hx : g x ≠ x) : - (g ^ m) x = (g ^ n) x ↔ - m % (g.cycleOf x).support.card = n % (g.cycleOf x).support.card := by + (g ^ m) x = (g ^ n) x ↔ m % #(g.cycleOf x).support = n % #(g.cycleOf x).support := by rw [Int.emod_eq_emod_iff_emod_sub_eq_zero] conv_lhs => rw [← Int.sub_add_cancel m n, Int.add_comm, zpow_add] simp only [coe_mul, Function.comp_apply, EmbeddingLike.apply_eq_iff_eq] diff --git a/Mathlib/GroupTheory/Perm/Support.lean b/Mathlib/GroupTheory/Perm/Support.lean index ea8f7d275ba904..a772723cfe3b30 100644 --- a/Mathlib/GroupTheory/Perm/Support.lean +++ b/Mathlib/GroupTheory/Perm/Support.lean @@ -264,8 +264,7 @@ theorem apply_zpow_apply_eq_iff (f : Perm α) (n : ℤ) {x : α} : variable [DecidableEq α] [Fintype α] {f g : Perm α} /-- The `Finset` of nonfixed points of a permutation. -/ -def support (f : Perm α) : Finset α := - univ.filter fun x => f x ≠ x +def support (f : Perm α) : Finset α := {x | f x ≠ x} @[simp] theorem mem_support {x : α} : x ∈ f.support ↔ f x ≠ x := by @@ -529,48 +528,48 @@ theorem support_extend_domain (f : α ≃ Subtype p) {g : Perm α} : exact pb (Subtype.prop _) theorem card_support_extend_domain (f : α ≃ Subtype p) {g : Perm α} : - (g.extendDomain f).support.card = g.support.card := by simp + #(g.extendDomain f).support = #g.support := by simp end ExtendDomain section Card -- @[simp] -- Porting note (#10618): simp can prove thisrove this -theorem card_support_eq_zero {f : Perm α} : f.support.card = 0 ↔ f = 1 := by +theorem card_support_eq_zero {f : Perm α} : #f.support = 0 ↔ f = 1 := by rw [Finset.card_eq_zero, support_eq_empty_iff] -theorem one_lt_card_support_of_ne_one {f : Perm α} (h : f ≠ 1) : 1 < f.support.card := by +theorem one_lt_card_support_of_ne_one {f : Perm α} (h : f ≠ 1) : 1 < #f.support := by simp_rw [one_lt_card_iff, mem_support, ← not_or] contrapose! h ext a specialize h (f a) a rwa [apply_eq_iff_eq, or_self_iff, or_self_iff] at h -theorem card_support_ne_one (f : Perm α) : f.support.card ≠ 1 := by +theorem card_support_ne_one (f : Perm α) : #f.support ≠ 1 := by by_cases h : f = 1 · exact ne_of_eq_of_ne (card_support_eq_zero.mpr h) zero_ne_one · exact ne_of_gt (one_lt_card_support_of_ne_one h) @[simp] -theorem card_support_le_one {f : Perm α} : f.support.card ≤ 1 ↔ f = 1 := by +theorem card_support_le_one {f : Perm α} : #f.support ≤ 1 ↔ f = 1 := by rw [le_iff_lt_or_eq, Nat.lt_succ_iff, Nat.le_zero, card_support_eq_zero, or_iff_not_imp_right, imp_iff_right f.card_support_ne_one] -theorem two_le_card_support_of_ne_one {f : Perm α} (h : f ≠ 1) : 2 ≤ f.support.card := +theorem two_le_card_support_of_ne_one {f : Perm α} (h : f ≠ 1) : 2 ≤ #f.support := one_lt_card_support_of_ne_one h theorem card_support_swap_mul {f : Perm α} {x : α} (hx : f x ≠ x) : - (swap x (f x) * f).support.card < f.support.card := + #(swap x (f x) * f).support < #f.support := Finset.card_lt_card ⟨fun _ hz => (mem_support_swap_mul_imp_mem_support_ne hz).left, fun h => absurd (h (mem_support.2 hx)) (mt mem_support.1 (by simp))⟩ -theorem card_support_swap {x y : α} (hxy : x ≠ y) : (swap x y).support.card = 2 := - show (swap x y).support.card = Finset.card ⟨x ::ₘ y ::ₘ 0, by simp [hxy]⟩ from +theorem card_support_swap {x y : α} (hxy : x ≠ y) : #(swap x y).support = 2 := + show #(swap x y).support = #⟨x ::ₘ y ::ₘ 0, by simp [hxy]⟩ from congr_arg card <| by simp [support_swap hxy, *, Finset.ext_iff] @[simp] -theorem card_support_eq_two {f : Perm α} : f.support.card = 2 ↔ IsSwap f := by +theorem card_support_eq_two {f : Perm α} : #f.support = 2 ↔ IsSwap f := by constructor <;> intro h · obtain ⟨x, t, hmem, hins, ht⟩ := card_eq_succ.1 h obtain ⟨y, rfl⟩ := card_eq_one.1 ht @@ -589,7 +588,7 @@ theorem card_support_eq_two {f : Perm α} : f.support.card = 2 ↔ IsSwap f := b exact card_support_swap hxy theorem Disjoint.card_support_mul (h : Disjoint f g) : - (f * g).support.card = f.support.card + g.support.card := by + #(f * g).support = #f.support + #g.support := by rw [← Finset.card_union_of_disjoint] · congr ext @@ -597,7 +596,7 @@ theorem Disjoint.card_support_mul (h : Disjoint f g) : · simpa using h.disjoint_support theorem card_support_prod_list_of_pairwise_disjoint {l : List (Perm α)} (h : l.Pairwise Disjoint) : - l.prod.support.card = (l.map (Finset.card ∘ support)).sum := by + #l.prod.support = (l.map (card ∘ support)).sum := by induction' l with a t ih · exact card_support_eq_zero.mpr rfl · obtain ⟨ha, ht⟩ := List.pairwise_cons.1 h @@ -610,10 +609,8 @@ end support @[simp] theorem support_subtype_perm [DecidableEq α] {s : Finset α} (f : Perm α) (h) : - ((f.subtypePerm h : Perm { x // x ∈ s }).support) = - (s.attach.filter ((fun x => decide (f x ≠ x))) : Finset { x // x ∈ s }) := by - ext - simp [Subtype.ext_iff] + (f.subtypePerm h : Perm s).support = ({x | f x ≠ x} : Finset s) := by + ext; simp [Subtype.ext_iff] end Equiv.Perm @@ -627,7 +624,7 @@ namespace Equiv.Perm variable {α : Type*} theorem fixed_point_card_lt_of_ne_one [DecidableEq α] [Fintype α] {σ : Perm α} (h : σ ≠ 1) : - (filter (fun x => σ x = x) univ).card < Fintype.card α - 1 := by + #{x | σ x = x} < Fintype.card α - 1 := by rw [Nat.lt_sub_iff_add_lt, ← Nat.lt_sub_iff_add_lt', ← Finset.card_compl, Finset.compl_filter] exact one_lt_card_support_of_ne_one h @@ -647,7 +644,7 @@ theorem support_conj : (σ * τ * σ⁻¹).support = τ.support.map σ.toEmbeddi simp only [mem_map_equiv, Perm.coe_mul, Function.comp_apply, Ne, Perm.mem_support, Equiv.eq_symm_apply, inv_def] -theorem card_support_conj : (σ * τ * σ⁻¹).support.card = τ.support.card := by simp +theorem card_support_conj : #(σ * τ * σ⁻¹).support = #τ.support := by simp end Equiv.Perm diff --git a/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean b/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean index 38b5ab4ee0b073..4a709390a2037b 100644 --- a/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean +++ b/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean @@ -278,11 +278,11 @@ open scoped Classical @[to_additive IsAddCyclic.card_nsmul_eq_zero_le] theorem IsCyclic.card_pow_eq_one_le [DecidableEq α] [Fintype α] [IsCyclic α] {n : ℕ} (hn0 : 0 < n) : - (univ.filter fun a : α => a ^ n = 1).card ≤ n := + #{a : α | a ^ n = 1} ≤ n := let ⟨g, hg⟩ := IsCyclic.exists_generator (α := α) calc - (univ.filter fun a : α => a ^ n = 1).card ≤ - (zpowers (g ^ (Fintype.card α / Nat.gcd n (Fintype.card α))) : Set α).toFinset.card := + #{a : α | a ^ n = 1} ≤ + #(zpowers (g ^ (Fintype.card α / Nat.gcd n (Fintype.card α))) : Set α).toFinset := card_le_card fun x hx => let ⟨m, hm⟩ := show x ∈ Submonoid.powers g from mem_powers_iff_mem_zpowers.2 <| hg x Set.mem_toFinset.2 @@ -382,21 +382,19 @@ end section Totient -variable [DecidableEq α] [Fintype α] - (hn : ∀ n : ℕ, 0 < n → (univ.filter fun a : α => a ^ n = 1).card ≤ n) +variable [DecidableEq α] [Fintype α] (hn : ∀ n : ℕ, 0 < n → #{a : α | a ^ n = 1} ≤ n) include hn @[to_additive] -private theorem card_pow_eq_one_eq_orderOf_aux (a : α) : - (Finset.univ.filter fun b : α => b ^ orderOf a = 1).card = orderOf a := +private theorem card_pow_eq_one_eq_orderOf_aux (a : α) : #{b : α | b ^ orderOf a = 1} = orderOf a := le_antisymm (hn _ (orderOf_pos a)) (calc orderOf a = @Fintype.card (zpowers a) (id _) := Fintype.card_zpowers.symm _ ≤ - @Fintype.card (↑(univ.filter fun b : α => b ^ orderOf a = 1) : Set α) + @Fintype.card (({b : α | b ^ orderOf a = 1} : Finset _) : Set α) (Fintype.ofFinset _ fun _ => Iff.rfl) := (@Fintype.card_le_of_injective (zpowers a) - (↑(univ.filter fun b : α => b ^ orderOf a = 1) : Set α) (id _) (id _) + (({b : α | b ^ orderOf a = 1} : Finset _) : Set α) (id _) (id _) (fun b => ⟨b.1, mem_filter.2 @@ -405,25 +403,21 @@ private theorem card_pow_eq_one_eq_orderOf_aux (a : α) : rw [← hi, ← zpow_natCast, ← zpow_mul, mul_comm, zpow_mul, zpow_natCast, pow_orderOf_eq_one, one_zpow]⟩⟩) fun _ _ h => Subtype.eq (Subtype.mk.inj h)) - _ = (univ.filter fun b : α => b ^ orderOf a = 1).card := Fintype.card_ofFinset _ _ + _ = #{b : α | b ^ orderOf a = 1} := Fintype.card_ofFinset _ _ ) -- Use φ for `Nat.totient` open Nat @[to_additive] -private theorem card_orderOf_eq_totient_aux₁ : - ∀ {d : ℕ}, - d ∣ Fintype.card α → - 0 < (univ.filter fun a : α => orderOf a = d).card → - (univ.filter fun a : α => orderOf a = d).card = φ d := by - intro d hd hpos +private theorem card_orderOf_eq_totient_aux₁ {d : ℕ} (hd : d ∣ Fintype.card α) + (hpos : 0 < #{a : α | orderOf a = d}) : #{a : α | orderOf a = d} = φ d := by induction' d using Nat.strongRec' with d IH rcases Decidable.eq_or_ne d 0 with (rfl | hd0) · cases Fintype.card_ne_zero (eq_zero_of_zero_dvd hd) - rcases card_pos.1 hpos with ⟨a, ha'⟩ + rcases Finset.card_pos.1 hpos with ⟨a, ha'⟩ have ha : orderOf a = d := (mem_filter.1 ha').2 have h1 : - (∑ m ∈ d.properDivisors, (univ.filter fun a : α => orderOf a = m).card) = + (∑ m ∈ d.properDivisors, #{a : α | orderOf a = m}) = ∑ m ∈ d.properDivisors, φ m := by refine Finset.sum_congr rfl fun m hm => ?_ simp only [mem_filter, mem_range, mem_properDivisors] at hm @@ -431,7 +425,7 @@ private theorem card_orderOf_eq_totient_aux₁ : 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) = + (∑ m ∈ d.divisors, #{a : α | orderOf a = m}) = ∑ m ∈ d.divisors, φ m := by rw [← filter_dvd_eq_divisors hd0, sum_card_orderOf_eq_card_pow_eq_one hd0, filter_dvd_eq_divisors hd0, sum_totient, ← ha, card_pow_eq_one_eq_orderOf_aux hn a] @@ -439,7 +433,7 @@ private theorem card_orderOf_eq_totient_aux₁ : @[to_additive] theorem card_orderOf_eq_totient_aux₂ {d : ℕ} (hd : d ∣ Fintype.card α) : - (univ.filter fun a : α => orderOf a = d).card = φ d := by + #{a : α | orderOf a = d} = φ d := by let c := Fintype.card α have hc0 : 0 < c := Fintype.card_pos_iff.2 ⟨1⟩ apply card_orderOf_eq_totient_aux₁ hn hd @@ -448,11 +442,11 @@ theorem card_orderOf_eq_totient_aux₂ {d : ℕ} (hd : d ∣ Fintype.card α) : simp_rw [not_lt, Nat.le_zero, Finset.card_eq_zero] at h0 apply lt_irrefl c calc - c = ∑ m ∈ c.divisors, (univ.filter fun a : α => orderOf a = m).card := by + c = ∑ m ∈ c.divisors, #{a : α | orderOf a = m} := by simp only [← filter_dvd_eq_divisors hc0.ne', sum_card_orderOf_eq_card_pow_eq_one hc0.ne'] apply congr_arg card simp [c] - _ = ∑ m ∈ c.divisors.erase d, (univ.filter fun a : α => orderOf a = m).card := by + _ = ∑ m ∈ c.divisors.erase d, #{a : α | orderOf a = m} := by rw [eq_comm] refine sum_subset (erase_subset _ _) fun m hm₁ hm₂ => ?_ have : m = d := by @@ -464,7 +458,7 @@ theorem card_orderOf_eq_totient_aux₂ {d : ℕ} (hd : d ∣ Fintype.card α) : have hmc : m ∣ c := by simp only [mem_erase, mem_divisors] at hm tauto - rcases (filter (fun a : α => orderOf a = m) univ).card.eq_zero_or_pos with (h1 | h1) + obtain h1 | h1 := (#{a : α | orderOf a = m}).eq_zero_or_pos · simp [h1] · simp [card_orderOf_eq_totient_aux₁ hn hmc h1] _ < ∑ m ∈ c.divisors, φ m := @@ -473,7 +467,7 @@ theorem card_orderOf_eq_totient_aux₂ {d : ℕ} (hd : d ∣ Fintype.card α) : @[to_additive isAddCyclic_of_card_nsmul_eq_zero_le] theorem isCyclic_of_card_pow_eq_one_le : IsCyclic α := - have : (univ.filter fun a : α => orderOf a = Fintype.card α).Nonempty := + have : Finset.Nonempty {a : α | orderOf a = Fintype.card α} := card_pos.1 <| by rw [card_orderOf_eq_totient_aux₂ hn dvd_rfl, totient_pos] apply Fintype.card_pos @@ -486,8 +480,8 @@ alias isAddCyclic_of_card_pow_eq_one_le := isAddCyclic_of_card_nsmul_eq_zero_le end Totient @[to_additive] -theorem IsCyclic.card_orderOf_eq_totient [IsCyclic α] [Fintype α] {d : ℕ} - (hd : d ∣ Fintype.card α) : (univ.filter fun a : α => orderOf a = d).card = totient d := by +lemma IsCyclic.card_orderOf_eq_totient [IsCyclic α] [Fintype α] {d : ℕ} (hd : d ∣ Fintype.card α) : + #{a : α | orderOf a = d} = totient d := by classical apply card_orderOf_eq_totient_aux₂ (fun n => IsCyclic.card_pow_eq_one_le) hd @[deprecated (since := "2024-02-21")] diff --git a/Mathlib/InformationTheory/Hamming.lean b/Mathlib/InformationTheory/Hamming.lean index 63b352144a9985..90ccdcdd7d04a4 100644 --- a/Mathlib/InformationTheory/Hamming.lean +++ b/Mathlib/InformationTheory/Hamming.lean @@ -34,8 +34,7 @@ variable {α ι : Type*} {β : ι → Type*} [Fintype ι] [∀ i, DecidableEq ( variable {γ : ι → Type*} [∀ i, DecidableEq (γ i)] /-- The Hamming distance function to the naturals. -/ -def hammingDist (x y : ∀ i, β i) : ℕ := - (univ.filter fun i => x i ≠ y i).card +def hammingDist (x y : ∀ i, β i) : ℕ := #{i | x i ≠ y i} /-- Corresponds to `dist_self`. -/ @[simp] @@ -133,8 +132,7 @@ section Zero variable [∀ i, Zero (β i)] [∀ i, Zero (γ i)] /-- The Hamming weight function to the naturals. -/ -def hammingNorm (x : ∀ i, β i) : ℕ := - (univ.filter (x · ≠ 0)).card +def hammingNorm (x : ∀ i, β i) : ℕ := #{i | x i ≠ 0} /-- Corresponds to `dist_zero_right`. -/ @[simp] diff --git a/Mathlib/LinearAlgebra/AffineSpace/Combination.lean b/Mathlib/LinearAlgebra/AffineSpace/Combination.lean index c3adfc82235fd8..33d82512d88eb2 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/Combination.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/Combination.lean @@ -202,18 +202,18 @@ theorem weightedVSubOfPoint_sdiff_sub [DecidableEq ι] {s₂ : Finset ι} (h : s s.weightedVSubOfPoint p b w := by rw [map_neg, sub_neg_eq_add, s.weightedVSubOfPoint_sdiff h] -/-- A weighted sum over `s.subtype pred` equals one over `s.filter pred`. -/ +/-- A weighted sum over `s.subtype pred` equals one over `{x ∈ s | pred x}`. -/ theorem weightedVSubOfPoint_subtype_eq_filter (w : ι → k) (p : ι → P) (b : P) (pred : ι → Prop) [DecidablePred pred] : ((s.subtype pred).weightedVSubOfPoint (fun i => p i) b fun i => w i) = - (s.filter pred).weightedVSubOfPoint p b w := by + {x ∈ s | pred x}.weightedVSubOfPoint p b w := by rw [weightedVSubOfPoint_apply, weightedVSubOfPoint_apply, ← sum_subtype_eq_sum_filter] -/-- A weighted sum over `s.filter pred` equals one over `s` if all the weights at indices in `s` +/-- A weighted sum over `{x ∈ s | pred x}` equals one over `s` if all the weights at indices in `s` not satisfying `pred` are zero. -/ theorem weightedVSubOfPoint_filter_of_ne (w : ι → k) (p : ι → P) (b : P) {pred : ι → Prop} [DecidablePred pred] (h : ∀ i ∈ s, w i ≠ 0 → pred i) : - (s.filter pred).weightedVSubOfPoint p b w = s.weightedVSubOfPoint p b w := by + {x ∈ s | pred x}.weightedVSubOfPoint p b w = s.weightedVSubOfPoint p b w := by rw [weightedVSubOfPoint_apply, weightedVSubOfPoint_apply, sum_filter_of_ne] intro i hi hne refine h i hi ?_ @@ -319,17 +319,17 @@ theorem weightedVSub_sdiff_sub [DecidableEq ι] {s₂ : Finset ι} (h : s₂ ⊆ (p : ι → P) : (s \ s₂).weightedVSub p w - s₂.weightedVSub p (-w) = s.weightedVSub p w := s.weightedVSubOfPoint_sdiff_sub h _ _ _ -/-- A weighted sum over `s.subtype pred` equals one over `s.filter pred`. -/ +/-- A weighted sum over `s.subtype pred` equals one over `{x ∈ s | pred x}`. -/ theorem weightedVSub_subtype_eq_filter (w : ι → k) (p : ι → P) (pred : ι → Prop) [DecidablePred pred] : ((s.subtype pred).weightedVSub (fun i => p i) fun i => w i) = - (s.filter pred).weightedVSub p w := + {x ∈ s | pred x}.weightedVSub p w := s.weightedVSubOfPoint_subtype_eq_filter _ _ _ _ -/-- A weighted sum over `s.filter pred` equals one over `s` if all the weights at indices in `s` +/-- A weighted sum over `{x ∈ s | pred x}` equals one over `s` if all the weights at indices in `s` not satisfying `pred` are zero. -/ theorem weightedVSub_filter_of_ne (w : ι → k) (p : ι → P) {pred : ι → Prop} [DecidablePred pred] - (h : ∀ i ∈ s, w i ≠ 0 → pred i) : (s.filter pred).weightedVSub p w = s.weightedVSub p w := + (h : ∀ i ∈ s, w i ≠ 0 → pred i) : {x ∈ s | pred x}.weightedVSub p w = s.weightedVSub p w := s.weightedVSubOfPoint_filter_of_ne _ _ _ h /-- A constant multiplier of the weights in `weightedVSub_of` may be moved outside the sum. -/ @@ -497,7 +497,7 @@ theorem affineCombination_sdiff_sub [DecidableEq ι] {s₂ : Finset ι} (h : s the affine combination of the other points with the given weights. -/ theorem affineCombination_eq_of_weightedVSub_eq_zero_of_eq_neg_one {w : ι → k} {p : ι → P} (hw : s.weightedVSub p w = (0 : V)) {i : ι} [DecidablePred (· ≠ i)] (his : i ∈ s) - (hwi : w i = -1) : (s.filter (· ≠ i)).affineCombination k p w = p i := by + (hwi : w i = -1) : {x ∈ s | x ≠ i}.affineCombination k p w = p i := by classical rw [← @vsub_eq_zero_iff_eq V, ← hw, ← s.affineCombination_sdiff_sub (singleton_subset_iff.2 his), sdiff_singleton_eq_erase, @@ -507,18 +507,18 @@ theorem affineCombination_eq_of_weightedVSub_eq_zero_of_eq_neg_one {w : ι → k · simp [hwi] · simp -/-- An affine combination over `s.subtype pred` equals one over `s.filter pred`. -/ +/-- An affine combination over `s.subtype pred` equals one over `{x ∈ s | pred x}`. -/ theorem affineCombination_subtype_eq_filter (w : ι → k) (p : ι → P) (pred : ι → Prop) [DecidablePred pred] : ((s.subtype pred).affineCombination k (fun i => p i) fun i => w i) = - (s.filter pred).affineCombination k p w := by + {x ∈ s | pred x}.affineCombination k p w := by rw [affineCombination_apply, affineCombination_apply, weightedVSubOfPoint_subtype_eq_filter] -/-- An affine combination over `s.filter pred` equals one over `s` if all the weights at indices +/-- An affine combination over `{x ∈ s | pred x}` equals one over `s` if all the weights at indices in `s` not satisfying `pred` are zero. -/ theorem affineCombination_filter_of_ne (w : ι → k) (p : ι → P) {pred : ι → Prop} [DecidablePred pred] (h : ∀ i ∈ s, w i ≠ 0 → pred i) : - (s.filter pred).affineCombination k p w = s.affineCombination k p w := by + {x ∈ s | pred x}.affineCombination k p w = s.affineCombination k p w := by rw [affineCombination_apply, affineCombination_apply, s.weightedVSubOfPoint_filter_of_ne _ _ _ h] @@ -712,29 +712,29 @@ variable [AffineSpace V P] {ι : Type*} (s : Finset ι) {ι₂ : Type*} (s₂ : /-- The weights for the centroid of some points. -/ def centroidWeights : ι → k := - Function.const ι (card s : k)⁻¹ + Function.const ι (#s : k)⁻¹ /-- `centroidWeights` at any point. -/ @[simp] -theorem centroidWeights_apply (i : ι) : s.centroidWeights k i = (card s : k)⁻¹ := +theorem centroidWeights_apply (i : ι) : s.centroidWeights k i = (#s : k)⁻¹ := rfl /-- `centroidWeights` equals a constant function. -/ -theorem centroidWeights_eq_const : s.centroidWeights k = Function.const ι (card s : k)⁻¹ := +theorem centroidWeights_eq_const : s.centroidWeights k = Function.const ι (#s : k)⁻¹ := rfl variable {k} /-- The weights in the centroid sum to 1, if the number of points, converted to `k`, is not zero. -/ -theorem sum_centroidWeights_eq_one_of_cast_card_ne_zero (h : (card s : k) ≠ 0) : +theorem sum_centroidWeights_eq_one_of_cast_card_ne_zero (h : (#s : k) ≠ 0) : ∑ i ∈ s, s.centroidWeights k i = 1 := by simp [h] variable (k) /-- In the characteristic zero case, the weights in the centroid sum to 1 if the number of points is not zero. -/ -theorem sum_centroidWeights_eq_one_of_card_ne_zero [CharZero k] (h : card s ≠ 0) : +theorem sum_centroidWeights_eq_one_of_card_ne_zero [CharZero k] (h : #s ≠ 0) : ∑ i ∈ s, s.centroidWeights k i = 1 := by -- Porting note: `simp` cannot find `mul_inv_cancel` and does not use `norm_cast` simp only [centroidWeights_apply, sum_const, nsmul_eq_mul, ne_eq, Nat.cast_eq_zero, card_eq_zero] @@ -749,7 +749,7 @@ theorem sum_centroidWeights_eq_one_of_nonempty [CharZero k] (h : s.Nonempty) : /-- In the characteristic zero case, the weights in the centroid sum to 1 if the number of points is `n + 1`. -/ -theorem sum_centroidWeights_eq_one_of_card_eq_add_one [CharZero k] {n : ℕ} (h : card s = n + 1) : +theorem sum_centroidWeights_eq_one_of_card_eq_add_one [CharZero k] {n : ℕ} (h : #s = n + 1) : ∑ i ∈ s, s.centroidWeights k i = 1 := s.sum_centroidWeights_eq_one_of_card_ne_zero k (h.symm ▸ Nat.succ_ne_zero n) @@ -780,7 +780,7 @@ theorem centroid_pair [DecidableEq ι] [Invertible (2 : k)] (p : ι → P) (i₁ ({i₁, i₂} : Finset ι).centroid k p = (2⁻¹ : k) • (p i₂ -ᵥ p i₁) +ᵥ p i₁ := by by_cases h : i₁ = i₂ · simp [h] - · have hc : (card ({i₁, i₂} : Finset ι) : k) ≠ 0 := by + · have hc : (#{i₁, i₂} : k) ≠ 0 := by rw [card_insert_of_not_mem (not_mem_singleton.2 h), card_singleton] norm_num exact Invertible.ne_zero _ @@ -825,7 +825,7 @@ theorem sum_centroidWeightsIndicator [Fintype ι] : indexed by a `Fintype` sum to 1 if the number of points is not zero. -/ theorem sum_centroidWeightsIndicator_eq_one_of_card_ne_zero [CharZero k] [Fintype ι] - (h : card s ≠ 0) : ∑ i, s.centroidWeightsIndicator k i = 1 := by + (h : #s ≠ 0) : ∑ i, s.centroidWeightsIndicator k i = 1 := by rw [sum_centroidWeightsIndicator] exact s.sum_centroidWeights_eq_one_of_card_ne_zero k h @@ -839,7 +839,7 @@ theorem sum_centroidWeightsIndicator_eq_one_of_nonempty [CharZero k] [Fintype ι /-- In the characteristic zero case, the weights in the centroid indexed by a `Fintype` sum to 1 if the number of points is `n + 1`. -/ theorem sum_centroidWeightsIndicator_eq_one_of_card_eq_add_one [CharZero k] [Fintype ι] {n : ℕ} - (h : card s = n + 1) : ∑ i, s.centroidWeightsIndicator k i = 1 := by + (h : #s = n + 1) : ∑ i, s.centroidWeightsIndicator k i = 1 := by rw [sum_centroidWeightsIndicator] exact s.sum_centroidWeights_eq_one_of_card_eq_add_one k h @@ -1099,7 +1099,7 @@ open Set Finset /-- The centroid lies in the affine span if the number of points, converted to `k`, is not zero. -/ theorem centroid_mem_affineSpan_of_cast_card_ne_zero {s : Finset ι} (p : ι → P) - (h : (card s : k) ≠ 0) : s.centroid k p ∈ affineSpan k (range p) := + (h : (#s : k) ≠ 0) : s.centroid k p ∈ affineSpan k (range p) := affineCombination_mem_affineSpan (s.sum_centroidWeights_eq_one_of_cast_card_ne_zero h) p variable (k) @@ -1107,7 +1107,7 @@ variable (k) /-- In the characteristic zero case, the centroid lies in the affine span if the number of points is not zero. -/ theorem centroid_mem_affineSpan_of_card_ne_zero [CharZero k] {s : Finset ι} (p : ι → P) - (h : card s ≠ 0) : s.centroid k p ∈ affineSpan k (range p) := + (h : #s ≠ 0) : s.centroid k p ∈ affineSpan k (range p) := affineCombination_mem_affineSpan (s.sum_centroidWeights_eq_one_of_card_ne_zero k h) p /-- In the characteristic zero case, the centroid lies in the affine @@ -1119,7 +1119,7 @@ theorem centroid_mem_affineSpan_of_nonempty [CharZero k] {s : Finset ι} (p : ι /-- In the characteristic zero case, the centroid lies in the affine span if the number of points is `n + 1`. -/ theorem centroid_mem_affineSpan_of_card_eq_add_one [CharZero k] {s : Finset ι} (p : ι → P) {n : ℕ} - (h : card s = n + 1) : s.centroid k p ∈ affineSpan k (range p) := + (h : #s = n + 1) : s.centroid k p ∈ affineSpan k (range p) := affineCombination_mem_affineSpan (s.sum_centroidWeights_eq_one_of_card_eq_add_one k h) p end DivisionRing diff --git a/Mathlib/LinearAlgebra/AffineSpace/FiniteDimensional.lean b/Mathlib/LinearAlgebra/AffineSpace/FiniteDimensional.lean index 935dc82abfb3b8..0ef068c764695e 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/FiniteDimensional.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/FiniteDimensional.lean @@ -23,6 +23,7 @@ subspaces of affine spaces. noncomputable section open Affine +open scoped Finset section AffineSpace' @@ -87,18 +88,18 @@ variable {k} /-- The `vectorSpan` of a finite subset of an affinely independent family has dimension one less than its cardinality. -/ theorem AffineIndependent.finrank_vectorSpan_image_finset [DecidableEq P] - {p : ι → P} (hi : AffineIndependent k p) {s : Finset ι} {n : ℕ} (hc : Finset.card s = n + 1) : + {p : ι → P} (hi : AffineIndependent k p) {s : Finset ι} {n : ℕ} (hc : #s = n + 1) : finrank k (vectorSpan k (s.image p : Set P)) = n := by classical have hi' := hi.range.mono (Set.image_subset_range p ↑s) - have hc' : (s.image p).card = n + 1 := by rwa [s.card_image_of_injective hi.injective] + have hc' : #(s.image p) = n + 1 := by rwa [s.card_image_of_injective hi.injective] have hn : (s.image p).Nonempty := by simp [hc', ← Finset.card_pos] rcases hn with ⟨p₁, hp₁⟩ have hp₁' : p₁ ∈ p '' s := by simpa using hp₁ rw [affineIndependent_set_iff_linearIndependent_vsub k hp₁', ← Finset.coe_singleton, ← Finset.coe_image, ← Finset.coe_sdiff, Finset.sdiff_singleton_eq_erase, ← Finset.coe_image] at hi' - have hc : (Finset.image (fun p : P => p -ᵥ p₁) ((Finset.image p s).erase p₁)).card = n := by + have hc : #(((s.image p).erase p₁).image (· -ᵥ p₁)) = n := by rw [Finset.card_image_of_injective _ (vsub_left_injective _), Finset.card_erase_of_mem hp₁] exact Nat.pred_eq_of_eq_succ hc' rwa [vectorSpan_eq_span_vsub_finset_right_ne k hp₁, finrank_span_finset_eq_card, hc] @@ -132,7 +133,7 @@ variable (k) /-- The `vectorSpan` of `n + 1` points in an indexed family has dimension at most `n`. -/ theorem finrank_vectorSpan_image_finset_le [DecidableEq P] (p : ι → P) (s : Finset ι) {n : ℕ} - (hc : Finset.card s = n + 1) : finrank k (vectorSpan k (s.image p : Set P)) ≤ n := by + (hc : #s = n + 1) : finrank k (vectorSpan k (s.image p : Set P)) ≤ n := by classical have hn : (s.image p).Nonempty := by rw [Finset.image_nonempty, ← Finset.card_pos, hc] @@ -214,7 +215,7 @@ open Finset in cardinality is at most the cardinality of that finset. -/ lemma AffineIndependent.card_le_card_of_subset_affineSpan {s t : Finset V} (hs : AffineIndependent k ((↑) : s → V)) (hst : (s : Set V) ⊆ affineSpan k (t : Set V)) : - s.card ≤ t.card := by + #s ≤ #t := by obtain rfl | hs' := s.eq_empty_or_nonempty · simp obtain rfl | ht' := t.eq_empty_or_nonempty @@ -234,7 +235,7 @@ open Finset in another finset, then its cardinality is strictly less than the cardinality of that finset. -/ lemma AffineIndependent.card_lt_card_of_affineSpan_lt_affineSpan {s t : Finset V} (hs : AffineIndependent k ((↑) : s → V)) - (hst : affineSpan k (s : Set V) < affineSpan k (t : Set V)) : s.card < t.card := by + (hst : affineSpan k (s : Set V) < affineSpan k (t : Set V)) : #s < #t := by obtain rfl | hs' := s.eq_empty_or_nonempty · simpa [card_pos] using hst obtain rfl | ht' := t.eq_empty_or_nonempty @@ -255,7 +256,7 @@ cardinality, it equals that submodule. -/ theorem AffineIndependent.vectorSpan_image_finset_eq_of_le_of_card_eq_finrank_add_one [DecidableEq P] {p : ι → P} (hi : AffineIndependent k p) {s : Finset ι} {sm : Submodule k V} [FiniteDimensional k sm] - (hle : vectorSpan k (s.image p : Set P) ≤ sm) (hc : Finset.card s = finrank k sm + 1) : + (hle : vectorSpan k (s.image p : Set P) ≤ sm) (hc : #s = finrank k sm + 1) : vectorSpan k (s.image p : Set P) = sm := Submodule.eq_of_le_of_finrank_eq hle <| hi.finrank_vectorSpan_image_finset hc @@ -275,7 +276,7 @@ theorem AffineIndependent.affineSpan_image_finset_eq_of_le_of_card_eq_finrank_ad [DecidableEq P] {p : ι → P} (hi : AffineIndependent k p) {s : Finset ι} {sp : AffineSubspace k P} [FiniteDimensional k sp.direction] (hle : affineSpan k (s.image p : Set P) ≤ sp) - (hc : Finset.card s = finrank k sp.direction + 1) : affineSpan k (s.image p : Set P) = sp := by + (hc : #s = finrank k sp.direction + 1) : affineSpan k (s.image p : Set P) = sp := by have hn : s.Nonempty := by rw [← Finset.card_pos, hc] apply Nat.succ_pos diff --git a/Mathlib/LinearAlgebra/AffineSpace/Independent.lean b/Mathlib/LinearAlgebra/AffineSpace/Independent.lean index 0fe67bc39f3f97..f8801772f8d029 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/Independent.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/Independent.lean @@ -810,19 +810,19 @@ add_decl_doc Affine.Simplex.ext_iff /-- A face of a simplex is a simplex with the given subset of points. -/ -def face {n : ℕ} (s : Simplex k P n) {fs : Finset (Fin (n + 1))} {m : ℕ} (h : fs.card = m + 1) : +def face {n : ℕ} (s : Simplex k P n) {fs : Finset (Fin (n + 1))} {m : ℕ} (h : #fs = m + 1) : Simplex k P m := ⟨s.points ∘ fs.orderEmbOfFin h, s.independent.comp_embedding (fs.orderEmbOfFin h).toEmbedding⟩ /-- The points of a face of a simplex are given by `mono_of_fin`. -/ theorem face_points {n : ℕ} (s : Simplex k P n) {fs : Finset (Fin (n + 1))} {m : ℕ} - (h : fs.card = m + 1) (i : Fin (m + 1)) : + (h : #fs = m + 1) (i : Fin (m + 1)) : (s.face h).points i = s.points (fs.orderEmbOfFin h i) := rfl /-- The points of a face of a simplex are given by `mono_of_fin`. -/ theorem face_points' {n : ℕ} (s : Simplex k P n) {fs : Finset (Fin (n + 1))} {m : ℕ} - (h : fs.card = m + 1) : (s.face h).points = s.points ∘ fs.orderEmbOfFin h := + (h : #fs = m + 1) : (s.face h).points = s.points ∘ fs.orderEmbOfFin h := rfl /-- A single-point face equals the 0-simplex constructed with @@ -838,7 +838,7 @@ theorem face_eq_mkOfPoint {n : ℕ} (s : Simplex k P n) (i : Fin (n + 1)) : /-- The set of points of a face. -/ @[simp] theorem range_face_points {n : ℕ} (s : Simplex k P n) {fs : Finset (Fin (n + 1))} {m : ℕ} - (h : fs.card = m + 1) : Set.range (s.face h).points = s.points '' ↑fs := by + (h : #fs = m + 1) : Set.range (s.face h).points = s.points '' ↑fs := by rw [face_points', Set.range_comp, Finset.range_orderEmbOfFin] /-- Remap a simplex along an `Equiv` of index types. -/ @@ -889,7 +889,7 @@ variable {k : Type*} {V : Type*} {P : Type*} [DivisionRing k] [AddCommGroup V] [ the points. -/ @[simp] theorem face_centroid_eq_centroid {n : ℕ} (s : Simplex k P n) {fs : Finset (Fin (n + 1))} {m : ℕ} - (h : fs.card = m + 1) : Finset.univ.centroid k (s.face h).points = fs.centroid k s.points := by + (h : #fs = m + 1) : Finset.univ.centroid k (s.face h).points = fs.centroid k s.points := by convert (Finset.univ.centroid_map k (fs.orderEmbOfFin h).toEmbedding s.points).symm rw [← Finset.coe_inj, Finset.coe_map, Finset.coe_univ, Set.image_univ] simp @@ -899,7 +899,7 @@ two subsets of the points of a simplex are equal if and only if those faces are given by the same subset of points. -/ @[simp] theorem centroid_eq_iff [CharZero k] {n : ℕ} (s : Simplex k P n) {fs₁ fs₂ : Finset (Fin (n + 1))} - {m₁ m₂ : ℕ} (h₁ : fs₁.card = m₁ + 1) (h₂ : fs₂.card = m₂ + 1) : + {m₁ m₂ : ℕ} (h₁ : #fs₁ = m₁ + 1) (h₂ : #fs₂ = m₂ + 1) : fs₁.centroid k s.points = fs₂.centroid k s.points ↔ fs₁ = fs₂ := by refine ⟨fun h => ?_, @congrArg _ _ fs₁ fs₂ (fun z => Finset.centroid k z s.points)⟩ rw [Finset.centroid_eq_affineCombination_fintype, @@ -924,7 +924,7 @@ theorem centroid_eq_iff [CharZero k] {n : ℕ} (s : Simplex k P n) {fs₁ fs₂ faces of a simplex are equal if and only if those faces are given by the same subset of points. -/ theorem face_centroid_eq_iff [CharZero k] {n : ℕ} (s : Simplex k P n) - {fs₁ fs₂ : Finset (Fin (n + 1))} {m₁ m₂ : ℕ} (h₁ : fs₁.card = m₁ + 1) (h₂ : fs₂.card = m₂ + 1) : + {fs₁ fs₂ : Finset (Fin (n + 1))} {m₁ m₂ : ℕ} (h₁ : #fs₁ = m₁ + 1) (h₂ : #fs₂ = m₂ + 1) : Finset.univ.centroid k (s.face h₁).points = Finset.univ.centroid k (s.face h₂).points ↔ fs₁ = fs₂ := by rw [face_centroid_eq_centroid, face_centroid_eq_centroid] diff --git a/Mathlib/LinearAlgebra/Lagrange.lean b/Mathlib/LinearAlgebra/Lagrange.lean index 9b87bad3dd5faa..b683703a0116d4 100644 --- a/Mathlib/LinearAlgebra/Lagrange.lean +++ b/Mathlib/LinearAlgebra/Lagrange.lean @@ -36,10 +36,11 @@ variable {R : Type*} [CommRing R] [IsDomain R] {f g : R[X]} section Finset open Function Fintype +open scoped Finset variable (s : Finset R) -theorem eq_zero_of_degree_lt_of_eval_finset_eq_zero (degree_f_lt : f.degree < s.card) +theorem eq_zero_of_degree_lt_of_eval_finset_eq_zero (degree_f_lt : f.degree < #s) (eval_f : ∀ x ∈ s, f.eval x = 0) : f = 0 := by rw [← mem_degreeLT] at degree_f_lt simp_rw [eval_eq_sum_degreeLTEquiv degree_f_lt] at eval_f @@ -49,15 +50,15 @@ theorem eq_zero_of_degree_lt_of_eval_finset_eq_zero (degree_f_lt : f.degree < s. (Injective.comp (Embedding.subtype _).inj' (equivFinOfCardEq (card_coe _)).symm.injective) fun _ => eval_f _ (Finset.coe_mem _) -theorem eq_of_degree_sub_lt_of_eval_finset_eq (degree_fg_lt : (f - g).degree < s.card) +theorem eq_of_degree_sub_lt_of_eval_finset_eq (degree_fg_lt : (f - g).degree < #s) (eval_fg : ∀ x ∈ s, f.eval x = g.eval x) : f = g := by rw [← sub_eq_zero] refine eq_zero_of_degree_lt_of_eval_finset_eq_zero _ degree_fg_lt ?_ simp_rw [eval_sub, sub_eq_zero] exact eval_fg -theorem eq_of_degrees_lt_of_eval_finset_eq (degree_f_lt : f.degree < s.card) - (degree_g_lt : g.degree < s.card) (eval_fg : ∀ x ∈ s, f.eval x = g.eval x) : f = g := by +theorem eq_of_degrees_lt_of_eval_finset_eq (degree_f_lt : f.degree < #s) + (degree_g_lt : g.degree < #s) (eval_fg : ∀ x ∈ s, f.eval x = g.eval x) : f = g := by rw [← mem_degreeLT] at degree_f_lt degree_g_lt refine eq_of_degree_sub_lt_of_eval_finset_eq _ ?_ eval_fg rw [← mem_degreeLT]; exact Submodule.sub_mem _ degree_f_lt degree_g_lt @@ -67,7 +68,7 @@ Two polynomials, with the same degree and leading coefficient, which have the sa on a set of distinct values with cardinality equal to the degree, are equal. -/ theorem eq_of_degree_le_of_eval_finset_eq - (h_deg_le : f.degree ≤ s.card) + (h_deg_le : f.degree ≤ #s) (h_deg_eq : f.degree = g.degree) (hlc : f.leadingCoeff = g.leadingCoeff) (h_eval : ∀ x ∈ s, f.eval x = g.eval x) : @@ -86,7 +87,7 @@ open Finset variable {ι : Type*} {v : ι → R} (s : Finset ι) theorem eq_zero_of_degree_lt_of_eval_index_eq_zero (hvs : Set.InjOn v s) - (degree_f_lt : f.degree < s.card) (eval_f : ∀ i ∈ s, f.eval (v i) = 0) : f = 0 := by + (degree_f_lt : f.degree < #s) (eval_f : ∀ i ∈ s, f.eval (v i) = 0) : f = 0 := by classical rw [← card_image_of_injOn hvs] at degree_f_lt refine eq_zero_of_degree_lt_of_eval_finset_eq_zero _ degree_f_lt ?_ @@ -95,21 +96,21 @@ theorem eq_zero_of_degree_lt_of_eval_index_eq_zero (hvs : Set.InjOn v s) exact eval_f _ hj theorem eq_of_degree_sub_lt_of_eval_index_eq (hvs : Set.InjOn v s) - (degree_fg_lt : (f - g).degree < s.card) (eval_fg : ∀ i ∈ s, f.eval (v i) = g.eval (v i)) : + (degree_fg_lt : (f - g).degree < #s) (eval_fg : ∀ i ∈ s, f.eval (v i) = g.eval (v i)) : f = g := by rw [← sub_eq_zero] refine eq_zero_of_degree_lt_of_eval_index_eq_zero _ hvs degree_fg_lt ?_ simp_rw [eval_sub, sub_eq_zero] exact eval_fg -theorem eq_of_degrees_lt_of_eval_index_eq (hvs : Set.InjOn v s) (degree_f_lt : f.degree < s.card) - (degree_g_lt : g.degree < s.card) (eval_fg : ∀ i ∈ s, f.eval (v i) = g.eval (v i)) : f = g := by +theorem eq_of_degrees_lt_of_eval_index_eq (hvs : Set.InjOn v s) (degree_f_lt : f.degree < #s) + (degree_g_lt : g.degree < #s) (eval_fg : ∀ i ∈ s, f.eval (v i) = g.eval (v i)) : f = g := by refine eq_of_degree_sub_lt_of_eval_index_eq _ hvs ?_ eval_fg rw [← mem_degreeLT] at degree_f_lt degree_g_lt ⊢ exact Submodule.sub_mem _ degree_f_lt degree_g_lt theorem eq_of_degree_le_of_eval_index_eq (hvs : Set.InjOn v s) - (h_deg_le : f.degree ≤ s.card) + (h_deg_le : f.degree ≤ #s) (h_deg_eq : f.degree = g.degree) (hlc : f.leadingCoeff = g.leadingCoeff) (h_eval : ∀ i ∈ s, f.eval (v i) = g.eval (v i)) : f = g := by @@ -233,7 +234,7 @@ theorem eval_basis_of_ne (hij : i ≠ j) (hj : j ∈ s) : (Lagrange.basis s v i) @[simp] theorem natDegree_basis (hvs : Set.InjOn v s) (hi : i ∈ s) : - (Lagrange.basis s v i).natDegree = s.card - 1 := by + (Lagrange.basis s v i).natDegree = #s - 1 := by have H : ∀ j, j ∈ s.erase i → basisDivisor (v i) (v j) ≠ 0 := by simp_rw [Ne, mem_erase, basisDivisor_eq_zero_iff] exact fun j ⟨hij₁, hj⟩ hij₂ => hij₁ (hvs hj hi hij₂.symm) @@ -244,14 +245,14 @@ theorem natDegree_basis (hvs : Set.InjOn v s) (hi : i ∈ s) : exact H _ hj theorem degree_basis (hvs : Set.InjOn v s) (hi : i ∈ s) : - (Lagrange.basis s v i).degree = ↑(s.card - 1) := by + (Lagrange.basis s v i).degree = ↑(#s - 1) := by rw [degree_eq_natDegree (basis_ne_zero hvs hi), natDegree_basis hvs hi] -- Porting note: Added `Nat.cast_withBot` rewrites theorem sum_basis (hvs : Set.InjOn v s) (hs : s.Nonempty) : ∑ j ∈ s, Lagrange.basis s v j = 1 := by refine eq_of_degrees_lt_of_eval_index_eq s hvs (lt_of_le_of_lt (degree_sum_le _ _) ?_) ?_ ?_ - · rw [Nat.cast_withBot, Finset.sup_lt_iff (WithBot.bot_lt_coe s.card)] + · rw [Nat.cast_withBot, Finset.sup_lt_iff (WithBot.bot_lt_coe #s)] intro i hi rw [degree_basis hvs hi, Nat.cast_withBot, WithBot.coe_lt_coe] exact Nat.pred_lt (card_ne_zero_of_mem hi) @@ -282,7 +283,7 @@ open Finset /-- Lagrange interpolation: given a finset `s : Finset ι`, a nodal map `v : ι → F` injective on `s` and a value function `r : ι → F`, `interpolate s v r` is the unique -polynomial of degree `< s.card` that takes value `r i` on `v i` for all `i` in `s`. -/ +polynomial of degree `< #s` that takes value `r i` on `v i` for all `i` in `s`. -/ @[simps] def interpolate (s : Finset ι) (v : ι → F) : (ι → F) →ₗ[F] F[X] where toFun r := ∑ i ∈ s, C (r i) * Lagrange.basis s v i @@ -316,7 +317,7 @@ theorem eval_interpolate_at_node (hvs : Set.InjOn v s) (hi : i ∈ s) : rw [eval_basis_of_ne (mem_erase.mp H).1 hi, mul_zero] theorem degree_interpolate_le (hvs : Set.InjOn v s) : - (interpolate s v r).degree ≤ ↑(s.card - 1) := by + (interpolate s v r).degree ≤ ↑(#s - 1) := by refine (degree_sum_le _ _).trans ?_ rw [Finset.sup_le_iff] intro i hi @@ -326,7 +327,7 @@ theorem degree_interpolate_le (hvs : Set.InjOn v s) : · rw [degree_C hr, zero_add] -- Porting note: Added `Nat.cast_withBot` rewrites -theorem degree_interpolate_lt (hvs : Set.InjOn v s) : (interpolate s v r).degree < s.card := by +theorem degree_interpolate_lt (hvs : Set.InjOn v s) : (interpolate s v r).degree < #s := by rw [Nat.cast_withBot] rcases eq_empty_or_nonempty s with (rfl | h) · rw [interpolate_empty, degree_zero, card_empty] @@ -336,7 +337,7 @@ theorem degree_interpolate_lt (hvs : Set.InjOn v s) : (interpolate s v r).degree exact Nat.sub_lt (Nonempty.card_pos h) zero_lt_one theorem degree_interpolate_erase_lt (hvs : Set.InjOn v s) (hi : i ∈ s) : - (interpolate (s.erase i) v r).degree < ↑(s.card - 1) := by + (interpolate (s.erase i) v r).degree < ↑(#s - 1) := by rw [← Finset.card_erase_of_mem hi] exact degree_interpolate_lt _ (Set.InjOn.mono (coe_subset.mpr (erase_subset _ _)) hvs) @@ -352,12 +353,12 @@ theorem interpolate_eq_iff_values_eq_on (hvs : Set.InjOn v s) : interpolate s v r = interpolate s v r' ↔ ∀ i ∈ s, r i = r' i := ⟨values_eq_on_of_interpolate_eq _ _ hvs, interpolate_eq_of_values_eq_on _ _⟩ -theorem eq_interpolate {f : F[X]} (hvs : Set.InjOn v s) (degree_f_lt : f.degree < s.card) : +theorem eq_interpolate {f : F[X]} (hvs : Set.InjOn v s) (degree_f_lt : f.degree < #s) : f = interpolate s v fun i => f.eval (v i) := eq_of_degrees_lt_of_eval_index_eq _ hvs degree_f_lt (degree_interpolate_lt _ hvs) fun _ hi => (eval_interpolate_at_node (fun x ↦ eval (v x) f) hvs hi).symm -theorem eq_interpolate_of_eval_eq {f : F[X]} (hvs : Set.InjOn v s) (degree_f_lt : f.degree < s.card) +theorem eq_interpolate_of_eval_eq {f : F[X]} (hvs : Set.InjOn v s) (degree_f_lt : f.degree < #s) (eval_f : ∀ i ∈ s, f.eval (v i) = r i) : f = interpolate s v r := by rw [eq_interpolate hvs degree_f_lt] exact interpolate_eq_of_values_eq_on _ _ eval_f @@ -366,7 +367,7 @@ theorem eq_interpolate_of_eval_eq {f : F[X]} (hvs : Set.InjOn v s) (degree_f_lt unique polynomial of `degree < Fintype.card ι` which takes the value of the `r i` on the `v i`. -/ theorem eq_interpolate_iff {f : F[X]} (hvs : Set.InjOn v s) : - (f.degree < s.card ∧ ∀ i ∈ s, eval (v i) f = r i) ↔ f = interpolate s v r := by + (f.degree < #s ∧ ∀ i ∈ s, eval (v i) f = r i) ↔ f = interpolate s v r := by constructor <;> intro h · exact eq_interpolate_of_eval_eq _ hvs h.1 h.2 · rw [h] @@ -374,7 +375,7 @@ theorem eq_interpolate_iff {f : F[X]} (hvs : Set.InjOn v s) : /-- Lagrange interpolation induces isomorphism between functions from `s` and polynomials of degree less than `Fintype.card ι`. -/ -def funEquivDegreeLT (hvs : Set.InjOn v s) : degreeLT F s.card ≃ₗ[F] s → F where +def funEquivDegreeLT (hvs : Set.InjOn v s) : degreeLT F #s ≃ₗ[F] s → F where toFun f i := f.1.eval (v i) map_add' _ _ := funext fun _ => eval_add map_smul' c f := funext <| by simp @@ -399,15 +400,15 @@ theorem interpolate_eq_sum_interpolate_insert_sdiff (hvt : Set.InjOn v t) (hs : interpolate t v r = ∑ i ∈ s, interpolate (insert i (t \ s)) v r * Lagrange.basis s v i := by symm refine eq_interpolate_of_eval_eq _ hvt (lt_of_le_of_lt (degree_sum_le _ _) ?_) fun i hi => ?_ - · simp_rw [Nat.cast_withBot, Finset.sup_lt_iff (WithBot.bot_lt_coe t.card), degree_mul] + · simp_rw [Nat.cast_withBot, Finset.sup_lt_iff (WithBot.bot_lt_coe #t), degree_mul] intro i hi - have hs : 1 ≤ s.card := Nonempty.card_pos ⟨_, hi⟩ - have hst' : s.card ≤ t.card := card_le_card hst - have H : t.card = 1 + (t.card - s.card) + (s.card - 1) := by + have hs : 1 ≤ #s := Nonempty.card_pos ⟨_, hi⟩ + have hst' : #s ≤ #t := card_le_card hst + have H : #t = 1 + (#t - #s) + (#s - 1) := by rw [add_assoc, tsub_add_tsub_cancel hst' hs, ← add_tsub_assoc_of_le (hs.trans hst'), Nat.succ_add_sub_one, zero_add] rw [degree_basis (Set.InjOn.mono hst hvt) hi, H, WithBot.coe_add, Nat.cast_withBot, - WithBot.add_lt_add_iff_right (@WithBot.coe_ne_bot _ (s.card - 1))] + WithBot.add_lt_add_iff_right (@WithBot.coe_ne_bot _ (#s - 1))] convert degree_interpolate_lt _ (hvt.mono (coe_subset.mpr (insert_subset_iff.mpr ⟨hst hi, sdiff_subset⟩))) rw [card_insert_of_not_mem (not_mem_sdiff_of_mem_right hi), card_sdiff hst, add_comm] @@ -471,7 +472,7 @@ theorem nodal_empty : nodal ∅ v = 1 := by rfl @[simp] -theorem natDegree_nodal [Nontrivial R] : (nodal s v).natDegree = s.card := by +theorem natDegree_nodal [Nontrivial R] : (nodal s v).natDegree = #s := by simp_rw [nodal, natDegree_prod_of_monic (h := fun i _ => monic_X_sub_C (v i)), natDegree_X_sub_C, sum_const, smul_eq_mul, mul_one] @@ -482,7 +483,7 @@ rcases s.eq_empty_or_nonempty with (rfl | h) simp only [natDegree_nodal, h.card_pos] @[simp] -theorem degree_nodal [Nontrivial R] : (nodal s v).degree = s.card := by +theorem degree_nodal [Nontrivial R] : (nodal s v).degree = #s := by simp_rw [degree_eq_natDegree nodal_ne_zero, natDegree_nodal] theorem nodal_monic : (nodal s v).Monic := diff --git a/Mathlib/LinearAlgebra/Matrix/AbsoluteValue.lean b/Mathlib/LinearAlgebra/Matrix/AbsoluteValue.lean index 1af383318a9ae0..201767b2c0744b 100644 --- a/Mathlib/LinearAlgebra/Matrix/AbsoluteValue.lean +++ b/Mathlib/LinearAlgebra/Matrix/AbsoluteValue.lean @@ -50,19 +50,19 @@ theorem det_le {A : Matrix n n R} {abv : AbsoluteValue R S} {x : S} (hx : ∀ i theorem det_sum_le {ι : Type*} (s : Finset ι) {A : ι → Matrix n n R} {abv : AbsoluteValue R S} {x : S} (hx : ∀ k i j, abv (A k i j) ≤ x) : abv (det (∑ k ∈ s, A k)) ≤ - Nat.factorial (Fintype.card n) • (Finset.card s • x) ^ Fintype.card n := + Nat.factorial (Fintype.card n) • (#s• x) ^ Fintype.card n := det_le fun i j => calc abv ((∑ k ∈ s, A k) i j) = abv (∑ k ∈ s, A k i j) := by simp only [sum_apply] _ ≤ ∑ k ∈ s, abv (A k i j) := abv.sum_le _ _ _ ≤ ∑ _k ∈ s, x := sum_le_sum fun k _ => hx k i j - _ = s.card • x := sum_const _ + _ = #s • x := sum_const _ theorem det_sum_smul_le {ι : Type*} (s : Finset ι) {c : ι → R} {A : ι → Matrix n n R} {abv : AbsoluteValue R S} {x : S} (hx : ∀ k i j, abv (A k i j) ≤ x) {y : S} (hy : ∀ k, abv (c k) ≤ y) : abv (det (∑ k ∈ s, c k • A k)) ≤ - Nat.factorial (Fintype.card n) • (Finset.card s • y * x) ^ Fintype.card n := by + Nat.factorial (Fintype.card n) • (#s • y * x) ^ Fintype.card n := by simpa only [smul_mul_assoc] using det_sum_le s fun k i j => calc diff --git a/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean b/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean index c33327a8858aeb..2780bc5cd41eae 100644 --- a/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean +++ b/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean @@ -131,12 +131,9 @@ theorem det_mul (M N : Matrix n n R) : det (M * N) = det M * det N := det (M * N) = ∑ p : n → n, ∑ σ : Perm n, ε σ * ∏ i, M (σ i) (p i) * N (p i) i := by simp only [det_apply', mul_apply, prod_univ_sum, mul_sum, Fintype.piFinset_univ] rw [Finset.sum_comm] - _ = - ∑ p ∈ (@univ (n → n) _).filter Bijective, - ∑ σ : 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, mem_filter, mem_univ] using hbij) + _ = ∑ p : n → n with Bijective p, ∑ σ : Perm n, ε σ * ∏ i, M (σ i) (p i) * N (p i) i := by + refine (sum_subset (filter_subset _ _) fun f _ hbij ↦ det_mul_aux ?_).symm + 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) @@ -543,8 +540,7 @@ theorem det_blockDiagonal {o : Type*} [Fintype o] [DecidableEq o] (M : o → Mat simp_rw [Finset.prod_attach_univ, Finset.univ_pi_univ] -- We claim that the only permutations contributing to the sum are those that -- preserve their second component. - let preserving_snd : Finset (Equiv.Perm (n × o)) := - Finset.univ.filter fun σ => ∀ x, (σ x).snd = x.snd + let preserving_snd : Finset (Equiv.Perm (n × o)) := {σ | ∀ x, (σ x).snd = x.snd} have mem_preserving_snd : ∀ {σ : Equiv.Perm (n × o)}, σ ∈ preserving_snd ↔ ∀ x, (σ x).snd = x.snd := fun {σ} => Finset.mem_filter.trans ⟨fun h => h.2, fun h => ⟨Finset.mem_univ _, h⟩⟩ diff --git a/Mathlib/LinearAlgebra/Multilinear/Basic.lean b/Mathlib/LinearAlgebra/Multilinear/Basic.lean index f37192686c3991..39f1d7e414a2b0 100644 --- a/Mathlib/LinearAlgebra/Multilinear/Basic.lean +++ b/Mathlib/LinearAlgebra/Multilinear/Basic.lean @@ -71,8 +71,7 @@ something similar, but of the form `Fin.decidableEq n = _inst`, which is much ea since `_inst` is a free variable and so the equality can just be substituted. -/ - -open Function Fin Set +open Fin Function Finset Set universe uR uS uι v v' v₁ v₂ v₃ @@ -280,7 +279,7 @@ the other ones equal to a given value `z`. It is denoted by `f.restr s hk z`, wh proof that the cardinality of `s` is `k`. The implicit identification between `Fin k` and `s` that we use is the canonical (increasing) bijection. -/ def restr {k n : ℕ} (f : MultilinearMap R (fun _ : Fin n => M') M₂) (s : Finset (Fin n)) - (hk : s.card = k) (z : M') : MultilinearMap R (fun _ : Fin k => M') M₂ where + (hk : #s = k) (z : M') : MultilinearMap R (fun _ : Fin k => M') M₂ where toFun v := f fun j => if h : j ∈ s then v ((DFunLike.coe (s.orderIsoOfFin hk).symm) ⟨j, h⟩) else z /- Porting note: The proofs of the following two lemmas used to only use `erw` followed by `simp`, but it seems `erw` no longer unfolds or unifies well enough to work without more help. -/ @@ -445,7 +444,7 @@ open Fintype Finset `r n ∈ Aₙ`. This follows from multilinearity by expanding successively with respect to each coordinate. Here, we give an auxiliary statement tailored for an inductive proof. Use instead `map_sum_finset`. -/ -theorem map_sum_finset_aux [DecidableEq ι] [Fintype ι] {n : ℕ} (h : (∑ i, (A i).card) = n) : +theorem map_sum_finset_aux [DecidableEq ι] [Fintype ι] {n : ℕ} (h : (∑ i, #(A i)) = n) : (f fun i => ∑ j ∈ A i, g i j) = ∑ r ∈ piFinset A, f fun i => g i (r i) := by letI := fun i => Classical.decEq (α i) induction' n using Nat.strong_induction_on with n IH generalizing A @@ -457,11 +456,11 @@ theorem map_sum_finset_aux [DecidableEq ι] [Fintype ι] {n : ℕ} (h : (∑ i, push_neg at Ai_empty -- Otherwise, if all sets are at most singletons, then they are exactly singletons and the result -- is again straightforward - by_cases Ai_singleton : ∀ i, (A i).card ≤ 1 - · have Ai_card : ∀ i, (A i).card = 1 := by + by_cases Ai_singleton : ∀ i, #(A i) ≤ 1 + · have Ai_card : ∀ i, #(A i) = 1 := by intro i - have pos : Finset.card (A i) ≠ 0 := by simp [Finset.card_eq_zero, Ai_empty i] - have : Finset.card (A i) ≤ 1 := Ai_singleton i + have pos : #(A i) ≠ 0 := by simp [Finset.card_eq_zero, Ai_empty i] + have : #(A i) ≤ 1 := Ai_singleton i exact le_antisymm this (Nat.succ_le_of_lt (_root_.pos_iff_ne_zero.mpr pos)) have : ∀ r : ∀ i, α i, r ∈ piFinset A → (f fun i => g i (r i)) = f fun i => ∑ j ∈ A i, g i j := by @@ -480,7 +479,7 @@ theorem map_sum_finset_aux [DecidableEq ι] [Fintype ι] {n : ℕ} (h : (∑ i, -- for `i ≠ i₀`, apply the inductive assumption to `B` and `C`, and add up the corresponding -- parts to get the sum for `A`. push_neg at Ai_singleton - obtain ⟨i₀, hi₀⟩ : ∃ i, 1 < (A i).card := Ai_singleton + obtain ⟨i₀, hi₀⟩ : ∃ i, 1 < #(A i) := Ai_singleton obtain ⟨j₁, j₂, _, hj₂, _⟩ : ∃ j₁ j₂, j₁ ∈ A i₀ ∧ j₂ ∈ A i₀ ∧ j₁ ≠ j₂ := Finset.one_lt_card_iff.1 hi₀ let B := Function.update A i₀ (A i₀ \ {j₂}) @@ -535,10 +534,8 @@ theorem map_sum_finset_aux [DecidableEq ι] [Fintype ι] {n : ℕ} (h : (∑ i, · simp only [C, hi, update_noteq, Ne, not_false_iff] -- Express the inductive assumption for `B` have Brec : (f fun i => ∑ j ∈ B i, g i j) = ∑ r ∈ piFinset B, f fun i => g i (r i) := by - have : (∑ i, Finset.card (B i)) < ∑ i, Finset.card (A i) := by - refine - Finset.sum_lt_sum (fun i _ => Finset.card_le_card (B_subset_A i)) - ⟨i₀, Finset.mem_univ _, ?_⟩ + have : ∑ i, #(B i) < ∑ i, #(A i) := by + refine sum_lt_sum (fun i _ => card_le_card (B_subset_A i)) ⟨i₀, mem_univ _, ?_⟩ have : {j₂} ⊆ A i₀ := by simp [hj₂] simp only [B, Finset.card_sdiff this, Function.update_same, Finset.card_singleton] exact Nat.pred_lt (ne_of_gt (lt_trans Nat.zero_lt_one hi₀)) @@ -546,7 +543,7 @@ theorem map_sum_finset_aux [DecidableEq ι] [Fintype ι] {n : ℕ} (h : (∑ i, exact IH _ this B rfl -- Express the inductive assumption for `C` have Crec : (f fun i => ∑ j ∈ C i, g i j) = ∑ r ∈ piFinset C, f fun i => g i (r i) := by - have : (∑ i, Finset.card (C i)) < ∑ i, Finset.card (A i) := + have : (∑ i, #(C i)) < ∑ i, #(A i) := Finset.sum_lt_sum (fun i _ => Finset.card_le_card (C_subset_A i)) ⟨i₀, Finset.mem_univ _, by simp [C, hi₀]⟩ rw [h] at this @@ -1319,10 +1316,9 @@ lemma map_piecewise_sub_map_piecewise [LinearOrder ι] (a b v : (i : ι) → M open Finset in lemma map_add_eq_map_add_linearDeriv_add [DecidableEq ι] [Fintype ι] (x h : (i : ι) → M₁ i) : - f (x + h) = f x + f.linearDeriv x h + - ∑ s ∈ univ.powerset.filter (2 ≤ ·.card), f (s.piecewise h x) := by + f (x + h) = f x + f.linearDeriv x h + ∑ s with 2 ≤ #s, f (s.piecewise h x) := by rw [add_comm, map_add_univ, ← Finset.powerset_univ, - ← sum_filter_add_sum_filter_not _ (2 ≤ ·.card)] + ← sum_filter_add_sum_filter_not _ (2 ≤ #·)] simp_rw [not_le, Nat.lt_succ, le_iff_lt_or_eq (b := 1), Nat.lt_one_iff, filter_or, ← powersetCard_eq_filter, sum_union (univ.pairwise_disjoint_powersetCard zero_ne_one), powersetCard_zero, powersetCard_one, sum_singleton, Finset.piecewise_empty, sum_map, @@ -1334,7 +1330,7 @@ at two points "close to `x`" in terms of the "derivative" of the multilinear map and of "second-order" terms. -/ lemma map_add_sub_map_add_sub_linearDeriv [DecidableEq ι] [Fintype ι] (x h h' : (i : ι) → M₁ i) : f (x + h) - f (x + h') - f.linearDeriv x (h - h') = - ∑ s ∈ univ.powerset.filter (2 ≤ ·.card), (f (s.piecewise h x) - f (s.piecewise h' x)) := by + ∑ s with 2 ≤ #s, (f (s.piecewise h x) - f (s.piecewise h' x)) := by simp_rw [map_add_eq_map_add_linearDeriv_add, add_assoc, add_sub_add_comm, sub_self, zero_add, ← LinearMap.map_sub, add_sub_cancel_left, sum_sub_distrib] @@ -1679,7 +1675,7 @@ variable (R M₂ M') `l`, then the space of multilinear maps on `fun i : Fin n => M'` is isomorphic to the space of multilinear maps on `fun i : Fin k => M'` taking values in the space of multilinear maps on `fun i : Fin l => M'`. -/ -def curryFinFinset {k l n : ℕ} {s : Finset (Fin n)} (hk : s.card = k) (hl : sᶜ.card = l) : +def curryFinFinset {k l n : ℕ} {s : Finset (Fin n)} (hk : #s = k) (hl : #sᶜ = l) : MultilinearMap R (fun _ : Fin n => M') M₂ ≃ₗ[R] MultilinearMap R (fun _ : Fin k => M') (MultilinearMap R (fun _ : Fin l => M') M₂) := (domDomCongrLinearEquiv R R M' M₂ (finSumEquivOfFinset hk hl).symm).trans @@ -1688,15 +1684,15 @@ def curryFinFinset {k l n : ℕ} {s : Finset (Fin n)} (hk : s.card = k) (hl : s variable {R M₂ M'} @[simp] -theorem curryFinFinset_apply {k l n : ℕ} {s : Finset (Fin n)} (hk : s.card = k) (hl : sᶜ.card = l) +theorem curryFinFinset_apply {k l n : ℕ} {s : Finset (Fin n)} (hk : #s = k) (hl : #sᶜ = l) (f : MultilinearMap R (fun _ : Fin n => M') M₂) (mk : Fin k → M') (ml : Fin l → M') : curryFinFinset R M₂ M' hk hl f mk ml = f fun i => Sum.elim mk ml ((finSumEquivOfFinset hk hl).symm i) := rfl @[simp] -theorem curryFinFinset_symm_apply {k l n : ℕ} {s : Finset (Fin n)} (hk : s.card = k) - (hl : sᶜ.card = l) +theorem curryFinFinset_symm_apply {k l n : ℕ} {s : Finset (Fin n)} (hk : #s = k) + (hl : #sᶜ = l) (f : MultilinearMap R (fun _ : Fin k => M') (MultilinearMap R (fun _ : Fin l => M') M₂)) (m : Fin n → M') : (curryFinFinset R M₂ M' hk hl).symm f m = @@ -1705,8 +1701,8 @@ theorem curryFinFinset_symm_apply {k l n : ℕ} {s : Finset (Fin n)} (hk : s.car rfl -- @[simp] -- Porting note: simpNF linter, lhs simplifies, added aux version below -theorem curryFinFinset_symm_apply_piecewise_const {k l n : ℕ} {s : Finset (Fin n)} (hk : s.card = k) - (hl : sᶜ.card = l) +theorem curryFinFinset_symm_apply_piecewise_const {k l n : ℕ} {s : Finset (Fin n)} (hk : #s = k) + (hl : #sᶜ = l) (f : MultilinearMap R (fun _ : Fin k => M') (MultilinearMap R (fun _ : Fin l => M') M₂)) (x y : M') : (curryFinFinset R M₂ M' hk hl).symm f (s.piecewise (fun _ => x) fun _ => y) = @@ -1721,7 +1717,7 @@ theorem curryFinFinset_symm_apply_piecewise_const {k l n : ℕ} {s : Finset (Fin @[simp] theorem curryFinFinset_symm_apply_piecewise_const_aux {k l n : ℕ} {s : Finset (Fin n)} - (hk : s.card = k) (hl : sᶜ.card = l) + (hk : #s = k) (hl : #sᶜ = l) (f : MultilinearMap R (fun _ : Fin k => M') (MultilinearMap R (fun _ : Fin l => M') M₂)) (x y : M') : ((⇑f fun _ => x) (fun i => (Finset.piecewise s (fun _ => x) (fun _ => y) @@ -1732,15 +1728,15 @@ theorem curryFinFinset_symm_apply_piecewise_const_aux {k l n : ℕ} {s : Finset exact this @[simp] -theorem curryFinFinset_symm_apply_const {k l n : ℕ} {s : Finset (Fin n)} (hk : s.card = k) - (hl : sᶜ.card = l) +theorem curryFinFinset_symm_apply_const {k l n : ℕ} {s : Finset (Fin n)} (hk : #s = k) + (hl : #sᶜ = l) (f : MultilinearMap R (fun _ : Fin k => M') (MultilinearMap R (fun _ : Fin l => M') M₂)) (x : M') : ((curryFinFinset R M₂ M' hk hl).symm f fun _ => x) = f (fun _ => x) fun _ => x := rfl -- @[simp] -- Porting note: simpNF, lhs simplifies, added aux version below -theorem curryFinFinset_apply_const {k l n : ℕ} {s : Finset (Fin n)} (hk : s.card = k) - (hl : sᶜ.card = l) (f : MultilinearMap R (fun _ : Fin n => M') M₂) (x y : M') : +theorem curryFinFinset_apply_const {k l n : ℕ} {s : Finset (Fin n)} (hk : #s = k) + (hl : #sᶜ = l) (f : MultilinearMap R (fun _ : Fin n => M') M₂) (x y : M') : (curryFinFinset R M₂ M' hk hl f (fun _ => x) fun _ => y) = f (s.piecewise (fun _ => x) fun _ => y) := by refine (curryFinFinset_symm_apply_piecewise_const hk hl _ _ _).symm.trans ?_ @@ -1748,8 +1744,8 @@ theorem curryFinFinset_apply_const {k l n : ℕ} {s : Finset (Fin n)} (hk : s.ca rw [LinearEquiv.symm_apply_apply] @[simp] -theorem curryFinFinset_apply_const_aux {k l n : ℕ} {s : Finset (Fin n)} (hk : s.card = k) - (hl : sᶜ.card = l) (f : MultilinearMap R (fun _ : Fin n => M') M₂) (x y : M') : +theorem curryFinFinset_apply_const_aux {k l n : ℕ} {s : Finset (Fin n)} (hk : #s = k) + (hl : #sᶜ = l) (f : MultilinearMap R (fun _ : Fin n => M') M₂) (x y : M') : (f fun i => Sum.elim (fun _ => x) (fun _ => y) ((⇑ (Equiv.symm (finSumEquivOfFinset hk hl))) i)) = f (s.piecewise (fun _ => x) fun _ => y) := by rw [← curryFinFinset_apply] diff --git a/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean b/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean index 39bc41327d5110..99076960053fe2 100644 --- a/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean +++ b/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean @@ -329,7 +329,7 @@ theorem choose_exists_companion : Q.exists_companion.choose = polarBilin Q := protected theorem map_sum {ι} [DecidableEq ι] (Q : QuadraticMap R M N) (s : Finset ι) (f : ι → M) : Q (∑ i ∈ s, f i) = ∑ i ∈ s, Q (f i) + - ∑ ij in s.sym2.filter (¬ ·.IsDiag), + ∑ ij ∈ s.sym2 with ¬ ij.IsDiag, Sym2.lift ⟨fun i j => polar Q (f i) (f j), fun _ _ => polar_comm _ _ _⟩ ij := by induction s using Finset.cons_induction with | empty => simp diff --git a/Mathlib/Logic/Denumerable.lean b/Mathlib/Logic/Denumerable.lean index cf0e64c97bd1dc..f96dccd8f839c8 100644 --- a/Mathlib/Logic/Denumerable.lean +++ b/Mathlib/Logic/Denumerable.lean @@ -30,7 +30,7 @@ class Denumerable (α : Type*) extends Encodable α where /-- `decode` and `encode` are inverses. -/ decode_inv : ∀ n, ∃ a ∈ decode n, encode a = n -open Nat +open Finset Nat namespace Denumerable @@ -274,7 +274,7 @@ private def toFunAux (x : s) : ℕ := (List.range x).countP (· ∈ s) private theorem toFunAux_eq {s : Set ℕ} [DecidablePred (· ∈ s)] (x : s) : - toFunAux x = ((Finset.range x).filter (· ∈ s)).card := by + toFunAux x = #{y ∈ Finset.range x | y ∈ s} := by rw [toFunAux, List.countP_eq_length_filter] rfl @@ -288,9 +288,9 @@ private theorem right_inverse_aux : ∀ n, toFunAux (ofNat s n) = n exact bot_le.not_lt (show (⟨n, hn.2⟩ : s) < ⊥ from hn.1) | n + 1 => by have ih : toFunAux (ofNat s n) = n := right_inverse_aux n - have h₁ : (ofNat s n : ℕ) ∉ (range (ofNat s n)).filter (· ∈ s) := by simp - have h₂ : (range (succ (ofNat s n))).filter (· ∈ s) = - insert ↑(ofNat s n) ((range (ofNat s n)).filter (· ∈ s)) := by + have h₁ : (ofNat s n : ℕ) ∉ {x ∈ range (ofNat s n) | x ∈ s} := by simp + have h₂ : {x ∈ range (succ (ofNat s n)) | x ∈ s} = + insert ↑(ofNat s n) {x ∈ range (ofNat s n) | x ∈ s} := by simp only [Finset.ext_iff, mem_insert, mem_range, mem_filter] exact fun m => ⟨fun h => by diff --git a/Mathlib/MeasureTheory/Function/SimpleFunc.lean b/Mathlib/MeasureTheory/Function/SimpleFunc.lean index 6a1c98e568221e..0a08303ca10a20 100644 --- a/Mathlib/MeasureTheory/Function/SimpleFunc.lean +++ b/Mathlib/MeasureTheory/Function/SimpleFunc.lean @@ -269,13 +269,13 @@ theorem map_const (g : β → γ) (b : β) : (const α b).map g = const α (g b) rfl theorem map_preimage (f : α →ₛ β) (g : β → γ) (s : Set γ) : - f.map g ⁻¹' s = f ⁻¹' ↑(f.range.filter fun b => g b ∈ s) := by + f.map g ⁻¹' s = f ⁻¹' ↑{b ∈ f.range | g b ∈ s} := by simp only [coe_range, sep_mem_eq, coe_map, Finset.coe_filter, ← mem_preimage, inter_comm, preimage_inter_range, ← Finset.mem_coe] exact preimage_comp theorem map_preimage_singleton (f : α →ₛ β) (g : β → γ) (c : γ) : - f.map g ⁻¹' {c} = f ⁻¹' ↑(f.range.filter fun b => g b = c) := + f.map g ⁻¹' {c} = f ⁻¹' ↑{b ∈ f.range | g b = c} := map_preimage _ _ _ /-- Composition of a `SimpleFun` and a measurable function is a `SimpleFunc`. -/ @@ -985,7 +985,7 @@ section FinMeasSupp open Finset Function theorem support_eq [MeasurableSpace α] [Zero β] (f : α →ₛ β) : - support f = ⋃ y ∈ f.range.filter fun y => y ≠ 0, f ⁻¹' {y} := + support f = ⋃ y ∈ {y ∈ f.range | y ≠ 0}, f ⁻¹' {y} := Set.ext fun x => by 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'] diff --git a/Mathlib/MeasureTheory/Integral/Bochner.lean b/Mathlib/MeasureTheory/Integral/Bochner.lean index 3c5f56fce5d8b4..ada8eefa261f7e 100644 --- a/Mathlib/MeasureTheory/Integral/Bochner.lean +++ b/Mathlib/MeasureTheory/Integral/Bochner.lean @@ -295,12 +295,12 @@ theorem integral_eq {m : MeasurableSpace α} (μ : Measure α) (f : α →ₛ F) theorem integral_eq_sum_filter [DecidablePred fun x : F => x ≠ 0] {m : MeasurableSpace α} (f : α →ₛ F) (μ : Measure α) : - f.integral μ = ∑ x ∈ f.range.filter fun x => x ≠ 0, (μ (f ⁻¹' {x})).toReal • x := by + f.integral μ = ∑ x ∈ {x ∈ f.range | x ≠ 0}, (μ (f ⁻¹' {x})).toReal • x := by simp_rw [integral_def, setToSimpleFunc_eq_sum_filter, weightedSMul_apply] /-- The Bochner integral is equal to a sum over any set that includes `f.range` (except `0`). -/ theorem integral_eq_sum_of_subset [DecidablePred fun x : F => x ≠ 0] {f : α →ₛ F} {s : Finset F} - (hs : (f.range.filter fun x => x ≠ 0) ⊆ s) : + (hs : {x ∈ f.range | x ≠ 0} ⊆ s) : f.integral μ = ∑ x ∈ s, (μ (f ⁻¹' {x})).toReal • x := by rw [SimpleFunc.integral_eq_sum_filter, Finset.sum_subset hs] rintro x - hx; rw [Finset.mem_filter, not_and_or, Ne, Classical.not_not] at hx diff --git a/Mathlib/MeasureTheory/Integral/SetToL1.lean b/Mathlib/MeasureTheory/Integral/SetToL1.lean index 3d16c95fed2c12..456069af75266f 100644 --- a/Mathlib/MeasureTheory/Integral/SetToL1.lean +++ b/Mathlib/MeasureTheory/Integral/SetToL1.lean @@ -277,9 +277,8 @@ theorem setToSimpleFunc_zero_apply {m : MeasurableSpace α} (T : Set α → F cases isEmpty_or_nonempty α <;> simp [setToSimpleFunc] theorem setToSimpleFunc_eq_sum_filter [DecidablePred fun x ↦ x ≠ (0 : F)] - {m : MeasurableSpace α} (T : Set α → F →L[ℝ] F') - (f : α →ₛ F) : - setToSimpleFunc T f = ∑ x ∈ f.range.filter fun x => x ≠ 0, (T (f ⁻¹' {x})) x := by + {m : MeasurableSpace α} (T : Set α → F →L[ℝ] F') (f : α →ₛ F) : + setToSimpleFunc T f = ∑ x ∈ f.range with x ≠ 0, T (f ⁻¹' {x}) x := by symm refine sum_filter_of_ne fun x _ => mt fun hx0 => ?_ rw [hx0] @@ -301,13 +300,13 @@ theorem map_setToSimpleFunc (T : Set α → F →L[ℝ] F') (h_add : FinMeasAddi rw [mem_filter] at hx rw [hx.2, ContinuousLinearMap.map_zero] have h_left_eq : - T (map g f ⁻¹' {g (f a)}) (g (f a)) = - T (f ⁻¹' (f.range.filter fun b => g b = g (f a))) (g (f a)) := by + T (map g f ⁻¹' {g (f a)}) (g (f a)) + = T (f ⁻¹' ({b ∈ f.range | g b = g (f a)} : Finset _)) (g (f a)) := by congr; rw [map_preimage_singleton] rw [h_left_eq] have h_left_eq' : - T (f ⁻¹' (filter (fun b : G => g b = g (f a)) f.range)) (g (f a)) = - T (⋃ y ∈ filter (fun b : G => g b = g (f a)) f.range, f ⁻¹' {y}) (g (f a)) := by + T (f ⁻¹' ({b ∈ f.range | g b = g (f a)} : Finset _)) (g (f a)) + = T (⋃ y ∈ {b ∈ f.range | g b = g (f a)}, f ⁻¹' {y}) (g (f a)) := by congr; rw [← Finset.set_biUnion_preimage_singleton] rw [h_left_eq'] rw [h_add.map_iUnion_fin_meas_set_eq_sum T T_empty] @@ -381,7 +380,7 @@ theorem setToSimpleFunc_add_left' (T T' T'' : Set α → E →L[ℝ] F) classical simp_rw [setToSimpleFunc_eq_sum_filter] suffices - ∀ x ∈ filter (fun x : E => x ≠ 0) f.range, T'' (f ⁻¹' {x}) = T (f ⁻¹' {x}) + T' (f ⁻¹' {x}) by + ∀ x ∈ {x ∈ f.range | x ≠ 0}, T'' (f ⁻¹' {x}) = T (f ⁻¹' {x}) + T' (f ⁻¹' {x}) by rw [← sum_add_distrib] refine Finset.sum_congr rfl fun x hx => ?_ rw [this x hx] @@ -402,7 +401,7 @@ theorem setToSimpleFunc_smul_left' (T T' : Set α → E →L[ℝ] F') (c : ℝ) setToSimpleFunc T' f = c • setToSimpleFunc T f := by classical simp_rw [setToSimpleFunc_eq_sum_filter] - suffices ∀ x ∈ filter (fun x : E => x ≠ 0) f.range, T' (f ⁻¹' {x}) = c • T (f ⁻¹' {x}) by + suffices ∀ x ∈ {x ∈ f.range | x ≠ 0}, T' (f ⁻¹' {x}) = c • T (f ⁻¹' {x}) by rw [smul_sum] refine Finset.sum_congr rfl fun x hx => ?_ rw [this x hx] diff --git a/Mathlib/NumberTheory/ArithmeticFunction.lean b/Mathlib/NumberTheory/ArithmeticFunction.lean index 71c09ddd1baad5..f954d1b9acb240 100644 --- a/Mathlib/NumberTheory/ArithmeticFunction.lean +++ b/Mathlib/NumberTheory/ArithmeticFunction.lean @@ -822,7 +822,7 @@ theorem sigma_one_apply_prime_pow {p i : ℕ} (hp : p.Prime) : σ 1 (p ^ i) = ∑ k in .range (i + 1), p ^ k := by simp [sigma_apply_prime_pow hp] -theorem sigma_zero_apply (n : ℕ) : σ 0 n = (divisors n).card := by simp [sigma_apply] +theorem sigma_zero_apply (n : ℕ) : σ 0 n = #n.divisors := by simp [sigma_apply] theorem sigma_zero_apply_prime_pow {p i : ℕ} (hp : p.Prime) : σ 0 (p ^ i) = i + 1 := by simp [sigma_apply_prime_pow hp] @@ -1286,13 +1286,13 @@ theorem prod_eq_iff_prod_pow_moebius_eq_on_of_nonzero [CommGroupWithZero R] end SpecialFunctions theorem _root_.Nat.card_divisors {n : ℕ} (hn : n ≠ 0) : - n.divisors.card = n.primeFactors.prod (n.factorization · + 1) := by + #n.divisors = n.primeFactors.prod (n.factorization · + 1) := by rw [← sigma_zero_apply, isMultiplicative_sigma.multiplicative_factorization _ hn] exact Finset.prod_congr n.support_factorization fun _ h => sigma_zero_apply_prime_pow <| Nat.prime_of_mem_primeFactors h @[deprecated (since := "2024-06-09")] theorem card_divisors (n : ℕ) (hn : n ≠ 0) : - n.divisors.card = n.primeFactors.prod (n.factorization · + 1) := Nat.card_divisors hn + #n.divisors = n.primeFactors.prod (n.factorization · + 1) := Nat.card_divisors hn theorem _root_.Nat.sum_divisors {n : ℕ} (hn : n ≠ 0) : ∑ d ∈ n.divisors, d = ∏ p ∈ n.primeFactors, ∑ k ∈ .range (n.factorization p + 1), p ^ k := by @@ -1307,7 +1307,7 @@ namespace Nat.Coprime open ArithmeticFunction theorem card_divisors_mul {m n : ℕ} (hmn : m.Coprime n) : - (m * n).divisors.card = m.divisors.card * n.divisors.card := by + #(m * n).divisors = #m.divisors * #n.divisors := by simp only [← sigma_zero_apply, isMultiplicative_sigma.map_mul_of_coprime hmn] theorem sum_divisors_mul {m n : ℕ} (hmn : m.Coprime n) : diff --git a/Mathlib/NumberTheory/DiophantineApproximation.lean b/Mathlib/NumberTheory/DiophantineApproximation.lean index 1d052c7a9f81e8..604980a89620e2 100644 --- a/Mathlib/NumberTheory/DiophantineApproximation.lean +++ b/Mathlib/NumberTheory/DiophantineApproximation.lean @@ -113,7 +113,7 @@ theorem exists_int_int_abs_mul_sub_le (ξ : ℝ) {n : ℕ} (n_pos : 0 < n) : simpa only [neg_add_cancel_comm_assoc] using hf' · -- Porting note(https://github.com/leanprover-community/mathlib4/issues/5127): added `not_and` simp_rw [not_exists, not_and] at H - have hD : (Ico (0 : ℤ) n).card < D.card := by rw [card_Icc, card_Ico]; exact lt_add_one n + have hD : #(Ico (0 : ℤ) n) < #D := by rw [card_Icc, card_Ico]; exact lt_add_one n have hfu' : ∀ m, f m ≤ n := fun m => lt_add_one_iff.mp (floor_lt.mpr (mod_cast hfu m)) have hwd : ∀ m : ℤ, m ∈ D → f m ∈ Ico (0 : ℤ) n := fun x hx => mem_Ico.mpr diff --git a/Mathlib/NumberTheory/Divisors.lean b/Mathlib/NumberTheory/Divisors.lean index d0f8d0cc61130e..e2bfeafb8e2171 100644 --- a/Mathlib/NumberTheory/Divisors.lean +++ b/Mathlib/NumberTheory/Divisors.lean @@ -36,30 +36,27 @@ namespace Nat variable (n : ℕ) /-- `divisors n` is the `Finset` of divisors of `n`. As a special case, `divisors 0 = ∅`. -/ -def divisors : Finset ℕ := - Finset.filter (fun x : ℕ => x ∣ n) (Finset.Ico 1 (n + 1)) +def divisors : Finset ℕ := {d ∈ Ico 1 (n + 1) | d ∣ n} /-- `properDivisors n` is the `Finset` of divisors of `n`, other than `n`. As a special case, `properDivisors 0 = ∅`. -/ -def properDivisors : Finset ℕ := - Finset.filter (fun x : ℕ => x ∣ n) (Finset.Ico 1 n) +def properDivisors : Finset ℕ := {d ∈ Ico 1 n | d ∣ n} /-- `divisorsAntidiagonal n` is the `Finset` of pairs `(x,y)` such that `x * y = n`. As a special case, `divisorsAntidiagonal 0 = ∅`. -/ def divisorsAntidiagonal : Finset (ℕ × ℕ) := - Finset.filter (fun x => x.fst * x.snd = n) (Ico 1 (n + 1) ×ˢ Ico 1 (n + 1)) + {x ∈ Ico 1 (n + 1) ×ˢ Ico 1 (n + 1) | x.fst * x.snd = n} variable {n} @[simp] -theorem filter_dvd_eq_divisors (h : n ≠ 0) : (Finset.range n.succ).filter (· ∣ n) = n.divisors := by +theorem filter_dvd_eq_divisors (h : n ≠ 0) : {d ∈ range n.succ | d ∣ n} = n.divisors := by ext simp only [divisors, mem_filter, mem_range, mem_Ico, and_congr_left_iff, iff_and_self] exact fun ha _ => succ_le_iff.mpr (pos_of_dvd_of_pos ha h.bot_lt) @[simp] -theorem filter_dvd_eq_properDivisors (h : n ≠ 0) : - (Finset.range n).filter (· ∣ n) = n.properDivisors := by +theorem filter_dvd_eq_properDivisors (h : n ≠ 0) : {d ∈ range n | d ∣ n} = n.properDivisors := by ext simp only [properDivisors, mem_filter, mem_range, mem_Ico, and_congr_left_iff, iff_and_self] exact fun ha _ => succ_le_iff.mpr (pos_of_dvd_of_pos ha h.bot_lt) @@ -147,7 +144,7 @@ theorem divisors_subset_properDivisors {m : ℕ} (hzero : n ≠ 0) (h : m ∣ n) (lt_of_le_of_ne (divisor_le (Nat.mem_divisors.2 ⟨h, hzero⟩)) hdiff)⟩ lemma divisors_filter_dvd_of_dvd {n m : ℕ} (hn : n ≠ 0) (hm : m ∣ n) : - (n.divisors.filter (· ∣ m)) = m.divisors := by + {d ∈ n.divisors | d ∣ m} = m.divisors := by ext k simp_rw [mem_filter, mem_divisors] exact ⟨fun ⟨_, hkm⟩ ↦ ⟨hkm, ne_zero_of_dvd_ne_zero hn hm⟩, fun ⟨hk, _⟩ ↦ ⟨⟨hk.trans hm, hn⟩, hk⟩⟩ @@ -471,7 +468,7 @@ theorem prod_divisorsAntidiagonal' {M : Type*} [CommMonoid M] (f : ℕ → ℕ /-- The factors of `n` are the prime divisors -/ theorem primeFactors_eq_to_filter_divisors_prime (n : ℕ) : - n.primeFactors = (divisors n).filter Prime := by + n.primeFactors = {p ∈ divisors n | p.Prime} := by rcases n.eq_zero_or_pos with (rfl | hn) · simp · ext q @@ -481,7 +478,7 @@ theorem primeFactors_eq_to_filter_divisors_prime (n : ℕ) : alias prime_divisors_eq_to_filter_divisors_prime := primeFactors_eq_to_filter_divisors_prime lemma primeFactors_filter_dvd_of_dvd {m n : ℕ} (hn : n ≠ 0) (hmn : m ∣ n) : - n.primeFactors.filter (· ∣ m) = m.primeFactors := by + {p ∈ n.primeFactors | p ∣ m} = m.primeFactors := by simp_rw [primeFactors_eq_to_filter_divisors_prime, filter_comm, divisors_filter_dvd_of_dvd hn hmn] diff --git a/Mathlib/NumberTheory/EulerProduct/Basic.lean b/Mathlib/NumberTheory/EulerProduct/Basic.lean index 1b6b6f551ed155..be623e1cac9575 100644 --- a/Mathlib/NumberTheory/EulerProduct/Basic.lean +++ b/Mathlib/NumberTheory/EulerProduct/Basic.lean @@ -81,7 +81,7 @@ lemma summable_and_hasSum_factoredNumbers_prod_filter_prime_tsum (hsum : ∀ {p : ℕ}, p.Prime → Summable (fun n : ℕ ↦ ‖f (p ^ n)‖)) (s : Finset ℕ) : Summable (fun m : factoredNumbers s ↦ ‖f m‖) ∧ HasSum (fun m : factoredNumbers s ↦ f m) - (∏ p ∈ s.filter Nat.Prime, ∑' n : ℕ, f (p ^ n)) := by + (∏ p ∈ s with p.Prime, ∑' n : ℕ, f (p ^ n)) := by induction' s using Finset.induction with p s hp ih · rw [factoredNumbers_empty] simp only [not_mem_empty, IsEmpty.forall_iff, forall_const, filter_true_of_mem, prod_empty] @@ -93,7 +93,7 @@ lemma summable_and_hasSum_factoredNumbers_prod_filter_prime_tsum equivProdNatFactoredNumbers_apply', factoredNumbers.map_prime_pow_mul hmul hpp hp] refine Summable.of_nonneg_of_le (fun _ ↦ norm_nonneg _) (fun _ ↦ norm_mul_le ..) ?_ apply Summable.mul_of_nonneg (hsum hpp) ih.1 <;> exact fun n ↦ norm_nonneg _ - · have hp' : p ∉ s.filter Nat.Prime := mt (mem_of_mem_filter p) hp + · have hp' : p ∉ {p ∈ s | p.Prime} := mt (mem_of_mem_filter p) hp rw [prod_insert hp', ← (equivProdNatFactoredNumbers hpp hp).hasSum_iff, Function.comp_def] conv => enter [1, x] @@ -108,7 +108,7 @@ include hf₁ hmul in /-- A version of `EulerProduct.summable_and_hasSum_factoredNumbers_prod_filter_prime_tsum` in terms of the value of the series. -/ lemma prod_filter_prime_tsum_eq_tsum_factoredNumbers (hsum : Summable (‖f ·‖)) (s : Finset ℕ) : - ∏ p ∈ s.filter Nat.Prime, ∑' n : ℕ, f (p ^ n) = ∑' m : factoredNumbers s, f m := + ∏ p ∈ s with p.Prime, ∑' n : ℕ, f (p ^ n) = ∑' m : factoredNumbers s, f m := (summable_and_hasSum_factoredNumbers_prod_filter_prime_tsum hf₁ hmul (fun hp ↦ hsum.comp_injective <| Nat.pow_right_injective hp.one_lt) _).2.tsum_eq.symm @@ -175,7 +175,7 @@ theorem eulerProduct_hasProd (hsum : Summable (‖f ·‖)) (hf₀ : f 0 = 0) : intro ε hε obtain ⟨N₀, hN₀⟩ := norm_tsum_factoredNumbers_sub_tsum_lt hsum.of_norm hf₀ hε refine ⟨range N₀, fun s hs ↦ ?_⟩ - have : ∏ p ∈ s, {p | Nat.Prime p}.mulIndicator F p = ∏ p ∈ s.filter Nat.Prime, F p := + have : ∏ p ∈ s, {p | Nat.Prime p}.mulIndicator F p = ∏ p ∈ s with p.Prime, F p := prod_mulIndicator_eq_prod_filter s (fun _ ↦ F) _ id rw [this, dist_eq_norm, prod_filter_prime_tsum_eq_tsum_factoredNumbers hf₁ hmul hsum, norm_sub_rev] @@ -292,10 +292,10 @@ we show that the sum involved converges absolutely. -/ lemma summable_and_hasSum_factoredNumbers_prod_filter_prime_geometric {f : ℕ →* F} (h : ∀ {p : ℕ}, p.Prime → ‖f p‖ < 1) (s : Finset ℕ) : Summable (fun m : factoredNumbers s ↦ ‖f m‖) ∧ - HasSum (fun m : factoredNumbers s ↦ f m) (∏ p ∈ s.filter Nat.Prime, (1 - f p)⁻¹) := by + HasSum (fun m : factoredNumbers s ↦ f m) (∏ p ∈ s with p.Prime, (1 - f p)⁻¹) := by have hmul {m n} (_ : Nat.Coprime m n) := f.map_mul m n have H₁ : - ∏ p ∈ s.filter Nat.Prime, ∑' n : ℕ, f (p ^ n) = ∏ p ∈ s.filter Nat.Prime, (1 - f p)⁻¹ := by + ∏ p ∈ s with p.Prime, ∑' n : ℕ, f (p ^ n) = ∏ p ∈ s with p.Prime, (1 - f p)⁻¹ := by refine prod_congr rfl fun p hp ↦ ?_ simp only [map_pow] exact tsum_geometric_of_norm_lt_one <| h (mem_filter.mp hp).2 @@ -310,7 +310,7 @@ lemma summable_and_hasSum_factoredNumbers_prod_filter_prime_geometric {f : ℕ in terms of the value of the series. -/ lemma prod_filter_prime_geometric_eq_tsum_factoredNumbers {f : ℕ →* F} (hsum : Summable f) (s : Finset ℕ) : - ∏ p ∈ s.filter Nat.Prime, (1 - f p)⁻¹ = ∑' m : factoredNumbers s, f m := by + ∏ p ∈ s with p.Prime, (1 - f p)⁻¹ = ∑' m : factoredNumbers s, f m := by refine (summable_and_hasSum_factoredNumbers_prod_filter_prime_geometric ?_ s).2.tsum_eq.symm exact fun {_} hp ↦ hsum.norm_lt_one hp.one_lt diff --git a/Mathlib/NumberTheory/JacobiSum/Basic.lean b/Mathlib/NumberTheory/JacobiSum/Basic.lean index 7e6fd3e597e49f..2de42f2e7d95bb 100644 --- a/Mathlib/NumberTheory/JacobiSum/Basic.lean +++ b/Mathlib/NumberTheory/JacobiSum/Basic.lean @@ -107,7 +107,7 @@ theorem jacobiSum_trivial_trivial : simpa only [isUnit_iff_ne_zero, mul_ne_zero_iff, ne_eq, sub_eq_zero, @eq_comm _ _ x] using hx calc ∑ x ∈ univ \ {0, 1}, (MulChar.trivial F R) x * (MulChar.trivial F R) (1 - x) _ = ∑ _ ∈ univ \ {0, 1}, 1 := sum_congr rfl this - _ = Finset.card (univ \ {0, 1}) := (cast_card _).symm + _ = #(univ \ {0, 1}) := (cast_card _).symm _ = Fintype.card F - 2 := by rw [card_sdiff (subset_univ _), card_univ, card_pair zero_ne_one, Nat.cast_sub <| Nat.add_one_le_of_lt Fintype.one_lt_card, Nat.cast_two] diff --git a/Mathlib/NumberTheory/LegendreSymbol/GaussEisensteinLemmas.lean b/Mathlib/NumberTheory/LegendreSymbol/GaussEisensteinLemmas.lean index edf39ca7e1be82..6927e41b5eeeff 100644 --- a/Mathlib/NumberTheory/LegendreSymbol/GaussEisensteinLemmas.lean +++ b/Mathlib/NumberTheory/LegendreSymbol/GaussEisensteinLemmas.lean @@ -57,10 +57,9 @@ theorem Ico_map_valMinAbs_natAbs_eq_Ico_map_id (p : ℕ) [hp : Fact p.Prime] (a (fun x _ => (a * x : ZMod p).valMinAbs.natAbs) hmem (inj_on_of_surj_on_of_card_le _ hmem hsurj le_rfl) hsurj (fun _ _ => rfl) -private theorem gauss_lemma_aux₁ (p : ℕ) [Fact p.Prime] {a : ℤ} - (hap : (a : ZMod p) ≠ 0) : (a ^ (p / 2) * (p / 2)! : ZMod p) = - (-1 : ZMod p) ^ ((Ico 1 (p / 2).succ).filter fun x : ℕ => - ¬(a * x : ZMod p).val ≤ p / 2).card * (p / 2)! := +private theorem gauss_lemma_aux₁ (p : ℕ) [Fact p.Prime] {a : ℤ} (hap : (a : ZMod p) ≠ 0) : + (a ^ (p / 2) * (p / 2)! : ZMod p) = + (-1 : ZMod p) ^ #{x ∈ Ico 1 (p / 2).succ | ¬ (a * x.cast : ZMod p).val ≤ p / 2} * (p / 2)! := calc (a ^ (p / 2) * (p / 2)! : ZMod p) = ∏ x ∈ Ico 1 (p / 2).succ, a * x := by rw [prod_mul_distrib, ← prod_natCast, prod_Ico_id_eq_factorial, prod_const, card_Ico, @@ -71,26 +70,25 @@ private theorem gauss_lemma_aux₁ (p : ℕ) [Fact p.Prime] {a : ℤ} (prod_congr rfl fun _ _ => by simp only [natCast_natAbs_valMinAbs] split_ifs <;> simp) - _ = (-1 : ZMod p) ^ ((Ico 1 (p / 2).succ).filter fun x : ℕ => - ¬(a * x : ZMod p).val ≤ p / 2).card * ∏ x ∈ Ico 1 (p / 2).succ, - ↑((a * x : ZMod p).valMinAbs.natAbs) := by + _ = (-1 : ZMod p) ^ #{x ∈ Ico 1 (p / 2).succ | ¬(a * x.cast : ZMod p).val ≤ p / 2} * + ∏ x ∈ Ico 1 (p / 2).succ, ↑((a * x : ZMod p).valMinAbs.natAbs) := by have : (∏ x ∈ Ico 1 (p / 2).succ, if (a * x : ZMod p).val ≤ p / 2 then (1 : ZMod p) else -1) = - ∏ x ∈ (Ico 1 (p / 2).succ).filter fun x : ℕ => ¬(a * x : ZMod p).val ≤ p / 2, -1 := + ∏ x ∈ Ico 1 (p / 2).succ with ¬(a * x.cast : ZMod p).val ≤ p / 2, -1 := prod_bij_ne_one (fun x _ _ => x) (fun x => by split_ifs <;> (dsimp; simp_all)) (fun _ _ _ _ _ _ => id) (fun b h _ => ⟨b, by simp_all [-not_le]⟩) (by intros; split_ifs at * <;> simp_all) rw [prod_mul_distrib, this, prod_const] - _ = (-1 : ZMod p) ^ ((Ico 1 (p / 2).succ).filter fun x : ℕ => - ¬(a * x : ZMod p).val ≤ p / 2).card * (p / 2)! := by + _ = (-1 : ZMod p) ^ #{x ∈ Ico 1 (p / 2).succ | ¬(a * x.cast : ZMod p).val ≤ p / 2} * + (p / 2)! := by rw [← prod_natCast, Finset.prod_eq_multiset_prod, Ico_map_valMinAbs_natAbs_eq_Ico_map_id p a hap, ← Finset.prod_eq_multiset_prod, prod_Ico_id_eq_factorial] -theorem gauss_lemma_aux (p : ℕ) [hp : Fact p.Prime] {a : ℤ} - (hap : (a : ZMod p) ≠ 0) : (↑a ^ (p / 2) : ZMod p) = - ((-1) ^ ((Ico 1 (p / 2).succ).filter fun x : ℕ => p / 2 < (a * x : ZMod p).val).card :) := +theorem gauss_lemma_aux (p : ℕ) [hp : Fact p.Prime] {a : ℤ} (hap : (a : ZMod p) ≠ 0) : + (a ^ (p / 2) : ZMod p) = + ((-1) ^ #{x ∈ Ico 1 (p / 2).succ | p / 2 < (a * x.cast : ZMod p).val} :) := (mul_left_inj' (show ((p / 2)! : ZMod p) ≠ 0 by rw [Ne, CharP.cast_eq_zero_iff (ZMod p) p, hp.1.dvd_factorial, not_le] exact Nat.div_lt_self hp.1.pos (by decide))).1 <| by @@ -99,21 +97,20 @@ theorem gauss_lemma_aux (p : ℕ) [hp : Fact p.Prime] {a : ℤ} /-- **Gauss' lemma**. The Legendre symbol can be computed by considering the number of naturals less than `p/2` such that `(a * x) % p > p / 2`. -/ theorem gauss_lemma {p : ℕ} [h : Fact p.Prime] {a : ℤ} (hp : p ≠ 2) (ha0 : (a : ZMod p) ≠ 0) : - legendreSym p a = (-1) ^ ((Ico 1 (p / 2).succ).filter fun x : ℕ => - p / 2 < (a * x : ZMod p).val).card := by + legendreSym p a = (-1) ^ #{x ∈ Ico 1 (p / 2).succ | p / 2 < (a * x.cast : ZMod p).val} := by replace hp : Odd p := h.out.odd_of_ne_two hp - have : (legendreSym p a : ZMod p) = (((-1) ^ ((Ico 1 (p / 2).succ).filter fun x : ℕ => - p / 2 < (a * x : ZMod p).val).card : ℤ) : ZMod p) := by + have : (legendreSym p a : ZMod p) = + (((-1) ^ #{x ∈ Ico 1 (p / 2).succ | p / 2 < (a * x.cast : ZMod p).val} : ℤ) : ZMod p) := by rw [legendreSym.eq_pow, gauss_lemma_aux p ha0] cases legendreSym.eq_one_or_neg_one p ha0 <;> - cases neg_one_pow_eq_or ℤ - ((Ico 1 (p / 2).succ).filter fun x : ℕ => p / 2 < (a * x : ZMod p).val).card <;> + cases neg_one_pow_eq_or ℤ #{x ∈ Ico 1 (p / 2).succ | p / 2 < (a * x.cast : ZMod p).val} <;> simp_all [ne_neg_self hp one_ne_zero, (ne_neg_self hp one_ne_zero).symm] private theorem eisenstein_lemma_aux₁ (p : ℕ) [Fact p.Prime] [hp2 : Fact (p % 2 = 1)] {a : ℕ} - (hap : (a : ZMod p) ≠ 0) : ((∑ x ∈ Ico 1 (p / 2).succ, a * x : ℕ) : ZMod 2) = - ((Ico 1 (p / 2).succ).filter fun x : ℕ => p / 2 < (a * x : ZMod p).val).card + - ∑ x ∈ Ico 1 (p / 2).succ, x + (∑ x ∈ Ico 1 (p / 2).succ, a * x / p : ℕ) := + (hap : (a : ZMod p) ≠ 0) : + ((∑ x ∈ Ico 1 (p / 2).succ, a * x : ℕ) : ZMod 2) = + #{x ∈ Ico 1 (p / 2).succ | p / 2 < (a * x.cast : ZMod p).val} + + ∑ x ∈ Ico 1 (p / 2).succ, x + (∑ x ∈ Ico 1 (p / 2).succ, a * x / p : ℕ) := have hp2 : (p : ZMod 2) = (1 : ℕ) := (eq_iff_modEq_nat _).2 hp2.1 calc ((∑ x ∈ Ico 1 (p / 2).succ, a * x : ℕ) : ZMod 2) = @@ -130,8 +127,8 @@ private theorem eisenstein_lemma_aux₁ (p : ℕ) [Fact p.Prime] [hp2 : Fact (p ∑ x ∈ Ico 1 (p / 2).succ, (((a * x : ZMod p).valMinAbs + if (a * x : ZMod p).val ≤ p / 2 then 0 else p : ℤ) : ZMod 2) := by simp only [(val_eq_ite_valMinAbs _).symm]; simp [Nat.cast_sum] - _ = ((Ico 1 (p / 2).succ).filter fun x : ℕ => p / 2 < (a * x : ZMod p).val).card + - (∑ x ∈ Ico 1 (p / 2).succ, (a * x : ZMod p).valMinAbs.natAbs : ℕ) := by + _ = #{x ∈ Ico 1 (p / 2).succ | p / 2 < (a * x.cast : ZMod p).val} + + (∑ x ∈ Ico 1 (p / 2).succ, (a * x.cast : ZMod p).valMinAbs.natAbs : ℕ) := by simp [add_comm, sum_add_distrib, Finset.sum_ite, hp2, Nat.cast_sum] _ = _ := by rw [Finset.sum_eq_multiset_sum, Ico_map_valMinAbs_natAbs_eq_Ico_map_id p a hap, ← @@ -140,7 +137,7 @@ private theorem eisenstein_lemma_aux₁ (p : ℕ) [Fact p.Prime] [hp2 : Fact (p theorem eisenstein_lemma_aux (p : ℕ) [Fact p.Prime] [Fact (p % 2 = 1)] {a : ℕ} (ha2 : a % 2 = 1) (hap : (a : ZMod p) ≠ 0) : - ((Ico 1 (p / 2).succ).filter fun x : ℕ => p / 2 < (a * x : ZMod p).val).card ≡ + #{x ∈ Ico 1 (p / 2).succ | p / 2 < (a * x.cast : ZMod p).val} ≡ ∑ x ∈ Ico 1 (p / 2).succ, x * a / p [MOD 2] := have ha2 : (a : ZMod 2) = (1 : ℕ) := (eq_iff_modEq_nat _).2 ha2 (eq_iff_modEq_nat 2).1 <| sub_eq_zero.1 <| by @@ -149,10 +146,10 @@ theorem eisenstein_lemma_aux (p : ℕ) [Fact p.Prime] [Fact (p % 2 = 1)] {a : Eq.symm (eisenstein_lemma_aux₁ p hap) theorem div_eq_filter_card {a b c : ℕ} (hb0 : 0 < b) (hc : a / b ≤ c) : - a / b = ((Ico 1 c.succ).filter fun x => x * b ≤ a).card := + a / b = #{x ∈ Ico 1 c.succ | x * b ≤ a} := calc - a / b = (Ico 1 (a / b).succ).card := by simp - _ = ((Ico 1 c.succ).filter fun x => x * b ≤ a).card := + a / b = #(Ico 1 (a / b).succ) := by simp + _ = #{x ∈ Ico 1 c.succ | x * b ≤ a} := congr_arg _ <| Finset.ext fun x => by have : x * b ≤ a → x ≤ c := fun h => le_trans (by rwa [le_div_iff_mul_le hb0]) hc simp [Nat.lt_succ_iff, le_div_iff_mul_le hb0]; tauto @@ -161,13 +158,12 @@ theorem div_eq_filter_card {a b c : ℕ} (hb0 : 0 < b) (hc : a / b ≤ c) : rectangle `(0, p/2) × (0, q/2)`. -/ private theorem sum_Ico_eq_card_lt {p q : ℕ} : ∑ a ∈ Ico 1 (p / 2).succ, a * q / p = - ((Ico 1 (p / 2).succ ×ˢ Ico 1 (q / 2).succ).filter fun x : ℕ × ℕ => - x.2 * p ≤ x.1 * q).card := + #{x ∈ Ico 1 (p / 2).succ ×ˢ Ico 1 (q / 2).succ | x.2 * p ≤ x.1 * q} := if hp0 : p = 0 then by simp [hp0, Finset.ext_iff] else calc ∑ a ∈ Ico 1 (p / 2).succ, a * q / p = - ∑ a ∈ Ico 1 (p / 2).succ, ((Ico 1 (q / 2).succ).filter fun x => x * p ≤ a * q).card := + ∑ a ∈ Ico 1 (p / 2).succ, #{x ∈ Ico 1 (q / 2).succ | x * p ≤ a * q} := Finset.sum_congr rfl fun x hx => div_eq_filter_card (Nat.pos_of_ne_zero hp0) <| calc x * q / p ≤ p / 2 * q / p := by have := le_of_lt_succ (mem_Ico.mp hx).2; gcongr @@ -187,16 +183,15 @@ theorem sum_mul_div_add_sum_mul_div_eq_mul (p q : ℕ) [hp : Fact p.Prime] (hq0 ∑ a ∈ Ico 1 (p / 2).succ, a * q / p + ∑ a ∈ Ico 1 (q / 2).succ, a * p / q = p / 2 * (q / 2) := by have hswap : - ((Ico 1 (q / 2).succ ×ˢ Ico 1 (p / 2).succ).filter fun x : ℕ × ℕ => x.2 * q ≤ x.1 * p).card = - ((Ico 1 (p / 2).succ ×ˢ Ico 1 (q / 2).succ).filter fun x : ℕ × ℕ => - x.1 * q ≤ x.2 * p).card := + #{x ∈ Ico 1 (q / 2).succ ×ˢ Ico 1 (p / 2).succ | x.2 * q ≤ x.1 * p} = + #{x ∈ Ico 1 (p / 2).succ ×ˢ Ico 1 (q / 2).succ | x.1 * q ≤ x.2 * p} := card_equiv (Equiv.prodComm _ _) (fun ⟨_, _⟩ => by simp (config := { contextual := true }) only [mem_filter, and_self_iff, Prod.swap_prod_mk, forall_true_iff, mem_product, Equiv.prodComm_apply, and_assoc, and_left_comm]) have hdisj : - Disjoint ((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) := by + Disjoint {x ∈ Ico 1 (p / 2).succ ×ˢ Ico 1 (q / 2).succ | x.2 * p ≤ x.1 * q} + {x ∈ Ico 1 (p / 2).succ ×ˢ Ico 1 (q / 2).succ | x.1 * q ≤ x.2 * p} := by apply disjoint_filter.2 fun x hx hpq hqp => ?_ have hxp : x.1 < p := lt_of_le_of_lt (show x.1 ≤ p / 2 by simp_all only [Nat.lt_succ_iff, mem_Ico, mem_product]) @@ -207,8 +202,8 @@ theorem sum_mul_div_add_sum_mul_div_eq_mul (p q : ℕ) [hp : Fact p.Prime] (hq0 rw [val_cast_of_lt hxp, val_zero] at this 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) = + {x ∈ Ico 1 (p / 2).succ ×ˢ Ico 1 (q / 2).succ | x.2 * p ≤ x.1 * q} ∪ + {x ∈ Ico 1 (p / 2).succ ×ˢ Ico 1 (q / 2).succ | x.1 * q ≤ x.2 * p} = Ico 1 (p / 2).succ ×ˢ Ico 1 (q / 2).succ := Finset.ext fun x => by have := le_total (x.2 * p) (x.1 * q) diff --git a/Mathlib/NumberTheory/LegendreSymbol/QuadraticChar/Basic.lean b/Mathlib/NumberTheory/LegendreSymbol/QuadraticChar/Basic.lean index 9ba104ca9690a6..a46bc6602b64dc 100644 --- a/Mathlib/NumberTheory/LegendreSymbol/QuadraticChar/Basic.lean +++ b/Mathlib/NumberTheory/LegendreSymbol/QuadraticChar/Basic.lean @@ -223,7 +223,7 @@ theorem quadraticChar_isNontrivial (hF : ringChar F ≠ 2) : (quadraticChar F).I open Finset in /-- The number of solutions to `x^2 = a` is determined by the quadratic character. -/ theorem quadraticChar_card_sqrts (hF : ringChar F ≠ 2) (a : F) : - ↑{x : F | x ^ 2 = a}.toFinset.card = quadraticChar F a + 1 := by + #{x : F | x ^ 2 = a}.toFinset = quadraticChar F a + 1 := by -- we consider the cases `a = 0`, `a` is a nonzero square and `a` is a nonsquare in turn by_cases h₀ : a = 0 · simp only [h₀, sq_eq_zero_iff, Set.setOf_eq_eq_singleton, Set.toFinset_card, diff --git a/Mathlib/NumberTheory/MaricaSchoenheim.lean b/Mathlib/NumberTheory/MaricaSchoenheim.lean index 4ce9d2d8bf9213..a88458695d67da 100644 --- a/Mathlib/NumberTheory/MaricaSchoenheim.lean +++ b/Mathlib/NumberTheory/MaricaSchoenheim.lean @@ -40,9 +40,9 @@ lemma grahamConjecture_of_squarefree {n : ℕ} (f : ℕ → ℕ) (hf' : ∀ k < fun i hi j ↦ (hf' _ hi).squarefree_of_dvd <| div_dvd_of_dvd <| gcd_dvd_left _ _ refine lt_irrefl n ?_ calc - n = 𝒜.card := ?_ - _ ≤ (𝒜 \\ 𝒜).card := 𝒜.card_le_card_diffs - _ ≤ (Ioo 0 n).card := card_le_card_of_injOn (fun s ↦ ∏ p ∈ s, p) ?_ ?_ + n = #𝒜 := ?_ + _ ≤ #(𝒜 \\ 𝒜) := 𝒜.card_le_card_diffs + _ ≤ #(Ioo 0 n) := card_le_card_of_injOn (fun s ↦ ∏ p ∈ s, p) ?_ ?_ _ = n - 1 := by rw [card_Ioo, tsub_zero] _ < n := tsub_lt_self hn.bot_lt zero_lt_one · rw [Finset.card_image_of_injOn, card_Iio] diff --git a/Mathlib/NumberTheory/NumberField/Embeddings.lean b/Mathlib/NumberTheory/NumberField/Embeddings.lean index d713d6bce14422..d5e8af77ec71d7 100644 --- a/Mathlib/NumberTheory/NumberField/Embeddings.lean +++ b/Mathlib/NumberTheory/NumberField/Embeddings.lean @@ -32,7 +32,7 @@ for `x ∈ K`, we have `Π_w ‖x‖_w = |norm(x)|` where the product is over th number field, embeddings, places, infinite places -/ -open scoped Classical +open scoped Classical Finset namespace NumberField.Embeddings @@ -432,8 +432,7 @@ theorem one_le_mult {w : InfinitePlace K} : (1 : ℝ) ≤ mult w := by rw [← Nat.cast_one, Nat.cast_le] exact mult_pos -theorem card_filter_mk_eq [NumberField K] (w : InfinitePlace K) : - (Finset.univ.filter fun φ => mk φ = w).card = mult w := by +theorem card_filter_mk_eq [NumberField K] (w : InfinitePlace K) : #{φ | mk φ = w} = mult w := by conv_lhs => congr; congr; ext rw [← mk_embedding w, mk_eq_iff, ComplexEmbedding.conjugate, star_involutive.eq_iff] @@ -490,11 +489,8 @@ theorem prod_eq_abs_norm (x : K) : · rw [map_prod, ← Fintype.prod_equiv RingHom.equivRatAlgHom (fun f => Complex.abs (f x)) (fun φ => Complex.abs (φ x)) fun _ => by simp [RingHom.equivRatAlgHom_apply]; rfl] rw [← Finset.prod_fiberwise Finset.univ mk (fun φ => Complex.abs (φ x))] - have : ∀ w : InfinitePlace K, ∀ φ ∈ Finset.filter (fun a ↦ mk a = w) Finset.univ, - Complex.abs (φ x) = w x := by - intro _ _ hφ - rw [← (Finset.mem_filter.mp hφ).2] - rfl + have (w : InfinitePlace K) (φ) (hφ : φ ∈ ({φ | mk φ = w} : Finset _)) : + Complex.abs (φ x) = w x := by rw [← (Finset.mem_filter.mp hφ).2, apply] simp_rw [Finset.prod_congr rfl (this _), Finset.prod_const, card_filter_mk_eq] · rw [eq_ratCast, Rat.cast_abs, ← Complex.abs_ofReal, Complex.ofReal_ratCast] @@ -571,8 +567,8 @@ theorem card_eq_nrRealPlaces_add_nrComplexPlaces : theorem card_complex_embeddings : card { φ : K →+* ℂ // ¬ComplexEmbedding.IsReal φ } = 2 * NrComplexPlaces K := by - suffices ∀ w : { w : InfinitePlace K // IsComplex w }, (Finset.univ.filter - fun φ : { φ // ¬ ComplexEmbedding.IsReal φ } => mkComplex φ = w).card = 2 by + suffices ∀ w : { w : InfinitePlace K // IsComplex w }, + #{φ : {φ //¬ ComplexEmbedding.IsReal φ} | mkComplex φ = w} = 2 by rw [Fintype.card, Finset.card_eq_sum_ones, ← Finset.sum_fiberwise _ (fun φ => mkComplex φ)] simp_rw [Finset.sum_const, this, smul_eq_mul, mul_one, Fintype.card, Finset.card_eq_sum_ones, Finset.mul_sum, Finset.sum_const, smul_eq_mul, mul_one] @@ -930,16 +926,16 @@ variable [NumberField K] open Finset in lemma card_isUnramified [NumberField k] [IsGalois k K] : - Finset.card (univ.filter <| IsUnramified k (K := K)) = - Finset.card (univ.filter <| IsUnramifiedIn K (k := k)) * (finrank k K) := by + #{w : InfinitePlace K | w.IsUnramified k} = + #{w : InfinitePlace k | w.IsUnramifiedIn K} * finrank k K := by letI := Module.Finite.of_restrictScalars_finite ℚ k K rw [← IsGalois.card_aut_eq_finrank, Finset.card_eq_sum_card_fiberwise (f := (comap · (algebraMap k K))) - (t := (univ.filter <| IsUnramifiedIn K (k := k))), ← smul_eq_mul, ← sum_const] + (t := {w : InfinitePlace k | w.IsUnramifiedIn K}), ← smul_eq_mul, ← sum_const] · refine sum_congr rfl (fun w hw ↦ ?_) obtain ⟨w, rfl⟩ := comap_surjective (K := K) w simp only [mem_univ, forall_true_left, mem_filter, true_and] at hw - trans Finset.card (MulAction.orbit (K ≃ₐ[k] K) w).toFinset + trans #(MulAction.orbit (K ≃ₐ[k] K) w).toFinset · congr; ext w' simp only [mem_univ, forall_true_left, filter_congr_decidable, mem_filter, true_and, Set.mem_toFinset, mem_orbit_iff, @eq_comm _ (comap w' _), and_iff_right_iff_imp] @@ -952,12 +948,12 @@ lemma card_isUnramified [NumberField k] [IsGalois k K] : open Finset in lemma card_isUnramified_compl [NumberField k] [IsGalois k K] : - Finset.card (univ.filter <| IsUnramified k (K := K))ᶜ = - Finset.card (univ.filter <| IsUnramifiedIn K (k := k))ᶜ * (finrank k K / 2) := by + #({w : InfinitePlace K | w.IsUnramified k} : Finset _)ᶜ = + #({w : InfinitePlace k | w.IsUnramifiedIn K} : Finset _)ᶜ * (finrank k K / 2) := by letI := Module.Finite.of_restrictScalars_finite ℚ k K rw [← IsGalois.card_aut_eq_finrank, Finset.card_eq_sum_card_fiberwise (f := (comap · (algebraMap k K))) - (t := (univ.filter <| IsUnramifiedIn K (k := k))ᶜ), ← smul_eq_mul, ← sum_const] + (t := ({w : InfinitePlace k | w.IsUnramifiedIn K}: Finset _)ᶜ), ← smul_eq_mul, ← sum_const] · refine sum_congr rfl (fun w hw ↦ ?_) obtain ⟨w, rfl⟩ := comap_surjective (K := K) w simp only [mem_univ, forall_true_left, compl_filter, not_not, mem_filter, true_and] at hw @@ -974,8 +970,8 @@ lemma card_isUnramified_compl [NumberField k] [IsGalois k K] : lemma card_eq_card_isUnramifiedIn [NumberField k] [IsGalois k K] : Fintype.card (InfinitePlace K) = - Finset.card (Finset.univ.filter <| IsUnramifiedIn K (k := k)) * finrank k K + - Finset.card (Finset.univ.filter <| IsUnramifiedIn K (k := k))ᶜ * (finrank k K / 2) := by + #{w : InfinitePlace k | w.IsUnramifiedIn K} * finrank k K + + #({w : InfinitePlace k | w.IsUnramifiedIn K} : Finset _)ᶜ * (finrank k K / 2) := by rw [← card_isUnramified, ← card_isUnramified_compl, Finset.card_add_card_compl] end NumberField.InfinitePlace diff --git a/Mathlib/NumberTheory/NumberField/EquivReindex.lean b/Mathlib/NumberTheory/NumberField/EquivReindex.lean index 0226ce9f8e5d5f..76ece2db3d216a 100644 --- a/Mathlib/NumberTheory/NumberField/EquivReindex.lean +++ b/Mathlib/NumberTheory/NumberField/EquivReindex.lean @@ -11,8 +11,9 @@ import Mathlib.NumberTheory.NumberField.CanonicalEmbedding.Basic This file introduces an equivalence between the set of embeddings of `K` into `ℂ` and the index set of the chosen basis of the ring of integers of `K`. -## Tagshouse -number field, algebraic number +## Tags + +house, number field, algebraic number -/ variable (K : Type*) [Field K] [NumberField K] diff --git a/Mathlib/NumberTheory/PrimeCounting.lean b/Mathlib/NumberTheory/PrimeCounting.lean index 3005d0cb7165a5..a488afee078abd 100644 --- a/Mathlib/NumberTheory/PrimeCounting.lean +++ b/Mathlib/NumberTheory/PrimeCounting.lean @@ -101,7 +101,7 @@ theorem prime_nth_prime (n : ℕ) : Prime (nth Prime n) := /-- The cardinality of the finset `primesBelow n` equals the counting function `primeCounting'` at `n`. -/ -theorem primesBelow_card_eq_primeCounting' (n : ℕ) : n.primesBelow.card = primeCounting' n := by +theorem primesBelow_card_eq_primeCounting' (n : ℕ) : #n.primesBelow = primeCounting' n := by simp only [primesBelow, primeCounting'] exact (count_eq_card_filter_range Prime n).symm @@ -109,13 +109,13 @@ theorem primesBelow_card_eq_primeCounting' (n : ℕ) : n.primesBelow.card = prim theorem primeCounting'_add_le {a k : ℕ} (h0 : 0 < a) (h1 : a < k) (n : ℕ) : π' (k + n) ≤ π' k + Nat.totient a * (n / a + 1) := calc - π' (k + n) ≤ ((range k).filter Prime).card + ((Ico k (k + n)).filter Prime).card := by + π' (k + n) ≤ #{p ∈ range k | p.Prime} + #{p ∈ Ico k (k + n) | p.Prime} := by rw [primeCounting', count_eq_card_filter_range, range_eq_Ico, ← Ico_union_Ico_eq_Ico (zero_le k) le_self_add, filter_union] apply card_union_le - _ ≤ π' k + ((Ico k (k + n)).filter Prime).card := by + _ ≤ π' k + #{p ∈ Ico k (k + n) | p.Prime} := by rw [primeCounting', count_eq_card_filter_range] - _ ≤ π' k + ((Ico k (k + n)).filter (Coprime a)).card := by + _ ≤ π' k + #{b ∈ Ico k (k + n) | a.Coprime b} := by refine add_le_add_left (card_le_card ?_) k.primeCounting' simp only [subset_iff, and_imp, mem_filter, mem_Ico] intro p succ_k_le_p p_lt_n p_prime diff --git a/Mathlib/NumberTheory/Primorial.lean b/Mathlib/NumberTheory/Primorial.lean index 886f7183cfa38b..7ef909c7b0799a 100644 --- a/Mathlib/NumberTheory/Primorial.lean +++ b/Mathlib/NumberTheory/Primorial.lean @@ -31,8 +31,7 @@ open Nat /-- The primorial `n#` of `n` is the product of the primes less than or equal to `n`. -/ -def primorial (n : ℕ) : ℕ := - ∏ p ∈ filter Nat.Prime (range (n + 1)), p +def primorial (n : ℕ) : ℕ := ∏ p ∈ range (n + 1) with p.Prime, p local notation x "#" => primorial x @@ -45,14 +44,14 @@ theorem primorial_succ {n : ℕ} (hn1 : n ≠ 1) (hn : Odd n) : (n + 1)# = n# := exact fun h ↦ h.even_sub_one <| mt succ.inj hn1 theorem primorial_add (m n : ℕ) : - (m + n)# = m# * ∏ p ∈ filter Nat.Prime (Ico (m + 1) (m + n + 1)), p := by + (m + n)# = m# * ∏ p ∈ Ico (m + 1) (m + n + 1) with p.Prime, p := by rw [primorial, primorial, ← Ico_zero_eq_range, ← prod_union, ← filter_union, Ico_union_Ico_eq_Ico] exacts [Nat.zero_le _, add_le_add_right (Nat.le_add_right _ _) _, disjoint_filter_filter <| Ico_disjoint_Ico_consecutive _ _ _] theorem primorial_add_dvd {m n : ℕ} (h : n ≤ m) : (m + n)# ∣ m# * choose (m + n) m := calc - (m + n)# = m# * ∏ p ∈ filter Nat.Prime (Ico (m + 1) (m + n + 1)), p := primorial_add _ _ + (m + n)# = m# * ∏ p ∈ Ico (m + 1) (m + n + 1) with p.Prime, p := primorial_add _ _ _ ∣ m# * choose (m + n) m := mul_dvd_mul_left _ <| prod_primes_dvd _ (fun _ hk ↦ (mem_filter.1 hk).2.prime) fun p hp ↦ by diff --git a/Mathlib/NumberTheory/SiegelsLemma.lean b/Mathlib/NumberTheory/SiegelsLemma.lean index 998f8940d66b97..22160797bf7a58 100644 --- a/Mathlib/NumberTheory/SiegelsLemma.lean +++ b/Mathlib/NumberTheory/SiegelsLemma.lean @@ -61,7 +61,7 @@ section preparation /- In order to apply Pigeonhole we need: # Step 1: ∀ v ∈ T, A *ᵥ v ∈ S and -# Step 2: S.card < T.card +# Step 2: #S < #T Pigeonhole will give different x and y in T with A.mulVec x = A.mulVec y in S Their difference is the solution we are looking for -/ @@ -91,12 +91,12 @@ private lemma image_T_subset_S [DecidableEq α] [DecidableEq β] (v) (hv : v ∈ -- # Preparation for Step 2 -private lemma card_T_eq [DecidableEq β] : (T).card = (B + 1) ^ n := by +private lemma card_T_eq [DecidableEq β] : #T = (B + 1) ^ n := by rw [Pi.card_Icc 0 B'] simp only [Pi.zero_apply, card_Icc, sub_zero, toNat_ofNat_add_one, prod_const, card_univ, add_pos_iff, zero_lt_one, or_true] --- This lemma is necessary to be able to apply the formula (Icc a b).card = b + 1 - a +-- This lemma is necessary to be able to apply the formula #(Icc a b) = b + 1 - a private lemma N_le_P_add_one (i : α) : N i ≤ P i + 1 := by calc N i _ ≤ 0 := by @@ -109,7 +109,7 @@ private lemma N_le_P_add_one (i : α) : N i ≤ P i + 1 := by intro j _ exact mul_nonneg (Nat.cast_nonneg B) (posPart_nonneg (A i j)) -private lemma card_S_eq [DecidableEq α] : (Finset.Icc N P).card = ∏ i : α, (P i - N i + 1) := by +private lemma card_S_eq [DecidableEq α] : #(Finset.Icc N P) = ∏ i : α, (P i - N i + 1) := by rw [Pi.card_Icc N P, Nat.cast_prod] congr ext i @@ -128,13 +128,13 @@ lemma one_le_norm_A_of_ne_zero (hA : A ≠ 0) : 1 ≤ ‖A‖ := by norm_cast at h exact Int.abs_lt_one_iff.1 h --- # Step 2: S.card < T.card +-- # Step 2: #S < #T open Real Nat private lemma card_S_lt_card_T [DecidableEq α] [DecidableEq β] (hn : Fintype.card α < Fintype.card β) (hm : 0 < Fintype.card α) : - (S).card < (T).card := by + #S < #T := by zify -- This is necessary to use card_S_eq rw [card_T_eq A, card_S_eq] rify -- This is necessary because ‖A‖ is a real number diff --git a/Mathlib/Order/Interval/Finset/Basic.lean b/Mathlib/Order/Interval/Finset/Basic.lean index 53ccb40f792652..4f194b24248628 100644 --- a/Mathlib/Order/Interval/Finset/Basic.lean +++ b/Mathlib/Order/Interval/Finset/Basic.lean @@ -269,66 +269,58 @@ def _root_.Set.fintypeOfMemBounds {s : Set α} [DecidablePred (· ∈ s)] (ha : section Filter theorem Ico_filter_lt_of_le_left [DecidablePred (· < c)] (hca : c ≤ a) : - (Ico a b).filter (· < c) = ∅ := + {x ∈ Ico a b | x < c} = ∅ := filter_false_of_mem fun _ hx => (hca.trans (mem_Ico.1 hx).1).not_lt theorem Ico_filter_lt_of_right_le [DecidablePred (· < c)] (hbc : b ≤ c) : - (Ico a b).filter (· < c) = Ico a b := + {x ∈ Ico a b | x < c} = Ico a b := filter_true_of_mem fun _ hx => (mem_Ico.1 hx).2.trans_le hbc theorem Ico_filter_lt_of_le_right [DecidablePred (· < c)] (hcb : c ≤ b) : - (Ico a b).filter (· < c) = Ico a c := by + {x ∈ Ico a b | x < c} = Ico a c := by ext x rw [mem_filter, mem_Ico, mem_Ico, and_right_comm] exact and_iff_left_of_imp fun h => h.2.trans_le hcb theorem Ico_filter_le_of_le_left {a b c : α} [DecidablePred (c ≤ ·)] (hca : c ≤ a) : - (Ico a b).filter (c ≤ ·) = Ico a b := + {x ∈ Ico a b | c ≤ x} = Ico a b := filter_true_of_mem fun _ hx => hca.trans (mem_Ico.1 hx).1 theorem Ico_filter_le_of_right_le {a b : α} [DecidablePred (b ≤ ·)] : - (Ico a b).filter (b ≤ ·) = ∅ := + {x ∈ Ico a b | b ≤ x} = ∅ := filter_false_of_mem fun _ hx => (mem_Ico.1 hx).2.not_le theorem Ico_filter_le_of_left_le {a b c : α} [DecidablePred (c ≤ ·)] (hac : a ≤ c) : - (Ico a b).filter (c ≤ ·) = Ico c b := by + {x ∈ Ico a b | c ≤ x} = Ico c b := by ext x rw [mem_filter, mem_Ico, mem_Ico, and_comm, and_left_comm] exact and_iff_right_of_imp fun h => hac.trans h.1 theorem Icc_filter_lt_of_lt_right {a b c : α} [DecidablePred (· < c)] (h : b < c) : - (Icc a b).filter (· < c) = Icc a b := + {x ∈ Icc a b | x < c} = Icc a b := filter_true_of_mem fun _ hx => lt_of_le_of_lt (mem_Icc.1 hx).2 h theorem Ioc_filter_lt_of_lt_right {a b c : α} [DecidablePred (· < c)] (h : b < c) : - (Ioc a b).filter (· < c) = Ioc a b := + {x ∈ Ioc a b | x < c} = Ioc a b := filter_true_of_mem fun _ hx => lt_of_le_of_lt (mem_Ioc.1 hx).2 h theorem Iic_filter_lt_of_lt_right {α} [Preorder α] [LocallyFiniteOrderBot α] {a c : α} - [DecidablePred (· < c)] (h : a < c) : (Iic a).filter (· < c) = Iic a := + [DecidablePred (· < c)] (h : a < c) : {x ∈ Iic a | x < c} = Iic a := filter_true_of_mem fun _ hx => lt_of_le_of_lt (mem_Iic.1 hx) h variable (a b) [Fintype α] theorem filter_lt_lt_eq_Ioo [DecidablePred fun j => a < j ∧ j < b] : - (univ.filter fun j => a < j ∧ j < b) = Ioo a b := by - ext - simp + ({j | a < j ∧ j < b} : Finset _) = Ioo a b := by ext; simp theorem filter_lt_le_eq_Ioc [DecidablePred fun j => a < j ∧ j ≤ b] : - (univ.filter fun j => a < j ∧ j ≤ b) = Ioc a b := by - ext - simp + ({j | a < j ∧ j ≤ b} : Finset _) = Ioc a b := by ext; simp theorem filter_le_lt_eq_Ico [DecidablePred fun j => a ≤ j ∧ j < b] : - (univ.filter fun j => a ≤ j ∧ j < b) = Ico a b := by - ext - simp + ({j | a ≤ j ∧ j < b} : Finset _) = Ico a b := by ext; simp theorem filter_le_le_eq_Icc [DecidablePred fun j => a ≤ j ∧ j ≤ b] : - (univ.filter fun j => a ≤ j ∧ j ≤ b) = Icc a b := by - ext - simp + ({j | a ≤ j ∧ j ≤ b} : Finset _) = Icc a b := by ext; simp end Filter @@ -430,13 +422,8 @@ theorem _root_.Set.Infinite.not_bddBelow {s : Set α} : s.Infinite → ¬BddBelo variable [Fintype α] -theorem filter_lt_eq_Ioi [DecidablePred (a < ·)] : univ.filter (a < ·) = Ioi a := by - ext - simp - -theorem filter_le_eq_Ici [DecidablePred (a ≤ ·)] : univ.filter (a ≤ ·) = Ici a := by - ext - simp +theorem filter_lt_eq_Ioi [DecidablePred (a < ·)] : ({x | a < x} : Finset _) = Ioi a := by ext; simp +theorem filter_le_eq_Ici [DecidablePred (a ≤ ·)] : ({x | a ≤ x} : Finset _) = Ici a := by ext; simp end LocallyFiniteOrderTop @@ -455,13 +442,8 @@ theorem _root_.Set.Infinite.not_bddAbove {s : Set α} : s.Infinite → ¬BddAbov variable [Fintype α] -theorem filter_gt_eq_Iio [DecidablePred (· < a)] : univ.filter (· < a) = Iio a := by - ext - simp - -theorem filter_ge_eq_Iic [DecidablePred (· ≤ a)] : univ.filter (· ≤ a) = Iic a := by - ext - simp +theorem filter_gt_eq_Iio [DecidablePred (· < a)] : ({x | x < a} : Finset _) = Iio a := by ext; simp +theorem filter_ge_eq_Iic [DecidablePred (· ≤ a)] : ({x | x ≤ a} : Finset _) = Iic a := by ext; simp end LocallyFiniteOrderBot @@ -561,32 +543,32 @@ theorem Ico_eq_cons_Ioo (h : a < b) : Ico a b = (Ioo a b).cons a left_not_mem_Io classical rw [cons_eq_insert, Ioo_insert_left h] theorem Ico_filter_le_left {a b : α} [DecidablePred (· ≤ a)] (hab : a < b) : - ((Ico a b).filter fun x => x ≤ a) = {a} := by + {x ∈ Ico a b | x ≤ a} = {a} := by ext x rw [mem_filter, mem_Ico, mem_singleton, and_right_comm, ← le_antisymm_iff, eq_comm] exact and_iff_left_of_imp fun h => h.le.trans_lt hab -theorem card_Ico_eq_card_Icc_sub_one (a b : α) : (Ico a b).card = (Icc a b).card - 1 := by +theorem card_Ico_eq_card_Icc_sub_one (a b : α) : #(Ico a b) = #(Icc a b) - 1 := by classical by_cases h : a ≤ b · rw [Icc_eq_cons_Ico h, card_cons] exact (Nat.add_sub_cancel _ _).symm · rw [Ico_eq_empty fun h' => h h'.le, Icc_eq_empty h, card_empty, Nat.zero_sub] -theorem card_Ioc_eq_card_Icc_sub_one (a b : α) : (Ioc a b).card = (Icc a b).card - 1 := +theorem card_Ioc_eq_card_Icc_sub_one (a b : α) : #(Ioc a b) = #(Icc a b) - 1 := @card_Ico_eq_card_Icc_sub_one αᵒᵈ _ _ _ _ -theorem card_Ioo_eq_card_Ico_sub_one (a b : α) : (Ioo a b).card = (Ico a b).card - 1 := by +theorem card_Ioo_eq_card_Ico_sub_one (a b : α) : #(Ioo a b) = #(Ico a b) - 1 := by classical by_cases h : a < b · rw [Ico_eq_cons_Ioo h, card_cons] exact (Nat.add_sub_cancel _ _).symm · rw [Ioo_eq_empty h, Ico_eq_empty h, card_empty, Nat.zero_sub] -theorem card_Ioo_eq_card_Ioc_sub_one (a b : α) : (Ioo a b).card = (Ioc a b).card - 1 := +theorem card_Ioo_eq_card_Ioc_sub_one (a b : α) : #(Ioo a b) = #(Ioc a b) - 1 := @card_Ioo_eq_card_Ico_sub_one αᵒᵈ _ _ _ _ -theorem card_Ioo_eq_card_Icc_sub_two (a b : α) : (Ioo a b).card = (Icc a b).card - 2 := by +theorem card_Ioo_eq_card_Icc_sub_two (a b : α) : #(Ioo a b) = #(Icc a b) - 2 := by rw [card_Ioo_eq_card_Ico_sub_one, card_Ico_eq_card_Icc_sub_one] rfl @@ -619,7 +601,7 @@ theorem not_mem_Ioi_self {b : α} : b ∉ Ioi b := fun h => lt_irrefl _ (mem_Ioi theorem Ici_eq_cons_Ioi (a : α) : Ici a = (Ioi a).cons a not_mem_Ioi_self := by classical rw [cons_eq_insert, Ioi_insert] -theorem card_Ioi_eq_card_Ici_sub_one (a : α) : (Ioi a).card = (Ici a).card - 1 := by +theorem card_Ioi_eq_card_Ici_sub_one (a : α) : #(Ioi a) = #(Ici a) - 1 := by rw [Ici_eq_cons_Ioi, card_cons, Nat.add_sub_cancel_right] end OrderTop @@ -647,7 +629,7 @@ theorem not_mem_Iio_self {b : α} : b ∉ Iio b := fun h => lt_irrefl _ (mem_Iio theorem Iic_eq_cons_Iio (b : α) : Iic b = (Iio b).cons b not_mem_Iio_self := by classical rw [cons_eq_insert, Iio_insert] -theorem card_Iio_eq_card_Iic_sub_one (a : α) : (Iio a).card = (Iic a).card - 1 := by +theorem card_Iio_eq_card_Iic_sub_one (a : α) : #(Iio a) = #(Iic a) - 1 := by rw [Iic_eq_cons_Iio, card_cons, Nat.add_sub_cancel_right] end OrderBot @@ -715,25 +697,25 @@ theorem Ico_inter_Ico {a b c d : α} : Ico a b ∩ Ico c d = Ico (max a c) (min Set.Ico_inter_Ico] @[simp] -theorem Ico_filter_lt (a b c : α) : ((Ico a b).filter fun x => x < c) = Ico a (min b c) := by +theorem Ico_filter_lt (a b c : α) : {x ∈ Ico a b | x < c} = Ico a (min b c) := by cases le_total b c with | inl h => rw [Ico_filter_lt_of_right_le h, min_eq_left h] | inr h => rw [Ico_filter_lt_of_le_right h, min_eq_right h] @[simp] -theorem Ico_filter_le (a b c : α) : ((Ico a b).filter fun x => c ≤ x) = Ico (max a c) b := by +theorem Ico_filter_le (a b c : α) : {x ∈ Ico a b | c ≤ x} = Ico (max a c) b := by cases le_total a c with | inl h => rw [Ico_filter_le_of_left_le h, max_eq_right h] | inr h => rw [Ico_filter_le_of_le_left h, max_eq_left h] @[simp] -theorem Ioo_filter_lt (a b c : α) : (Ioo a b).filter (· < c) = Ioo a (min b c) := by +theorem Ioo_filter_lt (a b c : α) : {x ∈ Ioo a b | x < c} = Ioo a (min b c) := by ext simp [and_assoc] @[simp] theorem Iio_filter_lt {α} [LinearOrder α] [LocallyFiniteOrderBot α] (a b : α) : - (Iio a).filter (· < b) = Iio (min a b) := by + {x ∈ Iio a | x < b} = Iio (min a b) := by ext simp [and_assoc] @@ -933,7 +915,7 @@ lemma transGen_wcovBy_of_le [Preorder α] [LocallyFiniteOrder α] {x y : α} (hx TransGen (· ⩿ ·) x y := by -- We proceed by well-founded induction on the cardinality of `Icc x y`. -- It's impossible for the cardinality to be zero since `x ≤ y` - have : (Ico x y).card < (Icc x y).card := card_lt_card <| + have : #(Ico x y) < #(Icc x y) := card_lt_card <| ⟨Ico_subset_Icc_self, not_subset.mpr ⟨y, ⟨right_mem_Icc.mpr hxy, right_not_mem_Ico⟩⟩⟩ by_cases hxy' : y ≤ x -- If `y ≤ x`, then `x ⩿ y` @@ -943,16 +925,15 @@ lemma transGen_wcovBy_of_le [Preorder α] [LocallyFiniteOrder α] {x y : α} (hx induction hypothesis to show that `Relation.TransGen (· ⩿ ·) x z`. -/ · have h_non : (Ico x y).Nonempty := ⟨x, mem_Ico.mpr ⟨le_rfl, lt_of_le_not_le hxy hxy'⟩⟩ obtain ⟨z, z_mem, hz⟩ := (Ico x y).exists_maximal h_non - have z_card : (Icc x z).card <(Icc x y).card := calc - (Icc x z).card ≤ (Ico x y).card := - card_le_card <| Icc_subset_Ico_right (mem_Ico.mp z_mem).2 - _ < (Icc x y).card := this + have z_card := calc + #(Icc x z) ≤ #(Ico x y) := card_le_card <| Icc_subset_Ico_right (mem_Ico.mp z_mem).2 + _ < #(Icc x y) := this have h₁ := transGen_wcovBy_of_le (mem_Ico.mp z_mem).1 have h₂ : z ⩿ y := by refine ⟨(mem_Ico.mp z_mem).2.le, fun c hzc hcy ↦ hz c ?_ hzc⟩ exact mem_Ico.mpr <| ⟨(mem_Ico.mp z_mem).1.trans hzc.le, hcy⟩ exact .tail h₁ h₂ -termination_by (Icc x y).card +termination_by #(Icc x y) /-- In a locally finite preorder, `≤` is the transitive closure of `⩿`. -/ lemma le_iff_transGen_wcovBy [Preorder α] [LocallyFiniteOrder α] {x y : α} : @@ -976,7 +957,7 @@ lemma transGen_covBy_of_lt [Preorder α] [LocallyFiniteOrder α] {x y : α} (hxy -- `Ico x y` is a nonempty finset and so contains a maximal element `z` and -- `Ico x z` has cardinality strictly less than the cardinality of `Ico x y` obtain ⟨z, z_mem, hz⟩ := (Ico x y).exists_maximal h_non - have z_card : (Ico x z).card < (Ico x y).card := card_lt_card <| ssubset_iff_of_subset + have z_card : #(Ico x z) < #(Ico x y) := card_lt_card <| ssubset_iff_of_subset (Ico_subset_Ico le_rfl (mem_Ico.mp z_mem).2.le) |>.mpr ⟨z, z_mem, right_not_mem_Ico⟩ /- Since `z` is maximal in `Ico x y`, `z ⋖ y`. -/ have hzy : z ⋖ y := by @@ -990,7 +971,7 @@ lemma transGen_covBy_of_lt [Preorder α] [LocallyFiniteOrder α] {x y : α} (hxy `x ≤ z`), and since `z ⋖ y` we conclude that `x ⋖ y` , then `Relation.TransGen.single`. -/ · simp only [lt_iff_le_not_le, not_and, not_not] at hxz exact .single (hzy.of_le_of_lt (hxz (mem_Ico.mp z_mem).1) hxy) -termination_by (Ico x y).card +termination_by #(Ico x y) /-- In a locally finite preorder, `<` is the transitive closure of `⋖`. -/ lemma lt_iff_transGen_covBy [Preorder α] [LocallyFiniteOrder α] {x y : α} : diff --git a/Mathlib/Order/Interval/Finset/Box.lean b/Mathlib/Order/Interval/Finset/Box.lean index ae26532a1a7358..902ac28210bbd2 100644 --- a/Mathlib/Order/Interval/Finset/Box.lean +++ b/Mathlib/Order/Interval/Finset/Box.lean @@ -65,9 +65,9 @@ variable {α β : Type*} [OrderedRing α] [OrderedRing β] [LocallyFiniteOrder [DecidableEq α] [DecidableEq β] [@DecidableRel (α × β) (· ≤ ·)] @[simp] lemma card_box_succ (n : ℕ) : - (box (n + 1) : Finset (α × β)).card = - (Icc (-n.succ : α) n.succ).card * (Icc (-n.succ : β) n.succ).card - - (Icc (-n : α) n).card * (Icc (-n : β) n).card := by + #(box (n + 1) : Finset (α × β)) = + #(Icc (-n.succ : α) n.succ) * #(Icc (-n.succ : β) n.succ) - + #(Icc (-n : α) n) * #(Icc (-n : β) n) := by rw [box_succ_eq_sdiff, card_sdiff (Icc_neg_mono n.le_succ), Finset.card_Icc_prod, Finset.card_Icc_prod] rfl @@ -81,7 +81,7 @@ variable {n : ℕ} {x : ℤ × ℤ} attribute [norm_cast] toNat_ofNat -lemma card_box : ∀ {n}, n ≠ 0 → (box n : Finset (ℤ × ℤ)).card = 8 * n +lemma card_box : ∀ {n}, n ≠ 0 → #(box n : Finset (ℤ × ℤ)) = 8 * n | n + 1, _ => by simp_rw [Prod.card_box_succ, card_Icc, sub_neg_eq_add] norm_cast diff --git a/Mathlib/Order/Interval/Finset/Defs.lean b/Mathlib/Order/Interval/Finset/Defs.lean index 4957140ea58934..625bd3be33b319 100644 --- a/Mathlib/Order/Interval/Finset/Defs.lean +++ b/Mathlib/Order/Interval/Finset/Defs.lean @@ -70,7 +70,7 @@ Provide the `LocallyFiniteOrder` instance for `α →₀ β` where `β` is local From `LinearOrder α`, `NoMaxOrder α`, `LocallyFiniteOrder α`, we can also define an order isomorphism `α ≃ ℕ` or `α ≃ ℤ`, depending on whether we have `OrderBot α` or -`NoMinOrder α` and `Nonempty α`. When `OrderBot α`, we can match `a : α` to `(Iio a).card`. +`NoMinOrder α` and `Nonempty α`. When `OrderBot α`, we can match `a : α` to `#(Iio a)`. We can provide `SuccOrder α` from `LinearOrder α` and `LocallyFiniteOrder α` using @@ -146,79 +146,76 @@ the ends. As opposed to `LocallyFiniteOrder.ofIcc`, this one requires `Decidable only `Preorder`. -/ def LocallyFiniteOrder.ofIcc' (α : Type*) [Preorder α] [DecidableRel ((· ≤ ·) : α → α → Prop)] (finsetIcc : α → α → Finset α) (mem_Icc : ∀ a b x, x ∈ finsetIcc a b ↔ a ≤ x ∧ x ≤ b) : - LocallyFiniteOrder α := - { finsetIcc - finsetIco := fun a b => (finsetIcc a b).filter fun x => ¬b ≤ x - finsetIoc := fun a b => (finsetIcc a b).filter fun x => ¬x ≤ a - finsetIoo := fun a b => (finsetIcc a b).filter fun x => ¬x ≤ a ∧ ¬b ≤ x - finset_mem_Icc := mem_Icc - finset_mem_Ico := fun a b x => by rw [Finset.mem_filter, mem_Icc, and_assoc, lt_iff_le_not_le] - finset_mem_Ioc := fun a b x => by - rw [Finset.mem_filter, mem_Icc, and_right_comm, lt_iff_le_not_le] - finset_mem_Ioo := fun a b x => by - rw [Finset.mem_filter, mem_Icc, and_and_and_comm, lt_iff_le_not_le, lt_iff_le_not_le] } + LocallyFiniteOrder α where + finsetIcc := finsetIcc + finsetIco a b := {x ∈ finsetIcc a b | ¬b ≤ x} + finsetIoc a b := {x ∈ finsetIcc a b | ¬x ≤ a} + finsetIoo a b := {x ∈ finsetIcc a b | ¬x ≤ a ∧ ¬b ≤ x} + finset_mem_Icc := mem_Icc + finset_mem_Ico a b x := by rw [Finset.mem_filter, mem_Icc, and_assoc, lt_iff_le_not_le] + finset_mem_Ioc a b x := by rw [Finset.mem_filter, mem_Icc, and_right_comm, lt_iff_le_not_le] + finset_mem_Ioo a b x := by + rw [Finset.mem_filter, mem_Icc, and_and_and_comm, lt_iff_le_not_le, lt_iff_le_not_le] /-- A constructor from a definition of `Finset.Icc` alone, the other ones being derived by removing the ends. As opposed to `LocallyFiniteOrder.ofIcc'`, this one requires `PartialOrder` but only `DecidableEq`. -/ def LocallyFiniteOrder.ofIcc (α : Type*) [PartialOrder α] [DecidableEq α] (finsetIcc : α → α → Finset α) (mem_Icc : ∀ a b x, x ∈ finsetIcc a b ↔ a ≤ x ∧ x ≤ b) : - LocallyFiniteOrder α := - { finsetIcc - finsetIco := fun a b => (finsetIcc a b).filter fun x => x ≠ b - finsetIoc := fun a b => (finsetIcc a b).filter fun x => a ≠ x - finsetIoo := fun a b => (finsetIcc a b).filter fun x => a ≠ x ∧ x ≠ b - finset_mem_Icc := mem_Icc - finset_mem_Ico := fun a b x => by rw [Finset.mem_filter, mem_Icc, and_assoc, lt_iff_le_and_ne] - finset_mem_Ioc := fun a b x => by - rw [Finset.mem_filter, mem_Icc, and_right_comm, lt_iff_le_and_ne] - finset_mem_Ioo := fun a b x => by - rw [Finset.mem_filter, mem_Icc, and_and_and_comm, lt_iff_le_and_ne, lt_iff_le_and_ne] } + LocallyFiniteOrder α where + finsetIcc := finsetIcc + finsetIco a b := {x ∈ finsetIcc a b | x ≠ b} + finsetIoc a b := {x ∈ finsetIcc a b | a ≠ x} + finsetIoo a b := {x ∈ finsetIcc a b | a ≠ x ∧ x ≠ b} + finset_mem_Icc := mem_Icc + finset_mem_Ico a b x := by rw [Finset.mem_filter, mem_Icc, and_assoc, lt_iff_le_and_ne] + finset_mem_Ioc a b x := by rw [Finset.mem_filter, mem_Icc, and_right_comm, lt_iff_le_and_ne] + finset_mem_Ioo a b x := by + rw [Finset.mem_filter, mem_Icc, and_and_and_comm, lt_iff_le_and_ne, lt_iff_le_and_ne] /-- A constructor from a definition of `Finset.Ici` alone, the other ones being derived by removing the ends. As opposed to `LocallyFiniteOrderTop.ofIci`, this one requires `DecidableRel (· ≤ ·)` but only `Preorder`. -/ def LocallyFiniteOrderTop.ofIci' (α : Type*) [Preorder α] [DecidableRel ((· ≤ ·) : α → α → Prop)] (finsetIci : α → Finset α) (mem_Ici : ∀ a x, x ∈ finsetIci a ↔ a ≤ x) : - LocallyFiniteOrderTop α := - { finsetIci - finsetIoi := fun a => (finsetIci a).filter fun x => ¬x ≤ a - finset_mem_Ici := mem_Ici - finset_mem_Ioi := fun a x => by rw [mem_filter, mem_Ici, lt_iff_le_not_le] } + LocallyFiniteOrderTop α where + finsetIci := finsetIci + finsetIoi a := {x ∈ finsetIci a | ¬x ≤ a} + finset_mem_Ici := mem_Ici + finset_mem_Ioi a x := by rw [mem_filter, mem_Ici, lt_iff_le_not_le] /-- A constructor from a definition of `Finset.Ici` alone, the other ones being derived by removing the ends. As opposed to `LocallyFiniteOrderTop.ofIci'`, this one requires `PartialOrder` but only `DecidableEq`. -/ def LocallyFiniteOrderTop.ofIci (α : Type*) [PartialOrder α] [DecidableEq α] (finsetIci : α → Finset α) (mem_Ici : ∀ a x, x ∈ finsetIci a ↔ a ≤ x) : - LocallyFiniteOrderTop α := - { finsetIci - finsetIoi := fun a => (finsetIci a).filter fun x => a ≠ x - finset_mem_Ici := mem_Ici - finset_mem_Ioi := fun a x => by rw [mem_filter, mem_Ici, lt_iff_le_and_ne] } + LocallyFiniteOrderTop α where + finsetIci := finsetIci + finsetIoi a := {x ∈ finsetIci a | a ≠ x} + finset_mem_Ici := mem_Ici + finset_mem_Ioi a x := by rw [mem_filter, mem_Ici, lt_iff_le_and_ne] /-- A constructor from a definition of `Finset.Iic` alone, the other ones being derived by removing the ends. As opposed to `LocallyFiniteOrderBot.ofIic`, this one requires `DecidableRel (· ≤ ·)` but only `Preorder`. -/ def LocallyFiniteOrderBot.ofIic' (α : Type*) [Preorder α] [DecidableRel ((· ≤ ·) : α → α → Prop)] (finsetIic : α → Finset α) (mem_Iic : ∀ a x, x ∈ finsetIic a ↔ x ≤ a) : - LocallyFiniteOrderBot α := - { finsetIic - finsetIio := fun a => (finsetIic a).filter fun x => ¬a ≤ x - finset_mem_Iic := mem_Iic - finset_mem_Iio := fun a x => by rw [mem_filter, mem_Iic, lt_iff_le_not_le] } + LocallyFiniteOrderBot α where + finsetIic := finsetIic + finsetIio a := {x ∈ finsetIic a | ¬a ≤ x} + finset_mem_Iic := mem_Iic + finset_mem_Iio a x := by rw [mem_filter, mem_Iic, lt_iff_le_not_le] /-- A constructor from a definition of `Finset.Iic` alone, the other ones being derived by removing the ends. As opposed to `LocallyFiniteOrderBot.ofIic'`, this one requires `PartialOrder` but only `DecidableEq`. -/ def LocallyFiniteOrderBot.ofIic (α : Type*) [PartialOrder α] [DecidableEq α] (finsetIic : α → Finset α) (mem_Iic : ∀ a x, x ∈ finsetIic a ↔ x ≤ a) : - LocallyFiniteOrderBot α := - { finsetIic - finsetIio := fun a => (finsetIic a).filter fun x => x ≠ a - finset_mem_Iic := mem_Iic - finset_mem_Iio := fun a x => by rw [mem_filter, mem_Iic, lt_iff_le_and_ne] } --- Note: this was in the wrong namespace in Mathlib 3. + LocallyFiniteOrderBot α where + finsetIic := finsetIic + finsetIio a := {x ∈ finsetIic a | x ≠ a} + finset_mem_Iic := mem_Iic + finset_mem_Iio a x := by rw [mem_filter, mem_Iic, lt_iff_le_and_ne] variable {α β : Type*} @@ -786,8 +783,8 @@ lemma Finset.Icc_prod_def (x y : α × β) : Icc x y = Icc x.1 y.1 ×ˢ Icc x.2 lemma Finset.Icc_product_Icc (a₁ a₂ : α) (b₁ b₂ : β) : Icc a₁ a₂ ×ˢ Icc b₁ b₂ = Icc (a₁, b₁) (a₂, b₂) := rfl -lemma Finset.card_Icc_prod (x y : α × β) : - (Icc x y).card = (Icc x.1 y.1).card * (Icc x.2 y.2).card := card_product _ _ +lemma Finset.card_Icc_prod (x y : α × β) : #(Icc x y) = #(Icc x.1 y.1) * #(Icc x.2 y.2) := + card_product .. end LocallyFiniteOrder @@ -800,7 +797,7 @@ instance Prod.instLocallyFiniteOrderTop : LocallyFiniteOrderTop (α × β) := lemma Finset.Ici_prod_def (x : α × β) : Ici x = Ici x.1 ×ˢ Ici x.2 := rfl lemma Finset.Ici_product_Ici (a : α) (b : β) : Ici a ×ˢ Ici b = Ici (a, b) := rfl -lemma Finset.card_Ici_prod (x : α × β) : (Ici x).card = (Ici x.1).card * (Ici x.2).card := +lemma Finset.card_Ici_prod (x : α × β) : #(Ici x) = #(Ici x.1) * #(Ici x.2) := card_product _ _ end LocallyFiniteOrderTop @@ -814,8 +811,7 @@ instance Prod.instLocallyFiniteOrderBot : LocallyFiniteOrderBot (α × β) := lemma Finset.Iic_prod_def (x : α × β) : Iic x = Iic x.1 ×ˢ Iic x.2 := rfl lemma Finset.Iic_product_Iic (a : α) (b : β) : Iic a ×ˢ Iic b = Iic (a, b) := rfl -lemma Finset.card_Iic_prod (x : α × β) : (Iic x).card = (Iic x.1).card * (Iic x.2).card := - card_product _ _ +lemma Finset.card_Iic_prod (x : α × β) : #(Iic x) = #(Iic x.1) * #(Iic x.2) := card_product .. end LocallyFiniteOrderBot end Preorder @@ -829,8 +825,8 @@ lemma Finset.uIcc_prod_def (x y : α × β) : uIcc x y = uIcc x.1 y.1 ×ˢ uIcc lemma Finset.uIcc_product_uIcc (a₁ a₂ : α) (b₁ b₂ : β) : uIcc a₁ a₂ ×ˢ uIcc b₁ b₂ = uIcc (a₁, b₁) (a₂, b₂) := rfl -lemma Finset.card_uIcc_prod (x y : α × β) : - (uIcc x y).card = (uIcc x.1 y.1).card * (uIcc x.2 y.2).card := card_product _ _ +lemma Finset.card_uIcc_prod (x y : α × β) : #(uIcc x y) = #(uIcc x.1 y.1) * #(uIcc x.2 y.2) := + card_product .. end Lattice diff --git a/Mathlib/Order/Interval/Finset/Fin.lean b/Mathlib/Order/Interval/Finset/Fin.lean index f0a6f5f7b6aef5..24bf50afbdc492 100644 --- a/Mathlib/Order/Interval/Finset/Fin.lean +++ b/Mathlib/Order/Interval/Finset/Fin.lean @@ -85,23 +85,19 @@ theorem map_subtype_embedding_uIcc : (uIcc a b).map valEmbedding = uIcc ↑a ↑ map_valEmbedding_Icc _ _ @[simp] -theorem card_Icc : (Icc a b).card = b + 1 - a := by - rw [← Nat.card_Icc, ← map_valEmbedding_Icc, card_map] +lemma card_Icc : #(Icc a b) = b + 1 - a := by rw [← Nat.card_Icc, ← map_valEmbedding_Icc, card_map] @[simp] -theorem card_Ico : (Ico a b).card = b - a := by - rw [← Nat.card_Ico, ← map_valEmbedding_Ico, card_map] +lemma card_Ico : #(Ico a b) = b - a := by rw [← Nat.card_Ico, ← map_valEmbedding_Ico, card_map] @[simp] -theorem card_Ioc : (Ioc a b).card = b - a := by - rw [← Nat.card_Ioc, ← map_valEmbedding_Ioc, card_map] +lemma card_Ioc : #(Ioc a b) = b - a := by rw [← Nat.card_Ioc, ← map_valEmbedding_Ioc, card_map] @[simp] -theorem card_Ioo : (Ioo a b).card = b - a - 1 := by - rw [← Nat.card_Ioo, ← map_valEmbedding_Ioo, card_map] +lemma card_Ioo : #(Ioo a b) = b - a - 1 := by rw [← Nat.card_Ioo, ← map_valEmbedding_Ioo, card_map] @[simp] -theorem card_uIcc : (uIcc a b).card = (b - a : ℤ).natAbs + 1 := by +theorem card_uIcc : #(uIcc a b) = (b - a : ℤ).natAbs + 1 := by rw [← Nat.card_uIcc, ← map_subtype_embedding_uIcc, card_map] -- Porting note (#10618): simp can prove this @@ -172,23 +168,20 @@ theorem map_valEmbedding_Iio : (Iio b).map Fin.valEmbedding = Iio ↑b := by simp [Iio_eq_finset_subtype, Finset.fin, Finset.map_map] @[simp] -theorem card_Ici : (Ici a).card = n - a := by +theorem card_Ici : #(Ici a) = n - a := by cases n with | zero => exact Fin.elim0 a | succ => rw [← card_map, map_valEmbedding_Ici, Nat.card_Icc, Nat.add_one_sub_one] @[simp] -theorem card_Ioi : (Ioi a).card = n - 1 - a := by - rw [← card_map, map_valEmbedding_Ioi, Nat.card_Ioc] +theorem card_Ioi : #(Ioi a) = n - 1 - a := by rw [← card_map, map_valEmbedding_Ioi, Nat.card_Ioc] @[simp] -theorem card_Iic : (Iic b).card = b + 1 := by - rw [← Nat.card_Iic b, ← map_valEmbedding_Iic, card_map] +theorem card_Iic : #(Iic b) = b + 1 := by rw [← Nat.card_Iic b, ← map_valEmbedding_Iic, card_map] @[simp] -theorem card_Iio : (Iio b).card = b := by - rw [← Nat.card_Iio b, ← map_valEmbedding_Iio, card_map] +theorem card_Iio : #(Iio b) = b := by rw [← Nat.card_Iio b, ← map_valEmbedding_Iio, card_map] -- Porting note (#10618): simp can prove this -- @[simp] diff --git a/Mathlib/Order/Interval/Finset/Nat.lean b/Mathlib/Order/Interval/Finset/Nat.lean index 8be56194d64812..6b580e5787485c 100644 --- a/Mathlib/Order/Interval/Finset/Nat.lean +++ b/Mathlib/Order/Interval/Finset/Nat.lean @@ -65,31 +65,20 @@ lemma range_eq_Icc_zero_sub_one (n : ℕ) (hn : n ≠ 0) : range n = Icc 0 (n - theorem _root_.Finset.range_eq_Ico : range = Ico 0 := Ico_zero_eq_range.symm -@[simp] -theorem card_Icc : (Icc a b).card = b + 1 - a := - List.length_range' _ _ _ - -@[simp] -theorem card_Ico : (Ico a b).card = b - a := - List.length_range' _ _ _ - -@[simp] -theorem card_Ioc : (Ioc a b).card = b - a := - List.length_range' _ _ _ - -@[simp] -theorem card_Ioo : (Ioo a b).card = b - a - 1 := - List.length_range' _ _ _ +@[simp] lemma card_Icc : #(Icc a b) = b + 1 - a := List.length_range' .. +@[simp] lemma card_Ico : #(Ico a b) = b - a := List.length_range' .. +@[simp] lemma card_Ioc : #(Ioc a b) = b - a := List.length_range' .. +@[simp] lemma card_Ioo : #(Ioo a b) = b - a - 1 := List.length_range' .. @[simp] -theorem card_uIcc : (uIcc a b).card = (b - a : ℤ).natAbs + 1 := +theorem card_uIcc : #(uIcc a b) = (b - a : ℤ).natAbs + 1 := (card_Icc _ _).trans <| by rw [← Int.natCast_inj, sup_eq_max, inf_eq_min, Int.ofNat_sub] <;> omega @[simp] -lemma card_Iic : (Iic b).card = b + 1 := by rw [Iic_eq_Icc, card_Icc, Nat.bot_eq_zero, Nat.sub_zero] +lemma card_Iic : #(Iic b) = b + 1 := by rw [Iic_eq_Icc, card_Icc, Nat.bot_eq_zero, Nat.sub_zero] @[simp] -theorem card_Iio : (Iio b).card = b := by rw [Iio_eq_Ico, card_Ico, Nat.bot_eq_zero, Nat.sub_zero] +theorem card_Iio : #(Iio b) = b := by rw [Iio_eq_Ico, card_Ico, Nat.bot_eq_zero, Nat.sub_zero] -- Porting note (#10618): simp can prove this -- @[simp] diff --git a/Mathlib/Order/Partition/Equipartition.lean b/Mathlib/Order/Partition/Equipartition.lean index 3969b84ca5eb03..4cd1f1579f3398 100644 --- a/Mathlib/Order/Partition/Equipartition.lean +++ b/Mathlib/Order/Partition/Equipartition.lean @@ -35,44 +35,41 @@ def IsEquipartition : Prop := theorem isEquipartition_iff_card_parts_eq_average : P.IsEquipartition ↔ - ∀ a : Finset α, - a ∈ P.parts → a.card = s.card / P.parts.card ∨ a.card = s.card / P.parts.card + 1 := by + ∀ a : Finset α, a ∈ P.parts → #a = #s / #P.parts ∨ #a = #s / #P.parts + 1 := by simp_rw [IsEquipartition, Finset.equitableOn_iff, P.sum_card_parts] variable {P} lemma not_isEquipartition : - ¬P.IsEquipartition ↔ ∃ a ∈ P.parts, ∃ b ∈ P.parts, b.card + 1 < a.card := - Set.not_equitableOn + ¬P.IsEquipartition ↔ ∃ a ∈ P.parts, ∃ b ∈ P.parts, #b + 1 < #a := Set.not_equitableOn theorem _root_.Set.Subsingleton.isEquipartition (h : (P.parts : Set (Finset α)).Subsingleton) : P.IsEquipartition := Set.Subsingleton.equitableOn h _ theorem IsEquipartition.card_parts_eq_average (hP : P.IsEquipartition) (ht : t ∈ P.parts) : - t.card = s.card / P.parts.card ∨ t.card = s.card / P.parts.card + 1 := + #t = #s / #P.parts ∨ #t = #s / #P.parts + 1 := P.isEquipartition_iff_card_parts_eq_average.1 hP _ ht theorem IsEquipartition.card_part_eq_average_iff (hP : P.IsEquipartition) (ht : t ∈ P.parts) : - t.card = s.card / P.parts.card ↔ t.card ≠ s.card / P.parts.card + 1 := by + #t = #s / #P.parts ↔ #t ≠ #s / #P.parts + 1 := by have a := hP.card_parts_eq_average ht - have b : ¬(t.card = s.card / P.parts.card ∧ t.card = s.card / P.parts.card + 1) := by + have b : ¬(#t = #s / #P.parts ∧ #t = #s / #P.parts + 1) := by by_contra h; exact absurd (h.1 ▸ h.2) (lt_add_one _).ne tauto theorem IsEquipartition.average_le_card_part (hP : P.IsEquipartition) (ht : t ∈ P.parts) : - s.card / P.parts.card ≤ t.card := by + #s / #P.parts ≤ #t := by rw [← P.sum_card_parts] exact Finset.EquitableOn.le hP ht theorem IsEquipartition.card_part_le_average_add_one (hP : P.IsEquipartition) (ht : t ∈ P.parts) : - t.card ≤ s.card / P.parts.card + 1 := by + #t ≤ #s / #P.parts + 1 := by rw [← P.sum_card_parts] exact Finset.EquitableOn.le_add_one hP ht theorem IsEquipartition.filter_ne_average_add_one_eq_average (hP : P.IsEquipartition) : - P.parts.filter (fun p ↦ ¬p.card = s.card / P.parts.card + 1) = - P.parts.filter (fun p ↦ p.card = s.card / P.parts.card) := by + {p ∈ P.parts | ¬#p = #s / #P.parts + 1} = {p ∈ P.parts | #p = #s / #P.parts} := by ext p simp only [mem_filter, and_congr_right_iff] exact fun hp ↦ (hP.card_part_eq_average_iff hp).symm @@ -80,75 +77,70 @@ theorem IsEquipartition.filter_ne_average_add_one_eq_average (hP : P.IsEquiparti /-- An equipartition of a finset with `n` elements into `k` parts has `n % k` parts of size `n / k + 1`. -/ theorem IsEquipartition.card_large_parts_eq_mod (hP : P.IsEquipartition) : - (P.parts.filter fun p ↦ p.card = s.card / P.parts.card + 1).card = s.card % P.parts.card := by + #{p ∈ P.parts | #p = #s / #P.parts + 1} = #s % #P.parts := by have z := P.sum_card_parts - rw [← sum_filter_add_sum_filter_not (s := P.parts) - (p := fun x ↦ x.card = s.card / P.parts.card + 1), - hP.filter_ne_average_add_one_eq_average, - sum_const_nat (m := s.card / P.parts.card + 1) (by simp), - sum_const_nat (m := s.card / P.parts.card) (by simp), - ← hP.filter_ne_average_add_one_eq_average, - mul_add, add_comm, ← add_assoc, ← add_mul, mul_one, add_comm (Finset.card _), + rw [← sum_filter_add_sum_filter_not (s := P.parts) (p := fun x ↦ #x = #s / #P.parts + 1), + hP.filter_ne_average_add_one_eq_average, sum_const_nat (m := #s / #P.parts + 1) (by simp), + sum_const_nat (m := #s / #P.parts) (by simp), ← hP.filter_ne_average_add_one_eq_average, + mul_add, add_comm, ← add_assoc, ← add_mul, mul_one, add_comm #_, filter_card_add_filter_neg_card_eq_card, add_comm] at z rw [← add_left_inj, Nat.mod_add_div, z] /-- An equipartition of a finset with `n` elements into `k` parts has `n - n % k` parts of size `n / k`. -/ theorem IsEquipartition.card_small_parts_eq_mod (hP : P.IsEquipartition) : - (P.parts.filter fun p ↦ p.card = s.card / P.parts.card).card = - P.parts.card - s.card % P.parts.card := by + #{p ∈ P.parts | #p = #s / #P.parts} = #P.parts - #s % #P.parts := by conv_rhs => arg 1 - rw [← filter_card_add_filter_neg_card_eq_card (p := fun p ↦ p.card = s.card / P.parts.card + 1)] + rw [← filter_card_add_filter_neg_card_eq_card (p := fun p ↦ #p = #s / #P.parts + 1)] rw [hP.card_large_parts_eq_mod, add_tsub_cancel_left, hP.filter_ne_average_add_one_eq_average] /-- There exists an enumeration of an equipartition's parts where larger parts map to smaller numbers and vice versa. -/ theorem IsEquipartition.exists_partsEquiv (hP : P.IsEquipartition) : - ∃ f : P.parts ≃ Fin P.parts.card, - ∀ t, t.1.card = s.card / P.parts.card + 1 ↔ f t < s.card % P.parts.card := by - let el := (P.parts.filter fun p ↦ p.card = s.card / P.parts.card + 1).equivFin - let es := (P.parts.filter fun p ↦ p.card = s.card / P.parts.card).equivFin + ∃ f : P.parts ≃ Fin #P.parts, ∀ t, #t.1 = #s / #P.parts + 1 ↔ f t < #s % #P.parts := by + let el := {p ∈ P.parts | #p = #s / #P.parts + 1}.equivFin + let es := {p ∈ P.parts | #p = #s / #P.parts}.equivFin simp_rw [mem_filter, hP.card_large_parts_eq_mod] at el simp_rw [mem_filter, hP.card_small_parts_eq_mod] at es - let sneg : { x // x ∈ P.parts ∧ ¬x.card = s.card / P.parts.card + 1 } ≃ - { x // x ∈ P.parts ∧ x.card = s.card / P.parts.card } := by + let sneg : + {x // x ∈ P.parts ∧ ¬#x = #s / #P.parts + 1} ≃ {x // x ∈ P.parts ∧ #x = #s / #P.parts} := by apply (Equiv.refl _).subtypeEquiv simp only [Equiv.refl_apply, and_congr_right_iff] exact fun _ ha ↦ by rw [hP.card_part_eq_average_iff ha, ne_eq] - replace el : { x : P.parts // x.1.card = s.card / P.parts.card + 1 } ≃ - Fin (s.card % P.parts.card) := (Equiv.Set.sep ..).symm.trans el - replace es : { x : P.parts // ¬x.1.card = s.card / P.parts.card + 1 } ≃ - Fin (P.parts.card - s.card % P.parts.card) := (Equiv.Set.sep ..).symm.trans (sneg.trans es) + replace el : { x : P.parts // #x.1 = #s / #P.parts + 1 } ≃ + Fin (#s % #P.parts) := (Equiv.Set.sep ..).symm.trans el + replace es : { x : P.parts // ¬#x.1 = #s / #P.parts + 1 } ≃ + Fin (#P.parts - #s % #P.parts) := (Equiv.Set.sep ..).symm.trans (sneg.trans es) let f := (Equiv.sumCompl _).symm.trans ((el.sumCongr es).trans finSumFinEquiv) use f.trans (finCongr (Nat.add_sub_of_le P.card_mod_card_parts_le)) intro ⟨p, _⟩ simp_rw [f, Equiv.trans_apply, Equiv.sumCongr_apply, finCongr_apply, Fin.coe_cast] - by_cases hc : p.card = s.card / P.parts.card + 1 <;> simp [hc] + by_cases hc : #p = #s / #P.parts + 1 <;> simp [hc] /-- Given a finset equipartitioned into `k` parts, its elements can be enumerated such that elements in the same part have congruent indices modulo `k`. -/ -theorem IsEquipartition.exists_partPreservingEquiv (hP : P.IsEquipartition) : ∃ f : s ≃ Fin s.card, - ∀ a b : s, P.part a = P.part b ↔ f a % P.parts.card = f b % P.parts.card := by +theorem IsEquipartition.exists_partPreservingEquiv (hP : P.IsEquipartition) : ∃ f : s ≃ Fin #s, + ∀ a b : s, P.part a = P.part b ↔ f a % #P.parts = f b % #P.parts := by obtain ⟨f, hf⟩ := P.exists_enumeration obtain ⟨g, hg⟩ := hP.exists_partsEquiv - let z := fun a ↦ P.parts.card * (f a).2 + g (f a).1 + let z := fun a ↦ #P.parts * (f a).2 + g (f a).1 have gl := fun a ↦ (g (f a).1).2 - have less : ∀ a, z a < s.card := fun a ↦ by + have less : ∀ a, z a < #s := fun a ↦ by rcases hP.card_parts_eq_average (f a).1.2 with (c | c) · calc - _ < P.parts.card * ((f a).2 + 1) := add_lt_add_left (gl a) _ - _ ≤ P.parts.card * (s.card / P.parts.card) := mul_le_mul_left' (c ▸ (f a).2.2) _ - _ ≤ P.parts.card * (s.card / P.parts.card) + s.card % P.parts.card := Nat.le_add_right .. + _ < #P.parts * ((f a).2 + 1) := add_lt_add_left (gl a) _ + _ ≤ #P.parts * (#s / #P.parts) := mul_le_mul_left' (c ▸ (f a).2.2) _ + _ ≤ #P.parts * (#s / #P.parts) + #s % #P.parts := Nat.le_add_right .. _ = _ := Nat.div_add_mod .. - · rw [← Nat.div_add_mod s.card P.parts.card] + · rw [← Nat.div_add_mod #s #P.parts] exact add_lt_add_of_le_of_lt (mul_le_mul_left' (by omega) _) ((hg (f a).1).mp c) - let z' : s → Fin s.card := fun a ↦ ⟨z a, less a⟩ + let z' : s → Fin #s := fun a ↦ ⟨z a, less a⟩ have bij : z'.Bijective := by refine (bijective_iff_injective_and_card z').mpr ⟨fun a b e ↦ ?_, by simp⟩ - simp_rw [z', z, Fin.mk.injEq, mul_comm P.parts.card] at e - haveI : NeZero P.parts.card := ⟨((Nat.zero_le _).trans_lt (gl a)).ne'⟩ - change P.parts.card.divModEquiv.symm (_, _) = P.parts.card.divModEquiv.symm (_, _) at e + simp_rw [z', z, Fin.mk.injEq, mul_comm #P.parts] at e + haveI : NeZero #P.parts := ⟨((Nat.zero_le _).trans_lt (gl a)).ne'⟩ + change (#P.parts).divModEquiv.symm (_, _) = (#P.parts).divModEquiv.symm (_, _) at e simp only [Equiv.apply_eq_iff_eq, Prod.mk.injEq] at e apply_fun f exact Sigma.ext e.2 <| (Fin.heq_ext_iff (by rw [e.2])).mpr e.1 diff --git a/Mathlib/Order/Partition/Finpartition.lean b/Mathlib/Order/Partition/Finpartition.lean index 0c1bbdebaec82e..f92a6209a1310c 100644 --- a/Mathlib/Order/Partition/Finpartition.lean +++ b/Mathlib/Order/Partition/Finpartition.lean @@ -338,7 +338,7 @@ theorem exists_le_of_le {a b : α} {P Q : Finpartition a} (h : P ≤ Q) (hb : b simp only [not_exists, not_and] at H exact H _ hc hcd -theorem card_mono {a : α} {P Q : Finpartition a} (h : P ≤ Q) : Q.parts.card ≤ P.parts.card := by +theorem card_mono {a : α} {P Q : Finpartition a} (h : P ≤ Q) : #Q.parts ≤ #P.parts := by classical have : ∀ b ∈ Q.parts, ∃ c ∈ P.parts, c ≤ b := fun b ↦ exists_le_of_le h choose f hP hf using this @@ -389,7 +389,7 @@ theorem mem_bind : b ∈ (P.bind Q).parts ↔ ∃ A hA, b ∈ (Q A hA).parts := exact ⟨⟨A, hA⟩, mem_attach _ ⟨A, hA⟩, h⟩ theorem card_bind (Q : ∀ i ∈ P.parts, Finpartition i) : - (P.bind Q).parts.card = ∑ A ∈ P.parts.attach, (Q _ A.2).parts.card := by + #(P.bind Q).parts = ∑ A ∈ P.parts.attach, #(Q _ A.2).parts := by apply card_biUnion rintro ⟨b, hb⟩ - ⟨c, hc⟩ - hbc rw [Finset.disjoint_left] @@ -414,7 +414,7 @@ def extend (P : Finpartition a) (hb : b ≠ ⊥) (hab : Disjoint a b) (hc : a not_bot_mem h := (mem_insert.1 h).elim hb.symm P.not_bot_mem theorem card_extend (P : Finpartition a) (b c : α) {hb : b ≠ ⊥} {hab : Disjoint a b} - {hc : a ⊔ b = c} : (P.extend hb hab hc).parts.card = P.parts.card + 1 := + {hc : a ⊔ b = c} : #(P.extend hb hab hc).parts = #P.parts + 1 := card_insert_of_not_mem fun h ↦ hb <| hab.symm.eq_bot_of_le <| P.le h end DistribLattice @@ -506,12 +506,12 @@ def equivSigmaParts : s ≃ Σ t : P.parts, t.1 where rw [P.part_eq_of_mem mp mf] · simp -lemma exists_enumeration : ∃ f : s ≃ Σ t : P.parts, Fin t.1.card, +lemma exists_enumeration : ∃ f : s ≃ Σ t : P.parts, Fin #t.1, ∀ a b : s, P.part a = P.part b ↔ (f a).1 = (f b).1 := by use P.equivSigmaParts.trans ((Equiv.refl _).sigmaCongr (fun t ↦ t.1.equivFin)) simp [equivSigmaParts, Equiv.sigmaCongr, Equiv.sigmaCongrLeft] -theorem sum_card_parts : ∑ i ∈ P.parts, i.card = s.card := by +theorem sum_card_parts : ∑ i ∈ P.parts, #i = #s := by convert congr_arg Finset.card P.biUnion_parts rw [card_biUnion P.supIndep.pairwiseDisjoint] rfl @@ -532,8 +532,7 @@ theorem parts_bot (s : Finset α) : (⊥ : Finpartition s).parts = s.map ⟨singleton, singleton_injective⟩ := rfl -theorem card_bot (s : Finset α) : (⊥ : Finpartition s).parts.card = s.card := - Finset.card_map _ +theorem card_bot (s : Finset α) : #(⊥ : Finpartition s).parts = #s := Finset.card_map _ theorem mem_bot_iff : t ∈ (⊥ : Finpartition s).parts ↔ ∃ a ∈ s, {a} = t := mem_map @@ -546,15 +545,15 @@ instance (s : Finset α) : OrderBot (Finpartition s) := obtain ⟨t, ht, hat⟩ := P.exists_mem ha exact ⟨t, ht, singleton_subset_iff.2 hat⟩ } -theorem card_parts_le_card : P.parts.card ≤ s.card := by +theorem card_parts_le_card : #P.parts ≤ #s := by rw [← card_bot s] exact card_mono bot_le -lemma card_mod_card_parts_le : s.card % P.parts.card ≤ P.parts.card := by - rcases P.parts.card.eq_zero_or_pos with h | h - · have h' := h - rw [Finset.card_eq_zero, parts_eq_empty_iff, bot_eq_empty, ← Finset.card_eq_zero] at h' - rw [h, h'] +lemma card_mod_card_parts_le : #s % #P.parts ≤ #P.parts := by + obtain h | h := (#P.parts).eq_zero_or_pos + · rw [h] + rw [Finset.card_eq_zero, parts_eq_empty_iff, bot_eq_empty, ← Finset.card_eq_zero] at h + rw [h] · exact (Nat.mod_lt _ h).le section Setoid @@ -564,7 +563,7 @@ variable [Fintype α] /-- A setoid over a finite type induces a finpartition of the type's elements, where the parts are the setoid's equivalence classes. -/ def ofSetoid (s : Setoid α) [DecidableRel s.r] : Finpartition (univ : Finset α) where - parts := univ.image fun a => univ.filter (s.r a) + parts := univ.image fun a ↦ ({b | s.r a b} : Finset α) supIndep := by simp only [mem_univ, forall_true_left, supIndep_iff_pairwiseDisjoint, Set.PairwiseDisjoint, Set.Pairwise, coe_image, coe_univ, Set.image_univ, Set.mem_range, ne_eq, @@ -605,7 +604,7 @@ section Atomise /-- Cuts `s` along the finsets in `F`: Two elements of `s` will be in the same part if they are in the same finsets of `F`. -/ def atomise (s : Finset α) (F : Finset (Finset α)) : Finpartition s := - ofErase (F.powerset.image fun Q ↦ s.filter fun i ↦ ∀ t ∈ F, t ∈ Q ↔ i ∈ t) + ofErase (F.powerset.image fun Q ↦ {i ∈ s | ∀ t ∈ F, t ∈ Q ↔ i ∈ t}) (Set.PairwiseDisjoint.supIndep fun x hx y hy h ↦ disjoint_left.mpr fun z hz1 hz2 ↦ h (by @@ -614,8 +613,7 @@ def atomise (s : Finset α) (F : Finset (Finset α)) : Finpartition s := obtain ⟨R, hR, rfl⟩ := hy suffices h' : Q = R by subst h' - exact of_eq_true (eq_self ( - filter (fun i ↦ ∀ (t : Finset α), t ∈ F → (t ∈ Q ↔ i ∈ t)) s)) + exact of_eq_true (eq_self {i ∈ s | ∀ t ∈ F, t ∈ Q ↔ i ∈ t}) rw [id, mem_filter] at hz1 hz2 rw [mem_powerset] at hQ hR ext i @@ -629,7 +627,7 @@ def atomise (s : Finset α) (F : Finset (Finset α)) : Finpartition s := exact s.filter_subset _ · rw [mem_sup] refine - ⟨s.filter fun i ↦ ∀ t, t ∈ F → ((t ∈ F.filter fun u ↦ a ∈ u) ↔ i ∈ t), + ⟨{i ∈ s | ∀ t ∈ F, t ∈ {u ∈ F | a ∈ u} ↔ i ∈ t}, mem_image_of_mem _ (mem_powerset.2 <| filter_subset _ _), mem_filter.2 ⟨ha, fun t ht ↦ ?_⟩⟩ rw [mem_filter] @@ -639,7 +637,7 @@ variable {F : Finset (Finset α)} theorem mem_atomise : t ∈ (atomise s F).parts ↔ - t.Nonempty ∧ ∃ Q ⊆ F, (s.filter fun i ↦ ∀ u ∈ F, u ∈ Q ↔ i ∈ u) = t := by + t.Nonempty ∧ ∃ Q ⊆ F, {i ∈ s | ∀ u ∈ F, u ∈ Q ↔ i ∈ u} = t := by simp only [atomise, ofErase, bot_eq_empty, mem_erase, mem_image, nonempty_iff_ne_empty, mem_singleton, and_comm, mem_powerset, exists_prop] @@ -648,11 +646,11 @@ theorem atomise_empty (hs : s.Nonempty) : (atomise s ∅).parts = {s} := by imp_true_iff, filter_True] exact erase_eq_of_not_mem (not_mem_singleton.2 hs.ne_empty.symm) -theorem card_atomise_le : (atomise s F).parts.card ≤ 2 ^ F.card := +theorem card_atomise_le : #(atomise s F).parts ≤ 2 ^ #F := (card_le_card <| erase_subset _ _).trans <| Finset.card_image_le.trans (card_powerset _).le theorem biUnion_filter_atomise (ht : t ∈ F) (hts : t ⊆ s) : - ((atomise s F).parts.filter fun u ↦ u ⊆ t ∧ u.Nonempty).biUnion id = t := by + {u ∈ (atomise s F).parts | u ⊆ t ∧ u.Nonempty}.biUnion id = t := by ext a refine mem_biUnion.trans ⟨fun ⟨u, hu, ha⟩ ↦ (mem_filter.1 hu).2.1 ha, fun ha ↦ ?_⟩ obtain ⟨u, hu, hau⟩ := (atomise s F).exists_mem (hts ha) @@ -662,10 +660,10 @@ theorem biUnion_filter_atomise (ht : t ∈ F) (hts : t ⊆ s) : rwa [← hb.2 _ ht, hau.2 _ ht] theorem card_filter_atomise_le_two_pow (ht : t ∈ F) : - ((atomise s F).parts.filter fun u ↦ u ⊆ t ∧ u.Nonempty).card ≤ 2 ^ (F.card - 1) := by + #{u ∈ (atomise s F).parts | u ⊆ t ∧ u.Nonempty} ≤ 2 ^ (#F - 1) := by suffices h : - ((atomise s F).parts.filter fun u ↦ u ⊆ t ∧ u.Nonempty) ⊆ - (F.erase t).powerset.image fun P ↦ s.filter fun i ↦ ∀ x ∈ F, x ∈ insert t P ↔ i ∈ x by + {u ∈ (atomise s F).parts | u ⊆ t ∧ u.Nonempty} ⊆ + (F.erase t).powerset.image fun P ↦ {i ∈ s | ∀ x ∈ F, x ∈ insert t P ↔ i ∈ x} by refine (card_le_card h).trans (card_image_le.trans ?_) rw [card_powerset, card_erase_of_mem ht] rw [subset_iff] diff --git a/Mathlib/Order/SupClosed.lean b/Mathlib/Order/SupClosed.lean index f1940a4250ed13..1e09bd963d6710 100644 --- a/Mathlib/Order/SupClosed.lean +++ b/Mathlib/Order/SupClosed.lean @@ -306,7 +306,7 @@ lemma supClosure_min : s ⊆ t → SupClosed t → supClosure s ⊆ t := supClos protected lemma Set.Finite.supClosure (hs : s.Finite) : (supClosure s).Finite := by lift s to Finset α using hs classical - refine ((s.powerset.filter Finset.Nonempty).attach.image + refine ({t ∈ s.powerset | t.Nonempty}.attach.image fun t ↦ t.1.sup' (mem_filter.1 t.2).2 id).finite_toSet.subset ?_ rintro _ ⟨t, ht, hts, rfl⟩ simp only [id_eq, coe_image, mem_image, mem_coe, mem_attach, true_and, Subtype.exists, @@ -378,7 +378,7 @@ lemma infClosure_min : s ⊆ t → InfClosed t → infClosure s ⊆ t := infClos protected lemma Set.Finite.infClosure (hs : s.Finite) : (infClosure s).Finite := by lift s to Finset α using hs classical - refine ((s.powerset.filter Finset.Nonempty).attach.image + refine ({t ∈ s.powerset | t.Nonempty}.attach.image fun t ↦ t.1.inf' (mem_filter.1 t.2).2 id).finite_toSet.subset ?_ rintro _ ⟨t, ht, hts, rfl⟩ simp only [id_eq, coe_image, mem_image, mem_coe, mem_attach, true_and, Subtype.exists, diff --git a/Mathlib/Probability/StrongLaw.lean b/Mathlib/Probability/StrongLaw.lean index fae0c8332ee2f1..4158b48564a823 100644 --- a/Mathlib/Probability/StrongLaw.lean +++ b/Mathlib/Probability/StrongLaw.lean @@ -420,8 +420,7 @@ theorem strong_law_aux1 {c : ℝ} (c_one : 1 < c) {ε : ℝ} (εpos : 0 < ε) : exact (hident j).aestronglyMeasurable_fst.memℒp_truncation · intro k _ l _ hkl exact (hindep hkl).comp (A k).measurable (A l).measurable - _ = ∑ j ∈ range (u (N - 1)), - (∑ i ∈ (range N).filter fun i => j < u i, ((u i : ℝ) ^ 2)⁻¹) * Var[Y j] := by + _ = ∑ j ∈ range (u (N - 1)), (∑ i ∈ range N with j < u i, ((u i : ℝ) ^ 2)⁻¹) * Var[Y j] := by simp_rw [mul_sum, sum_mul, sum_sigma'] refine sum_nbij' (fun p ↦ ⟨p.2, p.1⟩) (fun p ↦ ⟨p.2, p.1⟩) ?_ ?_ ?_ ?_ ?_ · simp only [mem_sigma, mem_range, filter_congr_decidable, mem_filter, and_imp, diff --git a/Mathlib/RingTheory/GradedAlgebra/Radical.lean b/Mathlib/RingTheory/GradedAlgebra/Radical.lean index 976b7748f818df..72e6ad6f51e2be 100644 --- a/Mathlib/RingTheory/GradedAlgebra/Radical.lean +++ b/Mathlib/RingTheory/GradedAlgebra/Radical.lean @@ -66,10 +66,10 @@ theorem Ideal.IsHomogeneous.isPrime_of_homogeneous_mem_or_mem {I : Ideal A} (hI This is a contradiction, because both `proj (max₁ + max₂) (x * y) ∈ I` and the sum on the right hand side is in `I` however `proj max₁ x * proj max₂ y` is not in `I`. -/ - set set₁ := (decompose 𝒜 x).support.filter (fun i => proj 𝒜 i x ∉ I) with set₁_eq - set set₂ := (decompose 𝒜 y).support.filter (fun i => proj 𝒜 i y ∉ I) with set₂_eq + set set₁ := {i ∈ (decompose 𝒜 x).support | proj 𝒜 i x ∉ I} with set₁_eq + set set₂ := {i ∈ (decompose 𝒜 y).support | proj 𝒜 i y ∉ I} with set₂_eq have nonempty : - ∀ x : A, x ∉ I → ((decompose 𝒜 x).support.filter (fun i => proj 𝒜 i x ∉ I)).Nonempty := by + ∀ x : A, x ∉ I → {i ∈ (decompose 𝒜 x).support | proj 𝒜 i x ∉ I}.Nonempty := by intro x hx rw [filter_nonempty_iff] contrapose! hx @@ -83,8 +83,8 @@ theorem Ideal.IsHomogeneous.isPrime_of_homogeneous_mem_or_mem {I : Ideal A} (hI replace hxy : proj 𝒜 (max₁ + max₂) (x * y) ∈ I := hI _ hxy have mem_I : proj 𝒜 max₁ x * proj 𝒜 max₂ y ∈ I := by set antidiag := - ((decompose 𝒜 x).support ×ˢ (decompose 𝒜 y).support).filter (fun z : ι × ι => - z.1 + z.2 = max₁ + max₂) with ha + {z ∈ (decompose 𝒜 x).support ×ˢ (decompose 𝒜 y).support | z.1 + z.2 = max₁ + max₂} + with ha have mem_antidiag : (max₁, max₂) ∈ antidiag := by simp only [antidiag, add_sum_erase, mem_filter, mem_product] exact ⟨⟨mem_of_mem_filter _ mem_max₁, mem_of_mem_filter _ mem_max₂⟩, trivial⟩ diff --git a/Mathlib/RingTheory/IntegralDomain.lean b/Mathlib/RingTheory/IntegralDomain.lean index 7126ac0db11453..b292e8b1c1ffc0 100644 --- a/Mathlib/RingTheory/IntegralDomain.lean +++ b/Mathlib/RingTheory/IntegralDomain.lean @@ -109,14 +109,12 @@ end Ring variable [CommRing R] [IsDomain R] [Group G] --- Porting note: Finset doesn't seem to have `{g ∈ univ | g^n = g₀}` notation anymore, --- so we have to use `Finset.filter` instead theorem card_nthRoots_subgroup_units [Fintype G] [DecidableEq G] (f : G →* R) (hf : Injective f) {n : ℕ} (hn : 0 < n) (g₀ : G) : - Finset.card (Finset.univ.filter (fun g ↦ g^n = g₀)) ≤ Multiset.card (nthRoots n (f g₀)) := by + #{g | g ^ n = g₀} ≤ Multiset.card (nthRoots n (f g₀)) := by haveI : DecidableEq R := Classical.decEq _ calc - _ ≤ (nthRoots n (f g₀)).toFinset.card := card_le_card_of_injOn f (by aesop) hf.injOn + _ ≤ #(nthRoots n (f g₀)).toFinset := card_le_card_of_injOn f (by aesop) hf.injOn _ ≤ _ := (nthRoots n (f g₀)).toFinset_card_le /-- A finite subgroup of the unit group of an integral domain is cyclic. -/ @@ -194,12 +192,11 @@ theorem sum_hom_units_eq_zero (f : G →* R) (hf : f ≠ 1) : ∑ g : G, f g = 0 eq_comm] at hn replace hx1 : (x.val : R) - 1 ≠ 0 := -- Porting note: was `(x : R)` fun h => hx1 (Subtype.eq (Units.ext (sub_eq_zero.1 h))) - let c := (univ.filter fun g => f.toHomUnits g = 1).card + let c := #{g | f.toHomUnits g = 1} calc ∑ g : G, f g = ∑ g : G, (f.toHomUnits g : R) := rfl - _ = ∑ u ∈ univ.image f.toHomUnits, - (univ.filter fun g => f.toHomUnits g = u).card • (u : R) := - (sum_comp ((↑) : Rˣ → R) f.toHomUnits) + _ = ∑ u ∈ univ.image f.toHomUnits, #{g | f.toHomUnits g = u} • (u : R) := + sum_comp ((↑) : Rˣ → R) f.toHomUnits _ = ∑ u ∈ univ.image f.toHomUnits, c • (u : R) := (sum_congr rfl fun u hu => congr_arg₂ _ ?_ rfl) -- remaining goal 1, proven below @@ -211,7 +208,7 @@ theorem sum_hom_units_eq_zero (f : G →* R) (hf : f ≠ 1) : ∑ g : G, f g = 0 -- remaining goal 2, proven below _ = (0 : R) := smul_zero _ · -- remaining goal 1 - show (univ.filter fun g : G => f.toHomUnits g = u).card = c + show #{g : G | f.toHomUnits g = u} = c apply MonoidHom.card_fiber_eq_of_mem_range f.toHomUnits · simpa only [mem_image, mem_univ, true_and, Set.mem_range] using hu · exact ⟨1, f.toHomUnits.map_one⟩ diff --git a/Mathlib/RingTheory/MvPolynomial/Homogeneous.lean b/Mathlib/RingTheory/MvPolynomial/Homogeneous.lean index a5847ddf565cd7..64a1bbcd81c419 100644 --- a/Mathlib/RingTheory/MvPolynomial/Homogeneous.lean +++ b/Mathlib/RingTheory/MvPolynomial/Homogeneous.lean @@ -457,8 +457,7 @@ theorem coeff_homogeneousComponent (d : σ →₀ ℕ) : convert coeff_weightedHomogeneousComponent n φ d theorem homogeneousComponent_apply : - homogeneousComponent n φ = - ∑ d ∈ φ.support.filter fun d => d.degree = n, monomial d (coeff d φ) := by + homogeneousComponent n φ = ∑ d ∈ φ.support with d.degree = n, monomial d (coeff d φ) := by simp_rw [degree_eq_weight_one] convert weightedHomogeneousComponent_apply n φ diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric/Defs.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric/Defs.lean index 2b4ebd09a9114e..093790197e63a5 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric/Defs.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric/Defs.lean @@ -198,7 +198,7 @@ theorem aeval_esymm_eq_multiset_esymm [Algebra R S] (n : ℕ) (f : σ → S) : /-- We can define `esymm σ R n` by summing over a subtype instead of over `powerset_len`. -/ theorem esymm_eq_sum_subtype (n : ℕ) : - esymm σ R n = ∑ t : { s : Finset σ // s.card = n }, ∏ i ∈ (t : Finset σ), X i := + esymm σ R n = ∑ t : {s : Finset σ // #s = n}, ∏ i ∈ (t : Finset σ), X i := sum_subtype _ (fun _ => mem_powersetCard_univ) _ /-- We can define `esymm σ R n` as a sum over explicit monomials -/ diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric/FundamentalTheorem.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric/FundamentalTheorem.lean index 8a1f28577910ec..b822840d8c9ce8 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric/FundamentalTheorem.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric/FundamentalTheorem.lean @@ -60,7 +60,7 @@ section accumulate /-- The `j`th entry of `accumulate n m t` is the sum of `t i` over all `i ≥ j`. -/ @[simps] def accumulate (n m : ℕ) : (Fin n → ℕ) →+ (Fin m → ℕ) where - toFun t j := ∑ i in univ.filter (fun i : Fin n ↦ (j : ℕ) ≤ i), t i + toFun t j := ∑ i : Fin n with j.val ≤ i.val, t i map_zero' := funext <| fun _ ↦ sum_eq_zero <| fun _ _ ↦ rfl map_add' t₁ t₂ := funext <| fun j ↦ by dsimp only; exact sum_add_distrib @@ -178,7 +178,7 @@ private lemma supDegree_monic_esymm [Nontrivial R] {i : ℕ} (him : i < m) : supDegree_single_ne_zero _ one_ne_zero, leadingCoeff_single toLex.injective] at this · exact mem_powersetCard.2 ⟨subset_univ _, Fin.card_Iic _⟩ intro t ht hne - have ht' : t.card = (Iic (⟨i, him⟩ : Fin m)).card := by + have ht' : #t = #(Iic (⟨i, him⟩ : Fin m)) := by rw [(mem_powersetCard.1 ht).2, Fin.card_Iic] simp_rw [← single_eq_monomial, supDegree_single_ne_zero _ one_ne_zero, ← Finsupp.indicator_eq_sum_single] diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric/NewtonIdentities.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric/NewtonIdentities.lean index 3aa09ead7b4abe..802d25319935b3 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric/NewtonIdentities.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric/NewtonIdentities.lean @@ -87,15 +87,15 @@ private theorem pairMap_involutive : (pairMap σ).Involutive := by variable [Fintype σ] private def pairs (k : ℕ) : Finset (Finset σ × σ) := - univ.filter (fun t ↦ card t.fst ≤ k ∧ (card t.fst = k → t.snd ∈ t.fst)) + {t | #t.1 ≤ k ∧ (#t.1 = k → t.snd ∈ t.fst)} @[simp] private lemma mem_pairs (k : ℕ) (t : Finset σ × σ) : - t ∈ pairs σ k ↔ card t.fst ≤ k ∧ (card t.fst = k → t.snd ∈ t.fst) := by + t ∈ pairs σ k ↔ #t.1 ≤ k ∧ (#t.1 = k → t.snd ∈ t.fst) := by simp [pairs] private def weight (k : ℕ) (t : Finset σ × σ) : MvPolynomial σ R := - (-1) ^ card t.fst * ((∏ a ∈ t.fst, X a) * X t.snd ^ (k - card t.fst)) + (-1) ^ #t.1 * ((∏ a ∈ t.fst, X a) * X t.snd ^ (k - #t.1)) private theorem pairMap_mem_pairs {k : ℕ} (t : Finset σ × σ) (h : t ∈ pairs σ k) : pairMap σ t ∈ pairs σ k := by @@ -126,16 +126,16 @@ private theorem weight_add_weight_pairMap {k : ℕ} (t : Finset σ × σ) (h : t mul_assoc (∏ a ∈ erase t.fst t.snd, X a), card_erase_of_mem h1] nth_rewrite 1 [← pow_one (X t.snd)] simp only [← pow_add, add_comm] - have h3 : 1 ≤ card t.fst := lt_iff_add_one_le.mp (card_pos.mpr ⟨t.snd, h1⟩) - rw [← tsub_tsub_assoc h.left h3, ← neg_neg ((-1 : MvPolynomial σ R) ^ (card t.fst - 1)), - h2 (card t.fst - 1), Nat.sub_add_cancel h3] + have h3 : 1 ≤ #t.1 := lt_iff_add_one_le.mp (card_pos.mpr ⟨t.snd, h1⟩) + rw [← tsub_tsub_assoc h.left h3, ← neg_neg ((-1 : MvPolynomial σ R) ^ (#t.1 - 1)), + h2 (#t.1 - 1), Nat.sub_add_cancel h3] simp · rw [pairMap_of_snd_nmem_fst σ h1] simp only [mul_comm, mul_assoc (∏ a ∈ t.fst, X a), card_cons, prod_cons] nth_rewrite 2 [← pow_one (X t.snd)] simp only [← pow_add, ← Nat.add_sub_assoc (Nat.lt_of_le_of_ne h.left (mt h.right h1)), add_comm, Nat.succ_eq_add_one, Nat.add_sub_add_right] - rw [← neg_neg ((-1 : MvPolynomial σ R) ^ card t.fst), h2] + rw [← neg_neg ((-1 : MvPolynomial σ R) ^ #t.1), h2] simp private theorem weight_sum (k : ℕ) : ∑ t ∈ pairs σ k, weight σ R k t = 0 := @@ -145,45 +145,40 @@ private theorem weight_sum (k : ℕ) : ∑ t ∈ pairs σ k, weight σ R k t = 0 private theorem sum_filter_pairs_eq_sum_powersetCard_sum (k : ℕ) (f : Finset σ × σ → MvPolynomial σ R) : - (∑ t ∈ filter (fun t ↦ card t.fst = k) (pairs σ k), f t) = - ∑ A ∈ powersetCard k univ, (∑ j ∈ A, f (A, j)) := by + ∑ t ∈ pairs σ k with #t.1 = k, f t = ∑ A ∈ powersetCard k univ, ∑ j ∈ A, f (A, j) := by apply sum_finset_product aesop private theorem sum_filter_pairs_eq_sum_powersetCard_mem_filter_antidiagonal_sum (k : ℕ) (a : ℕ × ℕ) - (ha : a ∈ (antidiagonal k).filter (fun a ↦ a.fst < k)) (f : Finset σ × σ → MvPolynomial σ R) : - (∑ t ∈ filter (fun t ↦ card t.fst = a.fst) (pairs σ k), f t) = - ∑ A ∈ powersetCard a.fst univ, (∑ j, f (A, j)) := by + (ha : a ∈ {a ∈ antidiagonal k | a.fst < k}) (f : Finset σ × σ → MvPolynomial σ R) : + ∑ t ∈ pairs σ k with #t.1 = a.1, f t = ∑ A ∈ powersetCard a.1 univ, ∑ j, f (A, j) := by apply sum_finset_product simp only [mem_filter, mem_powersetCard_univ, mem_univ, and_true, and_iff_right_iff_imp] rintro p hp - have : card p.fst ≤ k := by apply le_of_lt; aesop + have : #p.fst ≤ k := by apply le_of_lt; aesop aesop private lemma filter_pairs_lt (k : ℕ) : - (pairs σ k).filter (fun (s, _) ↦ s.card < k) = + (pairs σ k).filter (fun (s, _) ↦ #s < k) = (range k).disjiUnion (powersetCard · univ) ((pairwise_disjoint_powersetCard _).set_pairwise _) ×ˢ univ := by ext; aesop (add unsafe le_of_lt) private theorem sum_filter_pairs_eq_sum_filter_antidiagonal_powersetCard_sum (k : ℕ) (f : Finset σ × σ → MvPolynomial σ R) : - ∑ t ∈ (pairs σ k).filter fun t ↦ card t.fst < k, f t = - ∑ a ∈ (antidiagonal k).filter fun a ↦ a.fst < k, - ∑ A ∈ powersetCard a.fst univ, ∑ j, f (A, j) := by + ∑ t ∈ pairs σ k with #t.1 < k, f t = + ∑ a ∈ antidiagonal k with a.fst < k, ∑ A ∈ powersetCard a.fst univ, ∑ j, f (A, j) := by rw [filter_pairs_lt, sum_product, sum_disjiUnion] refine sum_nbij' (fun n ↦ (n, k - n)) Prod.fst ?_ ?_ ?_ ?_ ?_ <;> simp (config := { contextual := true }) [@eq_comm _ _ k, Nat.add_sub_cancel', le_of_lt] private theorem disjoint_filter_pairs_lt_filter_pairs_eq (k : ℕ) : - Disjoint (filter (fun t ↦ card t.fst < k) (pairs σ k)) - (filter (fun t ↦ card t.fst = k) (pairs σ k)) := by + Disjoint {t ∈ pairs σ k | #t.1 < k} {t ∈ pairs σ k | #t.1 = k} := by rw [disjoint_filter] exact fun _ _ h1 h2 ↦ lt_irrefl _ (h2.symm.subst h1) private theorem disjUnion_filter_pairs_eq_pairs (k : ℕ) : - disjUnion (filter (fun t ↦ card t.fst < k) (pairs σ k)) - (filter (fun t ↦ card t.fst = k) (pairs σ k)) (disjoint_filter_pairs_lt_filter_pairs_eq σ k) = - pairs σ k := by + disjUnion {t ∈ pairs σ k | #t.1 < k} {t ∈ pairs σ k | #t.1 = k} + (disjoint_filter_pairs_lt_filter_pairs_eq σ k) = pairs σ k := by simp only [disjUnion_eq_union, Finset.ext_iff, pairs, filter_filter, mem_filter] intro a rw [← filter_or, mem_filter] @@ -200,7 +195,7 @@ private theorem esymm_summand_to_weight (k : ℕ) (A : Finset σ) (h : A ∈ pow simp [weight, mem_powersetCard_univ.mp h, mul_assoc] private theorem esymm_to_weight [DecidableEq σ] (k : ℕ) : k * esymm σ R k = - (-1) ^ k * ∑ t ∈ filter (fun t ↦ card t.fst = k) (pairs σ k), weight σ R k t := by + (-1) ^ k * ∑ t ∈ pairs σ k with #t.1 = k, weight σ R k t := by rw [esymm, sum_filter_pairs_eq_sum_powersetCard_sum σ R k (fun t ↦ weight σ R k t), sum_congr rfl (esymm_summand_to_weight σ R k), mul_comm (k : MvPolynomial σ R) ((-1) ^ k), ← mul_sum, ← mul_assoc, ← mul_assoc, ← pow_add, Even.neg_one_pow ⟨k, rfl⟩, one_mul] @@ -216,9 +211,8 @@ private theorem esymm_mul_psum_summand_to_weight (k : ℕ) (a : ℕ × ℕ) (ha rw [mem_powersetCard_univ.mp hs, ← mem_antidiagonal.mp ha, add_sub_self_left] private theorem esymm_mul_psum_to_weight [DecidableEq σ] (k : ℕ) : - ∑ a ∈ (antidiagonal k).filter (fun a ↦ a.fst < k), - (-1) ^ a.fst * esymm σ R a.fst * psum σ R a.snd = - ∑ t ∈ filter (fun t ↦ card t.fst < k) (pairs σ k), weight σ R k t := by + ∑ a ∈ antidiagonal k with a.fst < k, (-1) ^ a.fst * esymm σ R a.fst * psum σ R a.snd = + ∑ t ∈ pairs σ k with #t.1 < k, weight σ R k t := by rw [← sum_congr rfl (fun a ha ↦ esymm_mul_psum_summand_to_weight σ R k a (mem_filter.mp ha).left), sum_filter_pairs_eq_sum_filter_antidiagonal_powersetCard_sum σ R k] @@ -228,9 +222,9 @@ variable (σ : Type*) [Fintype σ] (R : Type*) [CommRing R] /-- **Newton's identities** give a recurrence relation for the kth elementary symmetric polynomial in terms of lower degree elementary symmetric polynomials and power sums. -/ -theorem mul_esymm_eq_sum (k : ℕ) : k * esymm σ R k = - (-1) ^ (k + 1) * ∑ a ∈ (antidiagonal k).filter (fun a ↦ a.fst < k), - (-1) ^ a.fst * esymm σ R a.fst * psum σ R a.snd := by +theorem mul_esymm_eq_sum (k : ℕ) : + k * esymm σ R k = (-1) ^ (k + 1) * + ∑ a ∈ antidiagonal k with a.1 < k, (-1) ^ a.1 * esymm σ R a.1 * psum σ R a.2 := by classical rw [NewtonIdentities.esymm_to_weight σ R k, NewtonIdentities.esymm_mul_psum_to_weight σ R k, eq_comm, ← sub_eq_zero, sub_eq_add_neg, neg_mul_eq_neg_mul, @@ -252,23 +246,22 @@ theorem sum_antidiagonal_card_esymm_psum_eq_zero : /-- A version of Newton's identities which may be more useful in the case that we know the values of the elementary symmetric polynomials and would like to calculate the values of the power sums. -/ -theorem psum_eq_mul_esymm_sub_sum (k : ℕ) (h : 0 < k) : psum σ R k = - (-1) ^ (k + 1) * k * esymm σ R k - - ∑ a ∈ (antidiagonal k).filter (fun a ↦ a.fst ∈ Set.Ioo 0 k), - (-1) ^ a.fst * esymm σ R a.fst * psum σ R a.snd := by +theorem psum_eq_mul_esymm_sub_sum (k : ℕ) (h : 0 < k) : + psum σ R k = (-1) ^ (k + 1) * k * esymm σ R k - + ∑ a ∈ antidiagonal k with a.1 ∈ Set.Ioo 0 k, (-1) ^ a.fst * esymm σ R a.1 * psum σ R a.2 := by simp only [Set.Ioo, Set.mem_setOf_eq, and_comm] have hesymm := mul_esymm_eq_sum σ R k - rw [← (sum_filter_add_sum_filter_not ((antidiagonal k).filter (fun a ↦ a.fst < k)) + rw [← (sum_filter_add_sum_filter_not {a ∈ antidiagonal k | a.fst < k} (fun a ↦ 0 < a.fst) (fun a ↦ (-1) ^ a.fst * esymm σ R a.fst * psum σ R a.snd))] at hesymm have sub_both_sides := congrArg (· - (-1 : MvPolynomial σ R) ^ (k + 1) * - ∑ a ∈ ((antidiagonal k).filter (fun a ↦ a.fst < k)).filter (fun a ↦ 0 < a.fst), + ∑ a ∈ {a ∈ antidiagonal k | a.fst < k} with 0 < a.fst, (-1) ^ a.fst * esymm σ R a.fst * psum σ R a.snd) hesymm simp only [left_distrib, add_sub_cancel_left] at sub_both_sides have sub_both_sides := congrArg ((-1 : MvPolynomial σ R) ^ (k + 1) * ·) sub_both_sides simp only [mul_sub_left_distrib, ← mul_assoc, ← pow_add, Even.neg_one_pow ⟨k + 1, rfl⟩, one_mul, not_le, lt_one_iff, filter_filter (fun a : ℕ × ℕ ↦ a.fst < k) (fun a ↦ ¬0 < a.fst)] at sub_both_sides - have : filter (fun a ↦ a.fst < k ∧ ¬0 < a.fst) (antidiagonal k) = {(0, k)} := by + have : {a ∈ antidiagonal k | a.fst < k ∧ ¬0 < a.fst} = {(0, k)} := by ext a rw [mem_filter, mem_antidiagonal, mem_singleton] refine ⟨?_, by rintro rfl; omega⟩ diff --git a/Mathlib/RingTheory/MvPolynomial/WeightedHomogeneous.lean b/Mathlib/RingTheory/MvPolynomial/WeightedHomogeneous.lean index d580a2fd428fcd..18bd357e9e1170 100644 --- a/Mathlib/RingTheory/MvPolynomial/WeightedHomogeneous.lean +++ b/Mathlib/RingTheory/MvPolynomial/WeightedHomogeneous.lean @@ -309,7 +309,7 @@ theorem coeff_weightedHomogeneousComponent [DecidableEq M] (d : σ →₀ ℕ) : theorem weightedHomogeneousComponent_apply [DecidableEq M] : weightedHomogeneousComponent w n φ = - ∑ d ∈ φ.support.filter fun d => weight w d = n, monomial d (coeff d φ) := + ∑ d ∈ φ.support with weight w d = n, monomial d (coeff d φ) := letI := Classical.decEq M Finsupp.filter_eq_sum (fun d : σ →₀ ℕ => weight w d = n) φ |>.trans <| by convert rfl diff --git a/Mathlib/RingTheory/MvPowerSeries/Basic.lean b/Mathlib/RingTheory/MvPowerSeries/Basic.lean index 3e649d881d0c8a..ab05350f3f87ac 100644 --- a/Mathlib/RingTheory/MvPowerSeries/Basic.lean +++ b/Mathlib/RingTheory/MvPowerSeries/Basic.lean @@ -671,7 +671,7 @@ theorem coeff_eq_zero_of_constantCoeff_nilpotent apply sum_eq_zero intro k hk rw [mem_finsuppAntidiag] at hk - set s := (range n).filter fun i ↦ k i = 0 with hs_def + set s := {i ∈ range n | k i = 0} with hs_def have hs : s ⊆ range n := filter_subset _ _ have hs' (i : ℕ) (hi : i ∈ s) : coeff R (k i) f = constantCoeff σ R f := by simp only [hs_def, mem_filter] at hi @@ -682,10 +682,10 @@ theorem coeff_eq_zero_of_constantCoeff_nilpotent rw [← prod_sdiff (s₁ := s) (filter_subset _ _)] apply mul_eq_zero_of_right rw [prod_congr rfl hs', prod_const] - suffices m ≤ s.card by + suffices m ≤ #s by obtain ⟨m', hm'⟩ := Nat.exists_eq_add_of_le this rw [hm', pow_add, hf, MulZeroClass.zero_mul] - rw [← Nat.add_le_add_iff_right, add_comm s.card, + rw [← Nat.add_le_add_iff_right, add_comm #s, Finset.card_sdiff_add_card_eq_card (filter_subset _ _), card_range] apply le_trans _ hn simp only [add_comm m, Nat.add_le_add_iff_right, ← hk.1, diff --git a/Mathlib/RingTheory/Polynomial/Vieta.lean b/Mathlib/RingTheory/Polynomial/Vieta.lean index 6ad52b3eb26ac9..6df36b563f3c0d 100644 --- a/Mathlib/RingTheory/Polynomial/Vieta.lean +++ b/Mathlib/RingTheory/Polynomial/Vieta.lean @@ -22,13 +22,10 @@ we derive `Polynomial.coeff_eq_esymm_roots_of_card`, the relationship between th the roots of `p` for a polynomial `p` that splits (i.e. having as many roots as its degree). -/ - -open Polynomial +open Finset Polynomial namespace Multiset -open Polynomial - section Semiring variable {R : Type*} [CommSemiring R] @@ -70,8 +67,8 @@ theorem prod_X_add_C_coeff' {σ} (s : Multiset σ) (r : σ → R) {k : ℕ} (h : (s.map fun i => X + C (r i)).prod.coeff k = (s.map r).esymm (Multiset.card s - k) := by erw [← map_map (fun r => X + C r) r, prod_X_add_C_coeff] <;> rw [s.card_map r]; assumption -theorem _root_.Finset.prod_X_add_C_coeff {σ} (s : Finset σ) (r : σ → R) {k : ℕ} (h : k ≤ s.card) : - (∏ i ∈ s, (X + C (r i))).coeff k = ∑ t ∈ s.powersetCard (s.card - k), ∏ i ∈ t, r i := by +theorem _root_.Finset.prod_X_add_C_coeff {σ} (s : Finset σ) (r : σ → R) {k : ℕ} (h : k ≤ #s) : + (∏ i ∈ s, (X + C (r i))).coeff k = ∑ t ∈ s.powersetCard (#s - k), ∏ i ∈ t, r i := by rw [Finset.prod, prod_X_add_C_coeff' _ r h, Finset.esymm_map_val] rfl diff --git a/Mathlib/RingTheory/Prime.lean b/Mathlib/RingTheory/Prime.lean index eba40fbe15828e..6fb8a1e0598db4 100644 --- a/Mathlib/RingTheory/Prime.lean +++ b/Mathlib/RingTheory/Prime.lean @@ -50,7 +50,7 @@ theorem mul_eq_mul_prime_pow {x y a p : R} {n : ℕ} (hp : Prime p) (hx : x * y rcases mul_eq_mul_prime_prod (fun _ _ ↦ hp) (show x * y = a * (range n).prod fun _ ↦ p by simpa) with ⟨t, u, b, c, htus, htu, rfl, rfl, rfl⟩ - exact ⟨t.card, u.card, b, c, by rw [← card_union_of_disjoint htu, htus, card_range], by simp⟩ + exact ⟨#t, #u, b, c, by rw [← card_union_of_disjoint htu, htus, card_range], by simp⟩ end CancelCommMonoidWithZero diff --git a/Mathlib/RingTheory/RootsOfUnity/Basic.lean b/Mathlib/RingTheory/RootsOfUnity/Basic.lean index a0b4d08c4b85dd..b90cc6d9efb531 100644 --- a/Mathlib/RingTheory/RootsOfUnity/Basic.lean +++ b/Mathlib/RingTheory/RootsOfUnity/Basic.lean @@ -856,14 +856,14 @@ theorem nthRoots_one_nodup {ζ : R} {n : ℕ} (h : IsPrimitiveRoot ζ n) : @[simp] theorem card_nthRootsFinset {ζ : R} {n : ℕ} (h : IsPrimitiveRoot ζ n) : - (nthRootsFinset n R).card = n := by + #(nthRootsFinset n R) = n := by rw [nthRootsFinset, ← Multiset.toFinset_eq (nthRoots_one_nodup h), card_mk, h.card_nthRoots_one] open scoped Nat /-- If an integral domain has a primitive `k`-th root of unity, then it has `φ k` of them. -/ theorem card_primitiveRoots {ζ : R} {k : ℕ} (h : IsPrimitiveRoot ζ k) : - (primitiveRoots k R).card = φ k := by + #(primitiveRoots k R) = φ k := by by_cases h0 : k = 0 · simp [h0] symm diff --git a/Mathlib/RingTheory/RootsOfUnity/Lemmas.lean b/Mathlib/RingTheory/RootsOfUnity/Lemmas.lean index 33d3afe29785af..75a39566c1fe5e 100644 --- a/Mathlib/RingTheory/RootsOfUnity/Lemmas.lean +++ b/Mathlib/RingTheory/RootsOfUnity/Lemmas.lean @@ -71,7 +71,7 @@ lemma self_sub_one_pow_dvd_order {k n : ℕ} (hn : k < n) {μ : R} (hμ : IsPrim rw [← this, mul_assoc, mul_assoc] congr 1 conv => enter [2, 2, 2]; rw [← card_range k] - rw [← prod_range_mul_prod_Ico _ (Nat.le_add_left k m), mul_comm _ (_ ^ card _), ← mul_assoc, + rw [← prod_range_mul_prod_Ico _ (Nat.le_add_left k m), mul_comm _ (_ ^ #_), ← mul_assoc, prod_mul_pow_card] conv => enter [2, 1, 2, j]; rw [← (Zdef _).2] diff --git a/Mathlib/Tactic/Positivity/Finset.lean b/Mathlib/Tactic/Positivity/Finset.lean index bef478ebd38c63..fdfd20e10b5548 100644 --- a/Mathlib/Tactic/Positivity/Finset.lean +++ b/Mathlib/Tactic/Positivity/Finset.lean @@ -20,7 +20,7 @@ namespace Mathlib.Meta.Positivity open Qq Lean Meta Finset -/-- Extension for `Finset.card`. `s.card` is positive if `s` is nonempty. +/-- Extension for `Finset.card`. `#s` is positive if `s` is nonempty. It calls `Mathlib.Meta.proveFinsetNonempty` to attempt proving that the finset is nonempty. -/ @[positivity Finset.card _] @@ -42,7 +42,7 @@ def evalFintypeCard : PositivityExt where eval {u α} _ _ e := do return .positive q(@Fintype.card_pos $β $instβ $instβno) | _ => throwError "not Fintype.card" -/-- Extension for `Finset.dens`. `s.card` is positive if `s` is nonempty. +/-- Extension for `Finset.dens`. `s.dens` is positive if `s` is nonempty. It calls `Mathlib.Meta.proveFinsetNonempty` to attempt proving that the finset is nonempty. -/ @[positivity Finset.dens _] @@ -91,8 +91,8 @@ def evalFinsetSum : PositivityExt where eval {u α} zα pα e := do variable {α : Type*} {s : Finset α} -example : 0 ≤ s.card := by positivity -example (hs : s.Nonempty) : 0 < s.card := by positivity +example : 0 ≤ #s := by positivity +example (hs : s.Nonempty) : 0 < #s := by positivity variable [Fintype α] @@ -100,13 +100,13 @@ example : 0 ≤ Fintype.card α := by positivity example : 0 ≤ dens s := by positivity example (hs : s.Nonempty) : 0 < dens s := by positivity example (hs : s.Nonempty) : dens s ≠ 0 := by positivity -example [Nonempty α] : 0 < (univ : Finset α).card := by positivity +example [Nonempty α] : 0 < #(univ : Finset α) := by positivity example [Nonempty α] : 0 < Fintype.card α := by positivity example [Nonempty α] : 0 < dens (univ : Finset α) := by positivity example [Nonempty α] : dens (univ : Finset α) ≠ 0 := by positivity example {G : Type*} {A : Finset G} : - let f := fun _ : G ↦ 1; (∀ s, f s ^ 2 = 1) → 0 ≤ A.card := by + let f := fun _ : G ↦ 1; (∀ s, f s ^ 2 = 1) → 0 ≤ #A := by intros positivity -- Should succeed despite failing to prove `A` is nonempty. diff --git a/Mathlib/Topology/Algebra/InfiniteSum/Constructions.lean b/Mathlib/Topology/Algebra/InfiniteSum/Constructions.lean index b47e95b56f4adf..47d09ed95c689d 100644 --- a/Mathlib/Topology/Algebra/InfiniteSum/Constructions.lean +++ b/Mathlib/Topology/Algebra/InfiniteSum/Constructions.lean @@ -86,7 +86,7 @@ theorem HasProd.sigma {γ : β → Type*} {f : (Σ b : β, γ b) → α} {g : β use u.image Sigma.fst, trivial intro bs hbs simp only [Set.mem_preimage, Finset.le_iff_subset] at hu - have : Tendsto (fun t : Finset (Σb, γ b) ↦ ∏ p ∈ t.filter fun p ↦ p.1 ∈ bs, f p) atTop + have : Tendsto (fun t : Finset (Σb, γ b) ↦ ∏ p ∈ t with p.1 ∈ bs, f p) atTop (𝓝 <| ∏ b ∈ bs, g b) := by simp only [← sigma_preimage_mk, prod_sigma] refine tendsto_finset_prod _ fun b _ ↦ ?_ @@ -161,17 +161,17 @@ theorem HasProd.of_sigma {γ : β → Type*} {f : (Σ b : β, γ b) → α} {g : obtain ⟨t0, st0, ht0⟩ : ∃ t0, ∏ i ∈ t0, g i ∈ v ∧ s.image Sigma.fst ⊆ t0 := by have A : ∀ᶠ t0 in (atTop : Filter (Finset β)), ∏ i ∈ t0, g i ∈ v := hg (v_open.mem_nhds hv) exact (A.and (Ici_mem_atTop _)).exists - have L : Tendsto (fun t : Finset (Σb, γ b) ↦ ∏ p ∈ t.filter fun p ↦ p.1 ∈ t0, f p) atTop + have L : Tendsto (fun t : Finset (Σb, γ b) ↦ ∏ p ∈ t with p.1 ∈ t0, f p) atTop (𝓝 <| ∏ b ∈ t0, g b) := by simp only [← sigma_preimage_mk, prod_sigma] refine tendsto_finset_prod _ fun b _ ↦ ?_ change Tendsto (fun t ↦ (fun t ↦ ∏ s ∈ t, f ⟨b, s⟩) (preimage t (Sigma.mk b) _)) atTop (𝓝 (g b)) exact (hf b).comp (tendsto_finset_preimage_atTop_atTop (sigma_mk_injective)) - have : ∃ t, ∏ p ∈ t.filter (fun p ↦ p.1 ∈ t0), f p ∈ v ∧ s ⊆ t := + have : ∃ t, ∏ p ∈ t with p.1 ∈ t0, f p ∈ v ∧ s ⊆ t := ((Tendsto.eventually_mem L (v_open.mem_nhds st0)).and (Ici_mem_atTop _)).exists obtain ⟨t, tv, st⟩ := this - refine ⟨t.filter (fun p ↦ p.1 ∈ t0), fun x hx ↦ ?_, vu tv⟩ + refine ⟨{p ∈ t | p.1 ∈ t0}, fun x hx ↦ ?_, vu tv⟩ simpa only [mem_filter, st hx, true_and] using ht0 (mem_image_of_mem Sigma.fst hx) variable [CompleteSpace α] diff --git a/Mathlib/Topology/Algebra/InfiniteSum/Group.lean b/Mathlib/Topology/Algebra/InfiniteSum/Group.lean index 9532cd4e53ab92..c4029fed54fcab 100644 --- a/Mathlib/Topology/Algebra/InfiniteSum/Group.lean +++ b/Mathlib/Topology/Algebra/InfiniteSum/Group.lean @@ -241,9 +241,9 @@ theorem Multipliable.multipliable_of_eq_one_or_self (hf : Multipliable f) exact multipliable_iff_vanishing.2 fun e he ↦ let ⟨s, hs⟩ := multipliable_iff_vanishing.1 hf e he ⟨s, fun t ht ↦ - have eq : ∏ b ∈ t.filter fun b ↦ g b = f b, f b = ∏ b ∈ t, g b := + have eq : ∏ b ∈ t with g b = f b, f b = ∏ b ∈ t, g b := calc - ∏ b ∈ t.filter fun b ↦ g b = f b, f b = ∏ b ∈ t.filter fun b ↦ g b = f b, g b := + ∏ b ∈ t with g b = f b, f b = ∏ b ∈ t with g b = f b, g b := Finset.prod_congr rfl fun b hb ↦ (Finset.mem_filter.1 hb).2.symm _ = ∏ b ∈ t, g b := by {refine Finset.prod_subset (Finset.filter_subset _ _) ?_ diff --git a/Mathlib/Topology/Algebra/InfiniteSum/Order.lean b/Mathlib/Topology/Algebra/InfiniteSum/Order.lean index 88fad0d5811364..612e335b225820 100644 --- a/Mathlib/Topology/Algebra/InfiniteSum/Order.lean +++ b/Mathlib/Topology/Algebra/InfiniteSum/Order.lean @@ -269,10 +269,10 @@ alias ⟨Summable.of_abs, Summable.abs⟩ := summable_abs_iff theorem Finite.of_summable_const [LinearOrderedAddCommGroup α] [TopologicalSpace α] [Archimedean α] [OrderClosedTopology α] {b : α} (hb : 0 < b) (hf : Summable fun _ : ι ↦ b) : Finite ι := by - have H : ∀ s : Finset ι, s.card • b ≤ ∑' _ : ι, b := fun s ↦ by + have H : ∀ s : Finset ι, #s • b ≤ ∑' _ : ι, b := fun s ↦ by simpa using sum_le_hasSum s (fun a _ ↦ hb.le) hf.hasSum obtain ⟨n, hn⟩ := Archimedean.arch (∑' _ : ι, b) hb - have : ∀ s : Finset ι, s.card ≤ n := fun s ↦ by + have : ∀ s : Finset ι, #s ≤ n := fun s ↦ by simpa [nsmul_le_nsmul_iff_left hb] using (H s).trans hn have : Fintype ι := fintypeOfFinsetCardLe n this infer_instance diff --git a/Mathlib/Topology/Instances/ENNReal.lean b/Mathlib/Topology/Instances/ENNReal.lean index bee9dc9f1a492a..c227c3915781de 100644 --- a/Mathlib/Topology/Instances/ENNReal.lean +++ b/Mathlib/Topology/Instances/ENNReal.lean @@ -19,7 +19,7 @@ import Mathlib.Topology.Metrizable.Uniformity noncomputable section open Set Filter Metric Function -open scoped Topology ENNReal NNReal +open scoped Finset Topology ENNReal NNReal variable {α : Type*} {β : Type*} {γ : Type*} @@ -828,11 +828,11 @@ theorem finite_const_le_of_tsum_ne_top {ι : Type*} {a : ι → ℝ≥0∞} (tsu /-- Markov's inequality for `Finset.card` and `tsum` in `ℝ≥0∞`. -/ theorem finset_card_const_le_le_of_tsum_le {ι : Type*} {a : ι → ℝ≥0∞} {c : ℝ≥0∞} (c_ne_top : c ≠ ∞) (tsum_le_c : ∑' i, a i ≤ c) {ε : ℝ≥0∞} (ε_ne_zero : ε ≠ 0) : - ∃ hf : { i : ι | ε ≤ a i }.Finite, ↑hf.toFinset.card ≤ c / ε := by + ∃ hf : { i : ι | ε ≤ a i }.Finite, #hf.toFinset ≤ c / ε := by have hf : { i : ι | ε ≤ a i }.Finite := finite_const_le_of_tsum_ne_top (ne_top_of_le_ne_top c_ne_top tsum_le_c) ε_ne_zero refine ⟨hf, (ENNReal.le_div_iff_mul_le (.inl ε_ne_zero) (.inr c_ne_top)).2 ?_⟩ - calc ↑hf.toFinset.card * ε = ∑ _i ∈ hf.toFinset, ε := by rw [Finset.sum_const, nsmul_eq_mul] + calc #hf.toFinset * ε = ∑ _i ∈ hf.toFinset, ε := by rw [Finset.sum_const, nsmul_eq_mul] _ ≤ ∑ i ∈ hf.toFinset, a i := Finset.sum_le_sum fun i => hf.mem_toFinset.1 _ ≤ ∑' i, a i := ENNReal.sum_le_tsum _ _ ≤ c := tsum_le_c diff --git a/test/antidiagonal.lean b/test/antidiagonal.lean index c3e22116f5fe74..ea8da80cfaae99 100644 --- a/test/antidiagonal.lean +++ b/test/antidiagonal.lean @@ -14,9 +14,9 @@ section -- `antidiagonalTuple` is faster than `finAntidiagonal` by a small constant factor /-- info: 23426 -/ -#guard_msgs in #eval (finAntidiagonal 4 50).card +#guard_msgs in #eval #(finAntidiagonal 4 50) /-- info: 23426 -/ -#guard_msgs in #eval (Finset.Nat.antidiagonalTuple 4 50).card +#guard_msgs in #eval #(Finset.Nat.antidiagonalTuple 4 50) end diff --git a/test/positivity.lean b/test/positivity.lean index 07295214c3bbe6..86a775d0e246cb 100644 --- a/test/positivity.lean +++ b/test/positivity.lean @@ -290,8 +290,8 @@ example (n : ℕ+) : 0 < (↑n : ℕ) := by positivity example (n : ℕ) : 0 < n ! := by positivity example (n k : ℕ) : 0 < (n+1).ascFactorial k := by positivity --- example {α : Type _} (s : Finset α) (hs : s.Nonempty) : 0 < s.card := by positivity --- example {α : Type _} [Fintype α] [Nonempty α] : 0 < Fintype.card α := by positivity +example {α : Type _} (s : Finset α) (hs : s.Nonempty) : 0 < #s := by positivity +example {α : Type _} [Fintype α] [Nonempty α] : 0 < Fintype.card α := by positivity example {r : ℝ} : 0 < Real.exp r := by positivity