| author | blanchet | 
| Wed, 05 Jun 2013 13:19:26 +0200 | |
| changeset 52303 | 16d7708aba40 | 
| parent 52138 | e21426f244aa | 
| child 53414 | dd64696d267a | 
| permissions | -rw-r--r-- | 
| 38970 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 1 | (* Title: Tools/Code/code_namespace.ML | 
| 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 2 | Author: Florian Haftmann, TU Muenchen | 
| 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 3 | |
| 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 4 | Mastering target language namespaces. | 
| 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 5 | *) | 
| 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 6 | |
| 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 7 | signature CODE_NAMESPACE = | 
| 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 8 | sig | 
| 39208 | 9 | type flat_program | 
| 52138 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
51930diff
changeset | 10 | val flat_program: Proof.context -> (string -> Code_Symbol.symbol) | 
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
51930diff
changeset | 11 |     -> { module_prefix: string, module_name: string,
 | 
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
51930diff
changeset | 12 | reserved: Name.context, identifiers: Code_Target.identifier_data, empty_nsp: 'a, | 
| 39208 | 13 | namify_stmt: Code_Thingol.stmt -> string -> 'a -> string * 'a, | 
| 14 | modify_stmt: Code_Thingol.stmt -> Code_Thingol.stmt option } | |
| 15 | -> Code_Thingol.program | |
| 16 |       -> { deresolver: string -> string -> string,
 | |
| 17 | flat_program: flat_program } | |
| 18 | ||
| 39017 
8cd5b6d688fa
generalized hierarchical data structure over statements
 haftmann parents: 
38970diff
changeset | 19 |   datatype ('a, 'b) node =
 | 
| 38970 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 20 | Dummy | 
| 39017 
8cd5b6d688fa
generalized hierarchical data structure over statements
 haftmann parents: 
