| author | wenzelm | 
| Tue, 07 Jan 2025 20:32:15 +0100 | |
| changeset 81740 | 9f0cee195ee9 | 
| parent 81706 | 7beb0cf38292 | 
| child 81875 | 7fe20d394593 | 
| permissions | -rw-r--r-- | 
| 
37745
 
6315b6426200
checking generated code for various target languages
 
haftmann 
parents: 
37528 
diff
changeset
 | 
1  | 
(* Title: Tools/Code/code_target.ML  | 
| 24219 | 2  | 
Author: Florian Haftmann, TU Muenchen  | 
3  | 
||
| 38910 | 4  | 
Generic infrastructure for target language data.  | 
| 24219 | 5  | 
*)  | 
6  | 
||
7  | 
signature CODE_TARGET =  | 
|
8  | 
sig  | 
|
| 55757 | 9  | 
val cert_tyco: Proof.context -> string -> string  | 
10  | 
val read_tyco: Proof.context -> string -> string  | 
|
| 36471 | 11  | 
|
| 
69623
 
ef02c5e051e5
explicit model concerning files of generated code
 
haftmann 
parents: 
69593 
diff
changeset
 | 
12  | 
datatype pretty_modules = Singleton of string * Pretty.T | Hierarchy of (string list * Pretty.T) list;  | 
| 70021 | 13  | 
val next_export: theory -> string * theory  | 
| 
69623
 
ef02c5e051e5
explicit model concerning files of generated code
 
haftmann 
parents: 
69593 
diff
changeset
 | 
14  | 
|
| 
70015
 
c8e08d8ffb93
clarified signature: more explicit type Path.binding;
 
wenzelm 
parents: 
70011 
diff
changeset
 | 
15  | 
  val export_code_for: ({physical: bool} * (Path.T * Position.T)) option -> string -> string
 | 
| 
 
c8e08d8ffb93
clarified signature: more explicit type Path.binding;
 
wenzelm 
parents: 
70011 
diff
changeset
 | 
16  | 
-> int option -> Token.T list -> Code_Thingol.program -> bool -> Code_Symbol.T list  | 
| 
 
c8e08d8ffb93
clarified signature: more explicit type Path.binding;
 
wenzelm 
parents: 
70011 
diff
changeset
 | 
17  | 
-> local_theory -> local_theory  | 
| 
69623
 
ef02c5e051e5
explicit model concerning files of generated code
 
haftmann 
parents: 
69593 
diff
changeset
 | 
18  | 
val produce_code_for: Proof.context -> string -> string -> int option -> Token.T list  | 
| 
75604
 
39df30349778
more scalable generated files and code export, using Bytes.T;
 
wenzelm 
parents: 
74887 
diff
changeset
 | 
19  | 
-> Code_Thingol.program -> bool -> Code_Symbol.T list -> (string list * Bytes.T) list * string option list  | 
| 
69623
 
ef02c5e051e5
explicit model concerning files of generated code
 
haftmann 
parents: 
69593 
diff
changeset
 | 
20  | 
val present_code_for: Proof.context -> string -> string -> int option -> Token.T list  | 
| 
75604
 
39df30349778
more scalable generated files and code export, using Bytes.T;
 
wenzelm 
parents: 
74887 
diff
changeset
 | 
21  | 
-> Code_Thingol.program -> Code_Symbol.T list * Code_Symbol.T list -> Bytes.T  | 
| 
69999
 
76554a0cafe3
export_code/check_code formally updates the theory -- this opens further possibilities concerning Generated_Files;
 
wenzelm 
parents: 
69998 
diff
changeset
 | 
22  | 
val check_code_for: string -> bool -> Token.T list  | 
| 70002 | 23  | 
-> Code_Thingol.program -> bool -> Code_Symbol.T list -> local_theory -> local_theory  | 
| 38918 | 24  | 
|
| 
69999
 
76554a0cafe3
export_code/check_code formally updates the theory -- this opens further possibilities concerning Generated_Files;
 
wenzelm 
parents: 
69998 
diff
changeset
 | 
25  | 
val export_code: bool -> string list  | 
| 
70015
 
c8e08d8ffb93
clarified signature: more explicit type Path.binding;
 
wenzelm 
parents: 
70011 
diff
changeset
 | 
26  | 
    -> (((string * string) * ({physical: bool} * (Path.T * Position.T)) option) * Token.T list) list
 | 
| 70002 | 27  | 
-> local_theory -> local_theory  | 
| 70001 | 28  | 
val export_code_cmd: bool -> string list  | 
| 72841 | 29  | 
    -> (((string * string) * ({physical: bool} * Input.source) option) * Token.T list) list
 | 
| 70002 | 30  | 
-> local_theory -> local_theory  | 
| 55757 | 31  | 
val produce_code: Proof.context -> bool -> string list  | 
| 
75604
 
39df30349778
more scalable generated files and code export, using Bytes.T;
 
wenzelm 
parents: 
74887 
diff
changeset
 | 
32  | 
-> string -> string -> int option -> Token.T list -> (string list * Bytes.T) list * string option list  | 
| 55757 | 33  | 
val present_code: Proof.context -> string list -> Code_Symbol.T list  | 
| 
75604
 
39df30349778
more scalable generated files and code export, using Bytes.T;
 
wenzelm 
parents: 
74887 
diff
changeset
 | 
34  | 
-> string -> string -> int option -> Token.T list -> Bytes.T  | 
| 70002 | 35  | 
val check_code: bool -> string list -> ((string * bool) * Token.T list) list  | 
36  | 
-> local_theory -> local_theory  | 
|
| 38918 | 37  | 
|
| 
70009
 
435fb018e8ee
"export_code ... file_prefix ..." is the preferred way to produce output within the logical file-system within the theory context, as well as session exports;
 
wenzelm 
parents: 
70002 
diff
changeset
 | 
38  | 
val codeN: string  | 
| 48568 | 39  | 
val generatedN: string  | 
| 70021 | 40  | 
val code_path: Path.T -> Path.T  | 
41  | 
val code_export_message: theory -> unit  | 
|
| 
75604
 
39df30349778
more scalable generated files and code export, using Bytes.T;
 
wenzelm 
parents: 
74887 
diff
changeset
 | 
42  | 
val export: Path.binding -> Bytes.T -> theory -> theory  | 
| 64957 | 43  | 
val compilation_text: Proof.context -> string -> Code_Thingol.program  | 
| 
55684
 
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
 
haftmann 
parents: 
55683 
diff
changeset
 | 
44  | 
-> Code_Symbol.T list -> bool -> ((string * class list) list * Code_Thingol.itype) * Code_Thingol.iterm  | 
| 
75604
 
39df30349778
more scalable generated files and code export, using Bytes.T;
 
wenzelm 
parents: 
74887 
diff
changeset
 | 
45  | 
-> (string list * Bytes.T) list * string  | 
| 64959 | 46  | 
val compilation_text': Proof.context -> string -> string option -> Code_Thingol.program  | 
47  | 
-> Code_Symbol.T list -> bool -> ((string * class list) list * Code_Thingol.itype) * Code_Thingol.iterm  | 
|
| 
75604
 
39df30349778
more scalable generated files and code export, using Bytes.T;
 
wenzelm 
parents: 
74887 
diff
changeset
 | 
48  | 
-> ((string list * Bytes.T) list * string) * (Code_Symbol.T -> string option)  | 
| 
41344
 
d990badc97a3
evaluator separating static and dynamic operations
 
haftmann 
parents: 
41342 
diff
changeset
 | 
49  | 
|
| 28054 | 50  | 
type serializer  | 
| 34152 | 51  | 
type literals = Code_Printer.literals  | 
| 59104 | 52  | 
type language  | 
53  | 
type ancestry  | 
|
| 59479 | 54  | 
val assert_target: theory -> string -> string  | 
| 59104 | 55  | 
val add_language: string * language -> theory -> theory  | 
56  | 
val add_derived_target: string * ancestry -> theory -> theory  | 
|
| 55757 | 57  | 
val the_literals: Proof.context -> string -> literals  | 
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36537 
diff
changeset
 | 
58  | 
val parse_args: 'a parser -> Token.T list -> 'a  | 
| 59324 | 59  | 
val default_code_width: int Config.T  | 
| 38917 | 60  | 
|
| 
52137
 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 
haftmann 
parents: 
52068 
diff
changeset
 | 
61  | 
  type ('a, 'b, 'c, 'd, 'e, 'f) symbol_attr_decl
 | 
| 
52138
 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 
haftmann 
parents: 
52137 
diff
changeset
 | 
62  | 
val set_identifiers: (string, string, string, string, string, string) symbol_attr_decl  | 
| 
 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 
haftmann 
parents: 
52137 
diff
changeset
 | 
63  | 
-> theory -> theory  | 
| 
69485
 
b734ff28e405
proper attach mechanism for any kind of symbols, not just constants
 
haftmann 
parents: 
69484 
diff
changeset
 | 
64  | 
val set_printings: (Code_Printer.raw_const_syntax, Code_Printer.tyco_syntax, string, unit, unit, string * Code_Symbol.T list) symbol_attr_decl  | 
| 
52137
 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 
haftmann 
parents: 
52068 
diff
changeset
 | 
65  | 
-> theory -> theory  | 
| 28054 | 66  | 
val add_reserved: string -> string -> theory -> theory  | 
| 24219 | 67  | 
end;  | 
68  | 
||
| 28054 | 69  | 
structure Code_Target : CODE_TARGET =  | 
| 24219 | 70  | 
struct  | 
71  | 
||
| 55150 | 72  | 
open Basic_Code_Symbol;  | 
| 28054 | 73  | 
open Basic_Code_Thingol;  | 
| 34152 | 74  | 
|
75  | 
type literals = Code_Printer.literals;  | 
|
| 
52137
 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 
haftmann 
parents: 
52068 
diff
changeset
 | 
76  | 
type ('a, 'b, 'c, 'd, 'e, 'f) symbol_attr_decl =
 | 
| 
 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 
haftmann 
parents: 
52068 
diff
changeset
 | 
