src/Tools/Code/code_namespace.ML
author haftmann
Sat, 25 Jan 2014 23:50:49 +0100
changeset 55145 2bb3cd36bcf7
parent 53414 dd64696d267a
child 55147 bce3dbc11f95
permissions -rw-r--r--
more abstract declaration of unqualified constant names in code printing context
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
52138
e21426f244aa bookkeeping and input syntax for exact specification of names of symbols in generated code
haftmann
parents: 51930
diff 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: 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
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
    16
      -> { deresolver: string -> string -> string,
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
39017
8cd5b6d688fa generalized hierarchical data structure over statements
haftmann
parents: 38970
diff changeset
    21
    | Stmt of 'a
39147
3c284a152bd6 printing combinator for hierarchical programs
haftmann
parents: 39055
diff changeset
    22
    | Module of ('b * (string * ('a, 'b) node) Graph.T)
39208
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
    23
  type ('a, 'b) hierarchical_program
52138
e21426f244aa bookkeeping and input syntax for exact specification of names of symbols in generated code
haftmann
parents: 51930
diff 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: 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,
ac7774a35bcf explicit modify_stmt parameter
haftmann
parents: 39018
diff changeset
    29
    cyclic_modules: bool, empty_data: 'b, memorize_data: string -> 'b -> 'b,
39024
30d5dd2f30b6 simultaneous modification of statements: statement names
haftmann
parents: 39023
diff 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
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,
3c284a152bd6 printing combinator for hierarchical programs
haftmann
parents: 39055
diff changeset
    35
    print_stmt: string list -> string * 'a -> 'c,
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
e21426f244aa bookkeeping and input syntax for exact specification of names of symbols in generated code
haftmann
parents: 51930
diff 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: 51930
diff changeset
    46
e21426f244aa bookkeeping and input syntax for exact specification of names of symbols in generated code
haftmann
parents: 51930
diff 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: 51930
diff 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: 51930
diff changeset
    49
  |> Option.map (split_last o Long_Name.explode);
39055
81e0368812ad removed namespace stuff from code_printer
haftmann
parents: 39029
diff changeset
    50
53414
dd64696d267a explicit module names have precedence over identifier declarations
haftmann
parents: 52138
diff changeset
    51
fun force_identifier symbol_of fragments_tab force_module identifiers name =
52138
e21426f244aa bookkeeping and input syntax for exact specification of names of symbols in generated code
haftmann
parents: 51930
diff 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: 51930
diff changeset
    53
      NONE => (apfst (the o Symtab.lookup fragments_tab) o split_name) name
53414
dd64696d267a explicit module names have precedence over identifier declarations
haftmann
parents: 52138
diff changeset
    54
    | SOME prefix_name => if null force_module then prefix_name
dd64696d267a explicit module names have precedence over identifier declarations
haftmann
parents: 52138
diff changeset
    55
        else (force_module, snd prefix_name);
39055
81e0368812ad removed namespace stuff from code_printer
haftmann
parents: 39029
diff changeset
    56
52138
e21426f244aa bookkeeping and input syntax for exact specification of names of symbols in generated code
haftmann
parents: 51930
diff changeset
    57
fun build_module_namespace { module_prefix, module_identifiers, reserved } program =
39203
b2f9a6f4b84b factored out build_module_namespace
haftmann
parents: 39147
diff changeset
    58
  let
52138
e21426f244aa bookkeeping and input syntax for exact specification of names of symbols in generated code
haftmann
parents: 51930
diff changeset
    59
    fun alias_fragments name = case module_identifiers name
39203
b2f9a6f4b84b factored out build_module_namespace
haftmann
parents: 39147
diff changeset
    60
     of SOME name' => Long_Name.explode name'
43326
47cf4bc789aa simplified Name.variant -- discontinued builtin fold_map;
wenzelm
parents: 40705
diff changeset
    61
      | 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: 51930
diff changeset
    62
    val module_names = Graph.fold (insert (op =) o fst o split_name o fst) program [];
39203
b2f9a6f4b84b factored out build_module_namespace
haftmann
parents: 39147
diff changeset
    63
  in
39208
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
    64
    fold (fn name => Symtab.update (name, Long_Name.explode module_prefix @ alias_fragments name))
39203
b2f9a6f4b84b factored out build_module_namespace
haftmann
parents: 39147
diff changeset
    65
      module_names Symtab.empty
b2f9a6f4b84b factored out build_module_namespace
haftmann
parents: 39147
diff changeset
    66
  end;
b2f9a6f4b84b factored out build_module_namespace
haftmann
parents: 39147
diff changeset
    67
39055
81e0368812ad removed namespace stuff from code_printer
haftmann
parents: 39029
diff changeset
    68
39208
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
    69
(** flat program structure **)
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
    70
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
    71
type flat_program = ((string * Code_Thingol.stmt option) Graph.T * (string * string list) list) Graph.T;
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
    72
52138
e21426f244aa bookkeeping and input syntax for exact specification of names of symbols in generated code
haftmann
parents: 51930
diff changeset
    73
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: 51930
diff changeset
    74
    identifiers, empty_nsp, namify_stmt, modify_stmt } program =
39208
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
    75
  let
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
    76
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
    77
    (* building module name hierarchy *)
52138
e21426f244aa bookkeeping and input syntax for exact specification of names of symbols in generated code
haftmann
parents: 51930
diff changeset
    78
    val module_identifiers = if module_name = ""
e21426f244aa bookkeeping and input syntax for exact specification of names of symbols in generated code
haftmann
parents: 51930
diff changeset
    79
      then Code_Symbol.lookup_module_data identifiers
e21426f244aa bookkeeping and input syntax for exact specification of names of symbols in generated code
haftmann
parents: 51930
diff changeset
    80
      else K (SOME module_name);
e21426f244aa bookkeeping and input syntax for exact specification of names of symbols in generated code
haftmann
parents: 51930
diff changeset
    81
    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: 51930
diff changeset
    82
      module_identifiers = module_identifiers, reserved = reserved } program;