38970diff
changeset | 21 | | Stmt of 'a | 
| 39147 | 22 |     | Module of ('b * (string * ('a, 'b) node) Graph.T)
 | 
| 39208 | 23 |   type ('a, 'b) hierarchical_program
 | 
| 52138 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
51930diff
changeset | 24 | val hierarchical_program: Proof.context -> (string -> Code_Symbol.symbol) | 
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
51930diff
changeset | 25 |     -> { module_name: string,
 | 
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
51930diff
changeset | 26 | reserved: Name.context, identifiers: Code_Target.identifier_data, | 
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
51930diff
changeset | 27 | empty_nsp: 'c, namify_module: string -> 'c -> string * 'c, | 
| 39022 | 28 | namify_stmt: Code_Thingol.stmt -> string -> 'c -> string * 'c, | 
| 29 | cyclic_modules: bool, empty_data: 'b, memorize_data: string -> 'b -> 'b, | |
| 39024 
30d5dd2f30b6
simultaneous modification of statements: statement names
 haftmann parents: 
39023diff
changeset | 30 | modify_stmts: (string * Code_Thingol.stmt) list -> 'a option list } | 
| 38970 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 31 | -> Code_Thingol.program | 
| 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 32 |       -> { deresolver: string list -> string -> string,
 | 
| 39147 | 33 |            hierarchical_program: ('a, 'b) hierarchical_program }
 | 
| 34 |   val print_hierarchical: { print_module: string list -> string -> 'b -> 'c list -> 'c,
 | |
| 35 | print_stmt: string list -> string * 'a -> 'c, | |
| 36 | lift_markup: (Pretty.T -> Pretty.T) -> 'c -> 'c } | |
| 37 |       -> ('a, 'b) hierarchical_program -> 'c list
 | |
| 38970 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 38 | end; | 
| 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 39 | |
| 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 40 | structure Code_Namespace : CODE_NAMESPACE = | 
| 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 41 | struct | 
| 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 42 | |
| 52138 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
51930diff
changeset | 43 | (** fundamental module name hierarchy **) | 
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
51930diff
changeset | 44 | |
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
51930diff
changeset | 45 | val split_name = apfst Long_Name.implode o split_last o fst o split_last o Long_Name.explode; | 
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
51930diff
changeset | 46 | |
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
51930diff
changeset | 47 | fun lookup_identifier symbol_of identifiers name = | 
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
51930diff
changeset | 48 | Code_Symbol.lookup identifiers (symbol_of name) | 
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
51930diff
changeset | 49 | |> Option.map (split_last o Long_Name.explode); | 
| 39055 | 50 | |
| 52138 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
51930diff
changeset | 51 | fun force_identifier symbol_of fragments_tab identifiers name = | 
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
51930diff
changeset | 52 | case lookup_identifier symbol_of identifiers name of | 
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
51930diff
changeset | 53 | NONE => (apfst (the o Symtab.lookup fragments_tab) o split_name) name | 
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
51930diff
changeset | 54 | | SOME name' => name'; | 
| 39055 | 55 | |
| 52138 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
51930diff
changeset | 56 | fun build_module_namespace { module_prefix, module_identifiers, reserved } program =
 | 
| 39203 | 57 | let | 
| 52138 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
51930diff
changeset | 58 | fun alias_fragments name = case module_identifiers name | 
| 39203 | 59 | of SOME name' => Long_Name.explode name' | 
| 43326 
47cf4bc789aa
simplified Name.variant -- discontinued builtin fold_map;
 wenzelm parents: 
40705diff
changeset | 60 | | NONE => map (fn name => fst (Name.variant name reserved)) (Long_Name.explode name); | 
| 52138 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
51930diff
changeset | 61 | val module_names = Graph.fold (insert (op =) o fst o split_name o fst) program []; | 
| 39203 | 62 | in | 
| 39208 | 63 | fold (fn name => Symtab.update (name, Long_Name.explode module_prefix @ alias_fragments name)) | 
| 39203 | 64 | module_names Symtab.empty | 
| 65 | end; | |
| 66 | ||
| 39055 | 67 | |
| 39208 | 68 | (** flat program structure **) | 
| 69 | ||
| 70 | type flat_program = ((string * Code_Thingol.stmt option) Graph.T * (string * string list) list) Graph.T; | |
| 71 | ||
| 52138 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
51930diff
changeset | 72 | fun flat_program ctxt symbol_of { module_prefix, module_name, reserved,
 | 
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
51930diff
changeset | 73 | identifiers, empty_nsp, namify_stmt, modify_stmt } program = | 
| 39208 | 74 | let | 
| 75 | ||
| 76 | (* building module name hierarchy *) | |
| 52138 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
51930diff
changeset | 77 | val module_identifiers = if module_name = "" | 
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
51930diff
changeset | 78 | then Code_Symbol.lookup_module_data identifiers | 
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
51930diff
changeset | 79 | else K (SOME module_name); | 
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
51930diff
changeset | 80 |     val fragments_tab = build_module_namespace { module_prefix = module_prefix,
 | 
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
51930diff
changeset | 81 | module_identifiers = module_identifiers, reserved = reserved } program; | 
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
51930diff
changeset | 82 | val prep_name = force_identifier symbol_of fragments_tab identifiers | 
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
51930diff
changeset | 83 | #>> Long_Name.implode; | 
| 39208 | 84 | |
| 85 | (* distribute statements over hierarchy *) | |
| 86 | fun add_stmt name stmt = | |
| 87 | let | |
| 52138 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
51930diff
changeset | 88 | val (module_name, base) = prep_name name; | 
| 39208 | 89 | in | 
| 90 | Graph.default_node (module_name, (Graph.empty, [])) | |
| 91 | #> (Graph.map_node module_name o apfst) (Graph.new_node (name, (base, stmt))) | |
| 92 | end; | |
| 93 | fun add_dependency name name' = | |
| 94 | let | |
| 52138 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
51930diff
changeset | 95 | val (module_name, _) = prep_name name; | 
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
51930diff
changeset | 96 | val (module_name', _) = prep_name name'; | 
| 39208 | 97 | in if module_name = module_name' | 
| 98 | then (Graph.map_node module_name o apfst) (Graph.add_edge (name, name')) | |
| 99 | else (Graph.map_node module_name o apsnd) (AList.map_default (op =) (module_name', []) (insert (op =) name')) | |
| 100 | end; | |
| 101 | val proto_program = Graph.empty | |
| 102 | |> Graph.fold (fn (name, (stmt, _)) => add_stmt name stmt) program | |
| 44338 
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
 wenzelm parents: 
43326diff
changeset | 103 | |> Graph.fold (fn (name, (_, (_, names))) => | 
| 
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
 wenzelm parents: 
43326diff
changeset | 104 | Graph.Keys.fold (add_dependency name) names) program; | 
| 39208 | 105 | |
| 106 | (* name declarations and statement modifications *) | |
| 107 | fun declare name (base, stmt) (gr, nsp) = | |
| 108 | let | |
| 109 | val (base', nsp') = namify_stmt stmt base nsp; | |
| 110 | val gr' = (Graph.map_node name o apfst) (K base') gr; | |
| 111 | in (gr', nsp') end; | |
| 112 | fun declarations gr = (gr, empty_nsp) | |
| 113 | |> fold (fn name => declare name (Graph.get_node gr name)) (Graph.keys gr) | |
| 114 | |> fst | |
| 115 | |> (Graph.map o K o apsnd) modify_stmt; | |
| 116 | val flat_program = proto_program | |
| 117 | |> (Graph.map o K o apfst) declarations; | |
| 118 | ||
| 119 | (* qualified and unqualified imports, deresolving *) | |
| 120 | fun base_deresolver name = fst (Graph.get_node | |
| 52138 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
51930diff
changeset | 121 | (fst (Graph.get_node flat_program (fst (prep_name name)))) name); | 
| 39208 | 122 | fun classify_names gr imports = | 
| 123 | let | |
| 124 | val import_tab = maps | |
| 125 | (fn (module_name, names) => map (rpair module_name) names) imports; | |
| 126 | val imported_names = map fst import_tab; | |
| 127 | val here_names = Graph.keys gr; | |
| 128 | in | |
| 129 | Symtab.empty | |
| 130 | |> fold (fn name => Symtab.update (name, base_deresolver name)) here_names | |
| 131 | |> fold (fn name => Symtab.update (name, | |
| 132 | Long_Name.append (the (AList.lookup (op =) import_tab name)) | |
| 133 | (base_deresolver name))) imported_names | |
| 134 | end; | |
| 40630 | 135 | val deresolver_tab = Symtab.make (AList.make | 
| 136 | (uncurry classify_names o Graph.get_node flat_program) | |
| 137 | (Graph.keys flat_program)); | |
| 40705 | 138 | fun deresolver "" name = | 
| 52138 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
51930diff
changeset | 139 | Long_Name.append (fst (prep_name name)) (base_deresolver name) | 
| 40705 | 140 | | deresolver module_name name = | 
| 141 | the (Symtab.lookup (the (Symtab.lookup deresolver_tab module_name)) name) | |
| 52138 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
51930diff
changeset | 142 |           handle Option.Option => error ("Unknown statement name: "
 | 
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
51930diff
changeset | 143 | ^ (Code_Symbol.quote_symbol ctxt o symbol_of) name); | 
| 39208 | 144 | |
| 145 |   in { deresolver = deresolver, flat_program = flat_program } end;
 | |
| 146 | ||
| 147 | ||
| 39055 | 148 | (** hierarchical program structure **) | 
| 38970 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 149 | |
| 39017 
8cd5b6d688fa
generalized hierarchical data structure over statements
 haftmann parents: 
38970diff
changeset | 150 | datatype ('a, 'b) node =
 | 
| 38970 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 151 | Dummy | 
| 39017 
8cd5b6d688fa
generalized hierarchical data structure over statements
 haftmann parents: 
38970diff
changeset | 152 | | Stmt of 'a | 
| 
8cd5b6d688fa
generalized hierarchical data structure over statements
 haftmann parents: 
38970diff
changeset | 153 |   | Module of ('b * (string * ('a, 'b) node) Graph.T);
 | 
| 38970 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 154 | |
| 39147 | 155 | type ('a, 'b) hierarchical_program = (string * ('a, 'b) node) Graph.T;
 | 
| 156 | ||
| 39018 | 157 | fun map_module_content f (Module content) = Module (f content); | 
| 158 | ||
| 159 | fun map_module [] = I | |
| 160 | | map_module (name_fragment :: name_fragments) = | |
| 161 | apsnd o Graph.map_node name_fragment o apsnd o map_module_content | |
| 162 | o map_module name_fragments; | |
| 163 | ||
| 52138 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
51930diff
changeset | 164 | fun hierarchical_program ctxt symbol_of { module_name, reserved, identifiers, empty_nsp,
 | 
| 39023 | 165 | namify_module, namify_stmt, cyclic_modules, empty_data, memorize_data, modify_stmts } program = | 
| 38970 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 166 | let | 
| 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 167 | |
| 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 168 | (* building module name hierarchy *) | 
| 52138 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
51930diff
changeset | 169 | val module_identifiers = if module_name = "" | 
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
51930diff
changeset | 170 | then Code_Symbol.lookup_module_data identifiers | 
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
51930diff
changeset | 171 | else K (SOME module_name); | 
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
51930diff
changeset | 172 |     val fragments_tab = build_module_namespace { module_prefix = "",
 | 
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
51930diff
changeset | 173 | module_identifiers = module_identifiers, reserved = reserved } program; | 
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
51930diff
changeset | 174 | val prep_name = force_identifier symbol_of fragments_tab identifiers; | 
| 38970 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 175 | |
| 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 176 | (* building empty module hierarchy *) | 
| 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 177 | val empty_module = (empty_data, Graph.empty); | 
| 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 178 | fun ensure_module name_fragment (data, nodes) = | 
| 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 179 | if can (Graph.get_node nodes) name_fragment then (data, nodes) | 
| 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 180 | else (data, | 
| 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 181 | nodes |> Graph.new_node (name_fragment, (name_fragment, Module empty_module))); | 
| 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 182 | fun allocate_module [] = I | 
| 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 183 | | allocate_module (name_fragment :: name_fragments) = | 
| 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 184 | ensure_module name_fragment | 
| 39018 | 185 | #> (apsnd o Graph.map_node name_fragment o apsnd o map_module_content o allocate_module) name_fragments; | 
| 52138 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
51930diff
changeset | 186 | val empty_program = | 
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
51930diff
changeset | 187 | empty_module | 
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
51930diff
changeset | 188 | |> Symtab.fold (fn (_, fragments) => allocate_module fragments) fragments_tab | 
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
51930diff
changeset | 189 | |> Graph.fold (allocate_module o these o Option.map fst | 
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
51930diff
changeset | 190 | o lookup_identifier symbol_of identifiers o fst) program; | 
| 38970 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 191 | |
| 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 192 | (* distribute statements over hierarchy *) | 
| 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 193 | fun add_stmt name stmt = | 
| 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 194 | let | 
| 52138 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
51930diff
changeset | 195 | val (name_fragments, base) = prep_name name; | 
| 38970 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 196 | in | 
| 39018 | 197 | (map_module name_fragments o apsnd) (Graph.new_node (name, (base, Stmt stmt))) | 
| 38970 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 198 | end; | 
| 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 199 | fun add_dependency name name' = | 
| 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 200 | let | 
| 52138 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
51930diff
changeset | 201 | val (name_fragments, _) = prep_name name; | 
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
51930diff
changeset | 202 | val (name_fragments', _) = prep_name name'; | 
| 38970 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 203 | val (name_fragments_common, (diff, diff')) = | 
| 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 204 | chop_prefix (op =) (name_fragments, name_fragments'); | 
| 39023 | 205 | val is_module = not (null diff andalso null diff'); | 
| 206 | val dep = pairself hd (diff @ [name], diff' @ [name']); | |
| 38970 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 207 | val add_edge = if is_module andalso not cyclic_modules | 
| 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 208 | then (fn node => Graph.add_edge_acyclic dep node | 
| 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 209 |             handle Graph.CYCLES _ => error ("Dependency "
 | 
| 52138 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
51930diff
changeset | 210 | ^ (Code_Symbol.quote_symbol ctxt o symbol_of) name ^ " -> " | 
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
51930diff
changeset | 211 | ^ (Code_Symbol.quote_symbol ctxt o symbol_of) name' | 
| 38970 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 212 | ^ " would result in module dependency cycle")) | 
| 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 213 | else Graph.add_edge dep | 
| 39018 | 214 | in (map_module name_fragments_common o apsnd) add_edge end; | 
| 38970 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 215 | val proto_program = empty_program | 
| 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 216 | |> Graph.fold (fn (name, (stmt, _)) => add_stmt name stmt) program | 
| 44338 
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
 wenzelm parents: 
43326diff
changeset | 217 | |> Graph.fold (fn (name, (_, (_, names))) => | 
| 
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
 wenzelm parents: 
43326diff
changeset | 218 | Graph.Keys.fold (add_dependency name) names) program; | 
| 38970 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 219 | |
| 39022 | 220 | (* name declarations, data and statement modifications *) | 
| 38970 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 221 | fun make_declarations nsps (data, nodes) = | 
| 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 222 | let | 
| 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 223 | val (module_fragments, stmt_names) = List.partition | 
| 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 224 | (fn name_fragment => case Graph.get_node nodes name_fragment | 
| 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 225 | of (_, Module _) => true | _ => false) (Graph.keys nodes); | 
| 39022 | 226 | fun declare namify name (nsps, nodes) = | 
| 38970 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 227 | let | 
| 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 228 | val (base, node) = Graph.get_node nodes name; | 
| 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 229 | val (base', nsps') = namify node base nsps; | 
| 39022 | 230 | val nodes' = Graph.map_node name (K (base', node)) nodes; | 
| 38970 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 231 | in (nsps', nodes') end; | 
| 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 232 | val (nsps', nodes') = (nsps, nodes) | 
| 39022 | 233 | |> fold (declare (K namify_module)) module_fragments | 
| 234 | |> fold (declare (namify_stmt o (fn Stmt stmt => stmt))) stmt_names; | |
| 39029 
cef7b58555aa
skip empty name bunches; fill up trailing positions with NONEs
 haftmann parents: 
39024diff
changeset | 235 | fun zip_fillup xs ys = xs ~~ ys @ replicate (length xs - length ys) NONE; | 
| 
cef7b58555aa
skip empty name bunches; fill up trailing positions with NONEs
 haftmann parents: 
39024diff
changeset | 236 | fun select_names names = case filter (member (op =) stmt_names) names | 
| 
cef7b58555aa
skip empty name bunches; fill up trailing positions with NONEs
 haftmann parents: 
39024diff
changeset | 237 | of [] => NONE | 
| 
cef7b58555aa
skip empty name bunches; fill up trailing positions with NONEs
 haftmann parents: 
39024diff
changeset | 238 | | xs => SOME xs; | 
| 
cef7b58555aa
skip empty name bunches; fill up trailing positions with NONEs
 haftmann parents: 
39024diff
changeset | 239 | val modify_stmts' = AList.make (snd o Graph.get_node nodes) | 
| 39023 | 240 | #> split_list | 
| 241 | ##> map (fn Stmt stmt => stmt) | |
| 39029 
cef7b58555aa
skip empty name bunches; fill up trailing positions with NONEs
 haftmann parents: 
39024diff
changeset | 242 | #> (fn (names, stmts) => zip_fillup names (modify_stmts (names ~~ stmts))); | 
| 
cef7b58555aa
skip empty name bunches; fill up trailing positions with NONEs
 haftmann parents: 
39024diff
changeset | 243 | val stmtss' = (maps modify_stmts' o map_filter select_names o Graph.strong_conn) nodes; | 
| 39023 | 244 | val nodes'' = Graph.map (fn name => apsnd (fn Module content => Module (make_declarations nsps' content) | 
| 245 | | _ => case AList.lookup (op =) stmtss' name of SOME (SOME stmt) => Stmt stmt | _ => Dummy)) nodes'; | |
| 39018 | 246 | val data' = fold memorize_data stmt_names data; | 
| 247 | in (data', nodes'') end; | |
| 38970 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 248 | val (_, hierarchical_program) = make_declarations empty_nsp proto_program; | 
| 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 249 | |
| 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 250 | (* deresolving *) | 
| 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 251 | fun deresolver prefix_fragments name = | 
| 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 252 | let | 
| 52138 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
51930diff
changeset | 253 | val (name_fragments, _) = prep_name name; | 
| 38970 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 254 | val (_, (_, remainder)) = chop_prefix (op =) (prefix_fragments, name_fragments); | 
| 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 255 | val nodes = fold (fn name_fragment => fn nodes => case Graph.get_node nodes name_fragment | 
| 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 256 | of (_, Module (_, nodes)) => nodes) name_fragments hierarchical_program; | 
| 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 257 | val (base', _) = Graph.get_node nodes name; | 
| 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 258 | in Long_Name.implode (remainder @ [base']) end | 
| 52138 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
51930diff
changeset | 259 |         handle Graph.UNDEF _ => error ("Unknown statement name: "
 | 
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
51930diff
changeset | 260 | ^ (Code_Symbol.quote_symbol ctxt o symbol_of) name); | 
| 38970 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 261 | |
| 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 262 |   in { deresolver = deresolver, hierarchical_program = hierarchical_program } end;
 | 
| 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 263 | |
| 39147 | 264 | fun print_hierarchical { print_module, print_stmt, lift_markup } =
 | 
| 265 | let | |
| 266 | fun print_node _ (_, Dummy) = | |
| 267 | NONE | |
| 268 | | print_node prefix_fragments (name, Stmt stmt) = | |
| 269 | SOME (lift_markup (Code_Printer.markup_stmt name) | |
| 270 | (print_stmt prefix_fragments (name, stmt))) | |
| 271 | | print_node prefix_fragments (name_fragment, Module (data, nodes)) = | |
| 272 | let | |
| 273 | val prefix_fragments' = prefix_fragments @ [name_fragment] | |
| 274 | in | |
| 275 | Option.map (print_module prefix_fragments' | |
| 276 | name_fragment data) (print_nodes prefix_fragments' nodes) | |
| 277 | end | |
| 278 | and print_nodes prefix_fragments nodes = | |
| 279 | let | |
| 280 | val xs = (map_filter (fn name => print_node prefix_fragments | |
| 281 | (name, snd (Graph.get_node nodes name))) o rev o flat o Graph.strong_conn) nodes | |
| 282 | in if null xs then NONE else SOME xs end; | |
| 283 | in these o print_nodes [] end; | |
| 284 | ||
| 38970 
53d1ee3d98b8
factored out generic part of Scala serializer into code_namespace.ML
 haftmann parents: diff
changeset | 285 | end; |