Theory Print
section ‹Target Language debug messages›
theory Print
imports
"HOL-Decision_Procs.Approximation"
Affine_Code
Show.Show_Instances
"HOL-Library.Monad_Syntax"
Optimize_Float
begin
hide_const (open) floatarith.Max
subsection ‹Printing›
text ‹Just for debugging purposes›
definition print::"String.literal ⇒ unit" where "print x = ()"
context includes integer.lifting begin
end
code_printing constant print ⇀ (SML) "TextIO.print"
subsection ‹Write to File›
definition file_output::"String.literal ⇒ ((String.literal ⇒ unit) ⇒ 'a) ⇒ 'a" where
"file_output _ f = f (λ_. ())"
code_printing constant file_output ⇀ (SML) "(fn s => fn f => File'_Stream.open'_output (fn os => f (File'_Stream.output os)) (Path.explode s))"
subsection ‹Show for Floats›
definition showsp_float :: "float showsp"
where
"showsp_float p x = (
let m = mantissa x; e = exponent x in
if e = 0 then showsp_int p m else showsp_int p m o shows_string ''*2^'' o showsp_int p e)"
lemma show_law_float [show_law_intros]:
"show_law showsp_float r"
by (auto simp: showsp_float_def Let_def show_law_simps intro!: show_lawI)
lemma showsp_float_append [show_law_simps]:
"showsp_float p r (x @ y) = showsp_float p r x @ y"
by (intro show_lawD show_law_intros)
local_setup ‹Show_Generator.register_foreign_showsp @{typ float} @{term "showsp_float"} @{thm show_law_float}›
derive "show" float
subsection ‹Convert Float to Decimal number›
text ‹type for decimal floating point numbers
(currently just for printing, TODO? generalize theory Float for arbitrary base)›