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

Commit

Permalink
chore(data/multiset/sort): make multiset repr a meta instance (#18163)
Browse files Browse the repository at this point in the history
Co-authored-by: Chris Hughes <[email protected]>
Co-authored-by: Eric Wieser <[email protected]>
  • Loading branch information
3 people committed Jan 14, 2023
1 parent 80c4301 commit 0063817
Show file tree
Hide file tree
Showing 6 changed files with 18 additions and 15 deletions.
2 changes: 1 addition & 1 deletion src/data/finset/sort.lean
Original file line number Diff line number Diff line change
Expand Up @@ -211,6 +211,6 @@ by simp only [order_emb_of_card_le, rel_embedding.coe_trans, finset.order_emb_of

end sort_linear_order

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

end finset
14 changes: 7 additions & 7 deletions src/data/list/cycle.lean
Original file line number Diff line number Diff line change
Expand Up @@ -17,8 +17,8 @@ Based on this, we define the quotient of lists by the rotation relation, called
We also define a representation of concrete cycles, available when viewing them in a goal state or
via `#eval`, when over representatble types. For example, the cycle `(2 1 4 3)` will be shown
as `c[1, 4, 3, 2]`. The representation of the cycle sorts the elements by the string value of the
underlying element. This representation also supports cycles that can contain duplicates.
as `c[2, 1, 4, 3]`. Two equal cycles may be printed differently if their internal representation
is different.
-/

Expand Down Expand Up @@ -728,12 +728,12 @@ end decidable

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

/-- `chain R s` means that `R` holds between adjacent elements of `s`.
Expand Down
6 changes: 3 additions & 3 deletions src/data/multiset/sort.lean
Original file line number Diff line number Diff line change
Expand Up @@ -5,7 +5,6 @@ Authors: Mario Carneiro
-/
import data.list.sort
import data.multiset.basic
import data.string.basic

/-!
# Construct a sorted list from a multiset.
Expand Down Expand Up @@ -50,7 +49,8 @@ list.merge_sort_singleton r a

end sort

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

end multiset
5 changes: 4 additions & 1 deletion src/data/pnat/factors.lean
Original file line number Diff line number Diff line change
Expand Up @@ -24,12 +24,15 @@ the multiplicity of `p` in this factors multiset being the p-adic valuation of `
/-- The type of multisets of prime numbers. Unique factorization
gives an equivalence between this set and ℕ+, as we will formalize
below. -/
@[derive [inhabited, has_repr, canonically_ordered_add_monoid, distrib_lattice,
@[derive [inhabited, canonically_ordered_add_monoid, distrib_lattice,
semilattice_sup, order_bot, has_sub, has_ordered_sub]]
def prime_multiset := multiset nat.primes

namespace prime_multiset

-- `@[derive]` doesn't work for `meta` instances
meta instance : has_repr prime_multiset := by delta prime_multiset; apply_instance

/-- The multiset consisting of a single prime -/
def of_prime (p : nat.primes) : prime_multiset := ({p} : multiset nat.primes)

Expand Down
2 changes: 1 addition & 1 deletion src/group_theory/perm/cycle/concrete.lean
Original file line number Diff line number Diff line change
Expand Up @@ -537,7 +537,7 @@ def iso_cycle' : {f : perm α // is_cycle f} ≃ {s : cycle α // s.nodup ∧ s.
notation `c[` l:(foldr `, ` (h t, list.cons h t) list.nil `]`) :=
cycle.form_perm ↑l (cycle.nodup_coe_iff.mpr dec_trivial)

instance repr_perm [has_repr α] : has_repr (perm α) :=
meta instance repr_perm [has_repr α] : has_repr (perm α) :=
⟨λ f, repr (multiset.pmap (λ (g : perm α) (hg : g.is_cycle),
iso_cycle ⟨g, hg⟩) -- to_cycle is faster?
(perm.cycle_factors_finset f).val
Expand Down
4 changes: 2 additions & 2 deletions test/cycle.lean
Original file line number Diff line number Diff line change
@@ -1,5 +1,5 @@
import data.list.cycle

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

0 comments on commit 0063817

Please sign in to comment.