| author | wenzelm | 
| Wed, 17 Jun 2015 22:30:22 +0200 | |
| changeset 60511 | 5e67a223a141 | 
| parent 59936 | b8ffc3dc9e24 | 
| child 62549 | 9498623b27f0 | 
| permissions | -rw-r--r-- | 
| 37745 
6315b6426200
checking generated code for various target languages
 haftmann parents: 
37528diff
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 | |
| 55757 | 12 | val export_code_for: Proof.context -> Path.T option -> string -> int option -> string -> Token.T list | 
| 55683 | 13 | -> Code_Thingol.program -> bool -> Code_Symbol.T list -> unit | 
| 55757 | 14 | val produce_code_for: Proof.context -> string -> int option -> string -> Token.T list | 
| 55683 | 15 | -> Code_Thingol.program -> bool -> Code_Symbol.T list -> (string * string) list * string option list | 
| 55757 | 16 | val present_code_for: Proof.context -> string -> int option -> string -> Token.T list | 
| 55150 | 17 | -> Code_Thingol.program -> Code_Symbol.T list * Code_Symbol.T list -> string | 
| 55757 | 18 | val check_code_for: Proof.context -> string -> bool -> Token.T list | 
| 55683 | 19 | -> Code_Thingol.program -> bool -> Code_Symbol.T list -> unit | 
| 38918 | 20 | |
| 55757 | 21 | val export_code: Proof.context -> bool -> string list | 
| 38933 
bd77e092f67c
avoid strange special treatment of empty module names
 haftmann parents: 
38929diff
changeset | 22 | -> (((string * string) * Path.T option) * Token.T list) list -> unit | 
| 55757 | 23 | val produce_code: Proof.context -> bool -> string list | 
| 48568 | 24 | -> string -> int option -> string -> Token.T list -> (string * string) list * string option list | 
| 55757 | 25 | val present_code: Proof.context -> string list -> Code_Symbol.T list | 
| 38933 
bd77e092f67c
avoid strange special treatment of empty module names
 haftmann parents: 
38929diff
changeset | 26 | -> string -> int option -> string -> Token.T list -> string | 
| 55757 | 27 | val check_code: Proof.context -> bool -> string list | 
| 38918 | 28 | -> ((string * bool) * Token.T list) list -> unit | 
| 29 | ||
| 48568 | 30 | val generatedN: string | 
| 55757 | 31 | val evaluator: Proof.context -> string -> Code_Thingol.program | 
| 55684 
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
 haftmann parents: 
55683diff
changeset | 32 | -> Code_Symbol.T list -> bool -> ((string * class list) list * Code_Thingol.itype) * Code_Thingol.iterm | 
| 48568 | 33 | -> (string * string) list * string | 
| 41344 
d990badc97a3
evaluator separating static and dynamic operations
 haftmann parents: 
