lemma assoc_nat: "shows_prec d (n::nat) r @ s = shows_prec d n (r @ s)"
proof (induct n arbitrary: r s rule: nat_less_induct)
case (1 n)
show ?case
proof (cases "n < 10")
case True thus ?thesis unfolding shows_prec_nat_def by simp
next
let ?m = "n div 10"
case False
hence "?m < n" by simp
with 1 have "\<And>s t. shows_prec d ?m s @ t = shows_prec d ?m (s @ t)" by simp
thus ?thesis unfolding shows_prec_nat_def by auto
qed
standard_shows_list assoc_nat
qed
standard_shows_list assoc_nat
end
instantiation int :: "show" begin
definition "shows_prec (d::nat) (i::int) \<equiv>
(if i < 0
then shows (CHR ''-'') \<circ> shows (nat (- i))
else shows (nat i))"
instantiation int :: "show"
begin
definition "shows_prec (d::nat) (i::int) = (
if i < 0 then shows (CHR ''-'') \<circ> shows (nat (- i))
else shows (nat i))"
lemma assoc_int:
"shows_prec d (i::int) r @ s = shows_prec d i (r @ s)"
by (simp add: shows_prec_int_def)
standard_shows_list assoc_int
lemma assoc_int: "shows_prec d (i::int) r @ s = shows_prec d i (r @ s)"
by (simp add: shows_prec_int_def)
standard_shows_list assoc_int
end
instantiation rat :: "show" begin
definition show_rat :: "rat \<Rightarrow> string"
where "show_rat x \<equiv> case (quotient_of x) of (Pair den num) \<Rightarrow> if num = 1 then shows den '''' else ((shows den \<circ> shows(CHR ''/'') \<circ> shows num) '''')"
definition "shows_prec (d::nat) x \<equiv> shows_string (show_rat x)"
instantiation rat :: "show"
begin
definition "shows_prec (d::nat) (x::rat) =
(case quotient_of x of
Pair den num \<Rightarrow>
if num = 1 then shows den else shows den \<circ> shows (CHR ''/'') \<circ> shows num)"
lemma assoc_rat:
"shows_prec d (x::rat) r @ s = shows_prec d x (r @ s)"
unfolding shows_prec_rat_def by (cases "quotient_of x") auto
standard_shows_list assoc_rat
lemma assoc_rat: "shows_prec d (x::rat) r @ s = shows_prec d x (r @ s)" unfolding shows_prec_rat_def by auto