wenzelm@53164: (* Title: Tools/Spec_Check/gen_construction.ML bulwahn@52248: Author: Lukas Bulwahn and Nicolai Schaffroth, TU Muenchen bulwahn@52248: Author: Christopher League bulwahn@52248: bulwahn@52248: Constructing generators and pretty printing function for complex types. bulwahn@52248: *) bulwahn@52248: bulwahn@52248: signature GEN_CONSTRUCTION = bulwahn@52248: sig bulwahn@52248: val register : string * (string * string) -> theory -> theory bulwahn@52248: type mltype wenzelm@52254: val parse_pred : string -> string * mltype wenzelm@52254: val build_check : Proof.context -> string -> mltype * string -> string bulwahn@52248: (*val safe_check : string -> mltype * string -> string*) bulwahn@52248: val string_of_bool : bool -> string bulwahn@52248: val string_of_ref : ('a -> string) -> 'a Unsynchronized.ref -> string bulwahn@52248: end; bulwahn@52248: bulwahn@52248: structure Gen_Construction : GEN_CONSTRUCTION = bulwahn@52248: struct bulwahn@52248: bulwahn@52248: (* Parsing ML types *) bulwahn@52248: bulwahn@52248: datatype mltype = Var | Con of string * mltype list | Tuple of mltype list; bulwahn@52248: bulwahn@52248: (*Split string into tokens for parsing*) bulwahn@52248: fun split s = bulwahn@52248: let bulwahn@52248: fun split_symbol #"(" = "( " bulwahn@52248: | split_symbol #")" = " )" bulwahn@52248: | split_symbol #"," = " ," bulwahn@52248: | split_symbol #":" = " :" bulwahn@52248: | split_symbol c = Char.toString c bulwahn@52248: fun is_space c = c = #" " bulwahn@52248: in String.tokens is_space (String.translate split_symbol s) end; bulwahn@52248: bulwahn@52248: (*Accept anything that is not a recognized symbol*) bulwahn@52248: val scan_name = Scan.one (fn s => not (String.isSubstring s "(),*->;")); bulwahn@52248: bulwahn@52248: (*Turn a type list into a nested Con*) wenzelm@52257: fun make_con [] = raise Empty bulwahn@52248: | make_con [c] = c bulwahn@52248: | make_con (Con (s, _) :: cl) = Con (s, [make_con cl]); bulwahn@52248: bulwahn@52248: (*Parse a type*) bulwahn@52248: fun parse_type s = (parse_fun || parse_tuple || parse_type_single) s bulwahn@52248: bulwahn@52248: and parse_type_arg s = (parse_tuple || parse_type_single) s bulwahn@52248: bulwahn@52248: and parse_type_single s = (parse_con || parse_type_basic) s bulwahn@52248: bulwahn@52248: and parse_type_basic s = (parse_var || $$ "(" |-- parse_type --| $$ ")") s bulwahn@52248: bulwahn@52248: and parse_list s = bulwahn@52248: ($$ "(" |-- parse_type -- Scan.repeat1 ($$ "," |-- parse_type) --| $$ ")" >> op::) s bulwahn@52248: bulwahn@52248: and parse_var s = (Scan.one (String.isPrefix "'") >> (fn _ => Var)) s bulwahn@52248: bulwahn@52248: and parse_con s = ((parse_con_nest bulwahn@52248: || parse_type_basic -- parse_con_nest >> (fn (b, Con (t, _) :: tl) => Con (t, [b]) :: tl) bulwahn@52248: || parse_list -- parse_con_nest >> (fn (l, Con (t, _) :: tl) => Con (t, l) :: tl)) bulwahn@52248: >> (make_con o rev)) s bulwahn@52248: wenzelm@52257: and parse_con_nest s = Scan.unless parse_var (Scan.repeat1 (scan_name >> (fn t => Con (t, [])))) s bulwahn@52248: bulwahn@52248: and parse_fun s = (parse_type_arg -- $$ "->" -- parse_type >> (fn ((a, f), r) => Con (f, [a, r]))) s bulwahn@52248: bulwahn@52248: and parse_tuple s = (parse_type_single -- Scan.repeat1 ($$ "*" |-- parse_type_single) bulwahn@52248: >> (fn (t, tl) => Tuple (t :: tl))) s; bulwahn@52248: wenzelm@52253: (*Parse entire type + name*) bulwahn@52248: fun parse_function s = bulwahn@52248: let bulwahn@52248: val p = $$ "val" |-- scan_name --| ($$ "=" -- $$ "fn" -- $$ ":") bulwahn@52248: val (name, ty) = p (split s) bulwahn@52248: val stop = Scan.stopper (fn _ => ";") (fn s => s = ";"); bulwahn@52248: val (typ, _) = Scan.finite stop parse_type ty bulwahn@52248: in (name, typ) end; bulwahn@52248: wenzelm@52253: (*Create desired output*) bulwahn@52248: fun parse_pred s = bulwahn@52248: let bulwahn@52248: val (name, Con ("->", t :: _)) = parse_function s bulwahn@52248: in (name, t) end; bulwahn@52248: bulwahn@52248: (* Construct Generators and Pretty Printers *) bulwahn@52248: bulwahn@52248: (*copied from smt_config.ML *) bulwahn@52248: fun string_of_bool b = if b then "true" else "false" bulwahn@52248: bulwahn@52248: fun string_of_ref f r = f (!r) ^ " ref"; bulwahn@52248: bulwahn@52248: val initial_content = Symtab.make bulwahn@52248: [("bool", ("Generator.flip", "Gen_Construction.string_of_bool")), bulwahn@52248: ("option", ("Generator.option (Generator.flip' (1, 3))", "ML_Syntax.print_option")), bulwahn@52248: ("list", ("Generator.list (Generator.flip' (1, 3))", "ML_Syntax.print_list")), bulwahn@52248: ("unit", ("gen_unit", "fn () => \"()\"")), bulwahn@52248: ("int", ("Generator.int", "string_of_int")), bulwahn@52248: ("real", ("Generator.real", "string_of_real")), bulwahn@52248: ("char", ("Generator.char", "fn c => \"#'\" ^ (Char.toString c) ^ \"'\"")), bulwahn@52248: ("string", ("Generator.string (Generator.range (0, 100), Generator.char)", "ML_Syntax.print_string")), bulwahn@52248: ("->", ("Generator.function_lazy", "fn (_, _) => fn _ => \"fn\"")), bulwahn@52248: ("typ", ("Generator.typ 10", "ML_Syntax.print_typ")), bulwahn@52248: ("term", ("Generator.term 10", "ML_Syntax.print_term"))] bulwahn@52248: bulwahn@52248: structure Data = Theory_Data bulwahn@52248: ( bulwahn@52248: type T = (string * string) Symtab.table bulwahn@52248: val empty = initial_content bulwahn@52248: val extend = I bulwahn@52248: fun merge data : T = Symtab.merge (K true) data bulwahn@52248: ) bulwahn@52248: wenzelm@52254: fun data_of ctxt tycon = wenzelm@52254: (case Symtab.lookup (Data.get (Proof_Context.theory_of ctxt)) tycon of wenzelm@52253: SOME data => data wenzelm@52253: | NONE => error ("No generator and printer defined for ML type constructor " ^ quote tycon)) bulwahn@52248: bulwahn@52248: val generator_of = fst oo data_of bulwahn@52248: val printer_of = snd oo data_of bulwahn@52248: bulwahn@52248: fun register (ty, data) = Data.map (Symtab.update (ty, data)) bulwahn@52248: bulwahn@52248: (* bulwahn@52248: fun remove_gen ty = gen_table := AList.delete (op =) ty (!gen_table); bulwahn@52248: *) bulwahn@52248: bulwahn@52248: fun combine dict [] = dict bulwahn@52248: | combine dict dicts = enclose "(" ")" dict ^ " " ^ enclose "(" ")" (commas dicts) bulwahn@52248: bulwahn@52248: fun compose_generator _ Var = "Generator.int" wenzelm@52254: | compose_generator ctxt (Con (s, types)) = wenzelm@52254: combine (generator_of ctxt s) (map (compose_generator ctxt) types) wenzelm@52254: | compose_generator ctxt (Tuple t) = wenzelm@52254: let wenzelm@52254: fun tuple_body t = space_implode "" wenzelm@52254: (map (fn (ty, n) => "val (x" ^ string_of_int n ^ ", r" ^ string_of_int n ^ ") = " ^ wenzelm@52254: compose_generator ctxt ty ^ " r" ^ string_of_int (n - 1) ^ " ") (t ~~ (1 upto (length t)))) wenzelm@52254: fun tuple_ret a = commas (map (fn n => "x" ^ string_of_int n) (1 upto a)) wenzelm@52254: in wenzelm@52254: "fn r0 => let " ^ tuple_body t ^ wenzelm@52254: "in ((" ^ tuple_ret (length t) ^ "), r" ^ string_of_int (length t) ^ ") end" wenzelm@52254: end; bulwahn@52248: bulwahn@52248: fun compose_printer _ Var = "Int.toString" wenzelm@52254: | compose_printer ctxt (Con (s, types)) = wenzelm@52254: combine (printer_of ctxt s) (map (compose_printer ctxt) types) wenzelm@52254: | compose_printer ctxt (Tuple t) = wenzelm@52254: let wenzelm@52254: fun tuple_head a = commas (map (fn n => "x" ^ string_of_int n) (1 upto a)) wenzelm@52254: fun tuple_body t = space_implode " ^ \", \" ^ " wenzelm@52254: (map (fn (ty, n) => "(" ^ compose_printer ctxt ty ^ ") x" ^ string_of_int n) wenzelm@52254: (t ~~ (1 upto (length t)))) wenzelm@52254: in "fn (" ^ tuple_head (length t) ^ ") => \"(\" ^ " ^ tuple_body t ^ " ^ \")\"" end; bulwahn@52248: wenzelm@52253: (*produce compilable string*) wenzelm@52254: fun build_check ctxt name (ty, spec) = wenzelm@62876: "Spec_Check.checkGen (Context.the_local_context ()) (" wenzelm@52254: ^ compose_generator ctxt ty ^ ", SOME (" ^ compose_printer ctxt ty ^ ")) (\"" bulwahn@52248: ^ name ^ "\", Property.pred (" ^ spec ^ "));"; bulwahn@52248: wenzelm@52253: (*produce compilable string - non-eqtype functions*) bulwahn@52248: (* bulwahn@52248: fun safe_check name (ty, spec) = bulwahn@52248: let wenzelm@52254: val default = wenzelm@52254: (case AList.lookup (op =) (!gen_table) "->" of wenzelm@52254: NONE => ("gen_function_rand", "fn (_, _) => fn _ => \"fn\"") wenzelm@52254: | SOME entry => entry) wenzelm@52254: in wenzelm@52254: (gen_table := wenzelm@52254: AList.update (op =) ("->", ("gen_function_safe", "fn (_, _) => fn _ => \"fn\"")) (!gen_table); bulwahn@52248: build_check name (ty, spec) before wenzelm@52254: gen_table := AList.update (op =) ("->", default) (!gen_table)) wenzelm@52254: end; bulwahn@52248: *) bulwahn@52248: bulwahn@52248: end; bulwahn@52248: