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

Commit 0063817

Browse files
chore(data/multiset/sort): make multiset repr a meta instance (#18163)
Co-authored-by: Chris Hughes <[email protected]> Co-authored-by: Eric Wieser <[email protected]>
1 parent 80c4301 commit 0063817

File tree

6 files changed

+18
-15
lines changed

6 files changed

+18
-15
lines changed

src/data/finset/sort.lean

+1-1
Original file line numberDiff line numberDiff line change
@@ -211,6 +211,6 @@ by simp only [order_emb_of_card_le, rel_embedding.coe_trans, finset.order_emb_of
211211

212212
end sort_linear_order
213213

214-
instance [has_repr α] : has_repr (finset α) := ⟨λ s, repr s.1
214+
meta instance [has_repr α] : has_repr (finset α) := ⟨λ s, repr s.1
215215

216216
end finset

src/data/list/cycle.lean

+7-7
Original file line numberDiff line numberDiff line change
@@ -17,8 +17,8 @@ Based on this, we define the quotient of lists by the rotation relation, called
1717
1818
We also define a representation of concrete cycles, available when viewing them in a goal state or
1919
via `#eval`, when over representatble types. For example, the cycle `(2 1 4 3)` will be shown
20-
as `c[1, 4, 3, 2]`. The representation of the cycle sorts the elements by the string value of the
21-
underlying element. This representation also supports cycles that can contain duplicates.
20+
as `c[2, 1, 4, 3]`. Two equal cycles may be printed differently if their internal representation
21+
is different.
2222
2323
-/
2424

@@ -728,12 +728,12 @@ end decidable
728728

729729
/--
730730
We define a representation of concrete cycles, available when viewing them in a goal state or
731-
via `#eval`, when over representatble types. For example, the cycle `(2 1 4 3)` will be shown
732-
as `c[1, 4, 3, 2]`. The representation of the cycle sorts the elements by the string value of the
733-
underlying element. This representation also supports cycles that can contain duplicates.
731+
via `#eval`, when over representable types. For example, the cycle `(2 1 4 3)` will be shown
732+
as `c[2, 1, 4, 3]`. Two equal cycles may be printed differently if their internal representation
733+
is different.
734734
-/
735-
instance [has_repr α] : has_repr (cycle α) :=
736-
⟨λ s, "c[" ++ string.intercalate ", " ((s.map repr).lists.sort (≤)).head ++ "]"
735+
meta instance [has_repr α] : has_repr (cycle α) :=
736+
⟨λ s, "c[" ++ string.intercalate ", " ((s.map repr).lists.unquot).head ++ "]"
737737

738738
/-- `chain R s` means that `R` holds between adjacent elements of `s`.
739739

src/data/multiset/sort.lean

+3-3
Original file line numberDiff line numberDiff line change
@@ -5,7 +5,6 @@ Authors: Mario Carneiro
55
-/
66
import data.list.sort
77
import data.multiset.basic
8-
import data.string.basic
98

109
/-!
1110
# Construct a sorted list from a multiset.
@@ -50,7 +49,8 @@ list.merge_sort_singleton r a
5049

5150
end sort
5251

53-
instance [has_repr α] : has_repr (multiset α) :=
54-
⟨λ s, "{" ++ string.intercalate ", " ((s.map repr).sort (≤)) ++ "}"
52+
-- TODO: use a sort order if available, gh-18166
53+
meta instance [has_repr α] : has_repr (multiset α) :=
54+
⟨λ s, "{" ++ string.intercalate ", " (s.unquot.map repr) ++ "}"
5555

5656
end multiset

src/data/pnat/factors.lean

+4-1
Original file line numberDiff line numberDiff line change
@@ -24,12 +24,15 @@ the multiplicity of `p` in this factors multiset being the p-adic valuation of `
2424
/-- The type of multisets of prime numbers. Unique factorization
2525
gives an equivalence between this set and ℕ+, as we will formalize
2626
below. -/
27-
@[derive [inhabited, has_repr, canonically_ordered_add_monoid, distrib_lattice,
27+
@[derive [inhabited, canonically_ordered_add_monoid, distrib_lattice,
2828
semilattice_sup, order_bot, has_sub, has_ordered_sub]]
2929
def prime_multiset := multiset nat.primes
3030

3131
namespace prime_multiset
3232

33+
-- `@[derive]` doesn't work for `meta` instances
34+
meta instance : has_repr prime_multiset := by delta prime_multiset; apply_instance
35+
3336
/-- The multiset consisting of a single prime -/
3437
def of_prime (p : nat.primes) : prime_multiset := ({p} : multiset nat.primes)
3538

src/group_theory/perm/cycle/concrete.lean

+1-1
Original file line numberDiff line numberDiff line change
@@ -537,7 +537,7 @@ def iso_cycle' : {f : perm α // is_cycle f} ≃ {s : cycle α // s.nodup ∧ s.
537537
notation `c[` l:(foldr `, ` (h t, list.cons h t) list.nil `]`) :=
538538
cycle.form_perm ↑l (cycle.nodup_coe_iff.mpr dec_trivial)
539539

540-
instance repr_perm [has_repr α] : has_repr (perm α) :=
540+
meta instance repr_perm [has_repr α] : has_repr (perm α) :=
541541
⟨λ f, repr (multiset.pmap (λ (g : perm α) (hg : g.is_cycle),
542542
iso_cycle ⟨g, hg⟩) -- to_cycle is faster?
543543
(perm.cycle_factors_finset f).val

test/cycle.lean

+2-2
Original file line numberDiff line numberDiff line change
@@ -1,5 +1,5 @@
11
import data.list.cycle
22

33
run_cmd guard ("c[1, 4, 3, 2]" = repr (↑[1, 4, 3, 2] : cycle ℕ))
4-
run_cmd guard ("c[1, 4, 3, 2]" = repr (↑[2, 1, 4, 3] : cycle ℕ))
5-
run_cmd guard ("c[-1, 2, 1, 4]" = repr (↑[(2 : ℤ), 1, 4, -1] : cycle ℤ))
4+
run_cmd guard ("c[2, 1, 4, 3]" = repr (↑[2, 1, 4, 3] : cycle ℕ))
5+
run_cmd guard ("c[2, 1, 4, -1]" = repr (↑[(2 : ℤ), 1, 4, -1] : cycle ℤ))

0 commit comments

Comments
 (0)