41342diff
changeset | 34 | |
| 28054 | 35 | type serializer | 
| 34152 | 36 | type literals = Code_Printer.literals | 
| 59104 | 37 | type language | 
| 38 | type ancestry | |
| 59479 | 39 | val assert_target: theory -> string -> string | 
| 59104 | 40 | val add_language: string * language -> theory -> theory | 
| 41 | val add_derived_target: string * ancestry -> theory -> theory | |
| 55757 | 42 | val the_literals: Proof.context -> string -> literals | 
| 28054 | 43 | type serialization | 
| 36959 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 wenzelm parents: 
36537diff
changeset | 44 | val parse_args: 'a parser -> Token.T list -> 'a | 
| 38916 | 45 | val serialization: (int -> Path.T option -> 'a -> unit) | 
| 55150 | 46 | -> (Code_Symbol.T list -> int -> 'a -> (string * string) list * (Code_Symbol.T -> string option)) | 
| 38925 | 47 | -> 'a -> serialization | 
| 59324 | 48 | val default_code_width: int Config.T | 
| 38917 | 49 | |
| 52137 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 haftmann parents: 
52068diff
changeset | 50 |   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: 
52137diff
changeset | 51 | 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: 
52137diff
changeset | 52 | -> theory -> theory | 
| 55372 | 53 | val set_printings: (Code_Printer.raw_const_syntax, Code_Printer.tyco_syntax, string, unit, unit, (string * string list)) symbol_attr_decl | 
| 52137 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 haftmann parents: 
52068diff
changeset | 54 | -> theory -> theory | 
| 28054 | 55 | val add_reserved: string -> string -> theory -> theory | 
| 24219 | 56 | end; | 
| 57 | ||
| 28054 | 58 | structure Code_Target : CODE_TARGET = | 
| 24219 | 59 | struct | 
| 60 | ||
| 55150 | 61 | open Basic_Code_Symbol; | 
| 28054 | 62 | open Basic_Code_Thingol; | 
| 34152 | 63 | |
| 64 | type literals = Code_Printer.literals; | |
| 52137 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 haftmann parents: 
52068diff
changeset | 65 | 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: 
52068diff
changeset | 66 | (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: 
52068diff
changeset | 67 | 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: 
52068diff
changeset | 68 | (class * string) * (string * 'e option) list, | 
| 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 haftmann parents: 
52068diff
changeset | 69 | 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: 
52068diff
changeset | 70 | |
| 34152 | 71 | type tyco_syntax = Code_Printer.tyco_syntax; | 
| 55153 
eedd549de3ef
more suitable names, without any notion of "activating"
 haftmann parents: 
55150diff
changeset | 72 | type raw_const_syntax = Code_Printer.raw_const_syntax; | 
| 34152 | 73 | |
| 24219 | 74 | |
| 52377 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 haftmann parents: 
52218diff
changeset | 75 | (** checking and parsing of symbols **) | 
| 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 haftmann parents: 
52218diff
changeset | 76 | |
| 55757 | 77 | fun cert_const ctxt const = | 
| 52377 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 haftmann parents: 
52218diff
changeset | 78 | let | 
| 55757 | 79 | 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: 
52218diff
changeset | 80 |       else error ("No such constant: " ^ quote const);
 | 
| 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 haftmann parents: 
52218diff
changeset | 81 | in const end; | 
| 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 haftmann parents: 
52218diff
changeset | 82 | |
| 55757 | 83 | fun read_const ctxt = Code.read_const (Proof_Context.theory_of ctxt); | 
| 84 | ||
| 85 | fun cert_tyco ctxt tyco = | |
| 52377 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 haftmann parents: 
52218diff
changeset | 86 | let | 
| 55757 | 87 | 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: 
52218diff
changeset | 88 |       else error ("No such type constructor: " ^ quote tyco);
 | 
| 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 haftmann parents: 
52218diff
changeset | 89 | in tyco end; | 
| 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 haftmann parents: 
52218diff
changeset | 90 | |
| 55951 
c07d184aebe9
tuned signature -- more uniform check_type_name/read_type_name;
 wenzelm parents: 
55776diff
changeset | 91 | fun read_tyco ctxt = | 
| 56002 | 92 |   #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: 
52218diff
changeset | 93 | |
| 55757 | 94 | fun cert_class ctxt class = | 
| 52377 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 haftmann parents: 
52218diff
changeset | 95 | let | 
| 55757 | 96 | val _ = Axclass.get_info (Proof_Context.theory_of ctxt) class; | 
| 52377 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 haftmann parents: 
52218diff
changeset | 97 | in class end; | 
| 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 haftmann parents: 
52218diff
changeset | 98 | |
| 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 haftmann parents: 
52218diff
changeset | 99 | val parse_classrel_ident = Parse.class --| @{keyword "<"} -- Parse.class;
 | 
| 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 haftmann parents: 
52218diff
changeset | 100 | |
| 55757 | 101 | fun cert_inst ctxt (class, tyco) = | 
| 102 | (cert_class ctxt class, cert_tyco ctxt tyco); | |
| 52377 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 haftmann parents: 
52218diff
changeset | 103 | |
| 55757 | 104 | fun read_inst ctxt (raw_tyco, raw_class) = | 
| 105 | (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: 
52218diff
changeset | 106 | |
| 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 haftmann parents: 
52218diff
changeset | 107 | val parse_inst_ident = Parse.xname --| @{keyword "::"} -- Parse.class;
 | 
| 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 haftmann parents: 
52218diff
changeset | 108 | |
| 55757 | 109 | fun cert_syms ctxt = | 
| 110 | 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: 
58928diff
changeset | 111 | (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: 
52218diff
changeset | 112 | |
| 55757 | 113 | fun read_syms ctxt = | 
| 114 | 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: 
58928diff
changeset | 115 | (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: 
52218diff
changeset | 116 | |
| 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 haftmann parents: 
52218diff
changeset | 117 | fun check_name is_module s = | 
| 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 haftmann parents: 
52218diff
changeset | 118 | let | 
| 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 haftmann parents: 
52218diff
changeset | 119 | val _ = if s = "" then error "Bad empty code name" else (); | 
| 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 haftmann parents: 
52218diff
changeset | 120 | val xs = Long_Name.explode s; | 
| 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 haftmann parents: 
52218diff
changeset | 121 | 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: 
56811diff
changeset | 122 | then map (Name.desymbolize NONE) xs | 
| 52377 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 haftmann parents: 
52218diff
changeset | 123 | else if length xs < 2 | 
| 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 haftmann parents: 
52218diff
changeset | 124 |         then error ("Bad code name without module component: " ^ quote s)
 | 
| 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 haftmann parents: 
52218diff
changeset | 125 | else | 
| 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 haftmann parents: 
52218diff
changeset | 126 | let | 
| 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 haftmann parents: 
52218diff
changeset | 127 | 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: 
56811diff
changeset | 128 | 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: 
56811diff
changeset | 129 | val y' = Name.desymbolize NONE y; | 
| 52377 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 haftmann parents: 
52218diff
changeset | 130 | in ys' @ [y'] end; | 
| 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 haftmann parents: 
52218diff
changeset | 131 | in if xs' = xs | 
| 55291 | 132 | then if is_module then (xs, "") else split_last xs | 
| 52377 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 haftmann parents: 
52218diff
changeset | 133 |     else error ("Invalid code name: " ^ quote s ^ "\n"
 | 
| 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 haftmann parents: 
52218diff
changeset | 134 | ^ "better try " ^ quote (Long_Name.implode xs')) | 
| 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 haftmann parents: 
52218diff
changeset | 135 | end; | 
| 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 haftmann parents: 
52218diff
changeset | 136 | |
| 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 haftmann parents: 
52218diff
changeset | 137 | |
| 52137 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 haftmann parents: 
52068diff
changeset | 138 | (** serializations and serializer **) | 
| 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 haftmann parents: 
52068diff
changeset | 139 | |
| 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 haftmann parents: 
52068diff
changeset | 140 | (* serialization: abstract nonsense to cover different destinies for generated code *) | 
| 24219 | 141 | |
| 55150 | 142 | datatype destination = Export of Path.T option | Produce | Present of Code_Symbol.T list; | 
| 143 | type serialization = int -> destination -> ((string * string) list * (Code_Symbol.T -> string option)) option; | |
| 27014 
a5f53d9d2b60
yet another attempt to circumvent printmode problems
 haftmann parents: 
27001diff
changeset | 144 | |
| 39659 | 145 | fun serialization output _ content width (Export some_path) = | 
| 39661 
6381d18507ef
reverted cs 07549694e2f1 -- use re-printing with current print mode instead after code assembly, avoid Latex.output_typewriter
 haftmann parents: 
39659diff
changeset | 146 | (output width some_path content; NONE) | 
| 
6381d18507ef
reverted cs 07549694e2f1 -- use re-printing with current print mode instead after code assembly, avoid Latex.output_typewriter
 haftmann parents: 
39659diff
changeset | 147 | | serialization _ string content width Produce = | 
| 
6381d18507ef
reverted cs 07549694e2f1 -- use re-printing with current print mode instead after code assembly, avoid Latex.output_typewriter
 haftmann parents: 
39659diff
changeset | 148 | string [] width content |> SOME | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55146diff
changeset | 149 | | serialization _ string content width (Present syms) = | 
| 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55146diff
changeset | 150 | string syms width content | 
| 59430 | 151 | |> (apfst o map o apsnd) Output.output | 
| 39661 
6381d18507ef
reverted cs 07549694e2f1 -- use re-printing with current print mode instead after code assembly, avoid Latex.output_typewriter
 haftmann parents: 
39659diff
changeset | 152 | |> SOME; | 
| 39659 | 153 | |
| 39661 
6381d18507ef
reverted cs 07549694e2f1 -- use re-printing with current print mode instead after code assembly, avoid Latex.output_typewriter
 haftmann parents: 
39659diff
changeset | 154 | fun export some_path f = (f (Export some_path); ()); | 
| 
6381d18507ef
reverted cs 07549694e2f1 -- use re-printing with current print mode instead after code assembly, avoid Latex.output_typewriter
 haftmann parents: 
39659diff
changeset | 155 | fun produce f = the (f Produce); | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55146diff
changeset | 156 | fun present syms f = space_implode "\n\n" (map snd (fst (the (f (Present syms))))); | 
| 38917 | 157 | |
| 24219 | 158 | |
| 52137 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 haftmann parents: 
52068diff
changeset | 159 | (* serializers: functions producing serializations *) | 
| 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: 
52137diff
changeset | 162 | -> Proof.context | 
| 39142 | 163 |   -> {
 | 
| 52138 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 164 | module_name: string, | 
| 38926 | 165 | reserved_syms: string list, | 
| 59103 | 166 | identifiers: Code_Printer.identifiers, | 
| 38926 | 167 | includes: (string * Pretty.T) list, | 
| 168 | class_syntax: string -> string option, | |
| 169 | tyco_syntax: string -> Code_Printer.tyco_syntax option, | |
| 55153 
eedd549de3ef
more suitable names, without any notion of "activating"
 haftmann parents: 
55150diff
changeset | 170 | const_syntax: string -> Code_Printer.const_syntax option } | 
| 55683 | 171 | -> Code_Symbol.T list | 
| 41342 
3519e0dd8f75
more explicit structure for serializer invocation
 haftmann parents: 
41307diff
changeset | 172 | -> Code_Thingol.program | 
| 28054 | 173 | -> serialization; | 
| 27000 | 174 | |
| 59102 | 175 | |
| 52137 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 haftmann parents: 
52068diff
changeset | 176 | (** theory data **) | 
| 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 haftmann parents: 
52068diff
changeset | 177 | |
| 59102 | 178 | type language = { serializer: serializer, literals: literals,
 | 
| 179 |   check: { env_var: string, make_destination: Path.T -> Path.T, make_command: string -> string } }; 
 | |
| 180 | ||
| 59103 | 181 | type ancestry = (string * (Code_Thingol.program -> Code_Thingol.program)) list; | 
| 27103 | 182 | |
| 59103 | 183 | val merge_ancestry : ancestry * ancestry -> ancestry = AList.join (op =) (K snd); | 
| 27103 | 184 | |
| 59103 | 185 | type target = { serial: serial, language: language, ancestry: ancestry };
 | 
| 27437 | 186 | |
| 34248 | 187 | structure Targets = Theory_Data | 
| 34071 | 188 | ( | 
| 59324 | 189 | type T = (target * Code_Printer.data) Symtab.table; | 
| 190 | val empty = Symtab.empty; | |
| 34071 | 191 | val extend = I; | 
| 59324 | 192 | fun merge (targets1, targets2) : T = | 
| 193 | Symtab.join (fn target_name => fn ((target1, data1), (target2, data2)) => | |
| 59103 | 194 | if #serial target1 = #serial target2 then | 
| 195 |       ({ serial = #serial target1, language = #language target1,
 | |
| 196 | ancestry = merge_ancestry (#ancestry target1, #ancestry target2)}, | |
| 197 | Code_Printer.merge_data (data1, data2)) | |
| 59102 | 198 |       else error ("Incompatible targets: " ^ quote target_name) 
 | 
| 59324 | 199 | ) (targets1, targets2) | 
| 34071 | 200 | ); | 
| 27436 | 201 | |
| 59324 | 202 | fun exists_target thy = Symtab.defined (Targets.get thy); | 
| 203 | fun lookup_target_data thy = Symtab.lookup (Targets.get thy); | |
| 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 | |
| 59104 | 209 | fun fold1 f xs = fold f (tl xs) (hd xs); | 
| 210 | ||
| 59101 
6dc48c98303c
tuned target inheritance bookkeeping: ancestry is always fully maintained at current entry using canonical merge;
 haftmann parents: 
59100diff
changeset | 211 | fun join_ancestry thy target_name = | 
| 
6dc48c98303c
tuned target inheritance bookkeeping: ancestry is always fully maintained at current entry using canonical merge;
 haftmann parents: 
59100diff
changeset | 212 | let | 
| 59479 | 213 | val _ = assert_target thy target_name; | 
| 59103 | 214 | val the_target_data = the o lookup_target_data thy; | 
| 215 | val (target, this_data) = the_target_data target_name; | |
| 216 | val ancestry = #ancestry target; | |
| 59102 | 217 | val modifies = rev (map snd ancestry); | 
| 218 | val modify = fold (curry (op o)) modifies I; | |
| 59103 | 219 | val datas = rev (map (snd o the_target_data o fst) ancestry) @ [this_data]; | 
| 59104 | 220 | val data = fold1 (fn data' => fn data => Code_Printer.merge_data (data, data')) datas; | 
| 59103 | 221 | in (modify, (target, data)) end; | 
| 59101 
6dc48c98303c
tuned target inheritance bookkeeping: ancestry is always fully maintained at current entry using canonical merge;
 haftmann parents: 
59100diff
changeset | 222 | |
| 59479 | 223 | fun allocate_target target_name target thy = | 
| 27304 | 224 | let | 
| 59101 
6dc48c98303c
tuned target inheritance bookkeeping: ancestry is always fully maintained at current entry using canonical merge;
 haftmann parents: 
59100diff
changeset | 225 | 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: 
59100diff
changeset | 226 |       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: 
43324diff
changeset | 227 | else (); | 
| 28054 | 228 | in | 
| 229 | thy | |
| 59324 | 230 | |> (Targets.map o Symtab.update) (target_name, (target, Code_Printer.empty_data)) | 
| 28054 | 231 | end; | 
| 27436 | 232 | |
| 59104 | 233 | fun add_language (target_name, language) = | 
| 59102 | 234 |   allocate_target target_name { serial = serial (), language = language, ancestry = [] };
 | 
| 27436 | 235 | |
| 59104 | 236 | 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: 
59100diff
changeset | 237 | let | 
| 59104 | 238 | val _ = if null initial_ancestry | 
| 239 | then error "Must derive from existing target(s)" else (); | |
| 240 | fun the_target_data target_name' = case lookup_target_data thy target_name' of | |
| 241 |       NONE => error ("Unknown code target language: " ^ quote target_name')
 | |
| 242 | | SOME target_data' => target_data'; | |
| 243 | val targets = rev (map (fst o the_target_data o fst) initial_ancestry); | |
| 244 | val supremum = fold1 (fn target' => fn target => | |
| 245 | if #serial target = #serial target' | |
| 246 | then target else error "Incompatible targets") targets; | |
| 247 | val ancestries = map #ancestry targets @ [initial_ancestry]; | |
| 248 | val ancestry = fold1 (fn ancestry' => fn ancestry => | |
| 249 | merge_ancestry (ancestry, ancestry')) ancestries; | |
| 59101 
6dc48c98303c
tuned target inheritance bookkeeping: ancestry is always fully maintained at current entry using canonical merge;
 haftmann parents: 
59100diff
changeset | 250 | in | 
| 59104 | 251 |     allocate_target target_name { serial = #serial supremum, language = #language supremum,
 | 
| 252 | ancestry = ancestry } thy | |
| 59101 
6dc48c98303c
tuned target inheritance bookkeeping: ancestry is always fully maintained at current entry using canonical merge;
 haftmann parents: 
59100diff
changeset | 253 | end; | 
| 
6dc48c98303c
tuned target inheritance bookkeeping: ancestry is always fully maintained at current entry using canonical merge;
 haftmann parents: 
59100diff
changeset | 254 | |
| 59102 | 255 | fun map_data target_name f thy = | 
| 27436 | 256 | let | 
| 59479 | 257 | val _ = assert_target thy target_name; | 
| 28054 | 258 | in | 
| 259 | thy | |
| 59324 | 260 | |> (Targets.map o Symtab.map_entry target_name o apsnd o Code_Printer.map_data) f | 
| 28054 | 261 | end; | 
| 27304 | 262 | |
| 59099 | 263 | fun map_reserved target_name = | 
| 59102 | 264 |   map_data target_name o @{apply 3 (1)};
 | 
| 59099 | 265 | fun map_identifiers target_name = | 
| 59102 | 266 |   map_data target_name o @{apply 3 (2)};
 | 
| 59099 | 267 | fun map_printings target_name = | 
| 59102 | 268 |   map_data target_name o @{apply 3 (3)};
 | 
| 27000 | 269 | |
| 270 | ||
| 34021 | 271 | (** serializer usage **) | 
| 272 | ||
| 59324 | 273 | (* technical aside: pretty printing width *) | 
| 274 | ||
| 275 | val default_code_width = Attrib.setup_config_int @{binding "default_code_width"} (K 80);
 | |
| 276 | ||
| 277 | ||
| 34021 | 278 | (* montage *) | 
| 279 | ||
| 59102 | 280 | fun the_language ctxt = | 
| 59103 | 281 | #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: 
37821diff
changeset | 282 | |
| 59102 | 283 | fun the_literals ctxt = #literals o the_language ctxt; | 
| 34021 | 284 | |
| 39817 
5228c6b20273
check whole target hierarchy for existing reserved symbols
 haftmann parents: 
39750diff
changeset | 285 | local | 
| 
5228c6b20273
check whole target hierarchy for existing reserved symbols
 haftmann parents: 
39750diff
changeset | 286 | |
| 59099 | 287 | fun activate_target ctxt target_name = | 
| 39817 
5228c6b20273
check whole target hierarchy for existing reserved symbols
 haftmann parents: 
39750diff
changeset | 288 | let | 
| 59324 | 289 | val thy = Proof_Context.theory_of ctxt; | 
| 59103 | 290 | val (modify, target_data) = join_ancestry thy target_name; | 
| 59324 | 291 | in (target_data, modify) end; | 
| 38927 | 292 | |
| 55757 | 293 | fun project_program ctxt syms_hidden syms1 program2 = | 
| 38927 | 294 | let | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55146diff
changeset | 295 | val syms2 = subtract (op =) syms_hidden syms1; | 
| 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55146diff
changeset | 296 | val program3 = Code_Symbol.Graph.restrict (not o member (op =) syms_hidden) program2; | 
| 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55146diff
changeset | 297 | val syms4 = Code_Symbol.Graph.all_succs program3 syms2; | 
| 54890 
cb892d835803
fundamental treatment of undefined vs. universally partial replaces code_abort
 haftmann parents: 
54889diff
changeset | 298 | val unimplemented = Code_Thingol.unimplemented program3; | 
| 42359 | 299 | val _ = | 
| 54889 
4121d64fde90
explicit distinction between empty code equations and no code equations, including convenient declaration attributes
 haftmann parents: 
54312diff
changeset | 300 | if null unimplemented then () | 
| 42359 | 301 |       else error ("No code equations for " ^
 | 
| 55304 | 302 | commas (map (Proof_Context.markup_const ctxt) unimplemented)); | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55146diff
changeset | 303 | val program4 = Code_Symbol.Graph.restrict (member (op =) syms4) program3; | 
| 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55146diff
changeset | 304 | in (syms4, program4) end; | 
| 38927 | 305 | |
| 55757 | 306 | fun prepare_serializer ctxt (serializer : serializer) reserved identifiers | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55146diff
changeset | 307 | printings module_name args proto_program syms = | 
| 38927 | 308 | let | 
| 55149 
626d8f08d479
immediate "activation" of const syntax at declaration time
 haftmann parents: 
55147diff
changeset | 309 | val syms_hidden = Code_Symbol.symbols_of printings; | 
| 55757 | 310 | val (syms_all, program) = project_program ctxt syms_hidden syms proto_program; | 
| 39484 
505f95975a5a
closures separate serializer initialization from serializer invocation as far as appropriate
 haftmann parents: 
39480diff
changeset | 311 | fun select_include (name, (content, cs)) = | 
| 55150 | 312 | if null cs orelse exists (fn c => member (op =) syms_all (Constant c)) cs | 
| 39484 
505f95975a5a
closures separate serializer initialization from serializer invocation as far as appropriate
 haftmann parents: 
39480diff
changeset | 313 | then SOME (name, content) else NONE; | 
| 52137 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 haftmann parents: 
52068diff
changeset | 314 | val includes = map_filter select_include (Code_Symbol.dest_module_data printings); | 
| 34021 | 315 | in | 
| 55757 | 316 |     (serializer args ctxt {
 | 
| 52138 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 317 | module_name = module_name, | 
| 38926 | 318 | reserved_syms = reserved, | 
| 52138 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 319 | identifiers = identifiers, | 
| 38926 | 320 | includes = includes, | 
| 55149 
626d8f08d479
immediate "activation" of const syntax at declaration time
 haftmann parents: 
55147diff
changeset | 321 | const_syntax = Code_Symbol.lookup_constant_data printings, | 
| 
626d8f08d479
immediate "activation" of const syntax at declaration time
 haftmann parents: 
55147diff
changeset | 322 | tyco_syntax = Code_Symbol.lookup_type_constructor_data printings, | 
| 
626d8f08d479
immediate "activation" of const syntax at declaration time
 haftmann parents: 
55147diff
changeset | 323 | class_syntax = Code_Symbol.lookup_type_class_data printings }, | 
| 55776 
7dd1971b39c1
amended some slips, rolling back currently dysfunctional export minimimalisation for Scala
 haftmann parents: 
55757diff
changeset | 324 | (subtract (op =) syms_hidden syms, program)) | 
| 34021 | 325 | end; | 
| 326 | ||
| 59099 | 327 | fun mount_serializer ctxt target_name some_width module_name args program syms = | 
| 34021 | 328 | let | 
| 59324 | 329 | val default_width = Config.get ctxt default_code_width; | 
| 330 | val ((target, data), modify) = activate_target ctxt target_name; | |
| 59103 | 331 | val serializer = (#serializer o #language) target; | 
| 55683 | 332 | val (prepared_serializer, (prepared_syms, prepared_program)) = | 
| 55757 | 333 | prepare_serializer ctxt serializer | 
| 59103 | 334 | (Code_Printer.the_reserved data) (Code_Printer.the_identifiers data) | 
| 335 | (Code_Printer.the_printings data) | |
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55146diff
changeset | 336 | module_name args (modify program) syms | 
| 34071 | 337 | val width = the_default default_width some_width; | 
| 55683 | 338 | in (fn program => fn syms => prepared_serializer syms program width, (prepared_syms, prepared_program)) end; | 
| 41344 
d990badc97a3
evaluator separating static and dynamic operations
 haftmann parents: 
41342diff
changeset | 339 | |
| 59099 | 340 | fun invoke_serializer ctxt target_name some_width raw_module_name args program all_public syms = | 
| 41344 
d990badc97a3
evaluator separating static and dynamic operations
 haftmann parents: 
41342diff
changeset | 341 | let | 
| 55291 | 342 | val module_name = if raw_module_name = "" then "" | 
| 343 | else (check_name true raw_module_name; raw_module_name) | |
| 55757 | 344 | val (mounted_serializer, (prepared_syms, prepared_program)) = | 
| 59099 | 345 | mount_serializer ctxt target_name some_width module_name args program syms; | 
| 55776 
7dd1971b39c1
amended some slips, rolling back currently dysfunctional export minimimalisation for Scala
 haftmann parents: 
55757diff
changeset | 346 | in mounted_serializer prepared_program (if all_public then [] else prepared_syms) end; | 
| 34021 | 347 | |
| 52138 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 348 | fun assert_module_name "" = error "Empty module name not allowed here" | 
| 38933 
bd77e092f67c
avoid strange special treatment of empty module names
 haftmann parents: 
38929diff
changeset | 349 | | assert_module_name module_name = module_name; | 
| 
bd77e092f67c
avoid strange special treatment of empty module names
 haftmann parents: 
38929diff
changeset | 350 | |
| 55757 | 351 | fun using_master_directory ctxt = | 
| 56208 | 352 | Option.map (Path.append (File.pwd ()) o | 
| 353 | Path.append (Resources.master_directory (Proof_Context.theory_of ctxt))); | |
| 48371 | 354 | |
| 34021 | 355 | in | 
| 356 | ||
| 48568 | 357 | val generatedN = "Generated_Code"; | 
| 358 | ||
| 59099 | 359 | fun export_code_for ctxt some_path target_name some_width module_name args = | 
| 55757 | 360 | export (using_master_directory ctxt some_path) | 
| 59099 | 361 | ooo invoke_serializer ctxt target_name some_width module_name args; | 
| 38918 | 362 | |
| 59099 | 363 | fun produce_code_for ctxt target_name some_width module_name args = | 
| 39102 | 364 | let | 
| 59099 | 365 | val serializer = invoke_serializer ctxt target_name some_width (assert_module_name module_name) args; | 
| 55683 | 366 | in fn program => fn all_public => fn syms => | 
| 367 | produce (serializer program all_public syms) |> apsnd (fn deresolve => map deresolve syms) | |
| 39484 
505f95975a5a
closures separate serializer initialization from serializer invocation as far as appropriate
 haftmann parents: 
39480diff
changeset | 368 | end; | 
| 38929 
d9ac9dee764d
distinguish code production and code presentation
 haftmann parents: 
38928diff
changeset | 369 | |
| 59099 | 370 | fun present_code_for ctxt target_name some_width module_name args = | 
| 39484 
505f95975a5a
closures separate serializer initialization from serializer invocation as far as appropriate
 haftmann parents: 
39480diff
changeset | 371 | let | 
| 59099 | 372 | val serializer = invoke_serializer ctxt target_name some_width (assert_module_name module_name) args; | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55146diff
changeset | 373 | in fn program => fn (syms, selects) => | 
| 55683 | 374 | present selects (serializer program false syms) | 
| 39484 
505f95975a5a
closures separate serializer initialization from serializer invocation as far as appropriate
 haftmann parents: 
39480diff
changeset | 375 | end; | 
| 38918 | 376 | |
| 59099 | 377 | fun check_code_for ctxt target_name strict args program all_public syms = | 
| 37824 | 378 | let | 
| 379 |     val { env_var, make_destination, make_command } =
 | |
| 59102 | 380 | (#check o the_language ctxt) target_name; | 
| 41939 | 381 | fun ext_check p = | 
| 43564 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43324diff
changeset | 382 | let | 
| 37824 | 383 | val destination = make_destination p; | 
| 59099 | 384 | val _ = export (SOME destination) (invoke_serializer ctxt target_name (SOME 80) | 
| 55683 | 385 | generatedN args program all_public syms); | 
| 48568 | 386 | val cmd = make_command generatedN; | 
| 43850 
7f2cbc713344
moved bash operations to Isabelle_System (cf. Scala version);
 wenzelm parents: 
43564diff
changeset | 387 | in | 
| 
7f2cbc713344
moved bash operations to Isabelle_System (cf. Scala version);
 wenzelm parents: 
43564diff
changeset | 388 |         if Isabelle_System.bash ("cd " ^ File.shell_path p ^ " && " ^ cmd ^ " 2>&1") <> 0
 | 
| 59099 | 389 |         then error ("Code check failed for " ^ target_name ^ ": " ^ cmd)
 | 
| 37824 | 390 | else () | 
| 391 | end; | |
| 43850 
7f2cbc713344
moved bash operations to Isabelle_System (cf. Scala version);
 wenzelm parents: 
43564diff
changeset | 392 | in | 
| 
7f2cbc713344
moved bash operations to Isabelle_System (cf. Scala version);
 wenzelm parents: 
43564diff
changeset | 393 | if getenv env_var = "" | 
| 37825 | 394 | then if strict | 
| 59099 | 395 | then error (env_var ^ " not set; cannot check code for " ^ target_name) | 
| 396 | else warning (env_var ^ " not set; skipped checking code for " ^ target_name) | |
| 41939 | 397 | else Isabelle_System.with_tmp_dir "Code_Test" ext_check | 
| 37824 | 398 | end; | 
| 399 | ||
| 56920 
d651b944c67e
normalizing of type variables before evaluation with explicit resubstitution function: make nbe work with funny type variables like \<AA>;
 haftmann parents: 
56826diff
changeset | 400 | fun subevaluator mounted_serializer prepared_program syms all_public ((vs, ty), t) = | 
| 41344 
d990badc97a3
evaluator separating static and dynamic operations
 haftmann parents: 
41342diff
changeset | 401 | let | 
| 
d990badc97a3
evaluator separating static and dynamic operations
 haftmann parents: 
41342diff
changeset | 402 | val _ = if Code_Thingol.contains_dict_var t then | 
| 
d990badc97a3
evaluator separating static and dynamic operations
 haftmann parents: 
41342diff
changeset | 403 | error "Term to be evaluated contains free dictionaries" else (); | 
| 43324 
2b47822868e4
discontinued Name.variant to emphasize that this is old-style / indirect;
 wenzelm parents: 
42361diff
changeset | 404 | val v' = singleton (Name.variant_list (map fst vs)) "a"; | 
| 41344 
d990badc97a3
evaluator separating static and dynamic operations
 haftmann parents: 
41342diff
changeset | 405 | val vs' = (v', []) :: vs; | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55146diff
changeset | 406 | val ty' = ITyVar v' `-> ty; | 
| 41344 
d990badc97a3
evaluator separating static and dynamic operations
 haftmann parents: 
41342diff
changeset | 407 | val program = prepared_program | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55146diff
changeset | 408 | |> 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: 
56920diff
changeset | 409 | 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: 
55146diff
changeset | 410 | |> fold (curry (perhaps o try o | 
| 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55146diff
changeset | 411 | Code_Symbol.Graph.add_edge) Code_Symbol.value) syms; | 
| 55683 | 412 | val (program_code, deresolve) = | 
| 55684 
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
 haftmann parents: 
55683diff
changeset | 413 | produce (mounted_serializer program (if all_public then [] else [Code_Symbol.value])); | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55146diff
changeset | 414 | val value_name = the (deresolve Code_Symbol.value); | 
| 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55146diff
changeset | 415 | in (program_code, value_name) end; | 
| 41344 
d990badc97a3
evaluator separating static and dynamic operations
 haftmann parents: 
41342diff
changeset | 416 | |
| 59099 | 417 | fun evaluator ctxt target_name program syms = | 
| 41344 
d990badc97a3
evaluator separating static and dynamic operations
 haftmann parents: 
41342diff
changeset | 418 | let | 
| 55683 | 419 | val (mounted_serializer, (_, prepared_program)) = | 
| 59099 | 420 | mount_serializer ctxt target_name NONE generatedN [] program syms; | 
| 56920 
d651b944c67e
normalizing of type variables before evaluation with explicit resubstitution function: make nbe work with funny type variables like \<AA>;
 haftmann parents: 
56826diff
changeset | 421 | in subevaluator mounted_serializer prepared_program syms end; | 
| 41344 
d990badc97a3
evaluator separating static and dynamic operations
 haftmann parents: 
41342diff
changeset | 422 | |
| 34021 | 423 | end; (* local *) | 
| 424 | ||
| 425 | ||
| 426 | (* code generation *) | |
| 427 | ||
| 38918 | 428 | fun prep_destination "" = NONE | 
| 429 | | prep_destination s = SOME (Path.explode s); | |
| 430 | ||
| 55757 | 431 | fun export_code ctxt all_public cs seris = | 
| 34021 | 432 | let | 
| 55757 | 433 | val thy = Proof_Context.theory_of ctxt; | 
| 55188 
7ca0204ece66
reduced prominence of "permissive code generation"
 haftmann parents: 
55153diff
changeset | 434 | val program = Code_Thingol.consts_program thy cs; | 
| 59099 | 435 | val _ = map (fn (((target_name, module_name), some_path), args) => | 
| 436 | export_code_for ctxt some_path target_name NONE module_name args program all_public (map Constant cs)) seris; | |
| 34021 | 437 | in () end; | 
| 438 | ||
| 55757 | 439 | fun export_code_cmd all_public raw_cs seris ctxt = | 
| 440 | export_code ctxt all_public | |
| 441 | (Code_Thingol.read_const_exprs ctxt raw_cs) | |
| 55683 | 442 | ((map o apfst o apsnd) prep_destination seris); | 
| 38918 | 443 | |
| 59099 | 444 | fun produce_code ctxt all_public cs target_name some_width some_module_name args = | 
| 38918 | 445 | let | 
| 55757 | 446 | val thy = Proof_Context.theory_of ctxt; | 
| 55188 
7ca0204ece66
reduced prominence of "permissive code generation"
 haftmann parents: 
55153diff
changeset | 447 | val program = Code_Thingol.consts_program thy cs; | 
| 59099 | 448 | 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: 
38928diff
changeset | 449 | |
| 59099 | 450 | fun present_code ctxt cs syms target_name some_width some_module_name args = | 
| 38929 
d9ac9dee764d
distinguish code production and code presentation
 haftmann parents: 
38928diff
changeset | 451 | let | 
| 55757 | 452 | val thy = Proof_Context.theory_of ctxt; | 
| 55188 
7ca0204ece66
reduced prominence of "permissive code generation"
 haftmann parents: 
55153diff
changeset | 453 | val program = Code_Thingol.consts_program thy cs; | 
| 59099 | 454 | in present_code_for ctxt target_name some_width some_module_name args program (map Constant cs, syms) end; | 
| 34021 | 455 | |
| 55757 | 456 | fun check_code ctxt all_public cs seris = | 
| 37824 | 457 | let | 
| 55757 | 458 | val thy = Proof_Context.theory_of ctxt; | 
| 55188 
7ca0204ece66
reduced prominence of "permissive code generation"
 haftmann parents: 
55153diff
changeset | 459 | val program = Code_Thingol.consts_program thy cs; | 
| 59099 | 460 | val _ = map (fn ((target_name, strict), args) => | 
| 461 | check_code_for ctxt target_name strict args program all_public (map Constant cs)) seris; | |
| 37824 | 462 | in () end; | 
| 463 | ||
| 55757 | 464 | fun check_code_cmd all_public raw_cs seris ctxt = | 
| 465 | check_code ctxt all_public | |
| 466 | (Code_Thingol.read_const_exprs ctxt raw_cs) seris; | |
| 37824 | 467 | |
| 39480 
a2ed61449dcc
added code_stmts antiquotation from doc-src/more_antiquote.ML
 haftmann parents: 
39142diff
changeset | 468 | local | 
| 
a2ed61449dcc
added code_stmts antiquotation from doc-src/more_antiquote.ML
 haftmann parents: 
39142diff
changeset | 469 | |
| 
a2ed61449dcc
added code_stmts antiquotation from doc-src/more_antiquote.ML
 haftmann parents: 
39142diff
changeset | 470 | val parse_const_terms = Scan.repeat1 Args.term | 
| 55757 | 471 | >> (fn ts => fn ctxt => map (Code.check_const (Proof_Context.theory_of ctxt)) ts); | 
| 39480 
a2ed61449dcc
added code_stmts antiquotation from doc-src/more_antiquote.ML
 haftmann parents: 
39142diff
changeset | 472 | |
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55146diff
changeset | 473 | fun parse_names category parse internalize mark_symbol = | 
| 39480 
a2ed61449dcc
added code_stmts antiquotation from doc-src/more_antiquote.ML
 haftmann parents: 
39142diff
changeset | 474 | Scan.lift (Args.parens (Args.$$$ category)) |-- Scan.repeat1 parse | 
| 55757 | 475 | >> (fn xs => fn ctxt => map (mark_symbol o internalize ctxt) xs); | 
| 43564 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43324diff
changeset | 476 | |
| 39480 
a2ed61449dcc
added code_stmts antiquotation from doc-src/more_antiquote.ML
 haftmann parents: 
39142diff
changeset | 477 | val parse_consts = parse_names "consts" Args.term | 
| 55757 | 478 | (Code.check_const o Proof_Context.theory_of) Constant; | 
| 39480 
a2ed61449dcc
added code_stmts antiquotation from doc-src/more_antiquote.ML
 haftmann parents: 
39142diff
changeset | 479 | |
| 
a2ed61449dcc
added code_stmts antiquotation from doc-src/more_antiquote.ML
 haftmann parents: 
39142diff
changeset | 480 | val parse_types = parse_names "types" (Scan.lift Args.name) | 
| 55757 | 481 | (Sign.intern_type o Proof_Context.theory_of) Type_Constructor; | 
| 39480 
a2ed61449dcc
added code_stmts antiquotation from doc-src/more_antiquote.ML
 haftmann parents: 
39142diff
changeset | 482 | |
| 
a2ed61449dcc
added code_stmts antiquotation from doc-src/more_antiquote.ML
 haftmann parents: 
39142diff
changeset | 483 | val parse_classes = parse_names "classes" (Scan.lift Args.name) | 
| 55757 | 484 | (Sign.intern_class o Proof_Context.theory_of) Type_Class; | 
| 39480 
a2ed61449dcc
added code_stmts antiquotation from doc-src/more_antiquote.ML
 haftmann parents: 
39142diff
changeset | 485 | |
| 
a2ed61449dcc
added code_stmts antiquotation from doc-src/more_antiquote.ML
 haftmann parents: 
39142diff
changeset | 486 | val parse_instances = parse_names "instances" (Scan.lift (Args.name --| Args.$$$ "::" -- Args.name)) | 
| 55757 | 487 | (fn ctxt => fn (raw_tyco, raw_class) => | 
| 488 | let | |
| 489 | val thy = Proof_Context.theory_of ctxt; | |
| 490 | in (Sign.intern_class thy raw_tyco, Sign.intern_type thy raw_class) end) Class_Instance; | |
| 39480 
a2ed61449dcc
added code_stmts antiquotation from doc-src/more_antiquote.ML
 haftmann parents: 
39142diff
changeset | 491 | |
| 
a2ed61449dcc
added code_stmts antiquotation from doc-src/more_antiquote.ML
 haftmann parents: 
39142diff
changeset | 492 | in | 
| 
a2ed61449dcc
added code_stmts antiquotation from doc-src/more_antiquote.ML
 haftmann parents: 
39142diff
changeset | 493 | |
| 59323 | 494 | val _ = Theory.setup | 
| 495 |   (Thy_Output.antiquotation @{binding code_stmts}
 | |
| 43564 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43324diff
changeset | 496 | (parse_const_terms -- | 
| 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43324diff
changeset | 497 | Scan.repeat (parse_consts || parse_types || parse_classes || parse_instances) | 
| 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43324diff
changeset | 498 | -- Scan.lift (Args.parens (Args.name -- Scan.option Parse.int))) | 
| 59099 | 499 |     (fn {context = ctxt, ...} => fn ((mk_cs, mk_stmtss), (target_name, some_width)) =>
 | 
| 55757 | 500 | present_code ctxt (mk_cs ctxt) | 
| 501 | (maps (fn f => f ctxt) mk_stmtss) | |
| 59323 | 502 | target_name some_width "Example" [])); | 
| 39480 
a2ed61449dcc
added code_stmts antiquotation from doc-src/more_antiquote.ML
 haftmann parents: 
39142diff
changeset | 503 | |
| 
a2ed61449dcc
added code_stmts antiquotation from doc-src/more_antiquote.ML
 haftmann parents: 
39142diff
changeset | 504 | end; | 
| 
a2ed61449dcc
added code_stmts antiquotation from doc-src/more_antiquote.ML
 haftmann parents: 
39142diff
changeset | 505 | |
| 34021 | 506 | |
| 28054 | 507 | (** serializer configuration **) | 
| 27000 | 508 | |
| 52138 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 509 | (* reserved symbol names *) | 
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 510 | |
| 59099 | 511 | 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: 
52137diff
changeset | 512 | let | 
| 59103 | 513 | val (_, (_, data)) = join_ancestry thy target_name; | 
| 514 | 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: 
52137diff
changeset | 515 |       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: 
52137diff
changeset | 516 | else (); | 
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 517 | in | 
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 518 | thy | 
| 59099 | 519 | |> 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: 
52137diff
changeset | 520 | end; | 
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 521 | |
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 522 | |
| 52377 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 haftmann parents: 
52218diff
changeset | 523 | (* checking of syntax *) | 
| 52138 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 524 | |
| 59099 | 525 | fun check_const_syntax ctxt target_name c syn = | 
| 55757 | 526 | 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: 
52068diff
changeset | 527 |   then error ("Too many arguments in syntax for constant " ^ quote c)
 | 
| 59099 | 528 | 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: 
52068diff
changeset | 529 | |
| 59099 | 530 | fun check_tyco_syntax ctxt target_name tyco syn = | 
| 55757 | 531 | 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: 
52068diff
changeset | 532 |   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: 
52068diff
changeset | 533 | else syn; | 
| 34071 | 534 | |
| 52138 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 535 | |
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 536 | (* custom symbol names *) | 
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 537 | |
| 52218 | 538 | fun arrange_name_decls x = | 
| 52138 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 539 | let | 
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 540 | 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: 
52137diff
changeset | 541 | (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: 
52137diff
changeset | 542 | in | 
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 543 | Code_Symbol.maps_attr' (arrange false) (arrange false) (arrange false) | 
| 52218 | 544 | (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: 
52137diff
changeset | 545 | end; | 
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 546 | |
| 55757 | 547 | 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: 
52137diff
changeset | 548 | |
| 55757 | 549 | 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: 
52137diff
changeset | 550 | |
| 59099 | 551 | 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: 
52137diff
changeset | 552 | |
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 553 | fun gen_set_identifiers prep_name_decl raw_name_decls thy = | 
| 55757 | 554 | 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: 
52137diff
changeset | 555 | |
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 556 | 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: 
52137diff
changeset | 557 | 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: 
52137diff
changeset | 558 | |
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 559 | |
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 560 | (* custom printings *) | 
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 561 | |
| 55757 | 562 | fun arrange_printings prep_const ctxt = | 
| 52137 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 haftmann parents: 
52068diff
changeset | 563 | let | 
| 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 haftmann parents: 
52068diff
changeset | 564 | fun arrange check (sym, target_syns) = | 
| 59099 | 565 | map (fn (target_name, some_syn) => | 
| 566 | (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: 
52068diff
changeset | 567 | in | 
| 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 haftmann parents: 
52068diff
changeset | 568 | Code_Symbol.maps_attr' | 
| 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 haftmann parents: 
52068diff
changeset | 569 | (arrange check_const_syntax) (arrange check_tyco_syntax) | 
| 55149 
626d8f08d479
immediate "activation" of const syntax at declaration time
 haftmann parents: 
55147diff
changeset | 570 | (arrange ((K o K o K) I)) (arrange ((K o K o K) I)) (arrange ((K o K o K) I)) | 
| 55757 | 571 | (arrange (fn ctxt => fn _ => fn _ => fn (raw_content, raw_cs) => | 
| 59430 | 572 | (Pretty.blk (0, Pretty.fbreaks (map Code_Printer.str (split_lines raw_content))), | 
| 573 | map (prep_const ctxt) raw_cs))) | |
| 52137 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 haftmann parents: 
52068diff
changeset | 574 | end; | 
| 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 haftmann parents: 
52068diff
changeset | 575 | |
| 55757 | 576 | fun cert_printings ctxt = cert_syms ctxt #> arrange_printings cert_const ctxt; | 
| 24219 | 577 | |
| 55757 | 578 | fun read_printings ctxt = read_syms ctxt #> arrange_printings read_const ctxt; | 
| 52137 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 haftmann parents: 
52068diff
changeset | 579 | |
| 59099 | 580 | 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: 
52068diff
changeset | 581 | |
| 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 haftmann parents: 
52068diff
changeset | 582 | fun gen_set_printings prep_print_decl raw_print_decls thy = | 
| 55757 | 583 | 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: 
52068diff
changeset | 584 | |
| 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 haftmann parents: 
52068diff
changeset | 585 | 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: 
52068diff
changeset | 586 | 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: 
52068diff
changeset | 587 | |
| 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 haftmann parents: 
52068diff
changeset | 588 | |
| 27103 | 589 | (* concrete syntax *) | 
| 27000 | 590 | |
| 28054 | 591 | 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: 
36537diff
changeset | 592 | case Scan.read Token.stopper f args | 
| 28054 | 593 | of SOME x => x | 
| 594 | | NONE => error "Bad serializer arguments"; | |
| 595 | ||
| 596 | ||
| 27304 | 597 | (** Isar setup **) | 
| 27103 | 598 | |
| 52137 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 haftmann parents: 
52068diff
changeset | 599 | fun parse_single_symbol_pragma parse_keyword parse_isa parse_target = | 
| 52434 | 600 |   parse_keyword |-- Parse.!!! (parse_isa --| (@{keyword "\<rightharpoonup>"} || @{keyword "=>"})
 | 
| 601 |     -- Parse.and_list1 (@{keyword "("} |-- (Parse.name --| @{keyword ")"} -- Scan.option parse_target)));
 | |
| 52137 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 haftmann parents: 
52068diff
changeset | 602 | |
| 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 haftmann parents: 
52068diff
changeset | 603 | fun parse_symbol_pragma parse_const parse_tyco parse_class parse_classrel parse_inst parse_module = | 
| 52801 | 604 |   parse_single_symbol_pragma @{keyword "constant"} Parse.term parse_const
 | 
| 55150 | 605 | >> Constant | 
| 52377 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 haftmann parents: 
52218diff
changeset | 606 |   || parse_single_symbol_pragma @{keyword "type_constructor"} Parse.type_const parse_tyco
 | 
| 55150 | 607 | >> Type_Constructor | 
| 52377 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 haftmann parents: 
52218diff
changeset | 608 |   || parse_single_symbol_pragma @{keyword "type_class"} Parse.class parse_class
 | 
| 55150 | 609 | >> Type_Class | 
| 52137 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 haftmann parents: 
52068diff
changeset | 610 |   || parse_single_symbol_pragma @{keyword "class_relation"} parse_classrel_ident parse_classrel
 | 
| 55150 | 611 | >> Class_Relation | 
| 52137 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 haftmann parents: 
52068diff
changeset | 612 |   || parse_single_symbol_pragma @{keyword "class_instance"} parse_inst_ident parse_inst
 | 
| 55150 | 613 | >> Class_Instance | 
| 52137 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 haftmann parents: 
52068diff
changeset | 614 |   || parse_single_symbol_pragma @{keyword "code_module"} Parse.name parse_module
 | 
| 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 haftmann parents: 
52068diff
changeset | 615 | >> Code_Symbol.Module; | 
| 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 haftmann parents: 
52068diff
changeset | 616 | |
| 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 haftmann parents: 
52068diff
changeset | 617 | 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: 
52068diff
changeset | 618 | 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: 
52068diff
changeset | 619 | (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: 
52068diff
changeset | 620 | |
| 56201 | 621 | val code_expr_argsP = Scan.optional (@{keyword "("} |-- Parse.args --| @{keyword ")"}) [];
 | 
| 52138 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 622 | |
| 55683 | 623 | fun code_expr_inP all_public raw_cs = | 
| 52434 | 624 |   Scan.repeat (@{keyword "in"} |-- Parse.!!! (Parse.name
 | 
| 625 |     -- Scan.optional (@{keyword "module_name"} |-- Parse.name) ""
 | |
| 626 |     -- Scan.optional (@{keyword "file"} |-- Parse.name) ""
 | |
| 627 | -- code_expr_argsP)) | |
| 55683 | 628 | >> (fn seri_args => export_code_cmd all_public raw_cs seri_args); | 
| 52434 | 629 | |
| 55683 | 630 | fun code_expr_checkingP all_public raw_cs = | 
| 52434 | 631 |   (@{keyword "checking"} |-- Parse.!!!
 | 
| 632 |     (Scan.repeat (Parse.name -- ((@{keyword "?"} |-- Scan.succeed false) || Scan.succeed true)
 | |
| 633 | -- code_expr_argsP))) | |
| 55683 | 634 | >> (fn seri_args => check_code_cmd all_public raw_cs seri_args); | 
| 52434 | 635 | |
| 55683 | 636 | val code_exprP = (Scan.optional (@{keyword "open"} |-- Scan.succeed true) false
 | 
| 637 | -- Scan.repeat1 Parse.term) | |
| 638 | :|-- (fn (all_public, raw_cs) => (code_expr_checkingP all_public raw_cs || code_expr_inP all_public raw_cs)); | |
| 52138 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 639 | |
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 640 | val _ = | 
| 59936 
b8ffc3dc9e24
@{command_spec} is superseded by @{command_keyword};
 wenzelm parents: 
59480diff
changeset | 641 |   Outer_Syntax.command @{command_keyword code_reserved}
 | 
| 52138 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 642 | "declare words as reserved for target language" | 
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 643 | (Parse.name -- Scan.repeat1 Parse.name | 
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 644 | >> (fn (target, reserveds) => (Toplevel.theory o fold (add_reserved target)) reserveds)); | 
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 645 | |
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 646 | val _ = | 
| 59936 
b8ffc3dc9e24
@{command_spec} is superseded by @{command_keyword};
 wenzelm parents: 
59480diff
changeset | 647 |   Outer_Syntax.command @{command_keyword code_identifier} "declare mandatory names for code symbols"
 | 
| 52138 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 648 | (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: 
52137diff
changeset | 649 | >> (Toplevel.theory o fold set_identifiers_cmd)); | 
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 650 | |
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 651 | val _ = | 
| 59936 
b8ffc3dc9e24
@{command_spec} is superseded by @{command_keyword};
 wenzelm parents: 
59480diff
changeset | 652 |   Outer_Syntax.command @{command_keyword code_printing} "declare dedicated printing for code symbols"
 | 
| 52137 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 haftmann parents: 
52068diff
changeset | 653 | (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: 
52068diff
changeset | 654 | Parse.string (Parse.minus >> K ()) (Parse.minus >> K ()) | 
| 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 haftmann parents: 
52068diff
changeset | 655 |       (Parse.text -- Scan.optional (@{keyword "attach"} |-- Scan.repeat1 Parse.term) [])
 | 
| 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 haftmann parents: 
52068diff
changeset | 656 | >> (Toplevel.theory o fold set_printings_cmd)); | 
| 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 haftmann parents: 
52068diff
changeset | 657 | |
| 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 haftmann parents: 
52068diff
changeset | 658 | val _ = | 
| 59936 
b8ffc3dc9e24
@{command_spec} is superseded by @{command_keyword};
 wenzelm parents: 
59480diff
changeset | 659 |   Outer_Syntax.command @{command_keyword export_code} "generate executable code for constants"
 | 
| 55757 | 660 | (Parse.!!! code_exprP >> (fn f => Toplevel.keep (f o Toplevel.context_of))); | 
| 30494 
c150e6fa4e0d
consider exit status of code generation direcitve
 haftmann parents: 
30242diff
changeset | 661 | |
| 24219 | 662 | end; (*struct*) |