@[class]
Implement has_repr
if the output string is valid lean code that evaluates back to the original object.
If you just want to view the object as a string for a trace message, use has_to_string
.
Example: #
#eval to_string "hello world"
-- [Lean] "hello world"
#eval repr "hello world"
-- [Lean] "\"hello world\""
Reference: https://github.com/leanprover/lean/issues/1664
Instances of this typeclass
- slim_check.sampleable_ext.p_repr
- bool.has_repr
- decidable.has_repr
- list.has_repr
- unit.has_repr
- option.has_repr
- sum.has_repr
- prod.has_repr
- sigma.has_repr
- subtype.has_repr
- nat.has_repr
- char.has_repr
- string.has_repr
- fin.has_repr
- unsigned.has_repr
- ordering.has_repr
- name.has_repr
- binder_info.has_repr
- environment.has_repr
- occurrences.has_repr
- congr_arg_kind.has_repr
- tactic.interactive.case_tag.has_repr
- module_info.has_repr
- expr.coord.has_repr
- expr.address.has_repr
- int.has_repr
- array.has_repr
- native.float.has_repr
- json.has_repr
- widget.interactive_expression.sf.has_repr
- feature_search.feature.has_repr
- feature_search.feature_vec.has_repr
- feature_search.feature_stats.has_repr
- rbtree.has_repr
- rbmap.has_repr
- buffer.has_repr
- std_gen.has_repr
- widget_override.interactive_expression.sf.has_repr
- norm_cast.label.has_repr
- units.has_repr
- add_units.has_repr
- with_bot.has_repr
- with_top.has_repr
- with_one.with_zero.has_repr
- with_one.has_repr
- with_zero.has_repr
- rat.has_repr
- pnat.has_repr
- pi_fin.has_repr
- finsupp.has_repr
- multiset.has_repr
- finset.has_repr
- polynomial.has_repr
- pos_num.has_repr
- num.has_repr
- znum.has_repr
- tree.has_repr
- tactic.ring_exp.atom.has_repr
- tactic.ring_exp.coeff_has_repr
- tactic.ring_exp.ex.has_repr
- zmod.has_repr
- nat.primes.has_repr
- cycle.has_repr
- matrix.has_repr
- cau_seq.completion.Cauchy.has_repr
- real.has_repr
- lazy_list.has_repr
- bitvec.has_repr
- slim_check.use_has_to_string.has_repr
- polyrith.poly.has_repr
- omega.ee.has_repr
- omega.int.preform.has_repr
- omega.nat.preform.has_repr
- alexandroff.has_repr
- generalized_continued_fraction.pair.has_repr
- generalized_continued_fraction.int_fract_pair.has_repr
- gaussian_int.has_repr
- onote.has_repr
- nonote.has_repr
- free_magma.has_repr
- free_add_magma.has_repr
- slim_check.total_function.has_repr
- slim_check.injective_function.has_repr
- erased.has_repr
- prime_multiset.has_repr
- pnat.xgcd_type.has_repr
- ordnode.has_repr
- equiv.perm.repr_perm
- miu.miu_atom.has_repr
- miu.miurepr
Instances of other typeclasses for has_repr
- has_repr.has_sizeof_inst
@[protected, instance]
@[protected, instance]
@[protected]
Equations
- list.repr_aux bool.tt (x :: xs) = repr x ++ list.repr_aux bool.ff xs
- list.repr_aux bool.tt list.nil = ""
- list.repr_aux bool.ff (x :: xs) = ", " ++ repr x ++ list.repr_aux bool.ff xs
- list.repr_aux bool.ff list.nil = ""
@[protected, instance]
Equations
- unit.has_repr = {repr := λ (u : unit), "star"}
@[protected, instance]
Equations
- option.has_repr = {repr := λ (o : option α), option.has_repr._match_1 o}
- option.has_repr._match_1 (option.some a) = "(some " ++ repr a ++ ")"
- option.has_repr._match_1 option.none = "none"
@[protected, instance]
Equations
- n.digit_char = ite (n = 0) '0' (ite (n = 1) '1' (ite (n = 2) '2' (ite (n = 3) '3' (ite (n = 4) '4' (ite (n = 5) '5' (ite (n = 6) '6' (ite (n = 7) '7' (ite (n = 8) '8' (ite (n = 9) '9' (ite (n = 10) 'a' (ite (n = 11) 'b' (ite (n = 12) 'c' (ite (n = 13) 'd' (ite (n = 14) 'e' (ite (n = 15) 'f' '*')))))))))))))))
@[protected, instance]
Equations
- nat.has_repr = {repr := nat.repr}
Equations
Equations
- char_to_hex c = let n : ℕ := c.to_nat, d2 : ℕ := n / 16, d1 : ℕ := n % 16 in hex_digit_repr d2 ++ hex_digit_repr d1
@[protected, instance]
Equations
- char.has_repr = {repr := λ (c : char), "'" ++ c.quote_core ++ "'"}
Equations
- string.quote_aux (x :: xs) = x.quote_core ++ string.quote_aux xs
- string.quote_aux list.nil = ""
@[protected, instance]