src/Tools/Code/code_namespace.ML
author haftmann
Sun, 23 Feb 2014 10:33:43 +0100
changeset 55680 bc5edc5dbf18
parent 55679 59244fc1a7ca
child 55681 7714287dc044
permissions -rw-r--r--
tuned
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
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
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
     9
  type flat_program
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 53414
diff changeset
    10
  val flat_program: Proof.context
52138
e21426f244aa bookkeeping and input syntax for exact specification of names of symbols in generated code
haftmann
parents: 51930
diff 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: 51930
diff changeset
    12
    reserved: Name.context, identifiers: Code_Target.identifier_data, empty_nsp: 'a,
39208
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
    13
    namify_stmt: Code_Thingol.stmt -> string -> 'a -> string * 'a,
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
    14
    modify_stmt: Code_Thingol.stmt -> Code_Thingol.stmt option }
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
    15
      -> Code_Thingol.program
55150
0940309ed8f1 less clumsy namespace
haftmann
parents: 55147
diff changeset
    16
      -> { deresolver: string -> Code_Symbol.T -> string,
39208
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
    17
           flat_program: flat_program }
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
    18
39017
8cd5b6d688fa generalized hierarchical data structure over statements
haftmann
parents: 38970
diff 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
55679
59244fc1a7ca formal markup for public ingredients
haftmann
parents: 55605
diff changeset
    21
    | Stmt of bool * 'a
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 53414
diff changeset
    22
    | Module of ('b * (string * ('a, 'b) node) Code_Symbol.Graph.T)
39208
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
    23
  type ('a, 'b) hierarchical_program
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 53414
diff changeset
    24
  val hierarchical_program: Proof.context
52138
e21426f244aa bookkeeping and input syntax for exact specification of names of symbols in generated code
haftmann
parents: 51930
diff changeset
    25
    -> { module_name: string,
e21426f244aa bookkeeping and input syntax for exact specification of names of symbols in generated code
haftmann
parents: 51930
diff 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: 51930
diff changeset
    27
    empty_nsp: 'c, namify_module: string -> 'c -> string * 'c,
39022
ac7774a35bcf explicit modify_stmt parameter
haftmann
parents: 39018
diff changeset
    28
    namify_stmt: Code_Thingol.stmt -> string -> 'c -> string * 'c,
55150
0940309ed8f1 less clumsy namespace
haftmann
parents: 55147
diff changeset
    29
    cyclic_modules: bool, empty_data: 'b, memorize_data: Code_Symbol.T -> 'b -> 'b,
0940309ed8f1 less clumsy namespace
haftmann
parents: 55147
diff changeset
    30
    modify_stmts: (Code_Symbol.T * 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
55150
0940309ed8f1 less clumsy namespace
haftmann
parents: 55147
diff changeset
    32
      -> { deresolver: string list -> Code_Symbol.T -> string,
39147
3c284a152bd6 printing combinator for hierarchical programs
haftmann
parents: 39055
diff changeset
    33
           hierarchical_program: ('a, 'b) hierarchical_program }
3c284a152bd6 printing combinator for hierarchical programs
haftmann
parents: 39055
diff changeset
    34
  val print_hierarchical: { print_module: string list -> string -> 'b -> 'c list -> 'c,
55679
59244fc1a7ca formal markup for public ingredients
haftmann
parents: 55605
diff changeset
    35
    print_stmt: string list -> Code_Symbol.T * (bool * 'a) -> 'c,
39147
3c284a152bd6 printing combinator for hierarchical programs
haftmann
parents: 39055
diff changeset
    36
    lift_markup: (Pretty.T -> Pretty.T) -> 'c -> 'c }
3c284a152bd6 printing combinator for hierarchical programs
haftmann
parents: 39055
diff changeset
    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: 51930
diff changeset
    43
(** fundamental module name hierarchy **)
e21426f244aa bookkeeping and input syntax for exact specification of names of symbols in generated code
haftmann
parents: 51930
diff changeset
    44
55291
82780e5f7622 tuned storage of code identifiers
haftmann
parents: 55150
diff changeset
    45
fun module_fragments' { identifiers, reserved } name =
82780e5f7622 tuned storage of code identifiers
haftmann
parents: 55150
diff changeset
    46
  case Code_Symbol.lookup_module_data identifiers name of
82780e5f7622 tuned storage of code identifiers
haftmann
parents: 55150
diff changeset
    47
      SOME (fragments, _) => fragments
82780e5f7622 tuned storage of code identifiers
haftmann
parents: 55150
diff changeset
    48
    | NONE => map (fn fragment => fst (Name.variant fragment reserved)) (Long_Name.explode name);
82780e5f7622 tuned storage of code identifiers
haftmann
parents: 55150
diff changeset
    49
82780e5f7622 tuned storage of code identifiers
haftmann
parents: 55150
diff changeset
    50
fun module_fragments { module_name, identifiers, reserved } =
82780e5f7622 tuned storage of code identifiers
haftmann
parents: 55150
diff changeset
    51
  if module_name = ""
82780e5f7622 tuned storage of code identifiers
haftmann
parents: 55150
diff changeset
    52
  then module_fragments' { identifiers = identifiers, reserved = reserved }
82780e5f7622 tuned storage of code identifiers
haftmann
parents: 55150
diff changeset
    53
  else K (Long_Name.explode module_name);
39055
81e0368812ad removed namespace stuff from code_printer
haftmann
parents: 39029
diff changeset
    54
55291
82780e5f7622 tuned storage of code identifiers
haftmann
parents: 55150
diff changeset
    55
fun build_module_namespace ctxt { module_prefix, module_name, identifiers, reserved } program =
82780e5f7622 tuned storage of code identifiers
haftmann
parents: 55150
diff changeset
    56
  let
82780e5f7622 tuned storage of code identifiers
haftmann
parents: 55150
diff changeset
    57
    val module_names = Code_Symbol.Graph.fold (insert (op =) o Code_Symbol.default_prefix ctxt o fst) program [];
82780e5f7622 tuned storage of code identifiers
haftmann
parents: 55150
diff changeset
    58
    val module_fragments' = module_fragments
82780e5f7622 tuned storage of code identifiers
haftmann
parents: 55150
diff changeset
    59
      { module_name = module_name, identifiers = identifiers, reserved = reserved };
82780e5f7622 tuned storage of code identifiers
haftmann
parents: 55150
diff changeset
    60
  in
82780e5f7622 tuned storage of code identifiers
haftmann
parents: 55150
diff changeset
    61
    fold (fn name => Symtab.update (name, Long_Name.explode module_prefix @ module_fragments' name))
82780e5f7622 tuned storage of code identifiers
haftmann
parents: 55150
diff changeset
    62
      module_names Symtab.empty
82780e5f7622 tuned storage of code identifiers
haftmann
parents: 55150
diff changeset
    63
  end;
82780e5f7622 tuned storage of code identifiers
haftmann
parents: 55150
diff changeset
    64
55292
haftmann
parents: 55291
diff changeset
    65
fun prep_symbol ctxt { module_namespace, force_module, identifiers } sym =
55291
82780e5f7622 tuned storage of code identifiers
haftmann
parents: 55150
diff changeset
    66
  case Code_Symbol.lookup identifiers sym of
55292
haftmann
parents: 55291
diff changeset
    67
      NONE => ((the o Symtab.lookup module_namespace o Code_Symbol.default_prefix ctxt) sym,
haftmann
parents: 55291
diff changeset
    68
        Code_Symbol.default_base sym)
53414
dd64696d267a explicit module names have precedence over identifier declarations
haftmann
parents: 52138
diff changeset
    69
    | SOME prefix_name => if null force_module then prefix_name
dd64696d267a explicit module names have precedence over identifier declarations
haftmann
parents: 52138
diff changeset
    70
        else (force_module, snd prefix_name);
39055
81e0368812ad removed namespace stuff from code_printer
haftmann
parents: 39029
diff changeset
    71
55293
42cf5802d36a code generation: explicitly declared identifiers gain predence over implicit ones
haftmann
parents: 55292
diff changeset
    72
fun has_priority identifiers = is_some o Code_Symbol.lookup identifiers;
42cf5802d36a code generation: explicitly declared identifiers gain predence over implicit ones
haftmann
parents: 55292
diff changeset
    73
55292
haftmann
parents: 55291
diff changeset
    74
fun build_proto_program { empty, add_stmt, add_dep } program =
haftmann
parents: 55291
diff changeset
    75
  empty
haftmann
parents: 55291
diff changeset
    76
  |> Code_Symbol.Graph.fold (fn (sym, (stmt, _)) => add_stmt sym stmt) program
haftmann
parents: 55291
diff changeset
    77
  |> Code_Symbol.Graph.fold (fn (sym, (_, (_, syms))) =>
haftmann
parents: 55291
diff changeset
    78
      Code_Symbol.Graph.Keys.fold (add_dep sym) syms) program;
haftmann
parents: 55291
diff changeset
    79
55293
42cf5802d36a code generation: explicitly declared identifiers gain predence over implicit ones
haftmann
parents: 55292
diff changeset
    80
fun prioritize has_priority = uncurry append o List.partition has_priority;
42cf5802d36a code generation: explicitly declared identifiers gain predence over implicit ones
haftmann
parents: 55292
diff changeset
    81
39055
81e0368812ad removed namespace stuff from code_printer
haftmann
parents: 39029
diff changeset
    82
39208
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
    83
(** flat program structure **)
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
    84
55679
59244fc1a7ca formal markup for public ingredients
haftmann
parents: 55605
diff changeset
    85
type flat_program = ((string * (bool * Code_Thingol.stmt) option) Code_Symbol.Graph.T * (string * Code_Symbol.T list) list) Graph.T;
39208
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
    86
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 53414
diff changeset
    87
fun flat_program ctxt { module_prefix, module_name, reserved,
52138
e21426f244aa bookkeeping and input syntax for exact specification of names of symbols in generated code
haftmann
parents: 51930
diff changeset
    88
    identifiers, empty_nsp, namify_stmt, modify_stmt } program =
39208
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
    89
  let
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
    90
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
    91
    (* building module name hierarchy *)
55291
82780e5f7622 tuned storage of code identifiers
haftmann
parents: 55150
diff changeset
    92
    val module_namespace = build_module_namespace ctxt { module_prefix = module_prefix,
82780e5f7622 tuned storage of code identifiers
haftmann
parents: 55150
diff changeset
    93
      module_name = module_name, identifiers = identifiers, reserved = reserved } program;
55292
haftmann
parents: 55291
diff changeset
    94
    val prep_sym = prep_symbol ctxt { module_namespace = module_namespace,
haftmann
parents: 55291
diff changeset
    95
      force_module = Long_Name.explode module_name, identifiers = identifiers }
52138
e21426f244aa bookkeeping and input syntax for exact specification of names of symbols in generated code
haftmann
parents: 51930
diff changeset
    96
      #>> Long_Name.implode;
55293
42cf5802d36a code generation: explicitly declared identifiers gain predence over implicit ones
haftmann
parents: 55292
diff changeset
    97
    val sym_priority = has_priority identifiers;
39208
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
    98
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
    99
    (* distribute statements over hierarchy *)
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 53414
diff changeset
   100
    fun add_stmt sym stmt =
39208
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
   101
      let
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 53414
diff changeset
   102
        val (module_name, base) = prep_sym sym;
39208
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
   103
      in
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 53414
diff changeset
   104
        Graph.default_node (module_name, (Code_Symbol.Graph.empty, []))
55680
haftmann
parents: 55679
diff changeset
   105
        #> (Graph.map_node module_name o apfst)
haftmann
parents: 55679
diff changeset
   106
          (Code_Symbol.Graph.new_node (sym, (base, (true, stmt))))
39208
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
   107
      end;
55292
haftmann
parents: 55291
diff changeset
   108
    fun add_dep sym sym' =
39208
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
   109
      let
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 53414
diff changeset
   110
        val (module_name, _) = prep_sym sym;
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 53414
diff changeset
   111
        val (module_name', _) = prep_sym sym';
39208
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
   112
      in if module_name = module_name'
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 53414
diff changeset
   113
        then (Graph.map_node module_name o apfst) (Code_Symbol.Graph.add_edge (sym, sym'))
55680
haftmann
parents: 55679
diff changeset
   114
        else (Graph.map_node module_name o apsnd)
haftmann
parents: 55679
diff changeset
   115
          (AList.map_default (op =) (module_name', []) (insert (op =) sym'))
39208
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
   116
      end;
55292
haftmann
parents: 55291
diff changeset
   117
    val proto_program = build_proto_program
haftmann
parents: 55291
diff changeset
   118
      { empty = Graph.empty, add_stmt = add_stmt, add_dep = add_dep } program;
39208
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
   119
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
   120
    (* name declarations and statement modifications *)
55679
59244fc1a7ca formal markup for public ingredients
haftmann
parents: 55605
diff changeset
   121
    fun declare sym (base, (_, stmt)) (gr, nsp) = 
39208
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
   122
      let
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
   123
        val (base', nsp') = namify_stmt stmt base nsp;
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 53414
diff changeset
   124
        val gr' = (Code_Symbol.Graph.map_node sym o apfst) (K base') gr;
39208
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
   125
      in (gr', nsp') end;
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
   126
    fun declarations gr = (gr, empty_nsp)
55293
42cf5802d36a code generation: explicitly declared identifiers gain predence over implicit ones
haftmann
parents: 55292
diff changeset
   127
      |> fold (fn sym => declare sym (Code_Symbol.Graph.get_node gr sym))
42cf5802d36a code generation: explicitly declared identifiers gain predence over implicit ones
haftmann
parents: 55292
diff changeset
   128
          (prioritize sym_priority (Code_Symbol.Graph.keys gr))
39208
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
   129
      |> fst
55680
haftmann
parents: 55679
diff changeset
   130
      |> Code_Symbol.Graph.map_strong_conn (fn syms_bases_exports_stmts =>
haftmann
parents: 55679
diff changeset
   131
        map snd syms_bases_exports_stmts
haftmann
parents: 55679
diff changeset
   132
        |> (map o apsnd) (fn (export, stmt) => Option.map (pair export) (modify_stmt stmt)));
39208
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
   133
    val flat_program = proto_program
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
   134
      |> (Graph.map o K o apfst) declarations;
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
   135
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
   136
    (* qualified and unqualified imports, deresolving *)
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 53414
diff changeset
   137
    fun base_deresolver sym = fst (Code_Symbol.Graph.get_node
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 53414
diff changeset
   138
      (fst (Graph.get_node flat_program (fst (prep_sym sym)))) sym);
39208
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
   139
    fun classify_names gr imports =
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
   140
      let
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
   141
        val import_tab = maps
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 53414
diff changeset
   142
          (fn (module_name, syms) => map (rpair module_name) syms) imports;
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 53414
diff changeset
   143
        val imported_syms = map fst import_tab;
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 53414
diff changeset
   144
        val here_syms = Code_Symbol.Graph.keys gr;
39208
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
   145
      in
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 53414
diff changeset
   146
        Code_Symbol.Table.empty
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 53414
diff changeset
   147
        |> fold (fn sym => Code_Symbol.Table.update (sym, base_deresolver sym)) here_syms
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 53414
diff changeset
   148
        |> fold (fn sym => Code_Symbol.Table.update (sym,
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 53414
diff changeset
   149
            Long_Name.append (the (AList.lookup (op =) import_tab sym))
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 53414
diff changeset
   150
              (base_deresolver sym))) imported_syms
39208
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
   151
      end;
40630
haftmann
parents: 40629
diff changeset
   152
    val deresolver_tab = Symtab.make (AList.make
haftmann
parents: 40629
diff changeset
   153
      (uncurry classify_names o Graph.get_node flat_program)
haftmann
parents: 40629
diff changeset
   154
        (Graph.keys flat_program));
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 53414
diff changeset
   155
    fun deresolver "" sym =
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 53414
diff changeset
   156
          Long_Name.append (fst (prep_sym sym)) (base_deresolver sym)
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 53414
diff changeset
   157
      | deresolver module_name sym =
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 53414
diff changeset
   158
          the (Code_Symbol.Table.lookup (the (Symtab.lookup deresolver_tab module_name)) sym)
52138
e21426f244aa bookkeeping and input syntax for exact specification of names of symbols in generated code
haftmann
parents: 51930
diff changeset
   159
          handle Option.Option => error ("Unknown statement name: "
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 53414
diff changeset
   160
            ^ Code_Symbol.quote ctxt sym);
39208
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
   161
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
   162
  in { deresolver = deresolver, flat_program = flat_program } end;
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
   163
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
   164
39055
81e0368812ad removed namespace stuff from code_printer
haftmann
parents: 39029
diff changeset
   165
(** hierarchical program structure **)
38970
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   166
39017
8cd5b6d688fa generalized hierarchical data structure over statements
haftmann
parents: 38970
diff changeset
   167
datatype ('a, 'b) node =
38970
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   168
    Dummy
55679
59244fc1a7ca formal markup for public ingredients
haftmann
parents: 55605
diff changeset
   169
  | Stmt of bool * 'a
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 53414
diff changeset
   170
  | Module of ('b * (string * ('a, 'b) node) Code_Symbol.Graph.T);
38970
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   171
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 53414
diff changeset
   172
type ('a, 'b) hierarchical_program = (string * ('a, 'b) node) Code_Symbol.Graph.T;
39147
3c284a152bd6 printing combinator for hierarchical programs
haftmann
parents: 39055
diff changeset
   173
55679
59244fc1a7ca formal markup for public ingredients
haftmann
parents: 55605
diff changeset
   174
fun the_stmt (Stmt (export, stmt)) = (export, stmt);
59244fc1a7ca formal markup for public ingredients
haftmann
parents: 55605
diff changeset
   175
39018
haftmann
parents: 39017
diff changeset
   176
fun map_module_content f (Module content) = Module (f content);
haftmann
parents: 39017
diff changeset
   177
haftmann
parents: 39017
diff changeset
   178
fun map_module [] = I
haftmann
parents: 39017
diff changeset
   179
  | map_module (name_fragment :: name_fragments) =
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 53414
diff changeset
   180
      apsnd o Code_Symbol.Graph.map_node (Code_Symbol.Module name_fragment) o apsnd o map_module_content
39018
haftmann
parents: 39017
diff changeset
   181
        o map_module name_fragments;
haftmann
parents: 39017
diff changeset
   182
55680
haftmann
parents: 55679
diff changeset
   183
fun map_module_stmts f_module f_stmts sym_base_nodes =
haftmann
parents: 55679
diff changeset
   184
  let
haftmann
parents: 55679
diff changeset
   185
    val some_modules =
haftmann
parents: 55679
diff changeset
   186
      sym_base_nodes
haftmann
parents: 55679
diff changeset
   187
      |> map (fn (sym, (base, Module content)) => SOME (base, content) | _ => NONE)
haftmann
parents: 55679
diff changeset
   188
      |> (burrow_options o map o apsnd) f_module;
haftmann
parents: 55679
diff changeset
   189
    val some_export_stmts =
haftmann
parents: 55679
diff changeset
   190
      sym_base_nodes
haftmann
parents: 55679
diff changeset
   191
      |> map (fn (sym, (base, Stmt export_stmt)) => SOME ((sym, export_stmt), base) | _ => NONE)
haftmann
parents: 55679
diff changeset
   192
      |> (burrow_options o burrow_fst) (fn [] => [] | xs => f_stmts xs)
haftmann
parents: 55679
diff changeset
   193
  in
haftmann
parents: 55679
diff changeset
   194
    map2 (fn SOME (base, content) => (K (base, Module content))
haftmann
parents: 55679
diff changeset
   195
      | NONE => fn SOME (some_export_stmt, base) =>
haftmann
parents: 55679
diff changeset
   196
          (base, case some_export_stmt of SOME export_stmt => Stmt export_stmt | NONE => Dummy))
haftmann
parents: 55679
diff changeset
   197
      some_modules some_export_stmts
haftmann
parents: 55679
diff changeset
   198
  end;
haftmann
parents: 55679
diff changeset
   199
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 53414
diff changeset
   200
fun hierarchical_program ctxt { module_name, reserved, identifiers, empty_nsp,
39023
3f70c03e8282 simultaneous modification of statements
haftmann
parents: 39022
diff changeset
   201
      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
   202
  let
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   203
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   204
    (* building module name hierarchy *)
55291
82780e5f7622 tuned storage of code identifiers
haftmann
parents: 55150
diff changeset
   205
    val module_namespace = build_module_namespace ctxt { module_prefix = "",
82780e5f7622 tuned storage of code identifiers
haftmann
parents: 55150
diff changeset
   206
      module_name = module_name, identifiers = identifiers, reserved = reserved } program;
55292
haftmann
parents: 55291
diff changeset
   207
    val prep_sym = prep_symbol ctxt { module_namespace = module_namespace,
55293
42cf5802d36a code generation: explicitly declared identifiers gain predence over implicit ones
haftmann
parents: 55292
diff changeset
   208
      force_module = Long_Name.explode module_name, identifiers = identifiers }
42cf5802d36a code generation: explicitly declared identifiers gain predence over implicit ones
haftmann
parents: 55292
diff changeset
   209
    val sym_priority = has_priority identifiers;
38970
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   210
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   211
    (* building empty module hierarchy *)
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 53414
diff changeset
   212
    val empty_module = (empty_data, Code_Symbol.Graph.empty);
38970
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   213
    fun ensure_module name_fragment (data, nodes) =
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 53414
diff changeset
   214
      if can (Code_Symbol.Graph.get_node nodes) (Code_Symbol.Module name_fragment) then (data, nodes)
38970
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   215
      else (data,
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 53414
diff changeset
   216
        nodes |> Code_Symbol.Graph.new_node (Code_Symbol.Module name_fragment, (name_fragment, Module empty_module)));
38970
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   217
    fun allocate_module [] = I
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   218
      | allocate_module (name_fragment :: name_fragments) =
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   219
          ensure_module name_fragment
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 53414
diff changeset
   220
          #> (apsnd o Code_Symbol.Graph.map_node (Code_Symbol.Module 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: 51930
diff changeset
   221
    val empty_program =
e21426f244aa bookkeeping and input syntax for exact specification of names of symbols in generated code
haftmann
parents: 51930
diff changeset
   222
      empty_module
55291
82780e5f7622 tuned storage of code identifiers
haftmann
parents: 55150
diff changeset
   223
      |> Symtab.fold (fn (_, fragments) => allocate_module fragments) module_namespace
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 53414
diff changeset
   224
      |> Code_Symbol.Graph.fold (allocate_module o these o Option.map fst
55291
82780e5f7622 tuned storage of code identifiers
haftmann
parents: 55150
diff changeset
   225
          o Code_Symbol.lookup identifiers o fst) program;
38970
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   226
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   227
    (* distribute statements over hierarchy *)
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 53414
diff changeset
   228
    fun add_stmt sym stmt =
38970
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   229
      let
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 53414
diff changeset
   230
        val (name_fragments, base) = prep_sym sym;
38970
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   231
      in
55680
haftmann
parents: 55679
diff changeset
   232
        (map_module name_fragments o apsnd)
haftmann
parents: 55679
diff changeset
   233
          (Code_Symbol.Graph.new_node (sym, (base, Stmt (true, stmt))))
38970
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   234
      end;
55680
haftmann
parents: 55679
diff changeset
   235
    fun add_edge_acyclic_error error_msg dep gr =
haftmann
parents: 55679
diff changeset
   236
      Code_Symbol.Graph.add_edge_acyclic dep gr
haftmann
parents: 55679
diff changeset
   237
        handle Graph.CYCLES _ => error (error_msg ())
55292
haftmann
parents: 55291
diff changeset
   238
    fun add_dep sym sym' =
38970
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   239
      let
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 53414
diff changeset
   240
        val (name_fragments, _) = prep_sym sym;
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 53414
diff changeset
   241
        val (name_fragments', _) = prep_sym sym';
38970
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   242
        val (name_fragments_common, (diff, diff')) =
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   243
          chop_prefix (op =) (name_fragments, name_fragments');
55680
haftmann
parents: 55679
diff changeset
   244
        val is_cross_module = not (null diff andalso null diff');
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 53414
diff changeset
   245
        val dep = pairself hd (map Code_Symbol.Module diff @ [sym], map Code_Symbol.Module diff' @ [sym']);
55680
haftmann
parents: 55679
diff changeset
   246
        val add_edge = if is_cross_module andalso not cyclic_modules
haftmann
parents: 55679
diff changeset
   247
          then add_edge_acyclic_error (fn _ => "Dependency "
haftmann
parents: 55679
diff changeset
   248
            ^ Code_Symbol.quote ctxt sym ^ " -> "
haftmann
parents: 55679
diff changeset
   249
            ^ Code_Symbol.quote ctxt sym'
haftmann
parents: 55679
diff changeset
   250
            ^ " would result in module dependency cycle") dep
haftmann
parents: 55679
diff changeset
   251
          else Code_Symbol.Graph.add_edge dep;
39018
haftmann
parents: 39017
diff changeset
   252
      in (map_module name_fragments_common o apsnd) add_edge end;
55292
haftmann
parents: 55291
diff changeset
   253
    val proto_program = build_proto_program
haftmann
parents: 55291
diff changeset
   254
      { empty = empty_program, add_stmt = add_stmt, add_dep = add_dep } program;
38970
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   255
39022
ac7774a35bcf explicit modify_stmt parameter
haftmann
parents: 39018
diff changeset
   256
    (* name declarations, data and statement modifications *)
38970
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   257
    fun make_declarations nsps (data, nodes) =
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   258
      let
55293
42cf5802d36a code generation: explicitly declared identifiers gain predence over implicit ones
haftmann
parents: 55292
diff changeset
   259
        val (module_fragments, stmt_syms) =
42cf5802d36a code generation: explicitly declared identifiers gain predence over implicit ones
haftmann
parents: 55292
diff changeset
   260
          Code_Symbol.Graph.keys nodes
42cf5802d36a code generation: explicitly declared identifiers gain predence over implicit ones
haftmann
parents: 55292
diff changeset
   261
          |> List.partition
42cf5802d36a code generation: explicitly declared identifiers gain predence over implicit ones
haftmann
parents: 55292
diff changeset
   262
              (fn sym => case Code_Symbol.Graph.get_node nodes sym
42cf5802d36a code generation: explicitly declared identifiers gain predence over implicit ones
haftmann
parents: 55292
diff changeset
   263
                of (_, Module _) => true | _ => false)
42cf5802d36a code generation: explicitly declared identifiers gain predence over implicit ones
haftmann
parents: 55292
diff changeset
   264
          |> pairself (prioritize sym_priority)
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 53414
diff changeset
   265
        fun declare namify sym (nsps, nodes) =
38970
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   266
          let
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 53414
diff changeset
   267
            val (base, node) = Code_Symbol.Graph.get_node nodes sym;
38970
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   268
            val (base', nsps') = namify node base nsps;
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 53414
diff changeset
   269
            val nodes' = Code_Symbol.Graph.map_node sym (K (base', node)) nodes;
38970
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   270
          in (nsps', nodes') end;
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   271
        val (nsps', nodes') = (nsps, nodes)
39022
ac7774a35bcf explicit modify_stmt parameter
haftmann
parents: 39018
diff changeset
   272
          |> fold (declare (K namify_module)) module_fragments
55679
59244fc1a7ca formal markup for public ingredients
haftmann
parents: 55605
diff changeset
   273
          |> fold (declare (namify_stmt o snd o the_stmt)) stmt_syms;
55680
haftmann
parents: 55679
diff changeset
   274
        fun modify_stmts' syms_stmts =
55605
haftmann
parents: 55293
diff changeset
   275
          let
55680
haftmann
parents: 55679
diff changeset
   276
            val stmts' = modify_stmts syms_stmts
haftmann
parents: 55679
diff changeset
   277
          in stmts' @ replicate (length syms_stmts - length stmts') NONE end;
haftmann
parents: 55679
diff changeset
   278
        fun modify_stmts'' syms_exports_stmts =
haftmann
parents: 55679
diff changeset
   279
          syms_exports_stmts
haftmann
parents: 55679
diff changeset
   280
          |> map (fn (sym, (export, stmt)) => ((sym, stmt), export))
haftmann
parents: 55679
diff changeset
   281
          |> burrow_fst modify_stmts'
haftmann
parents: 55679
diff changeset
   282
          |> map (fn (SOME stmt, export) => SOME (export, stmt) | _ => NONE);
haftmann
parents: 55679
diff changeset
   283
        val nodes'' =
haftmann
parents: 55679
diff changeset
   284
          nodes'
haftmann
parents: 55679
diff changeset
   285
          |> Code_Symbol.Graph.map_strong_conn (map_module_stmts (make_declarations nsps') modify_stmts'');
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 53414
diff changeset
   286
        val data' = fold memorize_data stmt_syms data;
39018
haftmann
parents: 39017
diff changeset
   287
      in (data', nodes'') end;
38970
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   288
    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
   289
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   290
    (* deresolving *)
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 53414
diff changeset
   291
    fun deresolver prefix_fragments sym =
38970
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   292
      let
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 53414
diff changeset
   293
        val (name_fragments, _) = prep_sym sym;
38970
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   294
        val (_, (_, remainder)) = chop_prefix (op =) (prefix_fragments, name_fragments);
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 53414
diff changeset
   295
        val nodes = fold (fn name_fragment => fn nodes => case Code_Symbol.Graph.get_node nodes (Code_Symbol.Module name_fragment)
38970
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   296
         of (_, Module (_, nodes)) => nodes) name_fragments hierarchical_program;
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 53414
diff changeset
   297
        val (base', _) = Code_Symbol.Graph.get_node nodes sym;
38970
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   298
      in Long_Name.implode (remainder @ [base']) end
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 53414
diff changeset
   299
        handle Code_Symbol.Graph.UNDEF _ => error ("Unknown statement name: "
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 53414
diff changeset
   300
          ^ Code_Symbol.quote ctxt sym);
38970
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   301
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   302
  in { deresolver = deresolver, hierarchical_program = hierarchical_program } end;
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   303
39147
3c284a152bd6 printing combinator for hierarchical programs
haftmann
parents: 39055
diff changeset
   304
fun print_hierarchical { print_module, print_stmt, lift_markup } =
3c284a152bd6 printing combinator for hierarchical programs
haftmann
parents: 39055
diff changeset
   305
  let
3c284a152bd6 printing combinator for hierarchical programs
haftmann
parents: 39055
diff changeset
   306
    fun print_node _ (_, Dummy) =
3c284a152bd6 printing combinator for hierarchical programs
haftmann
parents: 39055
diff changeset
   307
          NONE
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 53414
diff changeset
   308
      | print_node prefix_fragments (sym, Stmt stmt) =
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 53414
diff changeset
   309
          SOME (lift_markup (Code_Printer.markup_stmt sym)
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 53414
diff changeset
   310
            (print_stmt prefix_fragments (sym, stmt)))
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 53414
diff changeset
   311
      | print_node prefix_fragments (Code_Symbol.Module name_fragment, Module (data, nodes)) =
39147
3c284a152bd6 printing combinator for hierarchical programs
haftmann
parents: 39055
diff changeset
   312
          let
3c284a152bd6 printing combinator for hierarchical programs
haftmann
parents: 39055
diff changeset
   313
            val prefix_fragments' = prefix_fragments @ [name_fragment]
3c284a152bd6 printing combinator for hierarchical programs
haftmann
parents: 39055
diff changeset
   314
          in
3c284a152bd6 printing combinator for hierarchical programs
haftmann
parents: 39055
diff changeset
   315
            Option.map (print_module prefix_fragments'
3c284a152bd6 printing combinator for hierarchical programs
haftmann
parents: 39055
diff changeset
   316
              name_fragment data) (print_nodes prefix_fragments' nodes)
3c284a152bd6 printing combinator for hierarchical programs
haftmann
parents: 39055
diff changeset
   317
          end
3c284a152bd6 printing combinator for hierarchical programs
haftmann
parents: 39055
diff changeset
   318
    and print_nodes prefix_fragments nodes =
3c284a152bd6 printing combinator for hierarchical programs
haftmann
parents: 39055
diff changeset
   319
      let
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 53414
diff changeset
   320
        val xs = (map_filter (fn sym => print_node prefix_fragments
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 53414
diff changeset
   321
          (sym, snd (Code_Symbol.Graph.get_node nodes sym))) o rev o flat o Code_Symbol.Graph.strong_conn) nodes
39147
3c284a152bd6 printing combinator for hierarchical programs
haftmann
parents: 39055
diff changeset
   322
      in if null xs then NONE else SOME xs end;
3c284a152bd6 printing combinator for hierarchical programs
haftmann
parents: 39055
diff changeset
   323
  in these o print_nodes [] end;
3c284a152bd6 printing combinator for hierarchical programs
haftmann
parents: 39055
diff changeset
   324
55147
bce3dbc11f95 prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents: 53414
diff changeset
   325
end;