53414
dd64696d267a explicit module names have precedence over identifier declarations
haftmann
parents: 52138
diff changeset
    83
    val prep_name = force_identifier symbol_of fragments_tab (Long_Name.explode module_name) identifiers
52138
e21426f244aa bookkeeping and input syntax for exact specification of names of symbols in generated code
haftmann
parents: 51930
diff changeset
    84
      #>> Long_Name.implode;
39208
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
    85
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
    86
    (* distribute statements over hierarchy *)
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
    87
    fun add_stmt name stmt =
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
    88
      let
52138
e21426f244aa bookkeeping and input syntax for exact specification of names of symbols in generated code
haftmann
parents: 51930
diff changeset
    89
        val (module_name, base) = prep_name name;
39208
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
    90
      in
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
    91
        Graph.default_node (module_name, (Graph.empty, []))
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
    92
        #> (Graph.map_node module_name o apfst) (Graph.new_node (name, (base, stmt)))
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
    93
      end;
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
    94
    fun add_dependency name name' =
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
    95
      let
52138
e21426f244aa bookkeeping and input syntax for exact specification of names of symbols in generated code
haftmann
parents: 51930
diff changeset
    96
        val (module_name, _) = prep_name name;
e21426f244aa bookkeeping and input syntax for exact specification of names of symbols in generated code
haftmann
parents: 51930
diff changeset
    97
        val (module_name', _) = prep_name name';
39208
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
    98
      in if module_name = module_name'
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
    99
        then (Graph.map_node module_name o apfst) (Graph.add_edge (name, name'))
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
   100
        else (Graph.map_node module_name o apsnd) (AList.map_default (op =) (module_name', []) (insert (op =) name'))
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
   101
      end;
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
   102
    val proto_program = Graph.empty
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
   103
      |> 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: 43326
diff changeset
   104
      |> 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: 43326
diff changeset
   105
          Graph.Keys.fold (add_dependency name) names) program;
39208
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
   106
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
   107
    (* name declarations and statement modifications *)
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
   108
    fun declare name (base, stmt) (gr, nsp) = 
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
   109
      let
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
   110
        val (base', nsp') = namify_stmt stmt base nsp;
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
   111
        val gr' = (Graph.map_node name o apfst) (K base') gr;
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
   112
      in (gr', nsp') end;
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
   113
    fun declarations gr = (gr, empty_nsp)
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
   114
      |> fold (fn name => declare name (Graph.get_node gr name)) (Graph.keys gr) 
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
   115
      |> fst
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
   116
      |> (Graph.map o K o apsnd) modify_stmt;
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
   117
    val flat_program = proto_program
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
   118
      |> (Graph.map o K o apfst) declarations;
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
   119
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
   120
    (* qualified and unqualified imports, deresolving *)
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
   121
    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: 51930
diff changeset
   122
      (fst (Graph.get_node flat_program (fst (prep_name name)))) name);
39208
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
   123
    fun classify_names gr imports =
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
   124
      let
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
   125
        val import_tab = maps
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
   126
          (fn (module_name, names) => map (rpair module_name) names) imports;
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
   127
        val imported_names = map fst import_tab;
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
   128
        val here_names = Graph.keys gr;
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
   129
      in
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
   130
        Symtab.empty
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
   131
        |> fold (fn name => Symtab.update (name, base_deresolver name)) here_names
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
   132
        |> fold (fn name => Symtab.update (name,
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
   133
            Long_Name.append (the (AList.lookup (op =) import_tab name))
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
   134
              (base_deresolver name))) imported_names
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
   135
      end;
40630
haftmann
parents: 40629
diff changeset
   136
    val deresolver_tab = Symtab.make (AList.make
haftmann
parents: 40629
diff changeset
   137
      (uncurry classify_names o Graph.get_node flat_program)
haftmann
parents: 40629
diff changeset
   138
        (Graph.keys flat_program));
40705
03f1266a066e toplevel deresolving for flat module name space
haftmann
parents: 40630
diff changeset
   139
    fun deresolver "" name =
52138
e21426f244aa bookkeeping and input syntax for exact specification of names of symbols in generated code
haftmann
parents: 51930
diff changeset
   140
          Long_Name.append (fst (prep_name name)) (base_deresolver name)
40705
03f1266a066e toplevel deresolving for flat module name space
haftmann
parents: 40630
diff changeset
   141
      | deresolver module_name name =
03f1266a066e toplevel deresolving for flat module name space
haftmann
parents: 40630
diff changeset
   142
          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: 51930
diff changeset
   143
          handle Option.Option => error ("Unknown statement name: "
e21426f244aa bookkeeping and input syntax for exact specification of names of symbols in generated code
haftmann
parents: 51930
diff changeset
   144
            ^ (Code_Symbol.quote_symbol ctxt o symbol_of) name);
39208
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
   145
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
   146
  in { deresolver = deresolver, flat_program = flat_program } end;
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
   147
fc1e02735438 added flat_program; tuned signature
haftmann
parents: 39203
diff changeset
   148
39055
81e0368812ad removed namespace stuff from code_printer
haftmann
parents: 39029
diff changeset
   149
(** hierarchical program structure **)
38970
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   150
39017
8cd5b6d688fa generalized hierarchical data structure over statements
haftmann
parents: 38970
diff changeset
   151
datatype ('a, 'b) node =
38970
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   152
    Dummy
39017
8cd5b6d688fa generalized hierarchical data structure over statements
haftmann
parents: 38970
diff changeset
   153
  | Stmt of 'a
8cd5b6d688fa generalized hierarchical data structure over statements
haftmann
parents: 38970
diff changeset
   154
  | 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
   155
39147
3c284a152bd6 printing combinator for hierarchical programs
haftmann
parents: 39055
diff changeset
   156
type ('a, 'b) hierarchical_program = (string * ('a, 'b) node) Graph.T;
3c284a152bd6 printing combinator for hierarchical programs
haftmann
parents: 39055
diff changeset
   157
39018
haftmann
parents: 39017
diff changeset
   158
fun map_module_content f (Module content) = Module (f content);
haftmann
parents: 39017
diff changeset
   159
haftmann
parents: 39017
diff changeset
   160
fun map_module [] = I
haftmann
parents: 39017
diff changeset
   161
  | map_module (name_fragment :: name_fragments) =
haftmann
parents: 39017
diff changeset
   162
      apsnd o Graph.map_node name_fragment o apsnd o map_module_content
haftmann
parents: 39017
diff changeset
   163
        o map_module name_fragments;
haftmann
parents: 39017
diff changeset
   164
52138
e21426f244aa bookkeeping and input syntax for exact specification of names of symbols in generated code
haftmann
parents: 51930
diff changeset
   165
fun hierarchical_program ctxt symbol_of { module_name, reserved, identifiers, empty_nsp,
39023
3f70c03e8282 simultaneous modification of statements
haftmann
parents: 39022
diff changeset
   166
      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
   167
  let
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   168
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   169
    (* building module name hierarchy *)
52138
e21426f244aa bookkeeping and input syntax for exact specification of names of symbols in generated code
haftmann
parents: 51930
diff changeset
   170
    val module_identifiers = if module_name = ""
e21426f244aa bookkeeping and input syntax for exact specification of names of symbols in generated code
haftmann
parents: 51930
diff changeset
   171
      then Code_Symbol.lookup_module_data identifiers
e21426f244aa bookkeeping and input syntax for exact specification of names of symbols in generated code
haftmann
parents: 51930
diff changeset
   172
      else K (SOME module_name);
e21426f244aa bookkeeping and input syntax for exact specification of names of symbols in generated code
haftmann
parents: 51930
diff changeset
   173
    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: 51930
diff changeset
   174
      module_identifiers = module_identifiers, reserved = reserved } program;
53414
dd64696d267a explicit module names have precedence over identifier declarations
haftmann
parents: 52138
diff changeset
   175
    val prep_name = force_identifier symbol_of fragments_tab (Long_Name.explode module_name) identifiers;
