bookkeeping and input syntax for exact specification of names of symbols in generated code
--- a/src/Tools/Code/code_haskell.ML Fri May 24 23:57:24 2013 +0200
+++ b/src/Tools/Code/code_haskell.ML Fri May 24 23:57:24 2013 +0200
@@ -269,7 +269,7 @@
end;
in print_stmt end;
-fun haskell_program_of_program labelled_name module_alias module_prefix reserved =
+fun haskell_program_of_program ctxt symbol_of module_prefix module_name reserved identifiers =
let
fun namify_fun upper base (nsp_fun, nsp_typ) =
let
@@ -297,9 +297,9 @@
| select_stmt (Code_Thingol.Classparam _) = false
| select_stmt (Code_Thingol.Classinst _) = true;
in
- Code_Namespace.flat_program labelled_name
- { module_alias = module_alias, module_prefix = module_prefix,
- reserved = reserved, empty_nsp = (reserved, reserved), namify_stmt = namify_stmt,
+ Code_Namespace.flat_program ctxt symbol_of
+ { module_prefix = module_prefix, module_name = module_name, reserved = reserved,
+ identifiers = identifiers, empty_nsp = (reserved, reserved), namify_stmt = namify_stmt,
modify_stmt = fn stmt => if select_stmt stmt then SOME stmt else NONE }
end;
@@ -324,14 +324,14 @@
("Maybe", ["Nothing", "Just"])
];
-fun serialize_haskell module_prefix string_classes { labelled_name, reserved_syms,
- includes, module_alias, class_syntax, tyco_syntax, const_syntax } program =
+fun serialize_haskell module_prefix string_classes ctxt { symbol_of, module_name,
+ reserved_syms, identifiers, includes, class_syntax, tyco_syntax, const_syntax } program =
let
(* build program *)
val reserved = fold (insert (op =) o fst) includes reserved_syms;
val { deresolver, flat_program = haskell_program } = haskell_program_of_program
- labelled_name module_alias module_prefix (Name.make_context reserved) program;
+ ctxt symbol_of module_prefix module_name (Name.make_context reserved) identifiers program;
(* print statements *)
fun deriving_show tyco =
--- a/src/Tools/Code/code_ml.ML Fri May 24 23:57:24 2013 +0200
+++ b/src/Tools/Code/code_ml.ML Fri May 24 23:57:24 2013 +0200
@@ -698,7 +698,7 @@
(** SML/OCaml generic part **)
-fun ml_program_of_program labelled_name reserved module_alias program =
+fun ml_program_of_program ctxt symbol_of module_name reserved identifiers program =
let
fun namify_const upper base (nsp_const, nsp_type) =
let
@@ -729,7 +729,8 @@
| ml_binding_of_stmt (name, Code_Thingol.Classinst (stmt as { vs, ... })) =
(ML_Instance (name, stmt), if forall (null o snd) vs then SOME (name, false) else NONE)
| ml_binding_of_stmt (name, _) =
- error ("Binding block containing illegal statement: " ^ labelled_name name)
+ error ("Binding block containing illegal statement: " ^
+ (Code_Symbol.quote_symbol ctxt o symbol_of) name)
fun modify_fun (name, stmt) =
let
val (binding, some_value_name) = ml_binding_of_stmt (name, stmt);
@@ -769,21 +770,22 @@
| modify_stmts ((stmts as (_, Code_Thingol.Classinst _)::_)) =
modify_funs stmts
| modify_stmts stmts = error ("Illegal mutual dependencies: " ^
- (Library.commas o map (labelled_name o fst)) stmts);
+ (Library.commas o map (Code_Symbol.quote_symbol ctxt o symbol_of o fst)) stmts);
in
- Code_Namespace.hierarchical_program labelled_name { module_alias = module_alias, reserved = reserved,
+ Code_Namespace.hierarchical_program ctxt symbol_of {
+ module_name = module_name, reserved = reserved, identifiers = identifiers,
empty_nsp = (reserved, reserved), namify_module = pair, namify_stmt = namify_stmt,
cyclic_modules = false, empty_data = (), memorize_data = K I, modify_stmts = modify_stmts } program
end;
-fun serialize_ml print_ml_module print_ml_stmt with_signatures
- { labelled_name, reserved_syms, includes, module_alias,
+fun serialize_ml print_ml_module print_ml_stmt with_signatures ctxt
+ { symbol_of, module_name, reserved_syms, identifiers, includes,
class_syntax, tyco_syntax, const_syntax } program =
let
(* build program *)
val { deresolver, hierarchical_program = ml_program } =
- ml_program_of_program labelled_name (Name.make_context reserved_syms) module_alias program;
+ ml_program_of_program ctxt symbol_of module_name (Name.make_context reserved_syms) identifiers program;
(* print statements *)
fun print_stmt prefix_fragments (_, stmt) = print_ml_stmt
--- a/src/Tools/Code/code_namespace.ML Fri May 24 23:57:24 2013 +0200
+++ b/src/Tools/Code/code_namespace.ML Fri May 24 23:57:24 2013 +0200
@@ -7,8 +7,9 @@
signature CODE_NAMESPACE =
sig
type flat_program
- val flat_program: (string -> string) -> { module_alias: string -> string option,
- module_prefix: string, reserved: Name.context, empty_nsp: 'a,
+ val flat_program: Proof.context -> (string -> Code_Symbol.symbol)
+ -> { module_prefix: string, module_name: string,
+ reserved: Name.context, identifiers: Code_Target.identifier_data, empty_nsp: 'a,
namify_stmt: Code_Thingol.stmt -> string -> 'a -> string * 'a,
modify_stmt: Code_Thingol.stmt -> Code_Thingol.stmt option }
-> Code_Thingol.program
@@ -20,8 +21,10 @@
| Stmt of 'a
| Module of ('b * (string * ('a, 'b) node) Graph.T)
type ('a, 'b) hierarchical_program
- val hierarchical_program: (string -> string) -> { module_alias: string -> string option,
- reserved: Name.context, empty_nsp: 'c, namify_module: string -> 'c -> string * 'c,
+ val hierarchical_program: Proof.context -> (string -> Code_Symbol.symbol)
+ -> { module_name: string,
+ reserved: Name.context, identifiers: Code_Target.identifier_data,
+ empty_nsp: 'c, namify_module: string -> 'c -> string * 'c,
namify_stmt: Code_Thingol.stmt -> string -> 'c -> string * 'c,
cyclic_modules: bool, empty_data: 'b, memorize_data: string -> 'b -> 'b,
modify_stmts: (string * Code_Thingol.stmt) list -> 'a option list }
@@ -37,17 +40,25 @@
structure Code_Namespace : CODE_NAMESPACE =
struct
-(** building module name hierarchy **)
+(** fundamental module name hierarchy **)
+
+val split_name = apfst Long_Name.implode o split_last o fst o split_last o Long_Name.explode;
+
+fun lookup_identifier symbol_of identifiers name =
+ Code_Symbol.lookup identifiers (symbol_of name)
+ |> Option.map (split_last o Long_Name.explode);
-val dest_name =
- apfst Long_Name.implode o split_last o fst o split_last o Long_Name.explode;
+fun force_identifier symbol_of fragments_tab identifiers name =
+ case lookup_identifier symbol_of identifiers name of
+ NONE => (apfst (the o Symtab.lookup fragments_tab) o split_name) name
+ | SOME name' => name';
-fun build_module_namespace { module_alias, module_prefix, reserved } program =
+fun build_module_namespace { module_prefix, module_identifiers, reserved } program =
let
- fun alias_fragments name = case module_alias name
+ fun alias_fragments name = case module_identifiers name
of SOME name' => Long_Name.explode name'
| NONE => map (fn name => fst (Name.variant name reserved)) (Long_Name.explode name);
- val module_names = Graph.fold (insert (op =) o fst o dest_name o fst) program [];
+ val module_names = Graph.fold (insert (op =) o fst o split_name o fst) program [];
in
fold (fn name => Symtab.update (name, Long_Name.explode module_prefix @ alias_fragments name))
module_names Symtab.empty
@@ -58,28 +69,31 @@
type flat_program = ((string * Code_Thingol.stmt option) Graph.T * (string * string list) list) Graph.T;
-fun flat_program labelled_name { module_alias, module_prefix, reserved,
- empty_nsp, namify_stmt, modify_stmt } program =
+fun flat_program ctxt symbol_of { module_prefix, module_name, reserved,
+ identifiers, empty_nsp, namify_stmt, modify_stmt } program =
let
(* building module name hierarchy *)
- val fragments_tab = build_module_namespace { module_alias = module_alias,
- module_prefix = module_prefix, reserved = reserved } program;
- val dest_name = dest_name
- #>> (Long_Name.implode o the o Symtab.lookup fragments_tab);
+ val module_identifiers = if module_name = ""
+ then Code_Symbol.lookup_module_data identifiers
+ else K (SOME module_name);
+ val fragments_tab = build_module_namespace { module_prefix = module_prefix,
+ module_identifiers = module_identifiers, reserved = reserved } program;
+ val prep_name = force_identifier symbol_of fragments_tab identifiers
+ #>> Long_Name.implode;
(* distribute statements over hierarchy *)
fun add_stmt name stmt =
let
- val (module_name, base) = dest_name name;
+ val (module_name, base) = prep_name name;
in
Graph.default_node (module_name, (Graph.empty, []))
#> (Graph.map_node module_name o apfst) (Graph.new_node (name, (base, stmt)))
end;
fun add_dependency name name' =
let
- val (module_name, _) = dest_name name;
- val (module_name', _) = dest_name name';
+ val (module_name, _) = prep_name name;
+ val (module_name', _) = prep_name name';
in if module_name = module_name'
then (Graph.map_node module_name o apfst) (Graph.add_edge (name, name'))
else (Graph.map_node module_name o apsnd) (AList.map_default (op =) (module_name', []) (insert (op =) name'))
@@ -104,7 +118,7 @@
(* qualified and unqualified imports, deresolving *)
fun base_deresolver name = fst (Graph.get_node
- (fst (Graph.get_node flat_program (fst (dest_name name)))) name);
+ (fst (Graph.get_node flat_program (fst (prep_name name)))) name);
fun classify_names gr imports =
let
val import_tab = maps
@@ -122,10 +136,11 @@
(uncurry classify_names o Graph.get_node flat_program)
(Graph.keys flat_program));
fun deresolver "" name =
- Long_Name.append (fst (dest_name name)) (base_deresolver name)
+ Long_Name.append (fst (prep_name name)) (base_deresolver name)
| deresolver module_name name =
the (Symtab.lookup (the (Symtab.lookup deresolver_tab module_name)) name)
- handle Option.Option => error ("Unknown statement name: " ^ labelled_name name);
+ handle Option.Option => error ("Unknown statement name: "
+ ^ (Code_Symbol.quote_symbol ctxt o symbol_of) name);
in { deresolver = deresolver, flat_program = flat_program } end;
@@ -146,14 +161,17 @@
apsnd o Graph.map_node name_fragment o apsnd o map_module_content
o map_module name_fragments;
-fun hierarchical_program labelled_name { module_alias, reserved, empty_nsp,
+fun hierarchical_program ctxt symbol_of { module_name, reserved, identifiers, empty_nsp,
namify_module, namify_stmt, cyclic_modules, empty_data, memorize_data, modify_stmts } program =
let
(* building module name hierarchy *)
- val fragments_tab = build_module_namespace { module_alias = module_alias,
- module_prefix = "", reserved = reserved } program;
- val dest_name = dest_name #>> (the o Symtab.lookup fragments_tab);
+ val module_identifiers = if module_name = ""
+ then Code_Symbol.lookup_module_data identifiers
+ else K (SOME module_name);
+ val fragments_tab = build_module_namespace { module_prefix = "",
+ module_identifiers = module_identifiers, reserved = reserved } program;
+ val prep_name = force_identifier symbol_of fragments_tab identifiers;
(* building empty module hierarchy *)
val empty_module = (empty_data, Graph.empty);
@@ -165,20 +183,23 @@
| allocate_module (name_fragment :: name_fragments) =
ensure_module name_fragment
#> (apsnd o Graph.map_node name_fragment o apsnd o map_module_content o allocate_module) name_fragments;
- val empty_program = Symtab.fold (fn (_, fragments) => allocate_module fragments)
- fragments_tab empty_module;
+ val empty_program =
+ empty_module
+ |> Symtab.fold (fn (_, fragments) => allocate_module fragments) fragments_tab
+ |> Graph.fold (allocate_module o these o Option.map fst
+ o lookup_identifier symbol_of identifiers o fst) program;
(* distribute statements over hierarchy *)
fun add_stmt name stmt =
let
- val (name_fragments, base) = dest_name name;
+ val (name_fragments, base) = prep_name name;
in
(map_module name_fragments o apsnd) (Graph.new_node (name, (base, Stmt stmt)))
end;
fun add_dependency name name' =
let
- val (name_fragments, _) = dest_name name;
- val (name_fragments', _) = dest_name name';
+ val (name_fragments, _) = prep_name name;
+ val (name_fragments', _) = prep_name name';
val (name_fragments_common, (diff, diff')) =
chop_prefix (op =) (name_fragments, name_fragments');
val is_module = not (null diff andalso null diff');
@@ -186,7 +207,8 @@
val add_edge = if is_module andalso not cyclic_modules
then (fn node => Graph.add_edge_acyclic dep node
handle Graph.CYCLES _ => error ("Dependency "
- ^ quote name ^ " -> " ^ quote name'
+ ^ (Code_Symbol.quote_symbol ctxt o symbol_of) name ^ " -> "
+ ^ (Code_Symbol.quote_symbol ctxt o symbol_of) name'
^ " would result in module dependency cycle"))
else Graph.add_edge dep
in (map_module name_fragments_common o apsnd) add_edge end;
@@ -228,13 +250,14 @@
(* deresolving *)
fun deresolver prefix_fragments name =
let
- val (name_fragments, _) = dest_name name;
+ val (name_fragments, _) = prep_name name;
val (_, (_, remainder)) = chop_prefix (op =) (prefix_fragments, name_fragments);
val nodes = fold (fn name_fragment => fn nodes => case Graph.get_node nodes name_fragment
of (_, Module (_, nodes)) => nodes) name_fragments hierarchical_program;
val (base', _) = Graph.get_node nodes name;
in Long_Name.implode (remainder @ [base']) end
- handle Graph.UNDEF _ => error ("Unknown statement name: " ^ labelled_name name);
+ handle Graph.UNDEF _ => error ("Unknown statement name: "
+ ^ (Code_Symbol.quote_symbol ctxt o symbol_of) name);
in { deresolver = deresolver, hierarchical_program = hierarchical_program } end;
--- a/src/Tools/Code/code_scala.ML Fri May 24 23:57:24 2013 +0200
+++ b/src/Tools/Code/code_scala.ML Fri May 24 23:57:24 2013 +0200
@@ -281,7 +281,7 @@
end;
in print_stmt end;
-fun scala_program_of_program labelled_name reserved module_alias program =
+fun scala_program_of_program ctxt symbol_of module_name reserved identifiers program =
let
fun namify_module name_fragment ((nsp_class, nsp_object), nsp_common) =
let
@@ -324,20 +324,20 @@
| modify_stmt (_, Code_Thingol.Classparam _) = NONE
| modify_stmt (_, stmt) = SOME stmt;
in
- Code_Namespace.hierarchical_program labelled_name
- { module_alias = module_alias, reserved = reserved,
+ Code_Namespace.hierarchical_program ctxt symbol_of
+ { module_name = module_name, reserved = reserved, identifiers = identifiers,
empty_nsp = ((reserved, reserved), reserved), namify_module = namify_module,
namify_stmt = namify_stmt, cyclic_modules = true, empty_data = [],
memorize_data = memorize_implicits, modify_stmts = map modify_stmt } program
end;
-fun serialize_scala { labelled_name, reserved_syms, includes,
- module_alias, class_syntax, tyco_syntax, const_syntax } program =
+fun serialize_scala ctxt { symbol_of, module_name, reserved_syms, identifiers,
+ includes, class_syntax, tyco_syntax, const_syntax } program =
let
(* build program *)
val { deresolver, hierarchical_program = scala_program } =
- scala_program_of_program labelled_name (Name.make_context reserved_syms) module_alias program;
+ scala_program_of_program ctxt symbol_of module_name (Name.make_context reserved_syms) identifiers program;
(* print statements *)
fun lookup_constr tyco constr = case Graph.get_node program tyco
--- a/src/Tools/Code/code_symbol.ML Fri May 24 23:57:24 2013 +0200
+++ b/src/Tools/Code/code_symbol.ML Fri May 24 23:57:24 2013 +0200
@@ -10,6 +10,7 @@
datatype ('a, 'b, 'c, 'd, 'e, 'f) attr =
Constant of 'a | Type_Constructor of 'b | Type_Class of 'c |
Class_Relation of 'd | Class_Instance of 'e | Module of 'f
+ type symbol = (string, string, class, string * class, class * class, string) attr
val map_attr: ('a -> 'g) -> ('b -> 'h) -> ('c -> 'i) -> ('d -> 'j) -> ('e -> 'k) -> ('f -> 'l)
-> ('a, 'b, 'c, 'd, 'e, 'f) attr -> ('g, 'h, 'i, 'j, 'k, 'l) attr
val maps_attr: ('a -> 'g list) -> ('b -> 'h list)
@@ -22,22 +23,23 @@
val empty_data: ('a, 'b, 'c, 'd, 'e, 'f) data
val merge_data: ('a, 'b, 'c, 'd, 'e, 'f) data * ('a, 'b, 'c, 'd, 'e, 'f) data
-> ('a, 'b, 'c, 'd, 'e, 'f) data
+ val set_data: (string * 'a option, string * 'b option, string * 'c option,
+ (class * class) * 'd option, (string * class) * 'e option, string * 'f option) attr
+ -> ('a, 'b, 'c, 'd, 'e, 'f) data -> ('a, 'b, 'c, 'd, 'e, 'f) data
val lookup_constant_data: ('a, 'b, 'c, 'd, 'e, 'f) data -> string -> 'a option
val lookup_type_constructor_data: ('a, 'b, 'c, 'd, 'e, 'f) data -> string -> 'b option
val lookup_type_class_data: ('a, 'b, 'c, 'd, 'e, 'f) data -> class -> 'c option
val lookup_class_relation_data: ('a, 'b, 'c, 'd, 'e, 'f) data -> class * class -> 'd option
val lookup_class_instance_data: ('a, 'b, 'c, 'd, 'e, 'f) data -> string * class -> 'e option
val lookup_module_data: ('a, 'b, 'c, 'd, 'e, 'f) data -> string -> 'f option
+ val lookup: ('a, 'a, 'a, 'a, 'a, 'a) data -> symbol -> 'a option
val dest_constant_data: ('a, 'b, 'c, 'd, 'e, 'f) data -> (string * 'a) list
val dest_type_constructor_data: ('a, 'b, 'c, 'd, 'e, 'f) data -> (string * 'b) list
val dest_type_class_data: ('a, 'b, 'c, 'd, 'e, 'f) data -> (class * 'c) list
val dest_class_relation_data: ('a, 'b, 'c, 'd, 'e, 'f) data -> ((class * class) * 'd) list
val dest_class_instance_data: ('a, 'b, 'c, 'd, 'e, 'f) data -> ((string * class) * 'e) list
val dest_module_data: ('a, 'b, 'c, 'd, 'e, 'f) data -> (string * 'f) list
- val set_data: (string * 'a option, string * 'b option, string * 'c option,
- (class * class) * 'd option, (string * class) * 'e option, string * 'f option) attr
- -> ('a, 'b, 'c, 'd, 'e, 'f) data -> ('a, 'b, 'c, 'd, 'e, 'f) data
- val quote_symbol: Proof.context -> (string, string, class, string * class, class * class, string) attr -> string
+ val quote_symbol: Proof.context -> symbol -> string
end;
structure Code_Symbol : CODE_SYMBOL =
@@ -48,6 +50,8 @@
datatype ('a, 'b, 'c, 'd, 'e, 'f) attr =
Constant of 'a | Type_Constructor of 'b | Type_Class of 'c | Class_Relation of 'd | Class_Instance of 'e | Module of 'f;
+type symbol = (string, string, class, string * class, class * class, string) attr;
+
fun map_attr const tyco class classrel inst module (Constant x) = Constant (const x)
| map_attr const tyco class classrel inst module (Type_Constructor x) = Type_Constructor (tyco x)
| map_attr const tyco class classrel inst module (Type_Class x) = Type_Class (class x)
@@ -97,20 +101,6 @@
(Symreltab.join (K snd) (class_instance1, class_instance2))
(Symtab.join (K snd) (module1, module2)); (*prefer later entries: K snd*)
-fun lookup_constant_data x = (Symtab.lookup o #constant o dest_data) x;
-fun lookup_type_constructor_data x = (Symtab.lookup o #type_constructor o dest_data) x;
-fun lookup_type_class_data x = (Symtab.lookup o #type_class o dest_data) x;
-fun lookup_class_relation_data x = (Symreltab.lookup o #class_relation o dest_data) x;
-fun lookup_class_instance_data x = (Symreltab.lookup o #class_instance o dest_data) x;
-fun lookup_module_data x = (Symtab.lookup o #module o dest_data) x;
-
-fun dest_constant_data x = (Symtab.dest o #constant o dest_data) x;
-fun dest_type_constructor_data x = (Symtab.dest o #type_constructor o dest_data) x;
-fun dest_type_class_data x = (Symtab.dest o #type_class o dest_data) x;
-fun dest_class_relation_data x = (Symreltab.dest o #class_relation o dest_data) x;
-fun dest_class_instance_data x = (Symreltab.dest o #class_instance o dest_data) x;
-fun dest_module_data x = (Symtab.dest o #module o dest_data) x;
-
fun set_sym (sym, NONE) = Symtab.delete_safe sym
| set_sym (sym, SOME y) = Symtab.update (sym, y);
fun set_symrel (symrel, NONE) = Symreltab.delete_safe symrel
@@ -123,6 +113,27 @@
| set_data (Class_Instance x) = map_data I I I I (set_symrel x) I
| set_data (Module x) = map_data I I I I I (set_sym x);
+fun lookup_constant_data x = (Symtab.lookup o #constant o dest_data) x;
+fun lookup_type_constructor_data x = (Symtab.lookup o #type_constructor o dest_data) x;
+fun lookup_type_class_data x = (Symtab.lookup o #type_class o dest_data) x;
+fun lookup_class_relation_data x = (Symreltab.lookup o #class_relation o dest_data) x;
+fun lookup_class_instance_data x = (Symreltab.lookup o #class_instance o dest_data) x;
+fun lookup_module_data x = (Symtab.lookup o #module o dest_data) x;
+
+fun lookup data (Constant x) = lookup_constant_data data x
+ | lookup data (Type_Constructor x) = lookup_type_constructor_data data x
+ | lookup data (Type_Class x) = lookup_type_class_data data x
+ | lookup data (Class_Relation x) = lookup_class_relation_data data x
+ | lookup data (Class_Instance x) = lookup_class_instance_data data x
+ | lookup data (Module x) = lookup_module_data data x;
+
+fun dest_constant_data x = (Symtab.dest o #constant o dest_data) x;
+fun dest_type_constructor_data x = (Symtab.dest o #type_constructor o dest_data) x;
+fun dest_type_class_data x = (Symtab.dest o #type_class o dest_data) x;
+fun dest_class_relation_data x = (Symreltab.dest o #class_relation o dest_data) x;
+fun dest_class_instance_data x = (Symreltab.dest o #class_instance o dest_data) x;
+fun dest_module_data x = (Symtab.dest o #module o dest_data) x;
+
fun quote_symbol ctxt (Constant c) = quote (Code.string_of_const (Proof_Context.theory_of ctxt) c)
| quote_symbol ctxt (Type_Constructor tyco) = "type " ^ quote (Proof_Context.extern_type ctxt tyco)
| quote_symbol ctxt (Type_Class class) = "class " ^ quote (Proof_Context.extern_class ctxt class)
--- a/src/Tools/Code/code_target.ML Fri May 24 23:57:24 2013 +0200
+++ b/src/Tools/Code/code_target.ML Fri May 24 23:57:24 2013 +0200
@@ -51,6 +51,9 @@
val set_default_code_width: int -> theory -> theory
type ('a, 'b, 'c, 'd, 'e, 'f) symbol_attr_decl
+ type identifier_data
+ val set_identifiers: (string, string, string, string, string, string) symbol_attr_decl
+ -> theory -> theory
type const_syntax = Code_Printer.const_syntax
type tyco_syntax = Code_Printer.tyco_syntax
val set_printings: (const_syntax, tyco_syntax, string, unit, unit, (string * string list)) symbol_attr_decl
@@ -79,6 +82,7 @@
class * (string * 'c option) list, (class * class) * (string * 'd option) list,
(class * string) * (string * 'e option) list,
string * (string * 'f option) list) Code_Symbol.attr;
+type identifier_data = (string, string, string, string, string, string) Code_Symbol.data;
type tyco_syntax = Code_Printer.tyco_syntax;
type const_syntax = Code_Printer.const_syntax;
@@ -108,11 +112,13 @@
(* serializers: functions producing serializations *)
type serializer = Token.T list
+ -> Proof.context
-> {
- labelled_name: string -> string,
+ symbol_of: string -> Code_Symbol.symbol,
+ module_name: string,
reserved_syms: string list,
+ identifiers: identifier_data,
includes: (string * Pretty.T) list,
- module_alias: string -> string option,
class_syntax: string -> string option,
tyco_syntax: string -> Code_Printer.tyco_syntax option,
const_syntax: string -> Code_Printer.activated_const_syntax option }
@@ -134,32 +140,30 @@
serial: serial,
description: description,
reserved: string list,
- module_alias: string Symtab.table,
+ identifiers: identifier_data,
printings: (Code_Printer.const_syntax, Code_Printer.tyco_syntax, string, unit, unit,
(Pretty.T * string list)) Code_Symbol.data
};
-fun make_target ((serial, description), ((reserved, module_alias), printings)) =
+fun make_target ((serial, description), (reserved, (identifiers, printings))) =
Target { serial = serial, description = description, reserved = reserved,
- module_alias = module_alias, printings = printings };
-fun map_target f ( Target { serial, description, reserved, module_alias, printings } ) =
- make_target (f ((serial, description), ((reserved, module_alias), printings)));
+ identifiers = identifiers, printings = printings };
+fun map_target f (Target { serial, description, reserved, identifiers, printings }) =
+ make_target (f ((serial, description), (reserved, (identifiers, printings))));
fun merge_target strict target (Target { serial = serial1, description = description,
- reserved = reserved1, module_alias = module_alias1, printings = printings1 },
+ reserved = reserved1, identifiers = identifiers1, printings = printings1 },
Target { serial = serial2, description = _,
- reserved = reserved2, module_alias = module_alias2, printings = printings2 }) =
+ reserved = reserved2, identifiers = identifiers2, printings = printings2 }) =
if serial1 = serial2 orelse not strict then
- make_target ((serial1, description),
- ((merge (op =) (reserved1, reserved2),
- Symtab.join (K snd) (module_alias1, module_alias2)),
- (Code_Symbol.merge_data (printings1, printings2))
- ))
+ make_target ((serial1, description), (merge (op =) (reserved1, reserved2),
+ (Code_Symbol.merge_data (identifiers1, identifiers2),
+ Code_Symbol.merge_data (printings1, printings2))))
else
error ("Incompatible targets: " ^ quote target);
fun the_description (Target { description, ... }) = description;
fun the_reserved (Target { reserved, ... }) = reserved;
-fun the_module_alias (Target { module_alias , ... }) = module_alias;
+fun the_identifiers (Target { identifiers , ... }) = identifiers;
fun the_printings (Target { printings, ... }) = printings;
structure Targets = Theory_Data
@@ -197,8 +201,8 @@
in
thy
|> (Targets.map o apfst o apfst o Symtab.update)
- (target, make_target ((serial (), seri), (([], Symtab.empty),
- Code_Symbol.empty_data)))
+ (target, make_target ((serial (), seri),
+ ([], (Code_Symbol.empty_data, Code_Symbol.empty_data))))
end;
fun add_target (target, seri) = put_target (target, Fundamental seri);
@@ -210,13 +214,13 @@
val _ = assert_target thy target;
in
thy
- |> (Targets.map o apfst o apfst o Symtab.map_entry target o map_target) f
+ |> (Targets.map o apfst o apfst o Symtab.map_entry target o map_target o apsnd) f
end;
fun map_reserved target =
- map_target_data target o apsnd o apfst o apfst;
-fun map_module_alias target =
- map_target_data target o apsnd o apfst o apsnd;
+ map_target_data target o apfst;
+fun map_identifiers target =
+ map_target_data target o apsnd o apfst;
fun map_printings target =
map_target_data target o apsnd o apsnd;
@@ -311,7 +315,7 @@
val program4 = Graph.restrict (member (op =) names4) program3;
in (names4, program4) end;
-fun prepare_serializer thy abortable (serializer : serializer) literals reserved module_alias
+fun prepare_serializer thy abortable (serializer : serializer) literals reserved identifiers
printings module_name args naming proto_program names =
let
val (names_hidden, (const_syntax, tyco_syntax, class_syntax)) =
@@ -324,11 +328,12 @@
then SOME (name, content) else NONE;
val includes = map_filter select_include (Code_Symbol.dest_module_data printings);
in
- (serializer args {
- labelled_name = Code_Thingol.labelled_name thy proto_program,
+ (serializer args (Proof_Context.init_global thy) {
+ symbol_of = Code_Thingol.symbol_of proto_program,
+ module_name = module_name,
reserved_syms = reserved,
+ identifiers = identifiers,
includes = includes,
- module_alias = if module_name = "" then Symtab.lookup module_alias else K (SOME module_name),
const_syntax = Symtab.lookup const_syntax,
tyco_syntax = Symtab.lookup tyco_syntax,
class_syntax = Symtab.lookup class_syntax },
@@ -340,11 +345,9 @@
val (default_width, abortable, data, modify) = activate_target thy target;
val serializer = case the_description data
of Fundamental seri => #serializer seri;
- val reserved = the_reserved data;
- val module_alias = the_module_alias data
- val literals = the_literals thy target;
- val (prepared_serializer, prepared_program) = prepare_serializer thy
- abortable serializer literals reserved module_alias (the_printings data)
+ val (prepared_serializer, prepared_program) =
+ prepare_serializer thy abortable serializer (the_literals thy target)
+ (the_reserved data) (the_identifiers data) (the_printings data)
module_name args naming (modify naming program) names
val width = the_default default_width some_width;
in (fn program => prepared_serializer program width, prepared_program) end;
@@ -355,7 +358,7 @@
target some_width module_name args naming program names;
in mounted_serializer prepared_program end;
-fun assert_module_name "" = error ("Empty module name not allowed.")
+fun assert_module_name "" = error "Empty module name not allowed here"
| assert_module_name module_name = module_name;
fun using_master_directory thy =
@@ -522,6 +525,20 @@
(** serializer configuration **)
+(* reserved symbol names *)
+
+fun add_reserved target sym thy =
+ let
+ val (_, data) = collapse_hierarchy thy target;
+ val _ = if member (op =) (the_reserved data) sym
+ then error ("Reserved symbol " ^ quote sym ^ " already declared")
+ else ();
+ in
+ thy
+ |> map_reserved target (insert (op =) sym)
+ end;
+
+
(* checking and parsing *)
fun cert_const thy const =
@@ -563,6 +580,26 @@
Code_Symbol.map_attr (apfst (Code.read_const thy)) (apfst (read_tyco thy))
(apfst (read_class thy)) ((apfst o pairself) (read_class thy)) (apfst (read_inst thy)) I;
+fun check_name is_module s =
+ let
+ val _ = if s = "" then error "Bad empty code name" else ();
+ val xs = Long_Name.explode s;
+ val xs' = if is_module
+ then map (Name.desymbolize true) xs
+ else if length xs < 2
+ then error ("Bad code name without module component: " ^ quote s)
+ else
+ let
+ val (ys, y) = split_last xs;
+ val ys' = map (Name.desymbolize true) ys;
+ val y' = Name.desymbolize false y;
+ in ys' @ [y'] end;
+ in if xs' = xs
+ then s
+ else error ("Invalid code name: " ^ quote s ^ "\n"
+ ^ "better try " ^ quote (Long_Name.implode xs'))
+ end;
+
fun check_const_syntax thy c syn =
if Code_Printer.requires_args syn > Code.args_number thy c
then error ("Too many arguments in syntax for constant " ^ quote c)
@@ -573,6 +610,36 @@
then error ("Number of arguments mismatch in syntax for type constructor " ^ quote tyco)
else syn;
+
+(* custom symbol names *)
+
+val arrange_name_decls =
+ let
+ fun arrange is_module (sym, target_names) = map (fn (target, some_name) =>
+ (target, (sym, Option.map (check_name is_module) some_name))) target_names;
+ in
+ Code_Symbol.maps_attr' (arrange false) (arrange false) (arrange false)
+ (arrange false) (arrange false) (arrange true)
+ end;
+
+fun cert_name_decls thy = cert_syms thy #> arrange_name_decls;
+
+fun read_name_decls thy = read_syms thy #> arrange_name_decls;
+
+fun set_identifier (target, sym_name) = map_identifiers target (Code_Symbol.set_data sym_name);
+
+fun gen_set_identifiers prep_name_decl raw_name_decls thy =
+ fold set_identifier (prep_name_decl thy raw_name_decls) thy;
+
+val set_identifiers = gen_set_identifiers cert_name_decls;
+val set_identifiers_cmd = gen_set_identifiers read_name_decls;
+
+fun add_module_alias_cmd target = fold (fn (sym, name) =>
+ set_identifier (target, Code_Symbol.Module (sym, if name = "" then NONE else SOME (check_name true name))));
+
+
+(* custom printings *)
+
fun arrange_printings prep_const thy =
let
fun arrange check (sym, target_syns) =
@@ -589,35 +656,6 @@
fun read_printings thy = read_syms thy #> arrange_printings Code.read_const thy;
-
-(* custom symbol names *)
-
-fun add_module_alias target (thyname, "") =
- map_module_alias target (Symtab.delete thyname)
- | add_module_alias target (thyname, modlname) =
- let
- val xs = Long_Name.explode modlname;
- val xs' = map (Name.desymbolize true) xs;
- in if xs' = xs
- then map_module_alias target (Symtab.update (thyname, modlname))
- else error ("Invalid module name: " ^ quote modlname ^ "\n"
- ^ "perhaps try " ^ quote (Long_Name.implode xs'))
- end;
-
-fun add_reserved target sym thy =
- let
- val (_, data) = collapse_hierarchy thy target;
- val _ = if member (op =) (the_reserved data) sym
- then error ("Reserved symbol " ^ quote sym ^ " already declared")
- else ();
- in
- thy
- |> map_reserved target (insert (op =) sym)
- end;
-
-
-(* custom printings *)
-
fun set_printing (target, sym_syn) = map_printings target (Code_Symbol.set_data sym_syn);
fun gen_set_printings prep_print_decl raw_print_decls thy =
@@ -699,18 +737,6 @@
(** Isar setup **)
-val code_expr_argsP = Scan.optional (@{keyword "("} |-- Args.parse --| @{keyword ")"}) [];
-
-val code_exprP =
- Scan.repeat1 Parse.term_group :|-- (fn raw_cs =>
- ((@{keyword "checking"} |-- Scan.repeat (Parse.name
- -- ((@{keyword "?"} |-- Scan.succeed false) || Scan.succeed true) -- code_expr_argsP))
- >> (fn seris => check_code_cmd raw_cs seris)
- || Scan.repeat (@{keyword "in"} |-- Parse.name
- -- Scan.optional (@{keyword "module_name"} |-- Parse.name) ""
- -- Scan.optional (@{keyword "file"} |-- Parse.name) ""
- -- code_expr_argsP) >> (fn seris => export_code_cmd raw_cs seris)));
-
fun parse_single_symbol_pragma parse_keyword parse_isa parse_target =
parse_keyword |-- Parse.!!! parse_isa --| (@{keyword "\<rightharpoonup>"} || @{keyword "=>"})
-- Parse.and_list1 (@{keyword "("} |-- (Parse.name --| @{keyword ")"} -- Scan.option parse_target));
@@ -733,6 +759,34 @@
Parse.enum1 "|" (Parse.group (fn () => "code symbol pragma")
(parse_symbol_pragma parse_const parse_tyco parse_class parse_classrel parse_inst parse_module));
+val code_expr_argsP = Scan.optional (@{keyword "("} |-- Args.parse --| @{keyword ")"}) [];
+
+val code_exprP =
+ Scan.repeat1 Parse.term_group :|-- (fn raw_cs =>
+ ((@{keyword "checking"} |-- Scan.repeat (Parse.name
+ -- ((@{keyword "?"} |-- Scan.succeed false) || Scan.succeed true) -- code_expr_argsP))
+ >> (fn seris => check_code_cmd raw_cs seris)
+ || Scan.repeat (@{keyword "in"} |-- Parse.name
+ -- Scan.optional (@{keyword "module_name"} |-- Parse.name) ""
+ -- Scan.optional (@{keyword "file"} |-- Parse.name) ""
+ -- code_expr_argsP) >> (fn seris => export_code_cmd raw_cs seris)));
+
+val _ =
+ Outer_Syntax.command @{command_spec "code_reserved"}
+ "declare words as reserved for target language"
+ (Parse.name -- Scan.repeat1 Parse.name
+ >> (fn (target, reserveds) => (Toplevel.theory o fold (add_reserved target)) reserveds));
+
+val _ =
+ Outer_Syntax.command @{command_spec "code_identifier"} "declare mandatory names for code symbols"
+ (parse_symbol_pragmas Parse.name Parse.name Parse.name Parse.name Parse.name Parse.name
+ >> (Toplevel.theory o fold set_identifiers_cmd));
+
+val _ =
+ Outer_Syntax.command @{command_spec "code_modulename"} "alias module to other name"
+ (Parse.name -- Scan.repeat1 (Parse.name -- Parse.name)
+ >> (fn (target, modlnames) => (Toplevel.theory o add_module_alias_cmd target) modlnames));
+
val _ =
Outer_Syntax.command @{command_spec "code_printing"} "declare dedicated printing for code symbols"
(parse_symbol_pragmas (Code_Printer.parse_const_syntax) (Code_Printer.parse_tyco_syntax)
@@ -761,12 +815,6 @@
add_instance_syntax_cmd);
val _ =
- Outer_Syntax.command @{command_spec "code_reserved"}
- "declare words as reserved for target language"
- (Parse.name -- Scan.repeat1 Parse.name
- >> (fn (target, reserveds) => (Toplevel.theory o fold (add_reserved target)) reserveds));
-
-val _ =
Outer_Syntax.command @{command_spec "code_include"}
"declare piece of code to be included in generated code"
(Parse.name -- Parse.name -- (Parse.text :|--
@@ -776,11 +824,6 @@
(Toplevel.theory o add_include_cmd target) (name, content_consts)));
val _ =
- Outer_Syntax.command @{command_spec "code_modulename"} "alias module to other name"
- (Parse.name -- Scan.repeat1 (Parse.name -- Parse.name)
- >> (fn (target, modlnames) => (Toplevel.theory o fold (add_module_alias target)) modlnames));
-
-val _ =
Outer_Syntax.command @{command_spec "code_abort"}
"permit constant to be implemented as program abort"
(Scan.repeat1 Parse.term_group >> (Toplevel.theory o fold allow_abort_cmd));
--- a/src/Tools/Code/code_thingol.ML Fri May 24 23:57:24 2013 +0200
+++ b/src/Tools/Code/code_thingol.ML Fri May 24 23:57:24 2013 +0200
@@ -86,7 +86,7 @@
val map_terms_stmt: (iterm -> iterm) -> stmt -> stmt
val is_constr: program -> string -> bool
val is_case: stmt -> bool
- val labelled_name: theory -> program -> string -> string
+ val symbol_of: program -> string -> Code_Symbol.symbol;
val group_stmts: theory -> program
-> ((string * stmt) list * (string * stmt) list
* ((string * stmt) list * (string * stmt) list)) list
@@ -471,29 +471,30 @@
fun is_case (Fun (_, (_, SOME _))) = true
| is_case _ = false;
-fun labelled_name thy program name =
- let val ctxt = Proof_Context.init_global thy in
- case Graph.get_node program name of
- Fun (c, _) => quote (Code.string_of_const thy c)
- | Datatype (tyco, _) => "type " ^ quote (Proof_Context.extern_type ctxt tyco)
- | Datatypecons (c, _) => quote (Code.string_of_const thy c)
- | Class (class, _) => "class " ^ quote (Proof_Context.extern_class ctxt class)
- | Classrel (sub, super) =>
- let
- val Class (sub, _) = Graph.get_node program sub;
- val Class (super, _) = Graph.get_node program super;
- in
- quote (Proof_Context.extern_class ctxt sub ^ " < " ^ Proof_Context.extern_class ctxt super)
- end
- | Classparam (c, _) => quote (Code.string_of_const thy c)
- | Classinst { class, tyco, ... } =>
- let
- val Class (class, _) = Graph.get_node program class;
- val Datatype (tyco, _) = Graph.get_node program tyco;
- in
- quote (Proof_Context.extern_type ctxt tyco ^ " :: " ^ Proof_Context.extern_class ctxt class)
- end
- end;
+fun symbol_of program name =
+ case try (Graph.get_node program) name of
+ NONE => Code_Symbol.Module "(unknown)"
+ (*FIXME: need to be liberal due to ad-hoc introduction of value for evaluation*)
+ | SOME stmt => case stmt of
+ Fun (c, _) => Code_Symbol.Constant c
+ | Datatype (tyco, _) => Code_Symbol.Type_Constructor tyco
+ | Datatypecons (c, _) => Code_Symbol.Constant c
+ | Class (class, _) => Code_Symbol.Type_Class class
+ | Classrel (sub, super) =>
+ let
+ val Class (sub, _) = Graph.get_node program sub;
+ val Class (super, _) = Graph.get_node program super;
+ in
+ Code_Symbol.Class_Relation (sub, super)
+ end
+ | Classparam (c, _) => Code_Symbol.Constant c
+ | Classinst { class, tyco, ... } =>
+ let
+ val Class (class, _) = Graph.get_node program class;
+ val Datatype (tyco, _) = Graph.get_node program tyco;
+ in
+ Code_Symbol.Class_Instance (tyco, class)
+ end;
fun linear_stmts program =
rev (Graph.strong_conn program)
@@ -515,8 +516,9 @@
then ([], filter is_class stmts, ([], []))
else if forall (is_fun orf is_classinst) stmts
then ([], [], List.partition is_fun stmts)
- else error ("Illegal mutual dependencies: " ^
- (commas o map (labelled_name thy program o fst)) stmts)
+ else error ("Illegal mutual dependencies: " ^ (commas
+ o map (Code_Symbol.quote_symbol (Proof_Context.init_global thy)
+ o symbol_of program o fst)) stmts);
in
linear_stmts program
|> map group
--- a/src/Tools/Code_Generator.thy Fri May 24 23:57:24 2013 +0200
+++ b/src/Tools/Code_Generator.thy Fri May 24 23:57:24 2013 +0200
@@ -8,7 +8,7 @@
imports Pure
keywords
"value" "print_codeproc" "code_thms" "code_deps" :: diag and
- "export_code" "code_printing" "code_class" "code_instance" "code_type"
+ "export_code" "code_identifier" "code_printing" "code_class" "code_instance" "code_type"
"code_const" "code_reserved" "code_include" "code_modulename"
"code_abort" "code_monad" "code_reflect" :: thy_decl and
"datatypes" "functions" "module_name" "file" "checking"