Theory Tabulate_Command
section ‹ Tabulating terms ›
theory Tabulate_Command
imports Main
keywords "tabulate" :: diag
begin
text ‹ The following little tool allows creation truth tables for predicates with a finite domain ›
definition "tabulate" :: "('a::enum ⇒ 'b) ⇒ ('a × 'b) list" where
"tabulate f = map (λ x. (x, f x)) (rev Enum.enum)"
ML ‹
structure Tabulate_Command =
struct
fun space_out cwidth str =
str ^ Library.replicate_string ((cwidth - length (Symbol.explode str))) Symbol.space;
fun print_row cwidths row =
String.concat (Library.separate " | " (map (fn (w, r) => space_out w r) (ListPair.zip (cwidths, row))));
fun print_table heads rows =
let
val cols = map (fn i => map (fn xs => nth xs i) (heads :: rows)) (0 upto length heads - 1)
val cwidths = map (fn c => foldr1 Int.max (map length (map Symbol.explode c))) cols
val llength = foldr1 (fn (c, x) => (c + 4) + x) cwidths + 2
in Print_Mode.with_modes [] (fn () =>
Pretty.paragraph ([Pretty.str (print_row cwidths heads)
, Pretty.fbrk
, Pretty.para ((Library.replicate_string llength "="))
, Pretty.fbrk]
@ Library.separate Pretty.fbrk (map (Pretty.str o print_row cwidths) rows))) () end;
fun tabulate_cmd raw_t state =
let
open Syntax
open Pretty
open HOLogic
val ctx = Toplevel.context_of state;
val t = Syntax.read_term ctx raw_t;
val ctx' = Proof_Context.augment t ctx;
val xs = Term.add_frees t []
val xp = foldr1 mk_prod (map Free (rev xs))
val t' = HOLogic.tupled_lambda xp t
val ts' = dest_list (Value_Command.value ctx (check_term ctx (const @{const_name "tabulate"} $ t')));
fun term_string t = XML.content_of (YXML.parse_body (symbolic_string_of (pretty_term ctx' t)));
val heads = (map term_string (map Free (rev xs)) @ [term_string t])
val rows = (map ((fn (x, y) => (map term_string (strip_tuple x) @ [term_string y])) o dest_prod) ts');
in Pretty.writeln ((print_table heads rows)) end;
end
val _ =
Outer_Syntax.command \<^command_keyword>‹tabulate› "tabulate the values of a term"
(Parse.term
>> (fn t => Toplevel.keep (Tabulate_Command.tabulate_cmd t)))
›
end