38970
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   176
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   177
    (* building empty module hierarchy *)
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   178
    val empty_module = (empty_data, Graph.empty);
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   179
    fun ensure_module name_fragment (data, nodes) =
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   180
      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
   181
      else (data,
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   182
        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
   183
    fun allocate_module [] = I
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   184
      | allocate_module (name_fragment :: name_fragments) =
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   185
          ensure_module name_fragment
39018
haftmann
parents: 39017
diff changeset
   186
          #> (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: 51930
diff changeset
   187
    val empty_program =
e21426f244aa bookkeeping and input syntax for exact specification of names of symbols in generated code
haftmann
parents: 51930
diff changeset
   188
      empty_module
e21426f244aa bookkeeping and input syntax for exact specification of names of symbols in generated code
haftmann
parents: 51930
diff changeset
   189
      |> 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: 51930
diff changeset
   190
      |> 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: 51930
diff changeset
   191
          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
   192
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   193
    (* distribute statements over hierarchy *)
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   194
    fun add_stmt name stmt =
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   195
      let
52138
e21426f244aa bookkeeping and input syntax for exact specification of names of symbols in generated code
haftmann
parents: 51930
diff changeset
   196
        val (name_fragments, base) = prep_name name;
38970
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   197
      in
39018
haftmann
parents: 39017
diff changeset
   198
        (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
   199
      end;
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   200
    fun add_dependency name name' =
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   201
      let
52138
e21426f244aa bookkeeping and input syntax for exact specification of names of symbols in generated code
haftmann
parents: 51930
diff changeset
   202
        val (name_fragments, _) = prep_name name;
e21426f244aa bookkeeping and input syntax for exact specification of names of symbols in generated code
haftmann
parents: 51930
diff changeset
   203
        val (name_fragments', _) = prep_name name';
38970
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   204
        val (name_fragments_common, (diff, diff')) =
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   205
          chop_prefix (op =) (name_fragments, name_fragments');
39023
3f70c03e8282 simultaneous modification of statements
haftmann
parents: 39022
diff changeset
   206
        val is_module = not (null diff andalso null diff');
3f70c03e8282 simultaneous modification of statements
haftmann
parents: 39022
diff changeset
   207
        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
   208
        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
   209
          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
   210
            handle Graph.CYCLES _ => error ("Dependency "
52138
e21426f244aa bookkeeping and input syntax for exact specification of names of symbols in generated code
haftmann
parents: 51930
diff changeset
   211
              ^ (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: 51930
diff changeset
   212
              ^ (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
   213
              ^ " would result in module dependency cycle"))
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   214
          else Graph.add_edge dep
39018
haftmann
parents: 39017
diff changeset
   215
      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
   216
    val proto_program = empty_program
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   217
      |> 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: 43326
diff changeset
   218
      |> 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: 43326
diff changeset
   219
          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
   220
39022
ac7774a35bcf explicit modify_stmt parameter
haftmann
parents: 39018
diff changeset
   221
    (* name declarations, data and statement modifications *)
38970
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   222
    fun make_declarations nsps (data, nodes) =
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   223
      let
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   224
        val (module_fragments, stmt_names) = List.partition
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   225
          (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
   226
            of (_, Module _) => true | _ => false) (Graph.keys nodes);
39022
ac7774a35bcf explicit modify_stmt parameter
haftmann
parents: 39018
diff changeset
   227
        fun declare namify name (nsps, nodes) =
38970
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   228
          let
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   229
            val (base, node) = Graph.get_node nodes name;
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   230
            val (base', nsps') = namify node base nsps;
39022
ac7774a35bcf explicit modify_stmt parameter
haftmann
parents: 39018
diff changeset
   231
            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
   232
          in (nsps', nodes') end;
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   233
        val (nsps', nodes') = (nsps, nodes)
39022
ac7774a35bcf explicit modify_stmt parameter
haftmann
parents: 39018
diff changeset
   234
          |> fold (declare (K namify_module)) module_fragments
ac7774a35bcf explicit modify_stmt parameter
haftmann
parents: 39018
diff changeset
   235
          |> 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: 39024
diff changeset
   236
        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: 39024
diff changeset
   237
        fun select_names names = case filter (member (op =) stmt_names) names
cef7b58555aa skip empty name bunches; fill up trailing positions with NONEs
haftmann
parents: 39024
diff changeset
   238
         of [] => NONE
cef7b58555aa skip empty name bunches; fill up trailing positions with NONEs
haftmann
parents: 39024
diff changeset
   239
          | xs => SOME xs;
cef7b58555aa skip empty name bunches; fill up trailing positions with NONEs
haftmann
parents: 39024
diff changeset
   240
        val modify_stmts' = AList.make (snd o Graph.get_node nodes)
39023
3f70c03e8282 simultaneous modification of statements
haftmann
parents: 39022
diff changeset
   241
          #> split_list
3f70c03e8282 simultaneous modification of statements
haftmann
parents: 39022
diff changeset
   242
          ##> map (fn Stmt stmt => stmt)
39029
cef7b58555aa skip empty name bunches; fill up trailing positions with NONEs
haftmann
parents: 39024
diff changeset
   243
          #> (fn (names, stmts) => zip_fillup names (modify_stmts (names ~~ stmts)));
cef7b58555aa skip empty name bunches; fill up trailing positions with NONEs
haftmann
parents: 39024
diff changeset
   244
        val stmtss' = (maps modify_stmts' o map_filter select_names o Graph.strong_conn) nodes;
39023
3f70c03e8282 simultaneous modification of statements
haftmann
parents: 39022
diff changeset
   245
        val nodes'' = Graph.map (fn name => apsnd (fn Module content => Module (make_declarations nsps' content)
3f70c03e8282 simultaneous modification of statements
haftmann
parents: 39022
diff changeset
   246
            | _ => case AList.lookup (op =) stmtss' name of SOME (SOME stmt) => Stmt stmt | _ => Dummy)) nodes';
39018
haftmann
parents: 39017
diff changeset
   247
        val data' = fold memorize_data stmt_names data;
haftmann
parents: 39017
diff changeset
   248
      in (data', nodes'') end;
38970
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   249
    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
   250
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   251
    (* deresolving *)
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   252
    fun deresolver prefix_fragments name =
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   253
      let
52138
e21426f244aa bookkeeping and input syntax for exact specification of names of symbols in generated code
haftmann
parents: 51930
diff changeset
   254
        val (name_fragments, _) = prep_name name;
38970
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   255
        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
   256
        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
   257
         of (_, Module (_, nodes)) => nodes) name_fragments hierarchical_program;
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   258
        val (base', _) = Graph.get_node nodes name;
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   259
      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: 51930
diff changeset
   260
        handle Graph.UNDEF _ => error ("Unknown statement name: "
e21426f244aa bookkeeping and input syntax for exact specification of names of symbols in generated code
haftmann
parents: 51930
diff changeset
   261
          ^ (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
   262
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   263
  in { deresolver = deresolver, hierarchical_program = hierarchical_program } end;
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   264
39147
3c284a152bd6 printing combinator for hierarchical programs
haftmann
parents: 39055
diff changeset
   265
fun print_hierarchical { print_module, print_stmt, lift_markup } =
3c284a152bd6 printing combinator for hierarchical programs
haftmann
parents: 39055
diff changeset
   266
  let
3c284a152bd6 printing combinator for hierarchical programs
haftmann
parents: 39055
diff changeset
   267
    fun print_node _ (_, Dummy) =
3c284a152bd6 printing combinator for hierarchical programs
haftmann
parents: 39055
diff changeset
   268
          NONE
3c284a152bd6 printing combinator for hierarchical programs
haftmann
parents: 39055
diff changeset
   269
      | print_node prefix_fragments (name, Stmt stmt) =
3c284a152bd6 printing combinator for hierarchical programs
haftmann
parents: 39055
diff changeset
   270
          SOME (lift_markup (Code_Printer.markup_stmt name)
3c284a152bd6 printing combinator for hierarchical programs
haftmann
parents: 39055
diff changeset
   271
            (print_stmt prefix_fragments (name, stmt)))
3c284a152bd6 printing combinator for hierarchical programs
haftmann
parents: 39055
diff changeset
   272
      | print_node prefix_fragments (name_fragment, Module (data, nodes)) =
3c284a152bd6 printing combinator for hierarchical programs
haftmann
parents: 39055
diff changeset
   273
          let
3c284a152bd6 printing combinator for hierarchical programs
haftmann
parents: 39055
diff changeset
   274
            val prefix_fragments' = prefix_fragments @ [name_fragment]
3c284a152bd6 printing combinator for hierarchical programs
haftmann
parents: 39055
diff changeset
   275
          in
3c284a152bd6 printing combinator for hierarchical programs
haftmann
parents: 39055
diff changeset
   276
            Option.map (print_module prefix_fragments'
3c284a152bd6 printing combinator for hierarchical programs
haftmann
parents: 39055
diff changeset
   277
              name_fragment data) (print_nodes prefix_fragments' nodes)
3c284a152bd6 printing combinator for hierarchical programs
haftmann
parents: 39055
diff changeset
   278
          end
3c284a152bd6 printing combinator for hierarchical programs
haftmann
parents: 39055
diff changeset
   279
    and print_nodes prefix_fragments nodes =
3c284a152bd6 printing combinator for hierarchical programs
haftmann
parents: 39055
diff changeset
   280
      let
3c284a152bd6 printing combinator for hierarchical programs
haftmann
parents: 39055
diff changeset
   281
        val xs = (map_filter (fn name => print_node prefix_fragments
3c284a152bd6 printing combinator for hierarchical programs
haftmann
parents: 39055
diff changeset
   282
          (name, snd (Graph.get_node nodes name))) o rev o flat o Graph.strong_conn) nodes
3c284a152bd6 printing combinator for hierarchical programs
haftmann
parents: 39055
diff changeset
   283
      in if null xs then NONE else SOME xs end;
3c284a152bd6 printing combinator for hierarchical programs
haftmann
parents: 39055
diff changeset
   284
  in these o print_nodes [] end;
3c284a152bd6 printing combinator for hierarchical programs
haftmann
parents: 39055
diff changeset
   285
38970
53d1ee3d98b8 factored out generic part of Scala serializer into code_namespace.ML
haftmann
parents:
diff changeset
   286
end;