Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit c3768cc

Browse files
fpvandoornFloris van Doorn
andcommitted
refactor(data/multiset/basic): remove sub lemmas (#9578)
* Remove the multiset sub lemmas that are special cases of lemmas in `algebra/order/sub` * [This](https://github.com/leanprover-community/mathlib/blob/14c3324143beef6e538f63da9c48d45f4a949604/src/data/multiset/basic.lean#L1513-L1538) gives the list of renamings. * Use `derive` in `pnat.factors`. Co-authored-by: Floris van Doorn <fpv@andrew.cmu.edu>
1 parent c400677 commit c3768cc

File tree

8 files changed

+21
-66
lines changed

8 files changed

+21
-66
lines changed

src/algebra/associated.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -554,7 +554,7 @@ theorem prod_le_prod {p q : multiset (associates α)} (h : p ≤ q) : p.prod ≤
554554
begin
555555
haveI := classical.dec_eq (associates α),
556556
haveI := classical.dec_eq α,
557-
suffices : p.prod ≤ (p + (q - p)).prod, { rwa [multiset.add_sub_of_le h] at this },
557+
suffices : p.prod ≤ (p + (q - p)).prod, { rwa [add_sub_cancel_of_le h] at this },
558558
suffices : p.prod * 1 ≤ p.prod * (q - p).prod, { simpa },
559559
exact mul_mono (le_refl p.prod) one_le
560560
end

src/data/finset/basic.lean

Lines changed: 2 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -1052,7 +1052,8 @@ lemma erase_inj_on (s : finset α) : set.inj_on s.erase s :=
10521052
/-! ### sdiff -/
10531053

10541054
/-- `s \ t` is the set consisting of the elements of `s` that are not in `t`. -/
1055-
instance : has_sdiff (finset α) := ⟨λs₁ s₂, ⟨s₁.1 - s₂.1, nodup_of_le (sub_le_self _ _) s₁.2⟩⟩
1055+
instance : has_sdiff (finset α) :=
1056+
⟨λs₁ s₂, ⟨s₁.1 - s₂.1, nodup_of_le sub_le_self' s₁.2⟩⟩
10561057

10571058
@[simp] lemma sdiff_val (s₁ s₂ : finset α) : (s₁ \ s₂).val = s₁.val - s₂.val := rfl
10581059

src/data/list/perm.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -704,7 +704,7 @@ end⟩
704704
lemma subperm.cons_right {α : Type*} {l l' : list α} (x : α) (h : l <+~ l') : l <+~ x :: l' :=
705705
h.trans (sublist_cons x l').subperm
706706

707-
/-- The list version of `multiset.add_sub_of_le`. -/
707+
/-- The list version of `add_sub_cancel_of_le` for multisets. -/
708708
lemma subperm_append_diff_self_of_count_le {l₁ l₂ : list α}
709709
(h : ∀ x ∈ l₁, count x l₁ ≤ count x l₂) : l₁ ++ l₂.diff l₁ ~ l₂ :=
710710
begin

src/data/multiset/antidiagonal.lean

Lines changed: 2 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -40,7 +40,8 @@ quotient.induction_on s $ λ l, begin
4040
haveI := classical.dec_eq α,
4141
simp [revzip_powerset_aux_lemma l revzip_powerset_aux, h.symm],
4242
cases x with x₁ x₂,
43-
exact ⟨_, le_add_right _ _, by rw add_sub_cancel_left _ _⟩
43+
dsimp only,
44+
exact ⟨x₁, le_add_right _ _, by rw add_sub_cancel_left x₁ x₂⟩
4445
end
4546

4647
@[simp] theorem antidiagonal_map_fst (s : multiset α) :

src/data/multiset/basic.lean

Lines changed: 7 additions & 35 deletions
Original file line numberDiff line numberDiff line change
@@ -1513,35 +1513,8 @@ quotient.induction_on₂ s t $ λ l₁ l₂,
15131513
show ↑(l₁.diff l₂) = foldl erase erase_comm ↑l₁ ↑l₂,
15141514
by { rw diff_eq_foldl l₁ l₂, symmetry, exact foldl_hom _ _ _ _ _ (λ x y, rfl) }
15151515

1516-
theorem add_sub_of_le (h : s ≤ t) : s + (t - s) = t :=
1517-
add_sub_cancel_of_le h
1518-
1519-
theorem sub_add' : s - (t + u) = s - t - u :=
1520-
sub_add_eq_sub_sub'
1521-
1522-
theorem sub_add_cancel (h : t ≤ s) : s - t + t = s :=
1523-
sub_add_cancel_of_le h
1524-
1525-
@[simp] theorem add_sub_cancel_left (s : multiset α) : ∀ t, s + t - s = t :=
1526-
add_sub_cancel_left s
1527-
1528-
@[simp] theorem add_sub_cancel (s t : multiset α) : s + t - t = s :=
1529-
add_sub_cancel_right s t
1530-
1531-
theorem sub_le_sub_right (h : s ≤ t) (u) : s - u ≤ t - u :=
1532-
sub_le_sub_right' h u
1533-
1534-
theorem sub_le_sub_left (h : s ≤ t) : ∀ u, u - t ≤ u - s :=
1535-
sub_le_sub_left' h
1536-
1537-
theorem le_sub_add (s t : multiset α) : s ≤ s - t + t :=
1538-
le_sub_add -- implicit args
1539-
1540-
theorem sub_le_self (s t : multiset α) : s - t ≤ s :=
1541-
sub_le_self' -- implicit args
1542-
15431516
@[simp] theorem card_sub {s t : multiset α} (h : t ≤ s) : card (s - t) = card s - card t :=
1544-
(nat.sub_eq_of_eq_add $ by rw [add_comm, ← card_add, sub_add_cancel h]).symm
1517+
(nat.sub_eq_of_eq_add $ by rw [add_comm, ← card_add, sub_add_cancel_of_le h]).symm
15451518

15461519
/-! ### Union -/
15471520

@@ -1554,20 +1527,20 @@ instance : has_union (multiset α) := ⟨union⟩
15541527

15551528
theorem union_def (s t : multiset α) : s ∪ t = s - t + t := rfl
15561529

1557-
theorem le_union_left (s t : multiset α) : s ≤ s ∪ t := le_sub_add _ _
1530+
theorem le_union_left (s t : multiset α) : s ≤ s ∪ t := le_sub_add
15581531

15591532
theorem le_union_right (s t : multiset α) : t ≤ s ∪ t := le_add_left _ _
15601533

1561-
theorem eq_union_left : t ≤ s → s ∪ t = s := sub_add_cancel
1534+
theorem eq_union_left : t ≤ s → s ∪ t = s := sub_add_cancel_of_le
15621535

15631536
theorem union_le_union_right (h : s ≤ t) (u) : s ∪ u ≤ t ∪ u :=
1564-
add_le_add_right (sub_le_sub_right h _) u
1537+
add_le_add_right (sub_le_sub_right' h _) u
15651538

15661539
theorem union_le (h₁ : s ≤ u) (h₂ : t ≤ u) : s ∪ t ≤ u :=
15671540
by rw ← eq_union_left h₂; exact union_le_union_right h₁ t
15681541

15691542
@[simp] theorem mem_union : a ∈ s ∪ t ↔ a ∈ s ∨ a ∈ t :=
1570-
⟨λ h, (mem_add.1 h).imp_left (mem_of_le $ sub_le_self _ _),
1543+
⟨λ h, (mem_add.1 h).imp_left (mem_of_le sub_le_self'),
15711544
or.rec (mem_of_le $ le_union_left _ _) (mem_of_le $ le_union_right _ _)⟩
15721545

15731546
@[simp] theorem map_union [decidable_eq β] {f : α → β} (finj : function.injective f)
@@ -1662,7 +1635,7 @@ union_le (le_add_right _ _) (le_add_left _ _)
16621635

16631636
theorem union_add_distrib (s t u : multiset α) : (s ∪ t) + u = (s + u) ∪ (t + u) :=
16641637
by simpa [(∪), union, eq_comm, add_assoc] using show s + u - (t + u) = s - t,
1665-
by rw [add_comm t, sub_add', add_sub_cancel]
1638+
by rw [add_comm t, sub_add_eq_sub_sub', add_sub_cancel_right]
16661639

16671640
theorem add_union_distrib (s t u : multiset α) : s + (t ∪ u) = (s + t) ∪ (s + u) :=
16681641
by rw [add_comm, union_add_distrib, add_comm s, add_comm s]
@@ -1709,8 +1682,7 @@ begin
17091682
end
17101683

17111684
theorem sub_inter (s t : multiset α) : s - (s ∩ t) = s - t :=
1712-
add_right_cancel $
1713-
by rw [sub_add_inter s t, sub_add_cancel (inter_le_left _ _)]
1685+
add_right_cancel $ by rw [sub_add_inter s t, sub_add_cancel_of_le (inter_le_left s t)]
17141686

17151687
end
17161688

src/data/multiset/nodup.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -191,10 +191,10 @@ theorem range_le {m n : ℕ} : range m ≤ range n ↔ m ≤ n :=
191191

192192
theorem mem_sub_of_nodup [decidable_eq α] {a : α} {s t : multiset α} (d : nodup s) :
193193
a ∈ s - t ↔ a ∈ s ∧ a ∉ t :=
194-
⟨λ h, ⟨mem_of_le (sub_le_self _ _) h, λ h',
194+
⟨λ h, ⟨mem_of_le sub_le_self' h, λ h',
195195
by refine count_eq_zero.1 _ h; rw [count_sub a s t, nat.sub_eq_zero_iff_le];
196196
exact le_trans (nodup_iff_count_le_one.1 d _) (count_pos.2 h')⟩,
197-
λ ⟨h₁, h₂⟩, or.resolve_right (mem_add.1 $ mem_of_le (le_sub_add _ _) h₁) h₂⟩
197+
λ ⟨h₁, h₂⟩, or.resolve_right (mem_add.1 $ mem_of_le le_sub_add h₁) h₂⟩
198198

199199
lemma map_eq_map_of_bij_of_nodup (f : α → γ) (g : β → γ) {s : multiset α} {t : multiset β}
200200
(hs : s.nodup) (ht : t.nodup) (i : Πa∈s, β)

src/data/pnat/factors.lean

Lines changed: 3 additions & 22 deletions
Original file line numberDiff line numberDiff line change
@@ -23,31 +23,12 @@ the multiplicity of `p` in this factors multiset being the p-adic valuation of `
2323
/-- The type of multisets of prime numbers. Unique factorization
2424
gives an equivalence between this set and ℕ+, as we will formalize
2525
below. -/
26+
@[derive [inhabited, has_repr, canonically_ordered_add_monoid, distrib_lattice,
27+
semilattice_sup_bot, has_sub, has_ordered_sub]]
2628
def prime_multiset := multiset nat.primes
2729

2830
namespace prime_multiset
2931

30-
instance : inhabited prime_multiset :=
31-
by unfold prime_multiset; apply_instance
32-
33-
instance : has_repr prime_multiset :=
34-
by { dsimp [prime_multiset], apply_instance }
35-
36-
instance : canonically_ordered_add_monoid prime_multiset :=
37-
by { dsimp [prime_multiset], apply_instance }
38-
39-
instance : distrib_lattice prime_multiset :=
40-
by { dsimp [prime_multiset], apply_instance }
41-
42-
instance : semilattice_sup_bot prime_multiset :=
43-
by { dsimp [prime_multiset], apply_instance }
44-
45-
instance : has_sub prime_multiset :=
46-
by { dsimp [prime_multiset], apply_instance }
47-
48-
theorem add_sub_of_le {u v : prime_multiset} : u ≤ v → u + (v - u) = v :=
49-
multiset.add_sub_of_le
50-
5132
/-- The multiset consisting of a single prime -/
5233
def of_prime (p : nat.primes) : prime_multiset := ({p} : multiset nat.primes)
5334

@@ -303,7 +284,7 @@ begin
303284
rw [← prod_factor_multiset m, ← prod_factor_multiset m],
304285
apply dvd.intro (n.factor_multiset - m.factor_multiset).prod,
305286
rw [← prime_multiset.prod_add, prime_multiset.factor_multiset_prod,
306-
prime_multiset.add_sub_of_le h, prod_factor_multiset] },
287+
add_sub_cancel_of_le h, prod_factor_multiset] },
307288
{ intro h,
308289
rw [← mul_div_exact h, factor_multiset_mul],
309290
exact le_self_add }

src/group_theory/perm/cycle_type.lean

Lines changed: 3 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -221,10 +221,10 @@ lemma cycle_type_mul_mem_cycle_factors_finset_eq_sub {f g : perm α}
221221
(g * f⁻¹).cycle_type = g.cycle_type - f.cycle_type :=
222222
begin
223223
suffices : (g * f⁻¹).cycle_type + f.cycle_type = g.cycle_type - f.cycle_type + f.cycle_type,
224-
{ rw multiset.sub_add_cancel (cycle_type_le_of_mem_cycle_factors_finset hf) at this,
224+
{ rw sub_add_cancel_of_le (cycle_type_le_of_mem_cycle_factors_finset hf) at this,
225225
simp [←this] },
226226
simp [←(disjoint_mul_inv_of_mem_cycle_factors_finset hf).cycle_type,
227-
multiset.sub_add_cancel (cycle_type_le_of_mem_cycle_factors_finset hf)]
227+
sub_add_cancel_of_le (cycle_type_le_of_mem_cycle_factors_finset hf)]
228228
end
229229

230230
theorem is_conj_of_cycle_type_eq {σ τ : perm α} (h : cycle_type σ = cycle_type τ) : is_conj σ τ :=
@@ -256,7 +256,7 @@ begin
256256
rw [hσ.cycle_type, ←hσ', hσ'l.left.cycle_type] },
257257
refine hστ.is_conj_mul (h1 hs) (h2 _) _,
258258
{ rw [cycle_type_mul_mem_cycle_factors_finset_eq_sub, ←hπ, add_comm, hs,
259-
multiset.add_sub_cancel],
259+
add_sub_cancel_right],
260260
rwa finset.mem_def },
261261
{ exact (disjoint_mul_inv_of_mem_cycle_factors_finset hσ'l).symm } } }
262262
end

0 commit comments

Comments
 (0)