77  | 
(string * (string * 'a option) list, string * (string * 'b option) list,  | 
| 
 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 
haftmann 
parents: 
52068 
diff
changeset
 | 
78  | 
class * (string * 'c option) list, (class * class) * (string * 'd option) list,  | 
| 
 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 
haftmann 
parents: 
52068 
diff
changeset
 | 
79  | 
(class * string) * (string * 'e option) list,  | 
| 
 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 
haftmann 
parents: 
52068 
diff
changeset
 | 
80  | 
string * (string * 'f option) list) Code_Symbol.attr;  | 
| 
 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 
haftmann 
parents: 
52068 
diff
changeset
 | 
81  | 
|
| 34152 | 82  | 
type tyco_syntax = Code_Printer.tyco_syntax;  | 
| 
55153
 
eedd549de3ef
more suitable names, without any notion of "activating"
 
haftmann 
parents: 
55150 
diff
changeset
 | 
83  | 
type raw_const_syntax = Code_Printer.raw_const_syntax;  | 
| 34152 | 84  | 
|
| 24219 | 85  | 
|
| 
52377
 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 
haftmann 
parents: 
52218 
diff
changeset
 | 
86  | 
(** checking and parsing of symbols **)  | 
| 
 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 
haftmann 
parents: 
52218 
diff
changeset
 | 
87  | 
|
| 55757 | 88  | 
fun cert_const ctxt const =  | 
| 
52377
 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 
haftmann 
parents: 
52218 
diff
changeset
 | 
89  | 
let  | 
| 55757 | 90  | 
val _ = if Sign.declared_const (Proof_Context.theory_of ctxt) const then ()  | 
| 
52377
 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 
haftmann 
parents: 
52218 
diff
changeset
 | 
91  | 
      else error ("No such constant: " ^ quote const);
 | 
| 
 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 
haftmann 
parents: 
52218 
diff
changeset
 | 
92  | 
in const end;  | 
| 
 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 
haftmann 
parents: 
52218 
diff
changeset
 | 
93  | 
|
| 55757 | 94  | 
fun read_const ctxt = Code.read_const (Proof_Context.theory_of ctxt);  | 
95  | 
||
96  | 
fun cert_tyco ctxt tyco =  | 
|
| 
52377
 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 
haftmann 
parents: 
52218 
diff
changeset
 | 
97  | 
let  | 
| 55757 | 98  | 
val _ = if Sign.declared_tyname (Proof_Context.theory_of ctxt) tyco then ()  | 
| 
52377
 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 
haftmann 
parents: 
52218 
diff
changeset
 | 
99  | 
      else error ("No such type constructor: " ^ quote tyco);
 | 
| 
 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 
haftmann 
parents: 
52218 
diff
changeset
 | 
100  | 
in tyco end;  | 
| 
 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 
haftmann 
parents: 
52218 
diff
changeset
 | 
101  | 
|
| 
55951
 
c07d184aebe9
tuned signature -- more uniform check_type_name/read_type_name;
 
wenzelm 
parents: 
55776 
diff
changeset
 | 
102  | 
fun read_tyco ctxt =  | 
| 56002 | 103  | 
  #1 o dest_Type o Proof_Context.read_type_name {proper = true, strict = true} ctxt;
 | 
| 
52377
 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 
haftmann 
parents: 
52218 
diff
changeset
 | 
104  | 
|
| 55757 | 105  | 
fun cert_class ctxt class =  | 
| 
52377
 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 
haftmann 
parents: 
52218 
diff
changeset
 | 
106  | 
let  | 
| 55757 | 107  | 
val _ = Axclass.get_info (Proof_Context.theory_of ctxt) class;  | 
| 
52377
 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 
haftmann 
parents: 
52218 
diff
changeset
 | 
108  | 
in class end;  | 
| 
 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 
haftmann 
parents: 
52218 
diff
changeset
 | 
109  | 
|
| 69593 | 110  | 
val parse_classrel_ident = Parse.class --| \<^keyword>\<open><\<close> -- Parse.class;  | 
| 
52377
 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 
haftmann 
parents: 
52218 
diff
changeset
 | 
111  | 
|
| 55757 | 112  | 
fun cert_inst ctxt (class, tyco) =  | 
113  | 
(cert_class ctxt class, cert_tyco ctxt tyco);  | 
|
| 
52377
 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 
haftmann 
parents: 
52218 
diff
changeset
 | 
114  | 
|
| 55757 | 115  | 
fun read_inst ctxt (raw_tyco, raw_class) =  | 
116  | 
(read_tyco ctxt raw_tyco, Proof_Context.read_class ctxt raw_class);  | 
|
| 
52377
 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 
haftmann 
parents: 
52218 
diff
changeset
 | 
117  | 
|
| 69593 | 118  | 
val parse_inst_ident = Parse.name --| \<^keyword>\<open>::\<close> -- Parse.class;  | 
| 
52377
 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 
haftmann 
parents: 
52218 
diff
changeset
 | 
119  | 
|
| 55757 | 120  | 
fun cert_syms ctxt =  | 
| 
69485
 
b734ff28e405
proper attach mechanism for any kind of symbols, not just constants
 
haftmann 
parents: 
69484 
diff
changeset
 | 
121  | 
Code_Symbol.map_attr (cert_const ctxt) (cert_tyco ctxt)  | 
| 
 
b734ff28e405
proper attach mechanism for any kind of symbols, not just constants
 
haftmann 
parents: 
69484 
diff
changeset
 | 
122  | 
(cert_class ctxt) (apply2 (cert_class ctxt)) (cert_inst ctxt) I;  | 
| 
 
b734ff28e405
proper attach mechanism for any kind of symbols, not just constants
 
haftmann 
parents: 
69484 
diff
changeset
 | 
123  | 
|
| 
 
b734ff28e405
proper attach mechanism for any kind of symbols, not just constants
 
haftmann 
parents: 
69484 
diff
changeset
 | 
124  | 
fun read_syms ctxt =  | 
| 
 
b734ff28e405
proper attach mechanism for any kind of symbols, not just constants
 
haftmann 
parents: 
69484 
diff
changeset
 | 
125  | 
Code_Symbol.map_attr (read_const ctxt) (read_tyco ctxt)  | 
| 
 
b734ff28e405
proper attach mechanism for any kind of symbols, not just constants
 
haftmann 
parents: 
69484 
diff
changeset
 | 
126  | 
(Proof_Context.read_class ctxt) (apply2 (Proof_Context.read_class ctxt)) (read_inst ctxt) I;  | 
| 
 
b734ff28e405
proper attach mechanism for any kind of symbols, not just constants
 
haftmann 
parents: 
69484 
diff
changeset
 | 
127  | 
|
| 
 
b734ff28e405
proper attach mechanism for any kind of symbols, not just constants
 
haftmann 
parents: 
69484 
diff
changeset
 | 
128  | 
fun cert_syms' ctxt =  | 
| 55757 | 129  | 
Code_Symbol.map_attr (apfst (cert_const ctxt)) (apfst (cert_tyco ctxt))  | 
| 
59058
 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 
wenzelm 
parents: 
58928 
diff
changeset
 | 
130  | 
(apfst (cert_class ctxt)) ((apfst o apply2) (cert_class ctxt)) (apfst (cert_inst ctxt)) I;  | 
| 
52377
 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 
haftmann 
parents: 
52218 
diff
changeset
 | 
131  | 
|
| 
69485
 
b734ff28e405
proper attach mechanism for any kind of symbols, not just constants
 
haftmann 
parents: 
69484 
diff
changeset
 | 
132  | 
fun read_syms' ctxt =  | 
| 55757 | 133  | 
Code_Symbol.map_attr (apfst (read_const ctxt)) (apfst (read_tyco ctxt))  | 
| 
59058
 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 
wenzelm 
parents: 
58928 
diff
changeset
 | 
134  | 
(apfst (Proof_Context.read_class ctxt)) ((apfst o apply2) (Proof_Context.read_class ctxt)) (apfst (read_inst ctxt)) I;  | 
| 
52377
 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 
haftmann 
parents: 
52218 
diff
changeset
 | 
135  | 
|
| 
 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 
haftmann 
parents: 
52218 
diff
changeset
 | 
136  | 
fun check_name is_module s =  | 
| 
 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 
haftmann 
parents: 
52218 
diff
changeset
 | 
137  | 
let  | 
| 
 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 
haftmann 
parents: 
52218 
diff
changeset
 | 
138  | 
val _ = if s = "" then error "Bad empty code name" else ();  | 
| 
 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 
haftmann 
parents: 
52218 
diff
changeset
 | 
139  | 
val xs = Long_Name.explode s;  | 
| 
 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 
haftmann 
parents: 
52218 
diff
changeset
 | 
140  | 
val xs' = if is_module  | 
| 
56826
 
ba18bd41e510
enforce case of identifiers only to accomodate strict language requirements (or clear separation of constructors from variables in the case of SML)
 
haftmann 
parents: 
56811 
diff
changeset
 | 
141  | 
then map (Name.desymbolize NONE) xs  | 
| 
52377
 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 
haftmann 
parents: 
52218 
diff
changeset
 | 
142  | 
else if length xs < 2  | 
| 
 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 
haftmann 
parents: 
52218 
diff
changeset
 | 
143  | 
        then error ("Bad code name without module component: " ^ quote s)
 | 
| 
 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 
haftmann 
parents: 
52218 
diff
changeset
 | 
144  | 
else  | 
| 
 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 
haftmann 
parents: 
52218 
diff
changeset
 | 
145  | 
let  | 
| 
 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 
haftmann 
parents: 
52218 
diff
changeset
 | 
146  | 
val (ys, y) = split_last xs;  | 
| 
56826
 
ba18bd41e510
enforce case of identifiers only to accomodate strict language requirements (or clear separation of constructors from variables in the case of SML)
 
haftmann 
parents: 
56811 
diff
changeset
 | 
147  | 
val ys' = map (Name.desymbolize NONE) ys;  | 
| 
 
ba18bd41e510
enforce case of identifiers only to accomodate strict language requirements (or clear separation of constructors from variables in the case of SML)
 
haftmann 
parents: 
56811 
diff
changeset
 | 
148  | 
val y' = Name.desymbolize NONE y;  | 
| 
52377
 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 
haftmann 
parents: 
52218 
diff
changeset
 | 
149  | 
in ys' @ [y'] end;  | 
| 
 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 
haftmann 
parents: 
52218 
diff
changeset
 | 
150  | 
in if xs' = xs  | 
| 55291 | 151  | 
then if is_module then (xs, "") else split_last xs  | 
| 
52377
 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 
haftmann 
parents: 
52218 
diff
changeset
 | 
152  | 
    else error ("Invalid code name: " ^ quote s ^ "\n"
 | 
| 
 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 
haftmann 
parents: 
52218 
diff
changeset
 | 
153  | 
^ "better try " ^ quote (Long_Name.implode xs'))  | 
| 
 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 
haftmann 
parents: 
52218 
diff
changeset
 | 
154  | 
end;  | 
| 
 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 
haftmann 
parents: 
52218 
diff
changeset
 | 
155  | 
|
| 
 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 
haftmann 
parents: 
52218 
diff
changeset
 | 
156  | 
|
| 
70009
 
435fb018e8ee
"export_code ... file_prefix ..." is the preferred way to produce output within the logical file-system within the theory context, as well as session exports;
 
wenzelm 
parents: 
70002 
diff
changeset
 | 
157  | 
(** theory data **)  | 
| 
27014
 
a5f53d9d2b60
yet another attempt to circumvent printmode problems
 
haftmann 
parents: 
27001 
diff
changeset
 | 
158  | 
|
| 
69623
 
ef02c5e051e5
explicit model concerning files of generated code
 
haftmann 
parents: 
69593 
diff
changeset
 | 
159  | 
datatype pretty_modules = Singleton of string * Pretty.T | Hierarchy of (string list * Pretty.T) list;  | 
| 28054 | 160  | 
|
| 39142 | 161  | 
type serializer = Token.T list  | 
| 
52138
 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 
haftmann 
parents: 
52137 
diff
changeset
 | 
162  | 
-> Proof.context  | 
| 39142 | 163  | 
  -> {
 | 
| 38926 | 164  | 
reserved_syms: string list,  | 
| 59103 | 165  | 
identifiers: Code_Printer.identifiers,  | 
| 38926 | 166  | 
includes: (string * Pretty.T) list,  | 
167  | 
class_syntax: string -> string option,  | 
|
168  | 
tyco_syntax: string -> Code_Printer.tyco_syntax option,  | 
|
| 
69623
 
ef02c5e051e5
explicit model concerning files of generated code
 
haftmann 
parents: 
69593 
diff
changeset
 | 
169  | 
const_syntax: string -> Code_Printer.const_syntax option,  | 
| 
 
ef02c5e051e5
explicit model concerning files of generated code
 
haftmann 
parents: 
69593 
diff
changeset
 | 
170  | 
module_name: string }  | 
| 
41342
 
3519e0dd8f75
more explicit structure for serializer invocation
 
haftmann 
parents: 
41307 
diff
changeset
 | 
171  | 
-> Code_Thingol.program  | 
| 
69623
 
ef02c5e051e5
explicit model concerning files of generated code
 
haftmann 
parents: 
69593 
diff
changeset
 | 
172  | 
-> Code_Symbol.T list  | 
| 
 
ef02c5e051e5
explicit model concerning files of generated code
 
haftmann 
parents: 
69593 
diff
changeset
 | 
173  | 
-> pretty_modules * (Code_Symbol.T -> string option);  | 
| 
 
ef02c5e051e5
explicit model concerning files of generated code
 
haftmann 
parents: 
69593 
diff
changeset
 | 
174  | 
|
| 
67207
 
ad538f6c5d2f
dedicated case option for code generation to Scala
 
haftmann 
parents: 
66586 
diff
changeset
 | 
175  | 
type language = {serializer: serializer, literals: literals,
 | 
| 
 
ad538f6c5d2f
dedicated case option for code generation to Scala
 
haftmann 
parents: 
66586 
diff
changeset
 | 
176  | 
  check: {env_var: string, make_destination: Path.T -> Path.T, make_command: string -> string},
 | 
| 69998 | 177  | 
evaluation_args: Token.T list};  | 
| 59102 | 178  | 
|
| 59103 | 179  | 
type ancestry = (string * (Code_Thingol.program -> Code_Thingol.program)) list;  | 
| 27103 | 180  | 
|
| 59103 | 181  | 
val merge_ancestry : ancestry * ancestry -> ancestry = AList.join (op =) (K snd);  | 
| 27103 | 182  | 
|
| 
67207
 
ad538f6c5d2f
dedicated case option for code generation to Scala
 
haftmann 
parents: 
66586 
diff
changeset
 | 
183  | 
type target = {serial: serial, language: language, ancestry: ancestry};
 | 
| 27437 | 184  | 
|
| 34248 | 185  | 
structure Targets = Theory_Data  | 
| 34071 | 186  | 
(  | 
| 
70009
 
435fb018e8ee
"export_code ... file_prefix ..." is the preferred way to produce output within the logical file-system within the theory context, as well as session exports;
 
wenzelm 
parents: 
70002 
diff
changeset
 | 
187  | 
type T = (target * Code_Printer.data) Symtab.table * int;  | 
| 
 
435fb018e8ee
"export_code ... file_prefix ..." is the preferred way to produce output within the logical file-system within the theory context, as well as session exports;
 
wenzelm 
parents: 
70002 
diff
changeset
 | 
188  | 
val empty = (Symtab.empty, 0);  | 
| 72057 | 189  | 
fun merge ((targets1, index1), (targets2, index2)) : T =  | 
190  | 
let  | 
|
191  | 
val targets' =  | 
|
192  | 
Symtab.join (fn target_name => fn ((target1, data1), (target2, data2)) =>  | 
|
193  | 
if #serial target1 = #serial target2 then  | 
|
194  | 
          ({serial = #serial target1, language = #language target1,
 | 
|
195  | 
ancestry = merge_ancestry (#ancestry target1, #ancestry target2)},  | 
|
196  | 
Code_Printer.merge_data (data1, data2))  | 
|
197  | 
          else error ("Incompatible targets: " ^ quote target_name)) (targets1, targets2)
 | 
|
198  | 
val index' = Int.max (index1, index2);  | 
|
199  | 
in (targets', index') end;  | 
|
| 34071 | 200  | 
);  | 
| 27436 | 201  | 
|
| 
70009
 
435fb018e8ee
"export_code ... file_prefix ..." is the preferred way to produce output within the logical file-system within the theory context, as well as session exports;
 
wenzelm 
parents: 
70002 
diff
changeset
 | 
202  | 
val exists_target = Symtab.defined o #1 o Targets.get;  | 
| 
 
435fb018e8ee
"export_code ... file_prefix ..." is the preferred way to produce output within the logical file-system within the theory context, as well as session exports;
 
wenzelm 
parents: 
70002 
diff
changeset
 | 
203  | 
val lookup_target_data = Symtab.lookup o #1 o Targets.get;  | 
| 59479 | 204  | 
fun assert_target thy target_name =  | 
205  | 
if exists_target thy target_name  | 
|
206  | 
then target_name  | 
|
207  | 
  else error ("Unknown code target language: " ^ quote target_name);
 | 
|
| 59100 | 208  | 
|
| 72057 | 209  | 
fun reset_index thy =  | 
210  | 
if #2 (Targets.get thy) = 0 then NONE  | 
|
211  | 
else SOME ((Targets.map o apsnd) (K 0) thy);  | 
|
212  | 
||
213  | 
val _ = Theory.setup (Theory.at_begin reset_index);  | 
|
214  | 
||
| 
70009
 
435fb018e8ee
"export_code ... file_prefix ..." is the preferred way to produce output within the logical file-system within the theory context, as well as session exports;
 
wenzelm 
parents: 
70002 
diff
changeset
 | 
215  | 
fun next_export thy =  | 
| 
 
435fb018e8ee
"export_code ... file_prefix ..." is the preferred way to produce output within the logical file-system within the theory context, as well as session exports;
 
wenzelm 
parents: 
70002 
diff
changeset
 | 
216  | 
let  | 
| 
 
435fb018e8ee
"export_code ... file_prefix ..." is the preferred way to produce output within the logical file-system within the theory context, as well as session exports;
 
wenzelm 
parents: 
70002 
diff
changeset
 | 
217  | 
val thy' = (Targets.map o apsnd) (fn i => i + 1) thy;  | 
| 
 
435fb018e8ee
"export_code ... file_prefix ..." is the preferred way to produce output within the logical file-system within the theory context, as well as session exports;
 
wenzelm 
parents: 
70002 
diff
changeset
 | 
218  | 
val i = #2 (Targets.get thy');  | 
| 
 
435fb018e8ee
"export_code ... file_prefix ..." is the preferred way to produce output within the logical file-system within the theory context, as well as session exports;
 
wenzelm 
parents: 
70002 
diff
changeset
 | 
219  | 
  in ("export" ^ string_of_int i, thy') end;
 | 
| 
 
435fb018e8ee
"export_code ... file_prefix ..." is the preferred way to produce output within the logical file-system within the theory context, as well as session exports;
 
wenzelm 
parents: 
70002 
diff
changeset
 | 
220  | 
|
| 59104 | 221  | 
fun fold1 f xs = fold f (tl xs) (hd xs);  | 
222  | 
||
| 
59101
 
6dc48c98303c
tuned target inheritance bookkeeping: ancestry is always fully maintained at current entry using canonical merge;
 
haftmann 
parents: 
59100 
diff
changeset
 | 
223  | 
fun join_ancestry thy target_name =  | 
| 
 
6dc48c98303c
tuned target inheritance bookkeeping: ancestry is always fully maintained at current entry using canonical merge;
 
haftmann 
parents: 
59100 
diff
changeset
 | 
224  | 
let  | 
| 69998 | 225  | 
val _ = assert_target thy target_name;  | 
| 59103 | 226  | 
val the_target_data = the o lookup_target_data thy;  | 
227  | 
val (target, this_data) = the_target_data target_name;  | 
|
228  | 
val ancestry = #ancestry target;  | 
|
| 59102 | 229  | 
val modifies = rev (map snd ancestry);  | 
230  | 
val modify = fold (curry (op o)) modifies I;  | 
|
| 59103 | 231  | 
val datas = rev (map (snd o the_target_data o fst) ancestry) @ [this_data];  | 
| 59104 | 232  | 
val data = fold1 (fn data' => fn data => Code_Printer.merge_data (data, data')) datas;  | 
| 59103 | 233  | 
in (modify, (target, data)) end;  | 
| 
59101
 
6dc48c98303c
tuned target inheritance bookkeeping: ancestry is always fully maintained at current entry using canonical merge;
 
haftmann 
parents: 
59100 
diff
changeset
 | 
234  | 
|
| 
69623
 
ef02c5e051e5
explicit model concerning files of generated code
 
haftmann 
parents: 
69593 
diff
changeset
 | 
235  | 
fun allocate_target target_name target thy =  | 
| 27304 | 236  | 
let  | 
| 
59101
 
6dc48c98303c
tuned target inheritance bookkeeping: ancestry is always fully maintained at current entry using canonical merge;
 
haftmann 
parents: 
59100 
diff
changeset
 | 
237  | 
val _ = if exists_target thy target_name  | 
| 
 
6dc48c98303c
tuned target inheritance bookkeeping: ancestry is always fully maintained at current entry using canonical merge;
 
haftmann 
parents: 
59100 
diff
changeset
 | 
238  | 
      then error ("Attempt to overwrite existing target " ^ quote target_name)
 | 
| 
43564
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
43324 
diff
changeset
 | 
239  | 
else ();  | 
| 28054 | 240  | 
in  | 
241  | 
thy  | 
|
| 
70009
 
435fb018e8ee
"export_code ... file_prefix ..." is the preferred way to produce output within the logical file-system within the theory context, as well as session exports;
 
wenzelm 
parents: 
70002 
diff
changeset
 | 
242  | 
|> (Targets.map o apfst o Symtab.update) (target_name, (target, Code_Printer.empty_data))  | 
| 28054 | 243  | 
end;  | 
| 27436 | 244  | 
|
| 59104 | 245  | 
fun add_language (target_name, language) =  | 
| 
67207
 
ad538f6c5d2f
dedicated case option for code generation to Scala
 
haftmann 
parents: 
66586 
diff
changeset
 | 
246  | 
  allocate_target target_name {serial = serial (), language = language,
 | 
| 
 
ad538f6c5d2f
dedicated case option for code generation to Scala
 
haftmann 
parents: 
66586 
diff
changeset
 | 
247  | 
ancestry = []};  | 
| 27436 | 248  | 
|
| 59104 | 249  | 
fun add_derived_target (target_name, initial_ancestry) thy =  | 
| 
59101
 
6dc48c98303c
tuned target inheritance bookkeeping: ancestry is always fully maintained at current entry using canonical merge;
 
haftmann 
parents: 
59100 
diff
changeset
 | 
250  | 
let  | 
| 59104 | 251  | 
val _ = if null initial_ancestry  | 
252  | 
then error "Must derive from existing target(s)" else ();  | 
|
253  | 
fun the_target_data target_name' = case lookup_target_data thy target_name' of  | 
|
254  | 
      NONE => error ("Unknown code target language: " ^ quote target_name')
 | 
|
255  | 
| SOME target_data' => target_data';  | 
|
256  | 
val targets = rev (map (fst o the_target_data o fst) initial_ancestry);  | 
|
257  | 
val supremum = fold1 (fn target' => fn target =>  | 
|
258  | 
if #serial target = #serial target'  | 
|
259  | 
then target else error "Incompatible targets") targets;  | 
|
260  | 
val ancestries = map #ancestry targets @ [initial_ancestry];  | 
|
261  | 
val ancestry = fold1 (fn ancestry' => fn ancestry =>  | 
|
262  | 
merge_ancestry (ancestry, ancestry')) ancestries;  | 
|
| 
59101
 
6dc48c98303c
tuned target inheritance bookkeeping: ancestry is always fully maintained at current entry using canonical merge;
 
haftmann 
parents: 
59100 
diff
changeset
 | 
263  | 
in  | 
| 
67207
 
ad538f6c5d2f
dedicated case option for code generation to Scala
 
haftmann 
parents: 
66586 
diff
changeset
 | 
264  | 
    allocate_target target_name {serial = #serial supremum, language = #language supremum,
 | 
| 69998 | 265  | 
ancestry = ancestry} thy  | 
| 
59101
 
6dc48c98303c
tuned target inheritance bookkeeping: ancestry is always fully maintained at current entry using canonical merge;
 
haftmann 
parents: 
59100 
diff
changeset
 | 
266  | 
end;  | 
| 69998 | 267  | 
|
| 59102 | 268  | 
fun map_data target_name f thy =  | 
| 27436 | 269  | 
let  | 
| 59479 | 270  | 
val _ = assert_target thy target_name;  | 
| 28054 | 271  | 
in  | 
272  | 
thy  | 
|
| 
70009
 
435fb018e8ee
"export_code ... file_prefix ..." is the preferred way to produce output within the logical file-system within the theory context, as well as session exports;
 
wenzelm 
parents: 
70002 
diff
changeset
 | 
273  | 
|> (Targets.map o apfst o Symtab.map_entry target_name o apsnd o Code_Printer.map_data) f  | 
| 28054 | 274  | 
end;  | 
| 27304 | 275  | 
|
| 69998 | 276  | 
fun map_reserved target_name = map_data target_name o @{apply 3(1)};
 | 
277  | 
fun map_identifiers target_name = map_data target_name o @{apply 3(2)};
 | 
|
278  | 
fun map_printings target_name = map_data target_name o @{apply 3(3)};
 | 
|
| 27000 | 279  | 
|
280  | 
||
| 
70009
 
435fb018e8ee
"export_code ... file_prefix ..." is the preferred way to produce output within the logical file-system within the theory context, as well as session exports;
 
wenzelm 
parents: 
70002 
diff
changeset
 | 
281  | 
(** serializers **)  | 
| 
 
435fb018e8ee
"export_code ... file_prefix ..." is the preferred way to produce output within the logical file-system within the theory context, as well as session exports;
 
wenzelm 
parents: 
70002 
diff
changeset
 | 
282  | 
|
| 
 
435fb018e8ee
"export_code ... file_prefix ..." is the preferred way to produce output within the logical file-system within the theory context, as well as session exports;
 
wenzelm 
parents: 
70002 
diff
changeset
 | 
283  | 
val codeN = "code";  | 
| 
 
435fb018e8ee
"export_code ... file_prefix ..." is the preferred way to produce output within the logical file-system within the theory context, as well as session exports;
 
wenzelm 
parents: 
70002 
diff
changeset
 | 
284  | 
val generatedN = "Generated_Code";  | 
| 
 
435fb018e8ee
"export_code ... file_prefix ..." is the preferred way to produce output within the logical file-system within the theory context, as well as session exports;
 
wenzelm 
parents: 
70002 
diff
changeset
 | 
285  | 
|
| 70021 | 286  | 
val code_path = Path.append (Path.basic codeN);  | 
287  | 
fun code_export_message thy = writeln (Export.message thy (Path.basic codeN));  | 
|
288  | 
||
289  | 
fun export binding content thy =  | 
|
290  | 
let  | 
|
291  | 
val thy' = thy |> Generated_Files.add_files (binding, content);  | 
|
| 
75604
 
39df30349778
more scalable generated files and code export, using Bytes.T;
 
wenzelm 
parents: 
74887 
diff
changeset
 | 
292  | 
val _ = Export.export thy' binding (Bytes.contents_blob content);  | 
| 70021 | 293  | 
in thy' end;  | 
294  | 
||
| 
70009
 
435fb018e8ee
"export_code ... file_prefix ..." is the preferred way to produce output within the logical file-system within the theory context, as well as session exports;
 
wenzelm 
parents: 
70002 
diff
changeset
 | 
295  | 
local  | 
| 
 
435fb018e8ee
"export_code ... file_prefix ..." is the preferred way to produce output within the logical file-system within the theory context, as well as session exports;
 
wenzelm 
parents: 
70002 
diff
changeset
 | 
296  | 
|
| 
70015
 
c8e08d8ffb93
clarified signature: more explicit type Path.binding;
 
wenzelm 
parents: 
70011 
diff
changeset
 | 
297  | 
fun export_logical (file_prefix, file_pos) format pretty_modules =  | 
| 
70009
 
435fb018e8ee
"export_code ... file_prefix ..." is the preferred way to produce output within the logical file-system within the theory context, as well as session exports;
 
wenzelm 
parents: 
70002 
diff
changeset
 | 
298  | 
let  | 
| 70021 | 299  | 
fun binding path = Path.binding (path, file_pos);  | 
300  | 
val prefix = code_path file_prefix;  | 
|
| 
70009
 
435fb018e8ee
"export_code ... file_prefix ..." is the preferred way to produce output within the logical file-system within the theory context, as well as session exports;
 
wenzelm 
parents: 
70002 
diff
changeset
 | 
301  | 
in  | 
| 
 
435fb018e8ee
"export_code ... file_prefix ..." is the preferred way to produce output within the logical file-system within the theory context, as well as session exports;
 
wenzelm 
parents: 
70002 
diff
changeset
 | 
302  | 
(case pretty_modules of  | 
| 70021 | 303  | 
Singleton (ext, p) => export (binding (Path.ext ext prefix)) (format p)  | 
| 70011 | 304  | 
| Hierarchy modules =>  | 
| 70021 | 305  | 
fold (fn (names, p) =>  | 
| 
72511
 
460d743010bc
clarified signature: overloaded "+" for Path.append;
 
wenzelm 
parents: 
72375 
diff
changeset
 | 
306  | 
export (binding (prefix + Path.make names)) (format p)) modules)  | 
| 70021 | 307  | 
#> tap code_export_message  | 
| 
70009
 
435fb018e8ee
"export_code ... file_prefix ..." is the preferred way to produce output within the logical file-system within the theory context, as well as session exports;
 
wenzelm 
parents: 
70002 
diff
changeset
 | 
308  | 
end;  | 
| 
 
435fb018e8ee
"export_code ... file_prefix ..." is the preferred way to produce output within the logical file-system within the theory context, as well as session exports;
 
wenzelm 
parents: 
70002 
diff
changeset
 | 
309  | 
|
| 
 
435fb018e8ee
"export_code ... file_prefix ..." is the preferred way to produce output within the logical file-system within the theory context, as well as session exports;
 
wenzelm 
parents: 
70002 
diff
changeset
 | 
310  | 
fun export_physical root format pretty_modules =  | 
| 
 
435fb018e8ee
"export_code ... file_prefix ..." is the preferred way to produce output within the logical file-system within the theory context, as well as session exports;
 
wenzelm 
parents: 
70002 
diff
changeset
 | 
311  | 
(case pretty_modules of  | 
| 
75604
 
39df30349778
more scalable generated files and code export, using Bytes.T;
 
wenzelm 
parents: 
74887 
diff
changeset
 | 
312  | 
Singleton (_, p) => Bytes.write root (format p)  | 
| 
70009
 
435fb018e8ee
"export_code ... file_prefix ..." is the preferred way to produce output within the logical file-system within the theory context, as well as session exports;
 
wenzelm 
parents: 
70002 
diff
changeset
 | 
313  | 
| Hierarchy code_modules =>  | 
| 
77859
 
a11e25bdd247
code_target: create subdirectories for export_code file
 
stuebinm <stuebinm@disroot.org> 
parents: 
76884 
diff
changeset
 | 
314  | 
List.app (fn (names, p) =>  | 
| 
 
a11e25bdd247
code_target: create subdirectories for export_code file
 
stuebinm <stuebinm@disroot.org> 
parents: 
76884 
diff
changeset
 | 
315  | 
let  | 
| 
 
a11e25bdd247
code_target: create subdirectories for export_code file
 
stuebinm <stuebinm@disroot.org> 
parents: 
76884 
diff
changeset
 | 
316  | 
val segments = map Path.basic names;  | 
| 
 
a11e25bdd247
code_target: create subdirectories for export_code file
 
stuebinm <stuebinm@disroot.org> 
parents: 
76884 
diff
changeset
 | 
317  | 
in  | 
| 
 
a11e25bdd247
code_target: create subdirectories for export_code file
 
stuebinm <stuebinm@disroot.org> 
parents: 
76884 
diff
changeset
 | 
318  | 
Isabelle_System.make_directory (Path.appends (root :: (fst (split_last segments))));  | 
| 
 
a11e25bdd247
code_target: create subdirectories for export_code file
 
stuebinm <stuebinm@disroot.org> 
parents: 
76884 
diff
changeset
 | 
319  | 
Bytes.write (Path.appends (root :: segments)) (format p)  | 
| 
 
a11e25bdd247
code_target: create subdirectories for export_code file
 
stuebinm <stuebinm@disroot.org> 
parents: 
76884 
diff
changeset
 | 
320  | 
end)  | 
| 
 
a11e25bdd247
code_target: create subdirectories for export_code file
 
stuebinm <stuebinm@disroot.org> 
parents: 
76884 
diff
changeset
 | 
321  | 
code_modules);  | 
| 
70009
 
435fb018e8ee
"export_code ... file_prefix ..." is the preferred way to produce output within the logical file-system within the theory context, as well as session exports;
 
wenzelm 
parents: 
70002 
diff
changeset
 | 
322  | 
|
| 
 
435fb018e8ee
"export_code ... file_prefix ..." is the preferred way to produce output within the logical file-system within the theory context, as well as session exports;
 
wenzelm 
parents: 
70002 
diff
changeset
 | 
323  | 
in  | 
| 
 
435fb018e8ee
"export_code ... file_prefix ..." is the preferred way to produce output within the logical file-system within the theory context, as well as session exports;
 
wenzelm 
parents: 
70002 
diff
changeset
 | 
324  | 
|
| 
 
435fb018e8ee
"export_code ... file_prefix ..." is the preferred way to produce output within the logical file-system within the theory context, as well as session exports;
 
wenzelm 
parents: 
70002 
diff
changeset
 | 
325  | 
fun export_result some_file format (pretty_code, _) thy =  | 
| 
 
435fb018e8ee
"export_code ... file_prefix ..." is the preferred way to produce output within the logical file-system within the theory context, as well as session exports;
 
wenzelm 
parents: 
70002 
diff
changeset
 | 
326  | 
(case some_file of  | 
| 
 
435fb018e8ee
"export_code ... file_prefix ..." is the preferred way to produce output within the logical file-system within the theory context, as well as session exports;
 
wenzelm 
parents: 
70002 
diff
changeset
 | 
327  | 
NONE =>  | 
| 
 
435fb018e8ee
"export_code ... file_prefix ..." is the preferred way to produce output within the logical file-system within the theory context, as well as session exports;
 
wenzelm 
parents: 
70002 
diff
changeset
 | 
328  | 
let val (file_prefix, thy') = next_export thy;  | 
| 
70015
 
c8e08d8ffb93
clarified signature: more explicit type Path.binding;
 
wenzelm 
parents: 
70011 
diff
changeset
 | 
329  | 
in export_logical (Path.basic file_prefix, Position.none) format pretty_code thy' end  | 
| 
70009
 
435fb018e8ee
"export_code ... file_prefix ..." is the preferred way to produce output within the logical file-system within the theory context, as well as session exports;
 
wenzelm 
parents: 
70002 
diff
changeset
 | 
330  | 
  | SOME ({physical = false}, file_prefix) =>
 | 
| 
 
435fb018e8ee
"export_code ... file_prefix ..." is the preferred way to produce output within the logical file-system within the theory context, as well as session exports;
 
wenzelm 
parents: 
70002 
diff
changeset
 | 
331  | 
export_logical file_prefix format pretty_code thy  | 
| 
70015
 
c8e08d8ffb93
clarified signature: more explicit type Path.binding;
 
wenzelm 
parents: 
70011 
diff
changeset
 | 
332  | 
  | SOME ({physical = true}, (file, _)) =>
 | 
| 
70009
 
435fb018e8ee
"export_code ... file_prefix ..." is the preferred way to produce output within the logical file-system within the theory context, as well as session exports;
 
wenzelm 
parents: 
70002 
diff
changeset
 | 
333  | 
let  | 
| 
 
435fb018e8ee
"export_code ... file_prefix ..." is the preferred way to produce output within the logical file-system within the theory context, as well as session exports;
 
wenzelm 
parents: 
70002 
diff
changeset
 | 
334  | 
val root = File.full_path (Resources.master_directory thy) file;  | 
| 
 
435fb018e8ee
"export_code ... file_prefix ..." is the preferred way to produce output within the logical file-system within the theory context, as well as session exports;
 
wenzelm 
parents: 
70002 
diff
changeset
 | 
335  | 
val _ = File.check_dir (Path.dir root);  | 
| 
 
435fb018e8ee
"export_code ... file_prefix ..." is the preferred way to produce output within the logical file-system within the theory context, as well as session exports;
 
wenzelm 
parents: 
70002 
diff
changeset
 | 
336  | 
val _ = export_physical root format pretty_code;  | 
| 
 
435fb018e8ee
"export_code ... file_prefix ..." is the preferred way to produce output within the logical file-system within the theory context, as well as session exports;
 
wenzelm 
parents: 
70002 
diff
changeset
 | 
337  | 
in thy end);  | 
| 
 
435fb018e8ee
"export_code ... file_prefix ..." is the preferred way to produce output within the logical file-system within the theory context, as well as session exports;
 
wenzelm 
parents: 
70002 
diff
changeset
 | 
338  | 
|
| 
 
435fb018e8ee
"export_code ... file_prefix ..." is the preferred way to produce output within the logical file-system within the theory context, as well as session exports;
 
wenzelm 
parents: 
70002 
diff
changeset
 | 
339  | 
fun produce_result syms width pretty_modules =  | 
| 
 
435fb018e8ee
"export_code ... file_prefix ..." is the preferred way to produce output within the logical file-system within the theory context, as well as session exports;
 
wenzelm 
parents: 
70002 
diff
changeset
 | 
340  | 
let val format = Code_Printer.format [] width in  | 
| 
 
435fb018e8ee
"export_code ... file_prefix ..." is the preferred way to produce output within the logical file-system within the theory context, as well as session exports;
 
wenzelm 
parents: 
70002 
diff
changeset
 | 
341  | 
(case pretty_modules of  | 
| 
 
435fb018e8ee
"export_code ... file_prefix ..." is the preferred way to produce output within the logical file-system within the theory context, as well as session exports;
 
wenzelm 
parents: 
70002 
diff
changeset
 | 
342  | 
(Singleton (_, p), deresolve) => ([([], format p)], map deresolve syms)  | 
| 
 
435fb018e8ee
"export_code ... file_prefix ..." is the preferred way to produce output within the logical file-system within the theory context, as well as session exports;
 
wenzelm 
parents: 
70002 
diff
changeset
 | 
343  | 
| (Hierarchy code_modules, deresolve) =>  | 
| 
 
435fb018e8ee
"export_code ... file_prefix ..." is the preferred way to produce output within the logical file-system within the theory context, as well as session exports;
 
wenzelm 
parents: 
70002 
diff
changeset
 | 
344  | 
((map o apsnd) format code_modules, map deresolve syms))  | 
| 
 
435fb018e8ee
"export_code ... file_prefix ..." is the preferred way to produce output within the logical file-system within the theory context, as well as session exports;
 
wenzelm 
parents: 
70002 
diff
changeset
 | 
345  | 
end;  | 
| 
 
435fb018e8ee
"export_code ... file_prefix ..." is the preferred way to produce output within the logical file-system within the theory context, as well as session exports;
 
wenzelm 
parents: 
70002 
diff
changeset
 | 
346  | 
|
| 
 
435fb018e8ee
"export_code ... file_prefix ..." is the preferred way to produce output within the logical file-system within the theory context, as well as session exports;
 
wenzelm 
parents: 
70002 
diff
changeset
 | 
347  | 
fun present_result selects width (pretty_modules, _) =  | 
| 
 
435fb018e8ee
"export_code ... file_prefix ..." is the preferred way to produce output within the logical file-system within the theory context, as well as session exports;
 
wenzelm 
parents: 
70002 
diff
changeset
 | 
348  | 
let val format = Code_Printer.format selects width in  | 
| 
 
435fb018e8ee
"export_code ... file_prefix ..." is the preferred way to produce output within the logical file-system within the theory context, as well as session exports;
 
wenzelm 
parents: 
70002 
diff
changeset
 | 
349  | 
(case pretty_modules of  | 
| 
 
435fb018e8ee
"export_code ... file_prefix ..." is the preferred way to produce output within the logical file-system within the theory context, as well as session exports;
 
wenzelm 
parents: 
70002 
diff
changeset
 | 
350  | 
Singleton (_, p) => format p  | 
| 
75604
 
39df30349778
more scalable generated files and code export, using Bytes.T;
 
wenzelm 
parents: 
74887 
diff
changeset
 | 
351  | 
| Hierarchy code_modules =>  | 
| 
 
39df30349778
more scalable generated files and code export, using Bytes.T;
 
wenzelm 
parents: 
74887 
diff
changeset
 | 
352  | 
Bytes.appends (separate (Bytes.string "\n\n") (map (format o #2) code_modules)))  | 
| 
70009
 
435fb018e8ee
"export_code ... file_prefix ..." is the preferred way to produce output within the logical file-system within the theory context, as well as session exports;
 
wenzelm 
parents: 
70002 
diff
changeset
 | 
353  | 
end;  | 
| 
 
435fb018e8ee
"export_code ... file_prefix ..." is the preferred way to produce output within the logical file-system within the theory context, as well as session exports;
 
wenzelm 
parents: 
70002 
diff
changeset
 | 
354  | 
|
| 
 
435fb018e8ee
"export_code ... file_prefix ..." is the preferred way to produce output within the logical file-system within the theory context, as well as session exports;
 
wenzelm 
parents: 
70002 
diff
changeset
 | 
355  | 
end;  | 
| 
 
435fb018e8ee
"export_code ... file_prefix ..." is the preferred way to produce output within the logical file-system within the theory context, as well as session exports;
 
wenzelm 
parents: 
70002 
diff
changeset
 | 
356  | 
|
| 
 
435fb018e8ee
"export_code ... file_prefix ..." is the preferred way to produce output within the logical file-system within the theory context, as well as session exports;
 
wenzelm 
parents: 
70002 
diff
changeset
 | 
357  | 
|
| 34021 | 358  | 
(** serializer usage **)  | 
359  | 
||
| 59324 | 360  | 
(* technical aside: pretty printing width *)  | 
361  | 
||
| 69593 | 362  | 
val default_code_width = Attrib.setup_config_int \<^binding>\<open>default_code_width\<close> (K 80);  | 
| 59324 | 363  | 
|
| 
69623
 
ef02c5e051e5
explicit model concerning files of generated code
 
haftmann 
parents: 
69593 
diff
changeset
 | 
364  | 
fun default_width ctxt = Config.get ctxt default_code_width;  | 
| 
 
ef02c5e051e5
explicit model concerning files of generated code
 
haftmann 
parents: 
69593 
diff
changeset
 | 
365  | 
|
| 
 
ef02c5e051e5
explicit model concerning files of generated code
 
haftmann 
parents: 
69593 
diff
changeset
 | 
366  | 
val the_width = the_default o default_width;  | 
| 
 
ef02c5e051e5
explicit model concerning files of generated code
 
haftmann 
parents: 
69593 
diff
changeset
 | 
367  | 
|
| 59324 | 368  | 
|
| 34021 | 369  | 
(* montage *)  | 
370  | 
||
| 59102 | 371  | 
fun the_language ctxt =  | 
| 59103 | 372  | 
#language o fst o the o lookup_target_data (Proof_Context.theory_of ctxt);  | 
| 
37822
 
cf3588177676
use generic description slot for formal code checking
 
haftmann 
parents: 
37821 
diff
changeset
 | 
373  | 
|
| 59102 | 374  | 
fun the_literals ctxt = #literals o the_language ctxt;  | 
| 34021 | 375  | 
|
| 
67207
 
ad538f6c5d2f
dedicated case option for code generation to Scala
 
haftmann 
parents: 
66586 
diff
changeset
 | 
376  | 
fun the_evaluation_args ctxt = #evaluation_args o the_language ctxt;  | 
| 
 
ad538f6c5d2f
dedicated case option for code generation to Scala
 
haftmann 
parents: 
66586 
diff
changeset
 | 
377  | 
|
| 
39817
 
5228c6b20273
check whole target hierarchy for existing reserved symbols
 
haftmann 
parents: 
39750 
diff
changeset
 | 
378  | 
local  | 
| 
 
5228c6b20273
check whole target hierarchy for existing reserved symbols
 
haftmann 
parents: 
39750 
diff
changeset
 | 
379  | 
|
| 59099 | 380  | 
fun activate_target ctxt target_name =  | 
| 
39817
 
5228c6b20273
check whole target hierarchy for existing reserved symbols
 
haftmann 
parents: 
39750 
diff
changeset
 | 
381  | 
let  | 
| 59324 | 382  | 
val thy = Proof_Context.theory_of ctxt;  | 
| 
69623
 
ef02c5e051e5
explicit model concerning files of generated code
 
haftmann 
parents: 
69593 
diff
changeset
 | 
383  | 
val (modify, (target, data)) = join_ancestry thy target_name;  | 
| 
 
ef02c5e051e5
explicit model concerning files of generated code
 
haftmann 
parents: 
69593 
diff
changeset
 | 
384  | 
val serializer = (#serializer o #language) target;  | 
| 
 
ef02c5e051e5
explicit model concerning files of generated code
 
haftmann 
parents: 
69593 
diff
changeset
 | 
385  | 
  in { serializer = serializer, data = data, modify = modify } end;
 | 
| 38927 | 386  | 
|
| 
69623
 
ef02c5e051e5
explicit model concerning files of generated code
 
haftmann 
parents: 
69593 
diff
changeset
 | 
387  | 
fun project_program_for_syms ctxt syms_hidden syms1 program1 =  | 
| 38927 | 388  | 
let  | 
| 
55147
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55146 
diff
changeset
 | 
389  | 
val syms2 = subtract (op =) syms_hidden syms1;  | 
| 69527 | 390  | 
val program2 = Code_Symbol.Graph.restrict (not o member (op =) syms_hidden) program1;  | 
391  | 
val unimplemented = Code_Thingol.unimplemented program2;  | 
|
| 42359 | 392  | 
val _ =  | 
| 
54889
 
4121d64fde90
explicit distinction between empty code equations and no code equations, including convenient declaration attributes
 
haftmann 
parents: 
54312 
diff
changeset
 | 
393  | 
if null unimplemented then ()  | 
| 42359 | 394  | 
      else error ("No code equations for " ^
 | 
| 55304 | 395  | 
commas (map (Proof_Context.markup_const ctxt) unimplemented));  | 
| 69527 | 396  | 
val syms3 = Code_Symbol.Graph.all_succs program2 syms2;  | 
397  | 
val program3 = Code_Symbol.Graph.restrict (member (op =) syms3) program2;  | 
|
398  | 
in program3 end;  | 
|
| 38927 | 399  | 
|
| 
69623
 
ef02c5e051e5
explicit model concerning files of generated code
 
haftmann 
parents: 
69593 
diff
changeset
 | 
400  | 
fun project_includes_for_syms syms includes =  | 
| 
 
ef02c5e051e5
explicit model concerning files of generated code
 
haftmann 
parents: 
69593 
diff
changeset
 | 
401  | 
let  | 
| 
 
ef02c5e051e5
explicit model concerning files of generated code
 
haftmann 
parents: 
69593 
diff
changeset
 | 
402  | 
fun select_include (name, (content, cs)) =  | 
| 
 
ef02c5e051e5
explicit model concerning files of generated code
 
haftmann 
parents: 
69593 
diff
changeset
 | 
403  | 
if null cs orelse exists (member (op =) syms) cs  | 
| 
 
ef02c5e051e5
explicit model concerning files of generated code
 
haftmann 
parents: 
69593 
diff
changeset
 | 
404  | 
then SOME (name, content) else NONE;  | 
| 
 
ef02c5e051e5
explicit model concerning files of generated code
 
haftmann 
parents: 
69593 
diff
changeset
 | 
405  | 
in map_filter select_include includes end;  | 
| 
 
ef02c5e051e5
explicit model concerning files of generated code
 
haftmann 
parents: 
69593 
diff
changeset
 | 
406  | 
|
| 
 
ef02c5e051e5
explicit model concerning files of generated code
 
haftmann 
parents: 
69593 
diff
changeset
 | 
407  | 
fun prepare_serializer ctxt target_name module_name args proto_program syms =  | 
| 38927 | 408  | 
let  | 
| 
69623
 
ef02c5e051e5
explicit model concerning files of generated code
 
haftmann 
parents: 
69593 
diff
changeset
 | 
409  | 
    val { serializer, data, modify } = activate_target ctxt target_name;
 | 
| 
 
ef02c5e051e5
explicit model concerning files of generated code
 
haftmann 
parents: 
69593 
diff
changeset
 | 
410  | 
val printings = Code_Printer.the_printings data;  | 
| 
 
ef02c5e051e5
explicit model concerning files of generated code
 
haftmann 
parents: 
69593 
diff
changeset
 | 
411  | 
val _ = if module_name = "" then () else (check_name true module_name; ())  | 
| 
 
ef02c5e051e5
explicit model concerning files of generated code
 
haftmann 
parents: 
69593 
diff
changeset
 | 
412  | 
val hidden_syms = Code_Symbol.symbols_of printings;  | 
| 
 
ef02c5e051e5
explicit model concerning files of generated code
 
haftmann 
parents: 
69593 
diff
changeset
 | 
413  | 
val prepared_program = project_program_for_syms ctxt hidden_syms syms proto_program;  | 
| 
 
ef02c5e051e5
explicit model concerning files of generated code
 
haftmann 
parents: 
69593 
diff
changeset
 | 
414  | 
val prepared_syms = subtract (op =) hidden_syms syms;  | 
| 
 
ef02c5e051e5
explicit model concerning files of generated code
 
haftmann 
parents: 
69593 
diff
changeset
 | 
415  | 
val all_syms = Code_Symbol.Graph.all_succs proto_program syms;  | 
| 
 
ef02c5e051e5
explicit model concerning files of generated code
 
haftmann 
parents: 
69593 
diff
changeset
 | 
416  | 
val includes = project_includes_for_syms all_syms  | 
| 
 
ef02c5e051e5
explicit model concerning files of generated code
 
haftmann 
parents: 
69593 
diff
changeset
 | 
417  | 
(Code_Symbol.dest_module_data printings);  | 
| 
 
ef02c5e051e5
explicit model concerning files of generated code
 
haftmann 
parents: 
69593 
diff
changeset
 | 
418  | 
    val prepared_serializer = serializer args ctxt {
 | 
| 
 
ef02c5e051e5
explicit model concerning files of generated code
 
haftmann 
parents: 
69593 
diff
changeset
 | 
419  | 
reserved_syms = Code_Printer.the_reserved data,  | 
| 
 
ef02c5e051e5
explicit model concerning files of generated code
 
haftmann 
parents: 
69593 
diff
changeset
 | 
420  | 
identifiers = Code_Printer.the_identifiers data,  | 
| 38926 | 421  | 
includes = includes,  | 
| 
55149
 
626d8f08d479
immediate "activation" of const syntax at declaration time
 
haftmann 
parents: 
55147 
diff
changeset
 | 
422  | 
const_syntax = Code_Symbol.lookup_constant_data printings,  | 
| 
 
626d8f08d479
immediate "activation" of const syntax at declaration time
 
haftmann 
parents: 
55147 
diff
changeset
 | 
423  | 
tyco_syntax = Code_Symbol.lookup_type_constructor_data printings,  | 
| 
69623
 
ef02c5e051e5
explicit model concerning files of generated code
 
haftmann 
parents: 
69593 
diff
changeset
 | 
424  | 
class_syntax = Code_Symbol.lookup_type_class_data printings,  | 
| 
 
ef02c5e051e5
explicit model concerning files of generated code
 
haftmann 
parents: 
69593 
diff
changeset
 | 
425  | 
module_name = module_name };  | 
| 
 
ef02c5e051e5
explicit model concerning files of generated code
 
haftmann 
parents: 
69593 
diff
changeset
 | 
426  | 
in  | 
| 
 
ef02c5e051e5
explicit model concerning files of generated code
 
haftmann 
parents: 
69593 
diff
changeset
 | 
427  | 
(prepared_serializer o modify, (prepared_program, prepared_syms))  | 
| 34021 | 428  | 
end;  | 
429  | 
||
| 
69623
 
ef02c5e051e5
explicit model concerning files of generated code
 
haftmann 
parents: 
69593 
diff
changeset
 | 
430  | 
fun invoke_serializer ctxt target_name module_name args program all_public syms =  | 
| 34021 | 431  | 
let  | 
| 
69623
 
ef02c5e051e5
explicit model concerning files of generated code
 
haftmann 
parents: 
69593 
diff
changeset
 | 
432  | 
val (prepared_serializer, (prepared_program, prepared_syms)) =  | 
| 
 
ef02c5e051e5
explicit model concerning files of generated code
 
haftmann 
parents: 
69593 
diff
changeset
 | 
433  | 
prepare_serializer ctxt target_name module_name args program syms;  | 
| 
 
ef02c5e051e5
explicit model concerning files of generated code
 
haftmann 
parents: 
69593 
diff
changeset
 | 
434  | 
val exports = if all_public then [] else prepared_syms;  | 
| 63164 | 435  | 
in  | 
436  | 
Code_Preproc.timed_exec "serializing"  | 
|
| 
69623
 
ef02c5e051e5
explicit model concerning files of generated code
 
haftmann 
parents: 
69593 
diff
changeset
 | 
437  | 
(fn () => prepared_serializer prepared_program exports) ctxt  | 
| 63164 | 438  | 
end;  | 
| 34021 | 439  | 
|
| 
52138
 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 
haftmann 
parents: 
52137 
diff
changeset
 | 
440  | 
fun assert_module_name "" = error "Empty module name not allowed here"  | 
| 
38933
 
bd77e092f67c
avoid strange special treatment of empty module names
 
haftmann 
parents: 
38929 
diff
changeset
 | 
441  | 
| assert_module_name module_name = module_name;  | 
| 
 
bd77e092f67c
avoid strange special treatment of empty module names
 
haftmann 
parents: 
38929 
diff
changeset
 | 
442  | 
|
| 34021 | 443  | 
in  | 
444  | 
||
| 70002 | 445  | 
fun export_code_for some_file target_name module_name some_width args program all_public cs lthy =  | 
| 
69999
 
76554a0cafe3
export_code/check_code formally updates the theory -- this opens further possibilities concerning Generated_Files;
 
wenzelm 
parents: 
69998 
diff
changeset
 | 
446  | 
let  | 
| 70002 | 447  | 
val format = Code_Printer.format [] (the_width lthy some_width);  | 
448  | 
val res = invoke_serializer lthy target_name module_name args program all_public cs;  | 
|
449  | 
in Local_Theory.background_theory (export_result some_file format res) lthy end;  | 
|
| 38918 | 450  | 
|
| 
69623
 
ef02c5e051e5
explicit model concerning files of generated code
 
haftmann 
parents: 
69593 
diff
changeset
 | 
451  | 
fun produce_code_for ctxt target_name module_name some_width args =  | 
| 39102 | 452  | 
let  | 
| 
69623
 
ef02c5e051e5
explicit model concerning files of generated code
 
haftmann 
parents: 
69593 
diff
changeset
 | 
453  | 
val serializer = invoke_serializer ctxt target_name (assert_module_name module_name) args;  | 
| 55683 | 454  | 
in fn program => fn all_public => fn syms =>  | 
| 
69623
 
ef02c5e051e5
explicit model concerning files of generated code
 
haftmann 
parents: 
69593 
diff
changeset
 | 
455  | 
produce_result syms (the_width ctxt some_width)  | 
| 
 
ef02c5e051e5
explicit model concerning files of generated code
 
haftmann 
parents: 
69593 
diff
changeset
 | 
456  | 
(serializer program all_public syms)  | 
| 
39484
 
505f95975a5a
closures separate serializer initialization from serializer invocation as far as appropriate
 
haftmann 
parents: 
39480 
diff
changeset
 | 
457  | 
end;  | 
| 
38929
 
d9ac9dee764d
distinguish code production and code presentation
 
haftmann 
parents: 
38928 
diff
changeset
 | 
458  | 
|
| 
69623
 
ef02c5e051e5
explicit model concerning files of generated code
 
haftmann 
parents: 
69593 
diff
changeset
 | 
459  | 
fun present_code_for ctxt target_name module_name some_width args =  | 
| 
39484
 
505f95975a5a
closures separate serializer initialization from serializer invocation as far as appropriate
 
haftmann 
parents: 
39480 
diff
changeset
 | 
460  | 
let  | 
| 
69623
 
ef02c5e051e5
explicit model concerning files of generated code
 
haftmann 
parents: 
69593 
diff
changeset
 | 
461  | 
val serializer = invoke_serializer ctxt target_name (assert_module_name module_name) args;  | 
| 
55147
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55146 
diff
changeset
 | 
462  | 
in fn program => fn (syms, selects) =>  | 
| 
69623
 
ef02c5e051e5
explicit model concerning files of generated code
 
haftmann 
parents: 
69593 
diff
changeset
 | 
463  | 
present_result selects (the_width ctxt some_width) (serializer program false syms)  | 
| 
39484
 
505f95975a5a
closures separate serializer initialization from serializer invocation as far as appropriate
 
haftmann 
parents: 
39480 
diff
changeset
 | 
464  | 
end;  | 
| 38918 | 465  | 
|
| 70002 | 466  | 
fun check_code_for target_name strict args program all_public syms lthy =  | 
| 37824 | 467  | 
let  | 
| 70002 | 468  | 
    val { env_var, make_destination, make_command } = #check (the_language lthy target_name);
 | 
| 70000 | 469  | 
val format = Code_Printer.format [] 80;  | 
| 41939 | 470  | 
fun ext_check p =  | 
| 
43564
 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
43324 
diff
changeset
 | 
471  | 
let  | 
| 37824 | 472  | 
val destination = make_destination p;  | 
| 70002 | 473  | 
val lthy' = lthy  | 
474  | 
|> Local_Theory.background_theory  | 
|
| 
70015
 
c8e08d8ffb93
clarified signature: more explicit type Path.binding;
 
wenzelm 
parents: 
70011 
diff
changeset
 | 
475  | 
            (export_result (SOME ({physical = true}, (destination, Position.none))) format
 | 
| 70002 | 476  | 
(invoke_serializer lthy target_name generatedN args program all_public syms));  | 
| 48568 | 477  | 
val cmd = make_command generatedN;  | 
| 
69906
 
55534affe445
migrated from Nums to Zarith as library for OCaml integer arithmetic
 
haftmann 
parents: 
69784 
diff
changeset
 | 
478  | 
val context_cmd = "cd " ^ File.bash_path p ^ " && " ^ cmd ^ " 2>&1";  | 
| 
43850
 
7f2cbc713344
moved bash operations to Isabelle_System (cf. Scala version);
 
wenzelm 
parents: 
43564 
diff
changeset
 | 
479  | 
in  | 
| 
69906
 
55534affe445
migrated from Nums to Zarith as library for OCaml integer arithmetic
 
haftmann 
parents: 
69784 
diff
changeset
 | 
480  | 
if Isabelle_System.bash context_cmd <> 0  | 
| 59099 | 481  | 
        then error ("Code check failed for " ^ target_name ^ ": " ^ cmd)
 | 
| 70002 | 482  | 
else lthy'  | 
| 37824 | 483  | 
end;  | 
| 
43850
 
7f2cbc713344
moved bash operations to Isabelle_System (cf. Scala version);
 
wenzelm 
parents: 
43564 
diff
changeset
 | 
484  | 
in  | 
| 
69999
 
76554a0cafe3
export_code/check_code formally updates the theory -- this opens further possibilities concerning Generated_Files;
 
wenzelm 
parents: 
69998 
diff
changeset
 | 
485  | 
if not (env_var = "") andalso getenv env_var = "" then  | 
| 
 
76554a0cafe3
export_code/check_code formally updates the theory -- this opens further possibilities concerning Generated_Files;
 
wenzelm 
parents: 
69998 
diff
changeset
 | 
486  | 
if strict  | 
| 59099 | 487  | 
then error (env_var ^ " not set; cannot check code for " ^ target_name)  | 
| 70002 | 488  | 
else (warning (env_var ^ " not set; skipped checking code for " ^ target_name); lthy)  | 
| 41939 | 489  | 
else Isabelle_System.with_tmp_dir "Code_Test" ext_check  | 
| 37824 | 490  | 
end;  | 
491  | 
||
| 
69623
 
ef02c5e051e5
explicit model concerning files of generated code
 
haftmann 
parents: 
69593 
diff
changeset
 | 
492  | 
fun dynamic_compilation_text prepared_serializer width prepared_program syms all_public ((vs, ty), t) =  | 
| 
41344
 
d990badc97a3
evaluator separating static and dynamic operations
 
haftmann 
parents: 
41342 
diff
changeset
 | 
493  | 
let  | 
| 
 
d990badc97a3
evaluator separating static and dynamic operations
 
haftmann 
parents: 
41342 
diff
changeset
 | 
494  | 
val _ = if Code_Thingol.contains_dict_var t then  | 
| 
 
d990badc97a3
evaluator separating static and dynamic operations
 
haftmann 
parents: 
41342 
diff
changeset
 | 
495  | 
error "Term to be evaluated contains free dictionaries" else ();  | 
| 
43324
 
2b47822868e4
discontinued Name.variant to emphasize that this is old-style / indirect;
 
wenzelm 
parents: 
42361 
diff
changeset
 | 
496  | 
val v' = singleton (Name.variant_list (map fst vs)) "a";  | 
| 
41344
 
d990badc97a3
evaluator separating static and dynamic operations
 
haftmann 
parents: 
41342 
diff
changeset
 | 
497  | 
val vs' = (v', []) :: vs;  | 
| 
55147
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55146 
diff
changeset
 | 
498  | 
val ty' = ITyVar v' `-> ty;  | 
| 
41344
 
d990badc97a3
evaluator separating static and dynamic operations
 
haftmann 
parents: 
41342 
diff
changeset
 | 
499  | 
val program = prepared_program  | 
| 
55147
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55146 
diff
changeset
 | 
500  | 
|> Code_Symbol.Graph.new_node (Code_Symbol.value,  | 
| 
57249
 
5c75baf68b77
formal variable name: IVar NONE is strictly spoken not supported on lhs of function definitions, e.g. in Scala
 
haftmann 
parents: 
56920 
diff
changeset
 | 
501  | 
Code_Thingol.Fun (((vs', ty'), [(([IVar (SOME "dummy")], t), (NONE, true))]), NONE))  | 
| 
55147
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55146 
diff
changeset
 | 
502  | 
|> fold (curry (perhaps o try o  | 
| 
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
55146 
diff
changeset
 | 
503  | 
Code_Symbol.Graph.add_edge) Code_Symbol.value) syms;  | 
| 
69623
 
ef02c5e051e5
explicit model concerning files of generated code
 
haftmann 
parents: 
69593 
diff
changeset
 | 
504  | 
val (pretty_code, deresolve) =  | 
| 
 
ef02c5e051e5
explicit model concerning files of generated code
 
haftmann 
parents: 
69593 
diff
changeset
 | 
505  | 
prepared_serializer program (if all_public then [] else [Code_Symbol.value]);  | 
| 
 
ef02c5e051e5
explicit model concerning files of generated code
 
haftmann 
parents: 
69593 
diff
changeset
 | 
506  | 
val (compilation_code, [SOME value_name]) =  | 
| 
 
ef02c5e051e5
explicit model concerning files of generated code
 
haftmann 
parents: 
69593 
diff
changeset
 | 
507  | 
produce_result [Code_Symbol.value] width (pretty_code, deresolve);  | 
| 
 
ef02c5e051e5
explicit model concerning files of generated code
 
haftmann 
parents: 
69593 
diff
changeset
 | 
508  | 
in ((compilation_code, value_name), deresolve) end;  | 
| 
41344
 
d990badc97a3
evaluator separating static and dynamic operations
 
haftmann 
parents: 
41342 
diff
changeset
 | 
509  | 
|
| 64959 | 510  | 
fun compilation_text' ctxt target_name some_module_name program syms =  | 
| 
41344
 
d990badc97a3
evaluator separating static and dynamic operations
 
haftmann 
parents: 
41342 
diff
changeset
 | 
511  | 
let  | 
| 
69623
 
ef02c5e051e5
explicit model concerning files of generated code
 
haftmann 
parents: 
69593 
diff
changeset
 | 
512  | 
val width = default_width ctxt;  | 
| 
67207
 
ad538f6c5d2f
dedicated case option for code generation to Scala
 
haftmann 
parents: 
66586 
diff
changeset
 | 
513  | 
val evaluation_args = the_evaluation_args ctxt target_name;  | 
| 
69623
 
ef02c5e051e5
explicit model concerning files of generated code
 
haftmann 
parents: 
69593 
diff
changeset
 | 
514  | 
val (prepared_serializer, (prepared_program, _)) =  | 
| 
 
ef02c5e051e5
explicit model concerning files of generated code
 
haftmann 
parents: 
69593 
diff
changeset
 | 
515  | 
prepare_serializer ctxt target_name (the_default generatedN some_module_name) evaluation_args program syms;  | 
| 63164 | 516  | 
in  | 
517  | 
Code_Preproc.timed_exec "serializing"  | 
|
| 
69623
 
ef02c5e051e5
explicit model concerning files of generated code
 
haftmann 
parents: 
69593 
diff
changeset
 | 
518  | 
(fn () => dynamic_compilation_text prepared_serializer width prepared_program syms) ctxt  | 
| 63164 | 519  | 
end;  | 
| 
41344
 
d990badc97a3
evaluator separating static and dynamic operations
 
haftmann 
parents: 
41342 
diff
changeset
 | 
520  | 
|
| 64959 | 521  | 
fun compilation_text ctxt target_name program syms =  | 
522  | 
fst oo compilation_text' ctxt target_name NONE program syms  | 
|
| 69998 | 523  | 
|
| 34021 | 524  | 
end; (* local *)  | 
525  | 
||
526  | 
||
527  | 
(* code generation *)  | 
|
528  | 
||
| 72841 | 529  | 
fun prep_destination (location, source) =  | 
530  | 
let  | 
|
531  | 
val s = Input.string_of source  | 
|
532  | 
val pos = Input.pos_of source  | 
|
533  | 
val delimited = Input.is_delimited source  | 
|
534  | 
in  | 
|
535  | 
    if location = {physical = false}
 | 
|
536  | 
then (location, Path.explode_pos (s, pos))  | 
|
537  | 
else  | 
|
538  | 
let  | 
|
539  | 
val _ =  | 
|
540  | 
if s = ""  | 
|
541  | 
          then error ("Bad bad empty " ^ Markup.markup Markup.keyword2 "file" ^ " argument")
 | 
|
542  | 
else ();  | 
|
543  | 
val _ =  | 
|
544  | 
legacy_feature  | 
|
545  | 
(Markup.markup Markup.keyword1 "export_code" ^ " with " ^  | 
|
546  | 
Markup.markup Markup.keyword2 "file" ^ " argument" ^ Position.here pos);  | 
|
547  | 
val _ = Position.report pos (Markup.language_path delimited);  | 
|
548  | 
val path = #1 (Path.explode_pos (s, pos));  | 
|
| 76884 | 549  | 
val _ = Position.report pos (Markup.path (File.symbolic_path path));  | 
| 72841 | 550  | 
in (location, (path, pos)) end  | 
551  | 
end;  | 
|
| 38918 | 552  | 
|
| 70002 | 553  | 
fun export_code all_public cs seris lthy =  | 
| 34021 | 554  | 
let  | 
| 70002 | 555  | 
val program = Code_Thingol.consts_program lthy cs;  | 
| 
69999
 
76554a0cafe3
export_code/check_code formally updates the theory -- this opens further possibilities concerning Generated_Files;
 
wenzelm 
parents: 
69998 
diff
changeset
 | 
556  | 
in  | 
| 70002 | 557  | 
(seris, lthy) |-> fold (fn (((target_name, module_name), some_file), args) =>  | 
| 
69999
 
76554a0cafe3
export_code/check_code formally updates the theory -- this opens further possibilities concerning Generated_Files;
 
wenzelm 
parents: 
69998 
diff
changeset
 | 
558  | 
export_code_for some_file target_name module_name NONE args  | 
| 
 
76554a0cafe3
export_code/check_code formally updates the theory -- this opens further possibilities concerning Generated_Files;
 
wenzelm 
parents: 
69998 
diff
changeset
 | 
559  | 
program all_public (map Constant cs))  | 
| 
 
76554a0cafe3
export_code/check_code formally updates the theory -- this opens further possibilities concerning Generated_Files;
 
wenzelm 
parents: 
69998 
diff
changeset
 | 
560  | 
end;  | 
| 34021 | 561  | 
|
| 70002 | 562  | 
fun export_code_cmd all_public raw_cs seris lthy =  | 
| 
69999
 
76554a0cafe3
export_code/check_code formally updates the theory -- this opens further possibilities concerning Generated_Files;
 
wenzelm 
parents: 
69998 
diff
changeset
 | 
563  | 
let  | 
| 70002 | 564  | 
val cs = Code_Thingol.read_const_exprs lthy raw_cs;  | 
565  | 
in export_code all_public cs ((map o apfst o apsnd o Option.map) prep_destination seris) lthy end;  | 
|
| 38918 | 566  | 
|
| 59099 | 567  | 
fun produce_code ctxt all_public cs target_name some_width some_module_name args =  | 
| 38918 | 568  | 
let  | 
| 63159 | 569  | 
val program = Code_Thingol.consts_program ctxt cs;  | 
| 59099 | 570  | 
in produce_code_for ctxt target_name some_width some_module_name args program all_public (map Constant cs) end;  | 
| 
38929
 
d9ac9dee764d
distinguish code production and code presentation
 
haftmann 
parents: 
38928 
diff
changeset
 | 
571  | 
|
| 59099 | 572  | 
fun present_code ctxt cs syms target_name some_width some_module_name args =  | 
| 
38929
 
d9ac9dee764d
distinguish code production and code presentation
 
haftmann 
parents: 
38928 
diff
changeset
 | 
573  | 
let  | 
| 63159 | 574  | 
val program = Code_Thingol.consts_program ctxt cs;  | 
| 59099 | 575  | 
in present_code_for ctxt target_name some_width some_module_name args program (map Constant cs, syms) end;  | 
| 34021 | 576  | 
|
| 70002 | 577  | 
fun check_code all_public cs seris lthy =  | 
| 37824 | 578  | 
let  | 
| 70002 | 579  | 
val program = Code_Thingol.consts_program lthy cs;  | 
| 
69999
 
76554a0cafe3
export_code/check_code formally updates the theory -- this opens further possibilities concerning Generated_Files;
 
wenzelm 
parents: 
69998 
diff
changeset
 | 
580  | 
in  | 
| 70002 | 581  | 
(seris, lthy) |-> fold (fn ((target_name, strict), args) =>  | 
| 
69999
 
76554a0cafe3
export_code/check_code formally updates the theory -- this opens further possibilities concerning Generated_Files;
 
wenzelm 
parents: 
69998 
diff
changeset
 | 
582  | 
check_code_for target_name strict args program all_public (map Constant cs))  | 
| 
 
76554a0cafe3
export_code/check_code formally updates the theory -- this opens further possibilities concerning Generated_Files;
 
wenzelm 
parents: 
69998 
diff
changeset
 | 
583  | 
end;  | 
| 37824 | 584  | 
|
| 70002 | 585  | 
fun check_code_cmd all_public raw_cs seris lthy =  | 
586  | 
check_code all_public (Code_Thingol.read_const_exprs lthy raw_cs) seris lthy;  | 
|
| 37824 | 587  | 
|
| 34021 | 588  | 
|
| 28054 | 589  | 
(** serializer configuration **)  | 
| 27000 | 590  | 
|
| 
52138
 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 
haftmann 
parents: 
52137 
diff
changeset
 | 
591  | 
(* reserved symbol names *)  | 
| 
 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 
haftmann 
parents: 
52137 
diff
changeset
 | 
592  | 
|
| 59099 | 593  | 
fun add_reserved target_name sym thy =  | 
| 
52138
 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 
haftmann 
parents: 
52137 
diff
changeset
 | 
594  | 
let  | 
| 59103 | 595  | 
val (_, (_, data)) = join_ancestry thy target_name;  | 
596  | 
val _ = if member (op =) (Code_Printer.the_reserved data) sym  | 
|
| 
52138
 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 
haftmann 
parents: 
52137 
diff
changeset
 | 
597  | 
      then error ("Reserved symbol " ^ quote sym ^ " already declared")
 | 
| 
 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 
haftmann 
parents: 
52137 
diff
changeset
 | 
598  | 
else ();  | 
| 
 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 
haftmann 
parents: 
52137 
diff
changeset
 | 
599  | 
in  | 
| 
 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 
haftmann 
parents: 
52137 
diff
changeset
 | 
600  | 
thy  | 
| 59099 | 601  | 
|> map_reserved target_name (insert (op =) sym)  | 
| 
52138
 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 
haftmann 
parents: 
52137 
diff
changeset
 | 
602  | 
end;  | 
| 
 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 
haftmann 
parents: 
52137 
diff
changeset
 | 
603  | 
|
| 
 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 
haftmann 
parents: 
52137 
diff
changeset
 | 
604  | 
|
| 
52377
 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 
haftmann 
parents: 
52218 
diff
changeset
 | 
605  | 
(* checking of syntax *)  | 
| 
52138
 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 
haftmann 
parents: 
52137 
diff
changeset
 | 
606  | 
|
| 59099 | 607  | 
fun check_const_syntax ctxt target_name c syn =  | 
| 55757 | 608  | 
if Code_Printer.requires_args syn > Code.args_number (Proof_Context.theory_of ctxt) c  | 
| 
52137
 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 
haftmann 
parents: 
52068 
diff
changeset
 | 
609  | 
  then error ("Too many arguments in syntax for constant " ^ quote c)
 | 
| 59099 | 610  | 
else Code_Printer.prep_const_syntax (Proof_Context.theory_of ctxt) (the_literals ctxt target_name) c syn;  | 
| 
52137
 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 
haftmann 
parents: 
52068 
diff
changeset
 | 
611  | 
|
| 59099 | 612  | 
fun check_tyco_syntax ctxt target_name tyco syn =  | 
| 55757 | 613  | 
if fst syn <> Sign.arity_number (Proof_Context.theory_of ctxt) tyco  | 
| 
52137
 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 
haftmann 
parents: 
52068 
diff
changeset
 | 
614  | 
  then error ("Number of arguments mismatch in syntax for type constructor " ^ quote tyco)
 | 
| 
 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 
haftmann 
parents: 
52068 
diff
changeset
 | 
615  | 
else syn;  | 
| 34071 | 616  | 
|
| 
52138
 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 
haftmann 
parents: 
52137 
diff
changeset
 | 
617  | 
|
| 
 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 
haftmann 
parents: 
52137 
diff
changeset
 | 
618  | 
(* custom symbol names *)  | 
| 
 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 
haftmann 
parents: 
52137 
diff
changeset
 | 
619  | 
|
| 52218 | 620  | 
fun arrange_name_decls x =  | 
| 
52138
 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 
haftmann 
parents: 
52137 
diff
changeset
 | 
621  | 
let  | 
| 
 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 
haftmann 
parents: 
52137 
diff
changeset
 | 
622  | 
fun arrange is_module (sym, target_names) = map (fn (target, some_name) =>  | 
| 
 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 
haftmann 
parents: 
52137 
diff
changeset
 | 
623  | 
(target, (sym, Option.map (check_name is_module) some_name))) target_names;  | 
| 
 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 
haftmann 
parents: 
52137 
diff
changeset
 | 
624  | 
in  | 
| 
 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 
haftmann 
parents: 
52137 
diff
changeset
 | 
625  | 
Code_Symbol.maps_attr' (arrange false) (arrange false) (arrange false)  | 
| 52218 | 626  | 
(arrange false) (arrange false) (arrange true) x  | 
| 
52138
 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 
haftmann 
parents: 
52137 
diff
changeset
 | 
627  | 
end;  | 
| 
 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 
haftmann 
parents: 
52137 
diff
changeset
 | 
628  | 
|
| 
69485
 
b734ff28e405
proper attach mechanism for any kind of symbols, not just constants
 
haftmann 
parents: 
69484 
diff
changeset
 | 
629  | 
fun cert_name_decls ctxt = cert_syms' ctxt #> arrange_name_decls;  | 
| 
52138
 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 
haftmann 
parents: 
52137 
diff
changeset
 | 
630  | 
|
| 
69485
 
b734ff28e405
proper attach mechanism for any kind of symbols, not just constants
 
haftmann 
parents: 
69484 
diff
changeset
 | 
631  | 
fun read_name_decls ctxt = read_syms' ctxt #> arrange_name_decls;  | 
| 
52138
 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 
haftmann 
parents: 
52137 
diff
changeset
 | 
632  | 
|
| 59099 | 633  | 
fun set_identifier (target_name, sym_name) = map_identifiers target_name (Code_Symbol.set_data sym_name);  | 
| 
52138
 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 
haftmann 
parents: 
52137 
diff
changeset
 | 
634  | 
|
| 
 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 
haftmann 
parents: 
52137 
diff
changeset
 | 
635  | 
fun gen_set_identifiers prep_name_decl raw_name_decls thy =  | 
| 55757 | 636  | 
fold set_identifier (prep_name_decl (Proof_Context.init_global thy) raw_name_decls) thy;  | 
| 
52138
 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 
haftmann 
parents: 
52137 
diff
changeset
 | 
637  | 
|
| 
 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 
haftmann 
parents: 
52137 
diff
changeset
 | 
638  | 
val set_identifiers = gen_set_identifiers cert_name_decls;  | 
| 
 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 
haftmann 
parents: 
52137 
diff
changeset
 | 
639  | 
val set_identifiers_cmd = gen_set_identifiers read_name_decls;  | 
| 
 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 
haftmann 
parents: 
52137 
diff
changeset
 | 
640  | 
|
| 
 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 
haftmann 
parents: 
52137 
diff
changeset
 | 
641  | 
|
| 
 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 
haftmann 
parents: 
52137 
diff
changeset
 | 
642  | 
(* custom printings *)  | 
| 
 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 
haftmann 
parents: 
52137 
diff
changeset
 | 
643  | 
|
| 
69485
 
b734ff28e405
proper attach mechanism for any kind of symbols, not just constants
 
haftmann 
parents: 
69484 
diff
changeset
 | 
644  | 
fun arrange_printings prep_syms ctxt =  | 
| 
52137
 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 
haftmann 
parents: 
52068 
diff
changeset
 | 
645  | 
let  | 
| 
81705
 
53fea2ccab19
explicit error message for non-existing code target
 
haftmann 
parents: 
80328 
diff
changeset
 | 
646  | 
val thy = Proof_Context.theory_of ctxt;  | 
| 
52137
 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 
haftmann 
parents: 
52068 
diff
changeset
 | 
647  | 
fun arrange check (sym, target_syns) =  | 
| 59099 | 648  | 
map (fn (target_name, some_syn) =>  | 
| 
81705
 
53fea2ccab19
explicit error message for non-existing code target
 
haftmann 
parents: 
80328 
diff
changeset
 | 
649  | 
(assert_target thy target_name, (sym, Option.map (check ctxt target_name sym) some_syn))) target_syns;  | 
| 
52137
 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 
haftmann 
parents: 
52068 
diff
changeset
 | 
650  | 
in  | 
| 
 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 
haftmann 
parents: 
52068 
diff
changeset
 | 
651  | 
Code_Symbol.maps_attr'  | 
| 
 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 
haftmann 
parents: 
52068 
diff
changeset
 | 
652  | 
(arrange check_const_syntax) (arrange check_tyco_syntax)  | 
| 
55149
 
626d8f08d479
immediate "activation" of const syntax at declaration time
 
haftmann 
parents: 
55147 
diff
changeset
 | 
653  | 
(arrange ((K o K o K) I)) (arrange ((K o K o K) I)) (arrange ((K o K o K) I))  | 
| 
69485
 
b734ff28e405
proper attach mechanism for any kind of symbols, not just constants
 
haftmann 
parents: 
69484 
diff
changeset
 | 
654  | 
(arrange (fn ctxt => fn _ => fn _ => fn (raw_content, raw_syms) =>  | 
| 80328 | 655  | 
(Pretty.block0 (Pretty.fbreaks (map Code_Printer.str (split_lines raw_content))),  | 
| 
69485
 
b734ff28e405
proper attach mechanism for any kind of symbols, not just constants
 
haftmann 
parents: 
69484 
diff
changeset
 | 
656  | 
map (prep_syms ctxt) raw_syms)))  | 
| 
52137
 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 
haftmann 
parents: 
52068 
diff
changeset
 | 
657  | 
end;  | 
| 
 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 
haftmann 
parents: 
52068 
diff
changeset
 | 
658  | 
|
| 
69485
 
b734ff28e405
proper attach mechanism for any kind of symbols, not just constants
 
haftmann 
parents: 
69484 
diff
changeset
 | 
659  | 
fun cert_printings ctxt = cert_syms' ctxt #> arrange_printings cert_syms ctxt;  | 
| 24219 | 660  | 
|
| 
69485
 
b734ff28e405
proper attach mechanism for any kind of symbols, not just constants
 
haftmann 
parents: 
69484 
diff
changeset
 | 
661  | 
fun read_printings ctxt = read_syms' ctxt #> arrange_printings read_syms ctxt;  | 
| 
52137
 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 
haftmann 
parents: 
52068 
diff
changeset
 | 
662  | 
|
| 59099 | 663  | 
fun set_printing (target_name, sym_syn) = map_printings target_name (Code_Symbol.set_data sym_syn);  | 
| 
52137
 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 
haftmann 
parents: 
52068 
diff
changeset
 | 
664  | 
|
| 
 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 
haftmann 
parents: 
52068 
diff
changeset
 | 
665  | 
fun gen_set_printings prep_print_decl raw_print_decls thy =  | 
| 55757 | 666  | 
fold set_printing (prep_print_decl (Proof_Context.init_global thy) raw_print_decls) thy;  | 
| 
52137
 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 
haftmann 
parents: 
52068 
diff
changeset
 | 
667  | 
|
| 
 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 
haftmann 
parents: 
52068 
diff
changeset
 | 
668  | 
val set_printings = gen_set_printings cert_printings;  | 
| 
 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 
haftmann 
parents: 
52068 
diff
changeset
 | 
669  | 
val set_printings_cmd = gen_set_printings read_printings;  | 
| 
 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 
haftmann 
parents: 
52068 
diff
changeset
 | 
670  | 
|
| 
 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 
haftmann 
parents: 
52068 
diff
changeset
 | 
671  | 
|
| 27103 | 672  | 
(* concrete syntax *)  | 
| 27000 | 673  | 
|
| 28054 | 674  | 
fun parse_args f args =  | 
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36537 
diff
changeset
 | 
675  | 
case Scan.read Token.stopper f args  | 
| 28054 | 676  | 
of SOME x => x  | 
677  | 
| NONE => error "Bad serializer arguments";  | 
|
678  | 
||
679  | 
||
| 27304 | 680  | 
(** Isar setup **)  | 
| 27103 | 681  | 
|
| 
69698
 
1a249d1c884b
more conventional parsing of code_stmts antiquotation
 
haftmann 
parents: 
69697 
diff
changeset
 | 
682  | 
val (constantK, type_constructorK, type_classK, class_relationK, class_instanceK, code_moduleK) =  | 
| 
 
1a249d1c884b
more conventional parsing of code_stmts antiquotation
 
haftmann 
parents: 
69697 
diff
changeset
 | 
683  | 
(\<^keyword>\<open>constant\<close>, \<^keyword>\<open>type_constructor\<close>, \<^keyword>\<open>type_class\<close>,  | 
| 
 
1a249d1c884b
more conventional parsing of code_stmts antiquotation
 
haftmann 
parents: 
69697 
diff
changeset
 | 
684  | 
\<^keyword>\<open>class_relation\<close>, \<^keyword>\<open>class_instance\<close>, \<^keyword>\<open>code_module\<close>);  | 
| 
 
1a249d1c884b
more conventional parsing of code_stmts antiquotation
 
haftmann 
parents: 
69697 
diff
changeset
 | 
685  | 
|
| 
69485
 
b734ff28e405
proper attach mechanism for any kind of symbols, not just constants
 
haftmann 
parents: 
69484 
diff
changeset
 | 
686  | 
local  | 
| 
 
b734ff28e405
proper attach mechanism for any kind of symbols, not just constants
 
haftmann 
parents: 
69484 
diff
changeset
 | 
687  | 
|
| 
 
b734ff28e405
proper attach mechanism for any kind of symbols, not just constants
 
haftmann 
parents: 
69484 
diff
changeset
 | 
688  | 
val parse_constants = constantK |-- Scan.repeat1 Parse.term >> map Constant;  | 
| 
 
b734ff28e405
proper attach mechanism for any kind of symbols, not just constants
 
haftmann 
parents: 
69484 
diff
changeset
 | 
689  | 
|
| 
 
b734ff28e405
proper attach mechanism for any kind of symbols, not just constants
 
haftmann 
parents: 
69484 
diff
changeset
 | 
690  | 
val parse_type_constructors = type_constructorK |-- Scan.repeat1 Parse.type_const >> map Type_Constructor;  | 
| 
 
b734ff28e405
proper attach mechanism for any kind of symbols, not just constants
 
haftmann 
parents: 
69484 
diff
changeset
 | 
691  | 
|
| 
 
b734ff28e405
proper attach mechanism for any kind of symbols, not just constants
 
haftmann 
parents: 
69484 
diff
changeset
 | 
692  | 
val parse_classes = type_classK |-- Scan.repeat1 Parse.class >> map Type_Class;  | 
| 
 
b734ff28e405
proper attach mechanism for any kind of symbols, not just constants
 
haftmann 
parents: 
69484 
diff
changeset
 | 
693  | 
|
| 
 
b734ff28e405
proper attach mechanism for any kind of symbols, not just constants
 
haftmann 
parents: 
69484 
diff
changeset
 | 
694  | 
val parse_class_relations = class_relationK |-- Scan.repeat1 parse_classrel_ident >> map Class_Relation;  | 
| 
 
b734ff28e405
proper attach mechanism for any kind of symbols, not just constants
 
haftmann 
parents: 
69484 
diff
changeset
 | 
695  | 
|
| 
 
b734ff28e405
proper attach mechanism for any kind of symbols, not just constants
 
haftmann 
parents: 
69484 
diff
changeset
 | 
696  | 
val parse_instances = class_instanceK |-- Scan.repeat1 parse_inst_ident >> map Class_Instance;  | 
| 
 
b734ff28e405
proper attach mechanism for any kind of symbols, not just constants
 
haftmann 
parents: 
69484 
diff
changeset
 | 
697  | 
|
| 
69698
 
1a249d1c884b
more conventional parsing of code_stmts antiquotation
 
haftmann 
parents: 
69697 
diff
changeset
 | 
698  | 
val parse_simple_symbols = Scan.repeats1 (parse_constants || parse_type_constructors || parse_classes  | 
| 
 
1a249d1c884b
more conventional parsing of code_stmts antiquotation
 
haftmann 
parents: 
69697 
diff
changeset
 | 
699  | 
|| parse_class_relations || parse_instances);  | 
| 
69485
 
b734ff28e405
proper attach mechanism for any kind of symbols, not just constants
 
haftmann 
parents: 
69484 
diff
changeset
 | 
700  | 
|
| 
52137
 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 
haftmann 
parents: 
52068 
diff
changeset
 | 
701  | 
fun parse_single_symbol_pragma parse_keyword parse_isa parse_target =  | 
| 69593 | 702  | 
parse_keyword |-- Parse.!!! (parse_isa --| (\<^keyword>\<open>\<rightharpoonup>\<close> || \<^keyword>\<open>=>\<close>)  | 
703  | 
-- Parse.and_list1 (\<^keyword>\<open>(\<close> |-- (Parse.name --| \<^keyword>\<open>)\<close> -- Scan.option parse_target)));  | 
|
| 
52137
 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 
haftmann 
parents: 
52068 
diff
changeset
 | 
704  | 
|
| 
 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 
haftmann 
parents: 
52068 
diff
changeset
 | 
705  | 
fun parse_symbol_pragma parse_const parse_tyco parse_class parse_classrel parse_inst parse_module =  | 
| 
69485
 
b734ff28e405
proper attach mechanism for any kind of symbols, not just constants
 
haftmann 
parents: 
69484 
diff
changeset
 | 
706  | 
parse_single_symbol_pragma constantK Parse.term parse_const  | 
| 55150 | 707  | 
>> Constant  | 
| 
69485
 
b734ff28e405
proper attach mechanism for any kind of symbols, not just constants
 
haftmann 
parents: 
69484 
diff
changeset
 | 
708  | 
|| parse_single_symbol_pragma type_constructorK Parse.type_const parse_tyco  | 
| 55150 | 709  | 
>> Type_Constructor  | 
| 
69485
 
b734ff28e405
proper attach mechanism for any kind of symbols, not just constants
 
haftmann 
parents: 
69484 
diff
changeset
 | 
710  | 
|| parse_single_symbol_pragma type_classK Parse.class parse_class  | 
| 55150 | 711  | 
>> Type_Class  | 
| 
69485
 
b734ff28e405
proper attach mechanism for any kind of symbols, not just constants
 
haftmann 
parents: 
69484 
diff
changeset
 | 
712  | 
|| parse_single_symbol_pragma class_relationK parse_classrel_ident parse_classrel  | 
| 55150 | 713  | 
>> Class_Relation  | 
| 
69485
 
b734ff28e405
proper attach mechanism for any kind of symbols, not just constants
 
haftmann 
parents: 
69484 
diff
changeset
 | 
714  | 
|| parse_single_symbol_pragma class_instanceK parse_inst_ident parse_inst  | 
| 55150 | 715  | 
>> Class_Instance  | 
| 
69485
 
b734ff28e405
proper attach mechanism for any kind of symbols, not just constants
 
haftmann 
parents: 
69484 
diff
changeset
 | 
716  | 
|| parse_single_symbol_pragma code_moduleK Parse.name parse_module  | 
| 
69698
 
1a249d1c884b
more conventional parsing of code_stmts antiquotation
 
haftmann 
parents: 
69697 
diff
changeset
 | 
717  | 
>> Module;  | 
| 
52137
 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 
haftmann 
parents: 
52068 
diff
changeset
 | 
718  | 
|
| 
 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 
haftmann 
parents: 
52068 
diff
changeset
 | 
719  | 
fun parse_symbol_pragmas parse_const parse_tyco parse_class parse_classrel parse_inst parse_module =  | 
| 
 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 
haftmann 
parents: 
52068 
diff
changeset
 | 
720  | 
Parse.enum1 "|" (Parse.group (fn () => "code symbol pragma")  | 
| 
 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 
haftmann 
parents: 
52068 
diff
changeset
 | 
721  | 
(parse_symbol_pragma parse_const parse_tyco parse_class parse_classrel parse_inst parse_module));  | 
| 
 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 
haftmann 
parents: 
52068 
diff
changeset
 | 
722  | 
|
| 69593 | 723  | 
val code_expr_argsP = Scan.optional (\<^keyword>\<open>(\<close> |-- Parse.args --| \<^keyword>\<open>)\<close>) [];  | 
| 
52138
 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 
haftmann 
parents: 
52137 
diff
changeset
 | 
724  | 
|
| 69645 | 725  | 
fun code_expr_inP (all_public, raw_cs) =  | 
| 69593 | 726  | 
Scan.repeat (\<^keyword>\<open>in\<close> |-- Parse.!!! (Parse.name  | 
727  | 
-- Scan.optional (\<^keyword>\<open>module_name\<close> |-- Parse.name) ""  | 
|
| 
70009
 
435fb018e8ee
"export_code ... file_prefix ..." is the preferred way to produce output within the logical file-system within the theory context, as well as session exports;
 
wenzelm 
parents: 
70002 
diff
changeset
 | 
728  | 
-- Scan.option  | 
| 
 
435fb018e8ee
"export_code ... file_prefix ..." is the preferred way to produce output within the logical file-system within the theory context, as well as session exports;
 
wenzelm 
parents: 
70002 
diff
changeset
 | 
729  | 
        ((\<^keyword>\<open>file_prefix\<close> >> K {physical = false} || \<^keyword>\<open>file\<close> >> K {physical = true})
 | 
| 72841 | 730  | 
-- Parse.!!! Parse.path_input)  | 
| 52434 | 731  | 
-- code_expr_argsP))  | 
| 55683 | 732  | 
>> (fn seri_args => export_code_cmd all_public raw_cs seri_args);  | 
| 52434 | 733  | 
|
| 69645 | 734  | 
fun code_expr_checkingP (all_public, raw_cs) =  | 
| 69593 | 735  | 
(\<^keyword>\<open>checking\<close> |-- Parse.!!!  | 
| 69645 | 736  | 
(Scan.repeat (Parse.name -- (Scan.optional (\<^keyword>\<open>?\<close> >> K false) true) -- code_expr_argsP)))  | 
| 55683 | 737  | 
>> (fn seri_args => check_code_cmd all_public raw_cs seri_args);  | 
| 52434 | 738  | 
|
| 
69698
 
1a249d1c884b
more conventional parsing of code_stmts antiquotation
 
haftmann 
parents: 
69697 
diff
changeset
 | 
739  | 
in  | 
| 
 
1a249d1c884b
more conventional parsing of code_stmts antiquotation
 
haftmann 
parents: 
69697 
diff
changeset
 | 
740  | 
|
| 
52138
 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 
haftmann 
parents: 
52137 
diff
changeset
 | 
741  | 
val _ =  | 
| 81706 | 742  | 
Outer_Syntax.command \<^command_keyword>\<open>code_reserved\<close> "declare words as reserved for target language"  | 
743  | 
(Parse.and_list1 (\<^keyword>\<open>(\<close> |-- (Parse.name --| \<^keyword>\<open>)\<close> -- Scan.repeat1 Parse.name))  | 
|
744  | 
>> (Toplevel.theory o fold (fn (target, reserveds) => fold (add_reserved target) reserveds)));  | 
|
| 
52138
 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 
haftmann 
parents: 
52137 
diff
changeset
 | 
745  | 
|
| 
 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 
haftmann 
parents: 
52137 
diff
changeset
 | 
746  | 
val _ =  | 
| 69593 | 747  | 
Outer_Syntax.command \<^command_keyword>\<open>code_identifier\<close> "declare mandatory names for code symbols"  | 
| 
52138
 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 
haftmann 
parents: 
52137 
diff
changeset
 | 
748  | 
(parse_symbol_pragmas Parse.name Parse.name Parse.name Parse.name Parse.name Parse.name  | 
| 
 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 
haftmann 
parents: 
52137 
diff
changeset
 | 
749  | 
>> (Toplevel.theory o fold set_identifiers_cmd));  | 
| 
 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 
haftmann 
parents: 
52137 
diff
changeset
 | 
750  | 
|
| 
 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 
haftmann 
parents: 
52137 
diff
changeset
 | 
751  | 
val _ =  | 
| 69593 | 752  | 
Outer_Syntax.command \<^command_keyword>\<open>code_printing\<close> "declare dedicated printing for code symbols"  | 
| 
52137
 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 
haftmann 
parents: 
52068 
diff
changeset
 | 
753  | 
(parse_symbol_pragmas (Code_Printer.parse_const_syntax) (Code_Printer.parse_tyco_syntax)  | 
| 
 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 
haftmann 
parents: 
52068 
diff
changeset
 | 
754  | 
Parse.string (Parse.minus >> K ()) (Parse.minus >> K ())  | 
| 74887 | 755  | 
(Parse.embedded -- Scan.optional (\<^keyword>\<open>for\<close> |-- parse_simple_symbols) [])  | 
| 
52137
 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 
haftmann 
parents: 
52068 
diff
changeset
 | 
756  | 
>> (Toplevel.theory o fold set_printings_cmd));  | 
| 
 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 
haftmann 
parents: 
52068 
diff
changeset
 | 
757  | 
|
| 
 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 
haftmann 
parents: 
52068 
diff
changeset
 | 
758  | 
val _ =  | 
| 70002 | 759  | 
Outer_Syntax.local_theory \<^command_keyword>\<open>export_code\<close> "generate executable code for constants"  | 
| 69646 | 760  | 
(Scan.optional (\<^keyword>\<open>open\<close> >> K true) false -- Scan.repeat1 Parse.term  | 
| 70002 | 761  | 
:|-- (fn args => (code_expr_checkingP args || code_expr_inP args)));  | 
| 
30494
 
c150e6fa4e0d
consider exit status of code generation direcitve
 
haftmann 
parents: 
30242 
diff
changeset
 | 
762  | 
|
| 
69698
 
1a249d1c884b
more conventional parsing of code_stmts antiquotation
 
haftmann 
parents: 
69697 
diff
changeset
 | 
763  | 
end;  | 
| 
 
1a249d1c884b
more conventional parsing of code_stmts antiquotation
 
haftmann 
parents: 
69697 
diff
changeset
 | 
764  | 
|
| 
 
1a249d1c884b
more conventional parsing of code_stmts antiquotation
 
haftmann 
parents: 
69697 
diff
changeset
 | 
765  | 
local  | 
| 
 
1a249d1c884b
more conventional parsing of code_stmts antiquotation
 
haftmann 
parents: 
69697 
diff
changeset
 | 
766  | 
|
| 
 
1a249d1c884b
more conventional parsing of code_stmts antiquotation
 
haftmann 
parents: 
69697 
diff
changeset
 | 
767  | 
val parse_const_terms = Args.theory -- Scan.repeat1 Args.term  | 
| 
 
1a249d1c884b
more conventional parsing of code_stmts antiquotation
 
haftmann 
parents: 
69697 
diff
changeset
 | 
768  | 
>> uncurry (fn thy => map (Code.check_const thy));  | 
| 
 
1a249d1c884b
more conventional parsing of code_stmts antiquotation
 
haftmann 
parents: 
69697 
diff
changeset
 | 
769  | 
|
| 
 
1a249d1c884b
more conventional parsing of code_stmts antiquotation
 
haftmann 
parents: 
69697 
diff
changeset
 | 
770  | 
fun parse_symbols keyword parse internalize mark_symbol =  | 
| 
 
1a249d1c884b
more conventional parsing of code_stmts antiquotation
 
haftmann 
parents: 
69697 
diff
changeset
 | 
771  | 
Scan.lift (keyword --| Args.colon) |-- Args.theory -- Scan.repeat1 parse  | 
| 
 
1a249d1c884b
more conventional parsing of code_stmts antiquotation
 
haftmann 
parents: 
69697 
diff
changeset
 | 
772  | 
>> uncurry (fn thy => map (mark_symbol o internalize thy));  | 
| 
 
1a249d1c884b
more conventional parsing of code_stmts antiquotation
 
haftmann 
parents: 
69697 
diff
changeset
 | 
773  | 
|
| 
 
1a249d1c884b
more conventional parsing of code_stmts antiquotation
 
haftmann 
parents: 
69697 
diff
changeset
 | 
774  | 
val parse_consts = parse_symbols constantK Args.term  | 
| 
 
1a249d1c884b
more conventional parsing of code_stmts antiquotation
 
haftmann 
parents: 
69697 
diff
changeset
 | 
775  | 
Code.check_const Constant;  | 
| 
 
1a249d1c884b
more conventional parsing of code_stmts antiquotation
 
haftmann 
parents: 
69697 
diff
changeset
 | 
776  | 
|
| 
 
1a249d1c884b
more conventional parsing of code_stmts antiquotation
 
haftmann 
parents: 
69697 
diff
changeset
 | 
777  | 
val parse_types = parse_symbols type_constructorK (Scan.lift Args.name)  | 
| 
 
1a249d1c884b
more conventional parsing of code_stmts antiquotation
 
haftmann 
parents: 
69697 
diff
changeset
 | 
778  | 
Sign.intern_type Type_Constructor;  | 
| 
 
1a249d1c884b
more conventional parsing of code_stmts antiquotation
 
haftmann 
parents: 
69697 
diff
changeset
 | 
779  | 
|
| 
 
1a249d1c884b
more conventional parsing of code_stmts antiquotation
 
haftmann 
parents: 
69697 
diff
changeset
 | 
780  | 
val parse_classes = parse_symbols type_classK (Scan.lift Args.name)  | 
| 
 
1a249d1c884b
more conventional parsing of code_stmts antiquotation
 
haftmann 
parents: 
69697 
diff
changeset
 | 
781  | 
Sign.intern_class Type_Class;  | 
| 
 
1a249d1c884b
more conventional parsing of code_stmts antiquotation
 
haftmann 
parents: 
69697 
diff
changeset
 | 
782  | 
|
| 
 
1a249d1c884b
more conventional parsing of code_stmts antiquotation
 
haftmann 
parents: 
69697 
diff
changeset
 | 
783  | 
val parse_instances = parse_symbols class_instanceK (Scan.lift (Args.name --| Args.$$$ "::" -- Args.name))  | 
| 
 
1a249d1c884b
more conventional parsing of code_stmts antiquotation
 
haftmann 
parents: 
69697 
diff
changeset
 | 
784  | 
(fn thy => fn (raw_tyco, raw_class) =>  | 
| 
 
1a249d1c884b
more conventional parsing of code_stmts antiquotation
 
haftmann 
parents: 
69697 
diff
changeset
 | 
785  | 
(Sign.intern_class thy raw_tyco, Sign.intern_type thy raw_class)) Class_Instance;  | 
| 
 
1a249d1c884b
more conventional parsing of code_stmts antiquotation
 
haftmann 
parents: 
69697 
diff
changeset
 | 
786  | 
|
| 
 
1a249d1c884b
more conventional parsing of code_stmts antiquotation
 
haftmann 
parents: 
69697 
diff
changeset
 | 
787  | 
in  | 
| 
 
1a249d1c884b
more conventional parsing of code_stmts antiquotation
 
haftmann 
parents: 
69697 
diff
changeset
 | 
788  | 
|
| 
 
1a249d1c884b
more conventional parsing of code_stmts antiquotation
 
haftmann 
parents: 
69697 
diff
changeset
 | 
789  | 
val _ = Theory.setup  | 
| 73761 | 790  | 
(Document_Output.antiquotation_raw \<^binding>\<open>code_stmts\<close>  | 
| 
69698
 
1a249d1c884b
more conventional parsing of code_stmts antiquotation
 
haftmann 
parents: 
69697 
diff
changeset
 | 
791  | 
(parse_const_terms --  | 
| 
 
1a249d1c884b
more conventional parsing of code_stmts antiquotation
 
haftmann 
parents: 
69697 
diff
changeset
 | 
792  | 
Scan.repeats (parse_consts || parse_types || parse_classes || parse_instances)  | 
| 
 
1a249d1c884b
more conventional parsing of code_stmts antiquotation
 
haftmann 
parents: 
69697 
diff
changeset
 | 
793  | 
-- Scan.lift (Args.parens (Args.name -- Scan.option Parse.int)))  | 
| 
 
1a249d1c884b
more conventional parsing of code_stmts antiquotation
 
haftmann 
parents: 
69697 
diff
changeset
 | 
794  | 
(fn ctxt => fn ((consts, symbols), (target_name, some_width)) =>  | 
| 
 
1a249d1c884b
more conventional parsing of code_stmts antiquotation
 
haftmann 
parents: 
69697 
diff
changeset
 | 
795  | 
present_code ctxt consts symbols  | 
| 
 
1a249d1c884b
more conventional parsing of code_stmts antiquotation
 
haftmann 
parents: 
69697 
diff
changeset
 | 
796  | 
target_name "Example" some_width []  | 
| 
75604
 
39df30349778
more scalable generated files and code export, using Bytes.T;
 
wenzelm 
parents: 
74887 
diff
changeset
 | 
797  | 
|> Bytes.content  | 
| 
69698
 
1a249d1c884b
more conventional parsing of code_stmts antiquotation
 
haftmann 
parents: 
69697 
diff
changeset
 | 
798  | 
|> trim_line  | 
| 73761 | 799  | 
|> Document_Output.verbatim (Config.put Document_Antiquotation.thy_output_display true ctxt)));  | 
| 
69698
 
1a249d1c884b
more conventional parsing of code_stmts antiquotation
 
haftmann 
parents: 
69697 
diff
changeset
 | 
800  | 
|
| 
 
1a249d1c884b
more conventional parsing of code_stmts antiquotation
 
haftmann 
parents: 
69697 
diff
changeset
 | 
801  | 
end;  | 
| 
 
1a249d1c884b
more conventional parsing of code_stmts antiquotation
 
haftmann 
parents: 
69697 
diff
changeset
 | 
802  | 
|
| 24219 | 803  | 
end; (*struct*)  |