| author | haftmann | 
| Sun, 05 Oct 2014 20:30:57 +0200 | |
| changeset 58558 | 30cc7ee5ac10 | 
| parent 57918 | f5d73caba4e5 | 
| child 58903 | 38c72f5f6c2e | 
| 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 | 
| 37822 
cf3588177676
use generic description slot for formal code checking
 haftmann parents: 
37821diff
changeset | 37 |   val add_target: string * { serializer: serializer, literals: literals,
 | 
| 41940 | 38 |     check: { env_var: string, make_destination: Path.T -> Path.T, make_command: string -> string } }
 | 
| 39 | -> theory -> theory | |
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28090diff
changeset | 40 | val extend_target: string * | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55146diff
changeset | 41 | (string * (Code_Thingol.program -> Code_Thingol.program)) | 
| 28054 | 42 | -> theory -> theory | 
| 55757 | 43 | val assert_target: Proof.context -> string -> string | 
| 44 | val the_literals: Proof.context -> string -> literals | |
| 28054 | 45 | 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 | 46 | val parse_args: 'a parser -> Token.T list -> 'a | 
| 38916 | 47 | val serialization: (int -> Path.T option -> 'a -> unit) | 
| 55150 | 48 | -> (Code_Symbol.T list -> int -> 'a -> (string * string) list * (Code_Symbol.T -> string option)) | 
| 38925 | 49 | -> 'a -> serialization | 
| 38918 | 50 | val set_default_code_width: int -> theory -> theory | 
| 38917 | 51 | |
| 52137 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 haftmann parents: 
52068diff
changeset | 52 |   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 | 53 | type identifier_data | 
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 54 | 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 | 55 | -> theory -> theory | 
| 55372 | 56 | 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 | 57 | -> theory -> theory | 
| 28054 | 58 | val add_reserved: string -> string -> theory -> theory | 
| 39646 | 59 | |
| 39750 
c0099428ca7b
consider quick_and_dirty option before loading theory
 haftmann parents: 
39679diff
changeset | 60 | val codegen_tool: string (*theory name*) -> string (*export_code expr*) -> unit | 
| 43564 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43324diff
changeset | 61 | |
| 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43324diff
changeset | 62 | val setup: theory -> theory | 
| 24219 | 63 | end; | 
| 64 | ||
| 28054 | 65 | structure Code_Target : CODE_TARGET = | 
| 24219 | 66 | struct | 
| 67 | ||
| 55150 | 68 | open Basic_Code_Symbol; | 
| 28054 | 69 | open Basic_Code_Thingol; | 
| 34152 | 70 | |
| 71 | 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 | 72 | 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 | 73 | (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 | 74 | 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 | 75 | (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 | 76 | string * (string * 'f option) list) Code_Symbol.attr; | 
| 55291 | 77 | type identifier_data = (string list * string, string list * string, string list * string, string list * string, | 
| 78 | string list * string, string list * string) Code_Symbol.data; | |
| 52137 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 haftmann parents: 
52068diff
changeset | 79 | |
| 34152 | 80 | type tyco_syntax = Code_Printer.tyco_syntax; | 
| 55153 
eedd549de3ef
more suitable names, without any notion of "activating"
 haftmann parents: 
55150diff
changeset | 81 | type raw_const_syntax = Code_Printer.raw_const_syntax; | 
| 34152 | 82 | |
| 24219 | 83 | |
| 52377 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 haftmann parents: 
52218diff
changeset | 84 | (** checking and parsing of symbols **) | 
| 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 haftmann parents: 
52218diff
changeset | 85 | |
| 55757 | 86 | fun cert_const ctxt const = | 
| 52377 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 haftmann parents: 
52218diff
changeset | 87 | let | 
| 55757 | 88 | 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 | 89 |       else error ("No such constant: " ^ quote const);
 | 
| 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 haftmann parents: 
52218diff
changeset | 90 | in const end; | 
| 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 haftmann parents: 
52218diff
changeset | 91 | |
| 55757 | 92 | fun read_const ctxt = Code.read_const (Proof_Context.theory_of ctxt); | 
| 93 | ||
| 94 | fun cert_tyco ctxt tyco = | |
| 52377 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 haftmann parents: 
52218diff
changeset | 95 | let | 
| 55757 | 96 | 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 | 97 |       else error ("No such type constructor: " ^ quote tyco);
 | 
| 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 haftmann parents: 
52218diff
changeset | 98 | in tyco end; | 
| 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 haftmann parents: 
52218diff
changeset | 99 | |
| 55951 
c07d184aebe9
tuned signature -- more uniform check_type_name/read_type_name;
 wenzelm parents: 
55776diff
changeset | 100 | fun read_tyco ctxt = | 
| 56002 | 101 |   #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 | 102 | |
| 55757 | 103 | fun cert_class ctxt class = | 
| 52377 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 haftmann parents: 
52218diff
changeset | 104 | let | 
| 55757 | 105 | 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 | 106 | in class end; | 
| 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 haftmann parents: 
52218diff
changeset | 107 | |
| 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 haftmann parents: 
52218diff
changeset | 108 | val parse_classrel_ident = Parse.class --| @{keyword "<"} -- Parse.class;
 | 
| 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 haftmann parents: 
52218diff
changeset | 109 | |
| 55757 | 110 | fun cert_inst ctxt (class, tyco) = | 
| 111 | (cert_class ctxt class, cert_tyco ctxt tyco); | |
| 52377 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 haftmann parents: 
52218diff
changeset | 112 | |
| 55757 | 113 | fun read_inst ctxt (raw_tyco, raw_class) = | 
| 114 | (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 | 115 | |
| 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 haftmann parents: 
52218diff
changeset | 116 | val parse_inst_ident = Parse.xname --| @{keyword "::"} -- Parse.class;
 | 
| 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 haftmann parents: 
52218diff
changeset | 117 | |
| 55757 | 118 | fun cert_syms ctxt = | 
| 119 | Code_Symbol.map_attr (apfst (cert_const ctxt)) (apfst (cert_tyco ctxt)) | |
| 120 | (apfst (cert_class ctxt)) ((apfst o pairself) (cert_class ctxt)) (apfst (cert_inst ctxt)) I; | |
| 52377 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 haftmann parents: 
52218diff
changeset | 121 | |
| 55757 | 122 | fun read_syms ctxt = | 
| 123 | Code_Symbol.map_attr (apfst (read_const ctxt)) (apfst (read_tyco ctxt)) | |
| 124 | (apfst (Proof_Context.read_class ctxt)) ((apfst o pairself) (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 | 125 | |
| 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 haftmann parents: 
52218diff
changeset | 126 | fun check_name is_module s = | 
| 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 haftmann parents: 
52218diff
changeset | 127 | let | 
| 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 haftmann parents: 
52218diff
changeset | 128 | 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 | 129 | val xs = Long_Name.explode s; | 
| 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 haftmann parents: 
52218diff
changeset | 130 | 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 | 131 | then map (Name.desymbolize NONE) xs | 
| 52377 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 haftmann parents: 
52218diff
changeset | 132 | else if length xs < 2 | 
| 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 haftmann parents: 
52218diff
changeset | 133 |         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 | 134 | else | 
| 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 haftmann parents: 
52218diff
changeset | 135 | let | 
| 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 haftmann parents: 
52218diff
changeset | 136 | 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 | 137 | 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 | 138 | val y' = Name.desymbolize NONE y; | 
| 52377 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 haftmann parents: 
52218diff
changeset | 139 | in ys' @ [y'] end; | 
| 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 haftmann parents: 
52218diff
changeset | 140 | in if xs' = xs | 
| 55291 | 141 | 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 | 142 |     else error ("Invalid code name: " ^ quote s ^ "\n"
 | 
| 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 haftmann parents: 
52218diff
changeset | 143 | ^ "better try " ^ quote (Long_Name.implode xs')) | 
| 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 haftmann parents: 
52218diff
changeset | 144 | end; | 
| 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 haftmann parents: 
52218diff
changeset | 145 | |
| 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 haftmann parents: 
52218diff
changeset | 146 | |
| 52137 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 haftmann parents: 
52068diff
changeset | 147 | (** serializations and serializer **) | 
| 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 haftmann parents: 
52068diff
changeset | 148 | |
| 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 haftmann parents: 
52068diff
changeset | 149 | (* serialization: abstract nonsense to cover different destinies for generated code *) | 
| 24219 | 150 | |
| 55150 | 151 | datatype destination = Export of Path.T option | Produce | Present of Code_Symbol.T list; | 
| 152 | 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 | 153 | |
| 39659 | 154 | 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 | 155 | (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 | 156 | | 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 | 157 | string [] width content |> SOME | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55146diff
changeset | 158 | | serialization _ string content width (Present syms) = | 
| 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55146diff
changeset | 159 | string syms width content | 
| 48568 | 160 | |> (apfst o map o apsnd) (Pretty.output (SOME width) o Pretty.str) | 
| 39661 
6381d18507ef
reverted cs 07549694e2f1 -- use re-printing with current print mode instead after code assembly, avoid Latex.output_typewriter
 haftmann parents: 
39659diff
changeset | 161 | |> SOME; | 
| 39659 | 162 | |
| 39661 
6381d18507ef
reverted cs 07549694e2f1 -- use re-printing with current print mode instead after code assembly, avoid Latex.output_typewriter
 haftmann parents: 
39659diff
changeset | 163 | 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 | 164 | fun produce f = the (f Produce); | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55146diff
changeset | 165 | fun present syms f = space_implode "\n\n" (map snd (fst (the (f (Present syms))))); | 
| 38917 | 166 | |
| 24219 | 167 | |
| 52137 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 haftmann parents: 
52068diff
changeset | 168 | (* serializers: functions producing serializations *) | 
| 28054 | 169 | |
| 39142 | 170 | 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 | 171 | -> Proof.context | 
| 39142 | 172 |   -> {
 | 
| 52138 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 173 | module_name: string, | 
| 38926 | 174 | reserved_syms: string list, | 
| 52138 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 175 | identifiers: identifier_data, | 
| 38926 | 176 | includes: (string * Pretty.T) list, | 
| 177 | class_syntax: string -> string option, | |
| 178 | tyco_syntax: string -> Code_Printer.tyco_syntax option, | |
| 55153 
eedd549de3ef
more suitable names, without any notion of "activating"
 haftmann parents: 
55150diff
changeset | 179 | const_syntax: string -> Code_Printer.const_syntax option } | 
| 55683 | 180 | -> Code_Symbol.T list | 
| 41342 
3519e0dd8f75
more explicit structure for serializer invocation
 haftmann parents: 
41307diff
changeset | 181 | -> Code_Thingol.program | 
| 28054 | 182 | -> serialization; | 
| 27000 | 183 | |
| 52137 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 haftmann parents: 
52068diff
changeset | 184 | datatype description = | 
| 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 haftmann parents: 
52068diff
changeset | 185 |     Fundamental of { serializer: serializer,
 | 
| 38910 | 186 | literals: literals, | 
| 37822 
cf3588177676
use generic description slot for formal code checking
 haftmann parents: 
37821diff
changeset | 187 |       check: { env_var: string, make_destination: Path.T -> Path.T,
 | 
| 41940 | 188 | make_command: string -> string } } | 
| 38909 | 189 | | Extension of string * | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55146diff
changeset | 190 | (Code_Thingol.program -> Code_Thingol.program); | 
| 28054 | 191 | |
| 52137 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 haftmann parents: 
52068diff
changeset | 192 | |
| 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 haftmann parents: 
52068diff
changeset | 193 | (** theory data **) | 
| 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 haftmann parents: 
52068diff
changeset | 194 | |
| 28054 | 195 | datatype target = Target of {
 | 
| 196 | serial: serial, | |
| 37821 | 197 | description: description, | 
| 28054 | 198 | reserved: string list, | 
| 52138 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 199 | identifiers: identifier_data, | 
| 55153 
eedd549de3ef
more suitable names, without any notion of "activating"
 haftmann parents: 
55150diff
changeset | 200 | printings: (Code_Printer.const_syntax, Code_Printer.tyco_syntax, string, unit, unit, | 
| 52137 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 haftmann parents: 
52068diff
changeset | 201 | (Pretty.T * string list)) Code_Symbol.data | 
| 28054 | 202 | }; | 
| 27103 | 203 | |
| 52138 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 204 | fun make_target ((serial, description), (reserved, (identifiers, printings))) = | 
| 43564 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43324diff
changeset | 205 |   Target { serial = serial, description = description, reserved = reserved,
 | 
| 52138 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 206 | identifiers = identifiers, printings = printings }; | 
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 207 | fun map_target f (Target { serial, description, reserved, identifiers, printings }) =
 | 
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 208 | make_target (f ((serial, description), (reserved, (identifiers, printings)))); | 
| 37821 | 209 | fun merge_target strict target (Target { serial = serial1, description = description,
 | 
| 52138 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 210 | reserved = reserved1, identifiers = identifiers1, printings = printings1 }, | 
| 37821 | 211 |     Target { serial = serial2, description = _,
 | 
| 52138 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 212 | reserved = reserved2, identifiers = identifiers2, printings = printings2 }) = | 
| 28054 | 213 | if serial1 = serial2 orelse not strict then | 
| 52138 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 214 | make_target ((serial1, description), (merge (op =) (reserved1, reserved2), | 
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 215 | (Code_Symbol.merge_data (identifiers1, identifiers2), | 
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 216 | Code_Symbol.merge_data (printings1, printings2)))) | 
| 28054 | 217 | else | 
| 37821 | 218 |     error ("Incompatible targets: " ^ quote target);
 | 
| 27103 | 219 | |
| 37821 | 220 | fun the_description (Target { description, ... }) = description;
 | 
| 28054 | 221 | fun the_reserved (Target { reserved, ... }) = reserved;
 | 
| 52138 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 222 | fun the_identifiers (Target { identifiers , ... }) = identifiers;
 | 
| 52137 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 haftmann parents: 
52068diff
changeset | 223 | fun the_printings (Target { printings, ... }) = printings;
 | 
| 27437 | 224 | |
| 34248 | 225 | structure Targets = Theory_Data | 
| 34071 | 226 | ( | 
| 54890 
cb892d835803
fundamental treatment of undefined vs. universally partial replaces code_abort
 haftmann parents: 
54889diff
changeset | 227 | type T = target Symtab.table * int; | 
| 
cb892d835803
fundamental treatment of undefined vs. universally partial replaces code_abort
 haftmann parents: 
54889diff
changeset | 228 | val empty = (Symtab.empty, 80); | 
| 34071 | 229 | val extend = I; | 
| 54890 
cb892d835803
fundamental treatment of undefined vs. universally partial replaces code_abort
 haftmann parents: 
54889diff
changeset | 230 | fun merge ((target1, width1), (target2, width2)) : T = | 
| 
cb892d835803
fundamental treatment of undefined vs. universally partial replaces code_abort
 haftmann parents: 
54889diff
changeset | 231 | (Symtab.join (merge_target true) (target1, target2), Int.max (width1, width2)); | 
| 34071 | 232 | ); | 
| 27436 | 233 | |
| 55757 | 234 | fun assert_target ctxt target = | 
| 235 | if Symtab.defined (fst (Targets.get (Proof_Context.theory_of ctxt))) target | |
| 33969 | 236 | then target | 
| 237 |   else error ("Unknown code target language: " ^ quote target);
 | |
| 28054 | 238 | |
| 239 | fun put_target (target, seri) thy = | |
| 27304 | 240 | let | 
| 54890 
cb892d835803
fundamental treatment of undefined vs. universally partial replaces code_abort
 haftmann parents: 
54889diff
changeset | 241 | val lookup_target = Symtab.lookup (fst (Targets.get thy)); | 
| 28054 | 242 | val _ = case seri | 
| 37821 | 243 | of Extension (super, _) => if is_some (lookup_target super) then () | 
| 28054 | 244 |           else error ("Unknown code target language: " ^ quote super)
 | 
| 245 | | _ => (); | |
| 37821 | 246 | val overwriting = case (Option.map the_description o lookup_target) target | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28090diff
changeset | 247 | of NONE => false | 
| 37821 | 248 | | SOME (Extension _) => true | 
| 249 | | SOME (Fundamental _) => (case seri | |
| 250 |          of Extension _ => error ("Will not overwrite existing target " ^ quote target)
 | |
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28090diff
changeset | 251 | | _ => true); | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28090diff
changeset | 252 | val _ = if overwriting | 
| 28054 | 253 |       then warning ("Overwriting existing target " ^ quote target)
 | 
| 43564 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43324diff
changeset | 254 | else (); | 
| 28054 | 255 | in | 
| 256 | thy | |
| 54890 
cb892d835803
fundamental treatment of undefined vs. universally partial replaces code_abort
 haftmann parents: 
54889diff
changeset | 257 | |> (Targets.map o apfst o Symtab.update) | 
| 52138 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 258 | (target, make_target ((serial (), seri), | 
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 259 | ([], (Code_Symbol.empty_data, Code_Symbol.empty_data)))) | 
| 28054 | 260 | end; | 
| 27436 | 261 | |
| 37821 | 262 | fun add_target (target, seri) = put_target (target, Fundamental seri); | 
| 28054 | 263 | fun extend_target (target, (super, modify)) = | 
| 37821 | 264 | put_target (target, Extension (super, modify)); | 
| 27436 | 265 | |
| 28054 | 266 | fun map_target_data target f thy = | 
| 27436 | 267 | let | 
| 55757 | 268 | val _ = assert_target (Proof_Context.init_global thy) target; | 
| 28054 | 269 | in | 
| 270 | thy | |
| 54890 
cb892d835803
fundamental treatment of undefined vs. universally partial replaces code_abort
 haftmann parents: 
54889diff
changeset | 271 | |> (Targets.map o apfst o Symtab.map_entry target o map_target o apsnd) f | 
| 28054 | 272 | end; | 
| 27304 | 273 | |
| 28054 | 274 | fun map_reserved target = | 
| 52138 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 275 | map_target_data target o apfst; | 
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 276 | fun map_identifiers target = | 
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 277 | map_target_data target o apsnd o apfst; | 
| 52137 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 haftmann parents: 
52068diff
changeset | 278 | fun map_printings target = | 
| 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 haftmann parents: 
52068diff
changeset | 279 | map_target_data target o apsnd o apsnd; | 
| 27000 | 280 | |
| 34248 | 281 | fun set_default_code_width k = (Targets.map o apsnd) (K k); | 
| 34071 | 282 | |
| 27000 | 283 | |
| 34021 | 284 | (** serializer usage **) | 
| 285 | ||
| 286 | (* montage *) | |
| 287 | ||
| 55757 | 288 | fun the_fundamental ctxt = | 
| 34021 | 289 | let | 
| 55757 | 290 | val (targets, _) = Targets.get (Proof_Context.theory_of ctxt); | 
| 37822 
cf3588177676
use generic description slot for formal code checking
 haftmann parents: 
37821diff
changeset | 291 | fun fundamental target = case Symtab.lookup targets target | 
| 37821 | 292 | of SOME data => (case the_description data | 
| 37822 
cf3588177676
use generic description slot for formal code checking
 haftmann parents: 
37821diff
changeset | 293 | of Fundamental data => data | 
| 
cf3588177676
use generic description slot for formal code checking
 haftmann parents: 
37821diff
changeset | 294 | | Extension (super, _) => fundamental super) | 
| 34021 | 295 |       | NONE => error ("Unknown code target language: " ^ quote target);
 | 
| 37822 
cf3588177676
use generic description slot for formal code checking
 haftmann parents: 
37821diff
changeset | 296 | in fundamental end; | 
| 
cf3588177676
use generic description slot for formal code checking
 haftmann parents: 
37821diff
changeset | 297 | |
| 55757 | 298 | fun the_literals ctxt = #literals o the_fundamental ctxt; | 
| 34021 | 299 | |
| 55757 | 300 | fun collapse_hierarchy ctxt = | 
| 38927 | 301 | let | 
| 55757 | 302 | val (targets, _) = Targets.get (Proof_Context.theory_of ctxt); | 
| 43564 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43324diff
changeset | 303 | fun collapse target = | 
| 38927 | 304 | let | 
| 305 | val data = case Symtab.lookup targets target | |
| 306 | of SOME data => data | |
| 307 |           | NONE => error ("Unknown code target language: " ^ quote target);
 | |
| 308 | in case the_description data | |
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55146diff
changeset | 309 | of Fundamental _ => (I, data) | 
| 38927 | 310 | | Extension (super, modify) => let | 
| 39817 
5228c6b20273
check whole target hierarchy for existing reserved symbols
 haftmann parents: 
39750diff
changeset | 311 | val (modify', data') = collapse super | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55146diff
changeset | 312 | in (modify' #> modify, merge_target false target (data', data)) end | 
| 38927 | 313 | end; | 
| 39817 
5228c6b20273
check whole target hierarchy for existing reserved symbols
 haftmann parents: 
39750diff
changeset | 314 | in collapse end; | 
| 
5228c6b20273
check whole target hierarchy for existing reserved symbols
 haftmann parents: 
39750diff
changeset | 315 | |
| 
5228c6b20273
check whole target hierarchy for existing reserved symbols
 haftmann parents: 
39750diff
changeset | 316 | local | 
| 
5228c6b20273
check whole target hierarchy for existing reserved symbols
 haftmann parents: 
39750diff
changeset | 317 | |
| 55757 | 318 | fun activate_target ctxt target = | 
| 39817 
5228c6b20273
check whole target hierarchy for existing reserved symbols
 haftmann parents: 
39750diff
changeset | 319 | let | 
| 55757 | 320 | val thy = Proof_Context.theory_of ctxt; | 
| 54890 
cb892d835803
fundamental treatment of undefined vs. universally partial replaces code_abort
 haftmann parents: 
54889diff
changeset | 321 | val (_, default_width) = Targets.get thy; | 
| 55757 | 322 | val (modify, data) = collapse_hierarchy ctxt target; | 
| 54890 
cb892d835803
fundamental treatment of undefined vs. universally partial replaces code_abort
 haftmann parents: 
54889diff
changeset | 323 | in (default_width, data, modify) end; | 
| 38927 | 324 | |
| 55757 | 325 | fun project_program ctxt syms_hidden syms1 program2 = | 
| 38927 | 326 | let | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55146diff
changeset | 327 | val syms2 = subtract (op =) syms_hidden syms1; | 
| 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55146diff
changeset | 328 | 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 | 329 | 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 | 330 | val unimplemented = Code_Thingol.unimplemented program3; | 
| 42359 | 331 | val _ = | 
| 54889 
4121d64fde90
explicit distinction between empty code equations and no code equations, including convenient declaration attributes
 haftmann parents: 
54312diff
changeset | 332 | if null unimplemented then () | 
| 42359 | 333 |       else error ("No code equations for " ^
 | 
| 55304 | 334 | commas (map (Proof_Context.markup_const ctxt) unimplemented)); | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55146diff
changeset | 335 | 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 | 336 | in (syms4, program4) end; | 
| 38927 | 337 | |
| 55757 | 338 | fun prepare_serializer ctxt (serializer : serializer) reserved identifiers | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55146diff
changeset | 339 | printings module_name args proto_program syms = | 
| 38927 | 340 | let | 
| 55149 
626d8f08d479
immediate "activation" of const syntax at declaration time
 haftmann parents: 
55147diff
changeset | 341 | val syms_hidden = Code_Symbol.symbols_of printings; | 
| 55757 | 342 | 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 | 343 | fun select_include (name, (content, cs)) = | 
| 55150 | 344 | 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 | 345 | 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 | 346 | val includes = map_filter select_include (Code_Symbol.dest_module_data printings); | 
| 34021 | 347 | in | 
| 55757 | 348 |     (serializer args ctxt {
 | 
| 52138 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 349 | module_name = module_name, | 
| 38926 | 350 | reserved_syms = reserved, | 
| 52138 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 351 | identifiers = identifiers, | 
| 38926 | 352 | includes = includes, | 
| 55149 
626d8f08d479
immediate "activation" of const syntax at declaration time
 haftmann parents: 
55147diff
changeset | 353 | const_syntax = Code_Symbol.lookup_constant_data printings, | 
| 
626d8f08d479
immediate "activation" of const syntax at declaration time
 haftmann parents: 
55147diff
changeset | 354 | tyco_syntax = Code_Symbol.lookup_type_constructor_data printings, | 
| 
626d8f08d479
immediate "activation" of const syntax at declaration time
 haftmann parents: 
55147diff
changeset | 355 | 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 | 356 | (subtract (op =) syms_hidden syms, program)) | 
| 34021 | 357 | end; | 
| 358 | ||
| 55757 | 359 | fun mount_serializer ctxt target some_width module_name args program syms = | 
| 34021 | 360 | let | 
| 55757 | 361 | val (default_width, data, modify) = activate_target ctxt target; | 
| 38921 | 362 | val serializer = case the_description data | 
| 38927 | 363 | of Fundamental seri => #serializer seri; | 
| 55683 | 364 | val (prepared_serializer, (prepared_syms, prepared_program)) = | 
| 55757 | 365 | prepare_serializer ctxt serializer | 
| 52138 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 366 | (the_reserved data) (the_identifiers data) (the_printings data) | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55146diff
changeset | 367 | module_name args (modify program) syms | 
| 34071 | 368 | val width = the_default default_width some_width; | 
| 55683 | 369 | 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 | 370 | |
| 55757 | 371 | fun invoke_serializer ctxt target some_width raw_module_name args program all_public syms = | 
| 41344 
d990badc97a3
evaluator separating static and dynamic operations
 haftmann parents: 
41342diff
changeset | 372 | let | 
| 55291 | 373 | val module_name = if raw_module_name = "" then "" | 
| 374 | else (check_name true raw_module_name; raw_module_name) | |
| 55757 | 375 | val (mounted_serializer, (prepared_syms, prepared_program)) = | 
| 376 | mount_serializer ctxt target some_width module_name args program syms; | |
| 55776 
7dd1971b39c1
amended some slips, rolling back currently dysfunctional export minimimalisation for Scala
 haftmann parents: 
55757diff
changeset | 377 | in mounted_serializer prepared_program (if all_public then [] else prepared_syms) end; | 
| 34021 | 378 | |
| 52138 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 379 | 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 | 380 | | assert_module_name module_name = module_name; | 
| 
bd77e092f67c
avoid strange special treatment of empty module names
 haftmann parents: 
38929diff
changeset | 381 | |
| 55757 | 382 | fun using_master_directory ctxt = | 
| 56208 | 383 | Option.map (Path.append (File.pwd ()) o | 
| 384 | Path.append (Resources.master_directory (Proof_Context.theory_of ctxt))); | |
| 48371 | 385 | |
| 34021 | 386 | in | 
| 387 | ||
| 48568 | 388 | val generatedN = "Generated_Code"; | 
| 389 | ||
| 55757 | 390 | fun export_code_for ctxt some_path target some_width module_name args = | 
| 391 | export (using_master_directory ctxt some_path) | |
| 392 | ooo invoke_serializer ctxt target some_width module_name args; | |
| 38918 | 393 | |
| 55757 | 394 | fun produce_code_for ctxt target some_width module_name args = | 
| 39102 | 395 | let | 
| 55757 | 396 | val serializer = invoke_serializer ctxt target some_width (assert_module_name module_name) args; | 
| 55683 | 397 | in fn program => fn all_public => fn syms => | 
| 398 | 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 | 399 | end; | 
| 38929 
d9ac9dee764d
distinguish code production and code presentation
 haftmann parents: 
38928diff
changeset | 400 | |
| 55757 | 401 | fun present_code_for ctxt target some_width module_name args = | 
| 39484 
505f95975a5a
closures separate serializer initialization from serializer invocation as far as appropriate
 haftmann parents: 
39480diff
changeset | 402 | let | 
| 55757 | 403 | val serializer = invoke_serializer ctxt target some_width (assert_module_name module_name) args; | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55146diff
changeset | 404 | in fn program => fn (syms, selects) => | 
| 55683 | 405 | present selects (serializer program false syms) | 
| 39484 
505f95975a5a
closures separate serializer initialization from serializer invocation as far as appropriate
 haftmann parents: 
39480diff
changeset | 406 | end; | 
| 38918 | 407 | |
| 55757 | 408 | fun check_code_for ctxt target strict args program all_public syms = | 
| 37824 | 409 | let | 
| 410 |     val { env_var, make_destination, make_command } =
 | |
| 55757 | 411 | (#check o the_fundamental ctxt) target; | 
| 41939 | 412 | fun ext_check p = | 
| 43564 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43324diff
changeset | 413 | let | 
| 37824 | 414 | val destination = make_destination p; | 
| 55757 | 415 | val _ = export (SOME destination) (invoke_serializer ctxt target (SOME 80) | 
| 55683 | 416 | generatedN args program all_public syms); | 
| 48568 | 417 | val cmd = make_command generatedN; | 
| 43850 
7f2cbc713344
moved bash operations to Isabelle_System (cf. Scala version);
 wenzelm parents: 
43564diff
changeset | 418 | in | 
| 
7f2cbc713344
moved bash operations to Isabelle_System (cf. Scala version);
 wenzelm parents: 
43564diff
changeset | 419 |         if Isabelle_System.bash ("cd " ^ File.shell_path p ^ " && " ^ cmd ^ " 2>&1") <> 0
 | 
| 37824 | 420 |         then error ("Code check failed for " ^ target ^ ": " ^ cmd)
 | 
| 421 | else () | |
| 422 | end; | |
| 43850 
7f2cbc713344
moved bash operations to Isabelle_System (cf. Scala version);
 wenzelm parents: 
43564diff
changeset | 423 | in | 
| 
7f2cbc713344
moved bash operations to Isabelle_System (cf. Scala version);
 wenzelm parents: 
43564diff
changeset | 424 | if getenv env_var = "" | 
| 37825 | 425 | then if strict | 
| 426 | then error (env_var ^ " not set; cannot check code for " ^ target) | |
| 427 | else warning (env_var ^ " not set; skipped checking code for " ^ target) | |
| 41939 | 428 | else Isabelle_System.with_tmp_dir "Code_Test" ext_check | 
| 37824 | 429 | end; | 
| 430 | ||
| 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 | 431 | fun subevaluator mounted_serializer prepared_program syms all_public ((vs, ty), t) = | 
| 41344 
d990badc97a3
evaluator separating static and dynamic operations
 haftmann parents: 
41342diff
changeset | 432 | let | 
| 
d990badc97a3
evaluator separating static and dynamic operations
 haftmann parents: 
41342diff
changeset | 433 | val _ = if Code_Thingol.contains_dict_var t then | 
| 
d990badc97a3
evaluator separating static and dynamic operations
 haftmann parents: 
41342diff
changeset | 434 | 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 | 435 | val v' = singleton (Name.variant_list (map fst vs)) "a"; | 
| 41344 
d990badc97a3
evaluator separating static and dynamic operations
 haftmann parents: 
41342diff
changeset | 436 | val vs' = (v', []) :: vs; | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55146diff
changeset | 437 | val ty' = ITyVar v' `-> ty; | 
| 41344 
d990badc97a3
evaluator separating static and dynamic operations
 haftmann parents: 
41342diff
changeset | 438 | val program = prepared_program | 
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55146diff
changeset | 439 | |> 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 | 440 | 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 | 441 | |> fold (curry (perhaps o try o | 
| 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55146diff
changeset | 442 | Code_Symbol.Graph.add_edge) Code_Symbol.value) syms; | 
| 55683 | 443 | val (program_code, deresolve) = | 
| 55684 
ee49b4f7edc8
keep only identifiers public which are explicitly requested or demanded by dependencies
 haftmann parents: 
55683diff
changeset | 444 | 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 | 445 | val value_name = the (deresolve Code_Symbol.value); | 
| 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55146diff
changeset | 446 | in (program_code, value_name) end; | 
| 41344 
d990badc97a3
evaluator separating static and dynamic operations
 haftmann parents: 
41342diff
changeset | 447 | |
| 55757 | 448 | fun evaluator ctxt target program syms = | 
| 41344 
d990badc97a3
evaluator separating static and dynamic operations
 haftmann parents: 
41342diff
changeset | 449 | let | 
| 55683 | 450 | val (mounted_serializer, (_, prepared_program)) = | 
| 55757 | 451 | mount_serializer ctxt target 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 | 452 | in subevaluator mounted_serializer prepared_program syms end; | 
| 41344 
d990badc97a3
evaluator separating static and dynamic operations
 haftmann parents: 
41342diff
changeset | 453 | |
| 34021 | 454 | end; (* local *) | 
| 455 | ||
| 456 | ||
| 457 | (* code generation *) | |
| 458 | ||
| 38918 | 459 | fun prep_destination "" = NONE | 
| 460 | | prep_destination s = SOME (Path.explode s); | |
| 461 | ||
| 55757 | 462 | fun export_code ctxt all_public cs seris = | 
| 34021 | 463 | let | 
| 55757 | 464 | val thy = Proof_Context.theory_of ctxt; | 
| 55188 
7ca0204ece66
reduced prominence of "permissive code generation"
 haftmann parents: 
55153diff
changeset | 465 | val program = Code_Thingol.consts_program thy cs; | 
| 38918 | 466 | val _ = map (fn (((target, module_name), some_path), args) => | 
| 55757 | 467 | export_code_for ctxt some_path target NONE module_name args program all_public (map Constant cs)) seris; | 
| 34021 | 468 | in () end; | 
| 469 | ||
| 55757 | 470 | fun export_code_cmd all_public raw_cs seris ctxt = | 
| 471 | export_code ctxt all_public | |
| 472 | (Code_Thingol.read_const_exprs ctxt raw_cs) | |
| 55683 | 473 | ((map o apfst o apsnd) prep_destination seris); | 
| 38918 | 474 | |
| 55757 | 475 | fun produce_code ctxt all_public cs target some_width some_module_name args = | 
| 38918 | 476 | let | 
| 55757 | 477 | val thy = Proof_Context.theory_of ctxt; | 
| 55188 
7ca0204ece66
reduced prominence of "permissive code generation"
 haftmann parents: 
55153diff
changeset | 478 | val program = Code_Thingol.consts_program thy cs; | 
| 55757 | 479 | in produce_code_for ctxt target 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 | 480 | |
| 55757 | 481 | fun present_code ctxt cs syms target some_width some_module_name args = | 
| 38929 
d9ac9dee764d
distinguish code production and code presentation
 haftmann parents: 
38928diff
changeset | 482 | let | 
| 55757 | 483 | val thy = Proof_Context.theory_of ctxt; | 
| 55188 
7ca0204ece66
reduced prominence of "permissive code generation"
 haftmann parents: 
55153diff
changeset | 484 | val program = Code_Thingol.consts_program thy cs; | 
| 55757 | 485 | in present_code_for ctxt target some_width some_module_name args program (map Constant cs, syms) end; | 
| 34021 | 486 | |
| 55757 | 487 | fun check_code ctxt all_public cs seris = | 
| 37824 | 488 | let | 
| 55757 | 489 | val thy = Proof_Context.theory_of ctxt; | 
| 55188 
7ca0204ece66
reduced prominence of "permissive code generation"
 haftmann parents: 
55153diff
changeset | 490 | val program = Code_Thingol.consts_program thy cs; | 
| 38918 | 491 | val _ = map (fn ((target, strict), args) => | 
| 55757 | 492 | check_code_for ctxt target strict args program all_public (map Constant cs)) seris; | 
| 37824 | 493 | in () end; | 
| 494 | ||
| 55757 | 495 | fun check_code_cmd all_public raw_cs seris ctxt = | 
| 496 | check_code ctxt all_public | |
| 497 | (Code_Thingol.read_const_exprs ctxt raw_cs) seris; | |
| 37824 | 498 | |
| 39480 
a2ed61449dcc
added code_stmts antiquotation from doc-src/more_antiquote.ML
 haftmann parents: 
39142diff
changeset | 499 | local | 
| 
a2ed61449dcc
added code_stmts antiquotation from doc-src/more_antiquote.ML
 haftmann parents: 
39142diff
changeset | 500 | |
| 
a2ed61449dcc
added code_stmts antiquotation from doc-src/more_antiquote.ML
 haftmann parents: 
39142diff
changeset | 501 | val parse_const_terms = Scan.repeat1 Args.term | 
| 55757 | 502 | >> (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 | 503 | |
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
55146diff
changeset | 504 | fun parse_names category parse internalize mark_symbol = | 
| 39480 
a2ed61449dcc
added code_stmts antiquotation from doc-src/more_antiquote.ML
 haftmann parents: 
39142diff
changeset | 505 | Scan.lift (Args.parens (Args.$$$ category)) |-- Scan.repeat1 parse | 
| 55757 | 506 | >> (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 | 507 | |
| 39480 
a2ed61449dcc
added code_stmts antiquotation from doc-src/more_antiquote.ML
 haftmann parents: 
39142diff
changeset | 508 | val parse_consts = parse_names "consts" Args.term | 
| 55757 | 509 | (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 | 510 | |
| 
a2ed61449dcc
added code_stmts antiquotation from doc-src/more_antiquote.ML
 haftmann parents: 
39142diff
changeset | 511 | val parse_types = parse_names "types" (Scan.lift Args.name) | 
| 55757 | 512 | (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 | 513 | |
| 
a2ed61449dcc
added code_stmts antiquotation from doc-src/more_antiquote.ML
 haftmann parents: 
39142diff
changeset | 514 | val parse_classes = parse_names "classes" (Scan.lift Args.name) | 
| 55757 | 515 | (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 | 516 | |
| 
a2ed61449dcc
added code_stmts antiquotation from doc-src/more_antiquote.ML
 haftmann parents: 
39142diff
changeset | 517 | val parse_instances = parse_names "instances" (Scan.lift (Args.name --| Args.$$$ "::" -- Args.name)) | 
| 55757 | 518 | (fn ctxt => fn (raw_tyco, raw_class) => | 
| 519 | let | |
| 520 | val thy = Proof_Context.theory_of ctxt; | |
| 521 | 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 | 522 | |
| 
a2ed61449dcc
added code_stmts antiquotation from doc-src/more_antiquote.ML
 haftmann parents: 
39142diff
changeset | 523 | in | 
| 
a2ed61449dcc
added code_stmts antiquotation from doc-src/more_antiquote.ML
 haftmann parents: 
39142diff
changeset | 524 | |
| 43564 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43324diff
changeset | 525 | val antiq_setup = | 
| 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43324diff
changeset | 526 |   Thy_Output.antiquotation @{binding code_stmts}
 | 
| 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43324diff
changeset | 527 | (parse_const_terms -- | 
| 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43324diff
changeset | 528 | 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 | 529 | -- Scan.lift (Args.parens (Args.name -- Scan.option Parse.int))) | 
| 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43324diff
changeset | 530 |     (fn {context = ctxt, ...} => fn ((mk_cs, mk_stmtss), (target, some_width)) =>
 | 
| 55757 | 531 | present_code ctxt (mk_cs ctxt) | 
| 532 | (maps (fn f => f ctxt) mk_stmtss) | |
| 533 | target some_width "Example" []); | |
| 39480 
a2ed61449dcc
added code_stmts antiquotation from doc-src/more_antiquote.ML
 haftmann parents: 
39142diff
changeset | 534 | |
| 
a2ed61449dcc
added code_stmts antiquotation from doc-src/more_antiquote.ML
 haftmann parents: 
39142diff
changeset | 535 | end; | 
| 
a2ed61449dcc
added code_stmts antiquotation from doc-src/more_antiquote.ML
 haftmann parents: 
39142diff
changeset | 536 | |
| 34021 | 537 | |
| 28054 | 538 | (** serializer configuration **) | 
| 27000 | 539 | |
| 52138 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 540 | (* reserved symbol names *) | 
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 541 | |
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 542 | fun add_reserved target sym thy = | 
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 543 | let | 
| 55757 | 544 | val (_, data) = collapse_hierarchy (Proof_Context.init_global thy) target; | 
| 52138 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 545 | val _ = if member (op =) (the_reserved data) sym | 
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 546 |       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 | 547 | else (); | 
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 548 | in | 
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 549 | thy | 
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 550 | |> map_reserved target (insert (op =) sym) | 
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 551 | end; | 
| 
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 | |
| 52377 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 haftmann parents: 
52218diff
changeset | 554 | (* checking of syntax *) | 
| 52138 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 555 | |
| 55757 | 556 | fun check_const_syntax ctxt target c syn = | 
| 557 | 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 | 558 |   then error ("Too many arguments in syntax for constant " ^ quote c)
 | 
| 55757 | 559 | else Code_Printer.prep_const_syntax (Proof_Context.theory_of ctxt) (the_literals ctxt target) c syn; | 
| 52137 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 haftmann parents: 
52068diff
changeset | 560 | |
| 55757 | 561 | fun check_tyco_syntax ctxt target tyco syn = | 
| 562 | 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 | 563 |   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 | 564 | else syn; | 
| 34071 | 565 | |
| 52138 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 566 | |
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 567 | (* custom symbol names *) | 
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 568 | |
| 52218 | 569 | 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 | 570 | let | 
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 571 | 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 | 572 | (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 | 573 | in | 
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 574 | Code_Symbol.maps_attr' (arrange false) (arrange false) (arrange false) | 
| 52218 | 575 | (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 | 576 | end; | 
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 577 | |
| 55757 | 578 | 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 | 579 | |
| 55757 | 580 | 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 | 581 | |
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 582 | fun set_identifier (target, sym_name) = map_identifiers target (Code_Symbol.set_data sym_name); | 
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 583 | |
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 584 | fun gen_set_identifiers prep_name_decl raw_name_decls thy = | 
| 55757 | 585 | 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 | 586 | |
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 587 | 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 | 588 | 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 | 589 | |
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 590 | |
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 591 | (* custom printings *) | 
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 592 | |
| 55757 | 593 | 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 | 594 | let | 
| 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 haftmann parents: 
52068diff
changeset | 595 | fun arrange check (sym, target_syns) = | 
| 55757 | 596 | map (fn (target, some_syn) => (target, (sym, Option.map (check ctxt target 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 | 597 | in | 
| 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 haftmann parents: 
52068diff
changeset | 598 | Code_Symbol.maps_attr' | 
| 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 haftmann parents: 
52068diff
changeset | 599 | (arrange check_const_syntax) (arrange check_tyco_syntax) | 
| 55149 
626d8f08d479
immediate "activation" of const syntax at declaration time
 haftmann parents: 
55147diff
changeset | 600 | (arrange ((K o K o K) I)) (arrange ((K o K o K) I)) (arrange ((K o K o K) I)) | 
| 55757 | 601 | (arrange (fn ctxt => fn _ => fn _ => fn (raw_content, raw_cs) => | 
| 602 | (Code_Printer.str raw_content, 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 | 603 | end; | 
| 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 haftmann parents: 
52068diff
changeset | 604 | |
| 55757 | 605 | fun cert_printings ctxt = cert_syms ctxt #> arrange_printings cert_const ctxt; | 
| 24219 | 606 | |
| 55757 | 607 | 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 | 608 | |
| 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 haftmann parents: 
52068diff
changeset | 609 | fun set_printing (target, sym_syn) = map_printings target (Code_Symbol.set_data sym_syn); | 
| 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 haftmann parents: 
52068diff
changeset | 610 | |
| 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 haftmann parents: 
52068diff
changeset | 611 | fun gen_set_printings prep_print_decl raw_print_decls thy = | 
| 55757 | 612 | 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 | 613 | |
| 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 haftmann parents: 
52068diff
changeset | 614 | 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 | 615 | 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 | 616 | |
| 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 haftmann parents: 
52068diff
changeset | 617 | |
| 27103 | 618 | (* concrete syntax *) | 
| 27000 | 619 | |
| 28054 | 620 | 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 | 621 | case Scan.read Token.stopper f args | 
| 28054 | 622 | of SOME x => x | 
| 623 | | NONE => error "Bad serializer arguments"; | |
| 624 | ||
| 625 | ||
| 27304 | 626 | (** Isar setup **) | 
| 27103 | 627 | |
| 52137 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 haftmann parents: 
52068diff
changeset | 628 | fun parse_single_symbol_pragma parse_keyword parse_isa parse_target = | 
| 52434 | 629 |   parse_keyword |-- Parse.!!! (parse_isa --| (@{keyword "\<rightharpoonup>"} || @{keyword "=>"})
 | 
| 630 |     -- 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 | 631 | |
| 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 haftmann parents: 
52068diff
changeset | 632 | fun parse_symbol_pragma parse_const parse_tyco parse_class parse_classrel parse_inst parse_module = | 
| 52801 | 633 |   parse_single_symbol_pragma @{keyword "constant"} Parse.term parse_const
 | 
| 55150 | 634 | >> Constant | 
| 52377 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 haftmann parents: 
52218diff
changeset | 635 |   || parse_single_symbol_pragma @{keyword "type_constructor"} Parse.type_const parse_tyco
 | 
| 55150 | 636 | >> Type_Constructor | 
| 52377 
afa72aaed518
more consistent parsing and reading of classes and type constructors
 haftmann parents: 
52218diff
changeset | 637 |   || parse_single_symbol_pragma @{keyword "type_class"} Parse.class parse_class
 | 
| 55150 | 638 | >> Type_Class | 
| 52137 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 haftmann parents: 
52068diff
changeset | 639 |   || parse_single_symbol_pragma @{keyword "class_relation"} parse_classrel_ident parse_classrel
 | 
| 55150 | 640 | >> Class_Relation | 
| 52137 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 haftmann parents: 
52068diff
changeset | 641 |   || parse_single_symbol_pragma @{keyword "class_instance"} parse_inst_ident parse_inst
 | 
| 55150 | 642 | >> Class_Instance | 
| 52137 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 haftmann parents: 
52068diff
changeset | 643 |   || 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 | 644 | >> Code_Symbol.Module; | 
| 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 haftmann parents: 
52068diff
changeset | 645 | |
| 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 haftmann parents: 
52068diff
changeset | 646 | 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 | 647 | 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 | 648 | (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 | 649 | |
| 56201 | 650 | 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 | 651 | |
| 55683 | 652 | fun code_expr_inP all_public raw_cs = | 
| 52434 | 653 |   Scan.repeat (@{keyword "in"} |-- Parse.!!! (Parse.name
 | 
| 654 |     -- Scan.optional (@{keyword "module_name"} |-- Parse.name) ""
 | |
| 655 |     -- Scan.optional (@{keyword "file"} |-- Parse.name) ""
 | |
| 656 | -- code_expr_argsP)) | |
| 55683 | 657 | >> (fn seri_args => export_code_cmd all_public raw_cs seri_args); | 
| 52434 | 658 | |
| 55683 | 659 | fun code_expr_checkingP all_public raw_cs = | 
| 52434 | 660 |   (@{keyword "checking"} |-- Parse.!!!
 | 
| 661 |     (Scan.repeat (Parse.name -- ((@{keyword "?"} |-- Scan.succeed false) || Scan.succeed true)
 | |
| 662 | -- code_expr_argsP))) | |
| 55683 | 663 | >> (fn seri_args => check_code_cmd all_public raw_cs seri_args); | 
| 52434 | 664 | |
| 55683 | 665 | val code_exprP = (Scan.optional (@{keyword "open"} |-- Scan.succeed true) false
 | 
| 666 | -- Scan.repeat1 Parse.term) | |
| 667 | :|-- (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 | 668 | |
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 669 | val _ = | 
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 670 |   Outer_Syntax.command @{command_spec "code_reserved"}
 | 
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 671 | "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 | 672 | (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 | 673 | >> (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 | 674 | |
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 675 | val _ = | 
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 676 |   Outer_Syntax.command @{command_spec "code_identifier"} "declare mandatory names for code symbols"
 | 
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 677 | (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 | 678 | >> (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 | 679 | |
| 
e21426f244aa
bookkeeping and input syntax for exact specification of names of symbols in generated code
 haftmann parents: 
52137diff
changeset | 680 | val _ = | 
| 52137 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 haftmann parents: 
52068diff
changeset | 681 |   Outer_Syntax.command @{command_spec "code_printing"} "declare dedicated printing for code symbols"
 | 
| 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 haftmann parents: 
52068diff
changeset | 682 | (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 | 683 | 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 | 684 |       (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 | 685 | >> (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 | 686 | |
| 
7f7337447b1b
use generic data for code symbols for unified "code_printing" syntax for custom serialisations
 haftmann parents: 
52068diff
changeset | 687 | val _ = | 
| 46961 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 wenzelm parents: 
46949diff
changeset | 688 |   Outer_Syntax.command @{command_spec "export_code"} "generate executable code for constants"
 | 
| 55757 | 689 | (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 | 690 | |
| 39646 | 691 | |
| 692 | (** external entrance point -- for codegen tool **) | |
| 693 | ||
| 39750 
c0099428ca7b
consider quick_and_dirty option before loading theory
 haftmann parents: 
39679diff
changeset | 694 | fun codegen_tool thyname cmd_expr = | 
| 39646 | 695 | let | 
| 55757 | 696 | val ctxt = Proof_Context.init_global (Thy_Info.get_theory thyname); | 
| 39646 | 697 | val parse = Scan.read Token.stopper (Parse.!!! code_exprP) o | 
| 57918 
f5d73caba4e5
tuned signature according to Scala version -- prefer explicit argument;
 wenzelm parents: 
57249diff
changeset | 698 | (filter Token.is_proper o Outer_Syntax.scan (Keyword.get_lexicons ()) Position.none); | 
| 39646 | 699 | in case parse cmd_expr | 
| 55757 | 700 | of SOME f => (writeln "Now generating code..."; f ctxt) | 
| 39646 | 701 |     | NONE => error ("Bad directive " ^ quote cmd_expr)
 | 
| 702 | end; | |
| 703 | ||
| 43564 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43324diff
changeset | 704 | |
| 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43324diff
changeset | 705 | (** theory setup **) | 
| 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43324diff
changeset | 706 | |
| 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43324diff
changeset | 707 | val setup = antiq_setup; | 
| 
9864182c6bad
document antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
43324diff
changeset | 708 | |
| 24219 | 709 | end; (*struct*) |