| author | bulwahn | 
| Fri, 21 Oct 2011 14:25:38 +0200 | |
| changeset 45236 | ac4a2a66707d | 
| parent 44802 | 65c397cc44ec | 
| child 45423 | 92f91f990165 | 
| permissions | -rw-r--r-- | 
| 18 | 1 | (* Title: Pure/Syntax/syntax.ML | 
| 0 | 2 | Author: Tobias Nipkow and Markus Wenzel, TU Muenchen | 
| 18 | 3 | |
| 24263 | 4 | Standard Isabelle syntax, based on arbitrary context-free grammars | 
| 5 | (specified by mixfix declarations). | |
| 0 | 6 | *) | 
| 7 | ||
| 8 | signature SYNTAX = | |
| 2383 | 9 | sig | 
| 42382 
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
 wenzelm parents: 
42360diff
changeset | 10 | type operations | 
| 
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
 wenzelm parents: 
42360diff
changeset | 11 | val install_operations: operations -> unit | 
| 42288 
2074b31650e6
discontinued special treatment of structure Syntax_Ext (formerly Syn_Ext);
 wenzelm parents: 
42287diff
changeset | 12 | val root: string Config.T | 
| 42268 | 13 | val positions_raw: Config.raw | 
| 14 | val positions: bool Config.T | |
| 15 | val ambiguity_enabled: bool Config.T | |
| 16 | val ambiguity_level_raw: Config.raw | |
| 17 | val ambiguity_level: int Config.T | |
| 18 | val ambiguity_limit: int Config.T | |
| 44069 | 19 | val ambiguity_warnings_raw: Config.raw | 
| 20 | val ambiguity_warnings: bool Config.T | |
| 30573 | 21 | val read_token: string -> Symbol_Pos.T list * Position.T | 
| 43731 
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
 wenzelm parents: 
43558diff
changeset | 22 | val parse_token: Proof.context -> (XML.tree list -> 'a) -> | 
| 
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
 wenzelm parents: 
43558diff
changeset | 23 | Markup.T -> (Symbol_Pos.T list * Position.T -> 'a) -> string -> 'a | 
| 24263 | 24 | val parse_sort: Proof.context -> string -> sort | 
| 25 | val parse_typ: Proof.context -> string -> typ | |
| 26 | val parse_term: Proof.context -> string -> term | |
| 27 | val parse_prop: Proof.context -> string -> term | |
| 24768 | 28 | val unparse_sort: Proof.context -> sort -> Pretty.T | 
| 24923 | 29 | val unparse_classrel: Proof.context -> class list -> Pretty.T | 
| 30 | val unparse_arity: Proof.context -> arity -> Pretty.T | |
| 24768 | 31 | val unparse_typ: Proof.context -> typ -> Pretty.T | 
| 32 | val unparse_term: Proof.context -> term -> Pretty.T | |
| 33 | val print_checks: Proof.context -> unit | |
| 42402 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42383diff
changeset | 34 | val context_typ_check: int -> string -> | 
| 25060 | 35 | (typ list -> Proof.context -> (typ list * Proof.context) option) -> | 
| 24768 | 36 | Context.generic -> Context.generic | 
| 42402 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42383diff
changeset | 37 | val context_term_check: int -> string -> | 
| 25060 | 38 | (term list -> Proof.context -> (term list * Proof.context) option) -> | 
| 24768 | 39 | Context.generic -> Context.generic | 
| 42402 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42383diff
changeset | 40 | val context_typ_uncheck: int -> string -> | 
| 25060 | 41 | (typ list -> Proof.context -> (typ list * Proof.context) option) -> | 
| 24768 | 42 | Context.generic -> Context.generic | 
| 42402 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42383diff
changeset | 43 | val context_term_uncheck: int -> string -> | 
| 25060 | 44 | (term list -> Proof.context -> (term list * Proof.context) option) -> | 
| 24768 | 45 | Context.generic -> Context.generic | 
| 42402 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42383diff
changeset | 46 | val add_typ_check: int -> string -> (Proof.context -> typ list -> typ list) -> | 
| 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42383diff
changeset | 47 | Context.generic -> Context.generic | 
| 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42383diff
changeset | 48 | val add_term_check: int -> string -> (Proof.context -> term list -> term list) -> | 
| 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42383diff
changeset | 49 | Context.generic -> Context.generic | 
| 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42383diff
changeset | 50 | val add_typ_uncheck: int -> string -> (Proof.context -> typ list -> typ list) -> | 
| 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42383diff
changeset | 51 | Context.generic -> Context.generic | 
| 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42383diff
changeset | 52 | val add_term_uncheck: int -> string -> (Proof.context -> term list -> term list) -> | 
| 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42383diff
changeset | 53 | Context.generic -> Context.generic | 
| 42382 
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
 wenzelm parents: 
42360diff
changeset | 54 | val typ_check: Proof.context -> typ list -> typ list | 
| 
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
 wenzelm parents: 
42360diff
changeset | 55 | val term_check: Proof.context -> term list -> term list | 
| 
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
 wenzelm parents: 
42360diff
changeset | 56 | val typ_uncheck: Proof.context -> typ list -> typ list | 
| 
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
 wenzelm parents: 
42360diff
changeset | 57 | val term_uncheck: Proof.context -> term list -> term list | 
| 24488 | 58 | val check_sort: Proof.context -> sort -> sort | 
| 24512 | 59 | val check_typ: Proof.context -> typ -> typ | 
| 60 | val check_term: Proof.context -> term -> term | |
| 61 | val check_prop: Proof.context -> term -> term | |
| 24488 | 62 | val check_typs: Proof.context -> typ list -> typ list | 
| 24263 | 63 | val check_terms: Proof.context -> term list -> term list | 
| 64 | val check_props: Proof.context -> term list -> term list | |
| 24923 | 65 | val uncheck_sort: Proof.context -> sort -> sort | 
| 66 | val uncheck_arity: Proof.context -> arity -> arity | |
| 67 | val uncheck_classrel: Proof.context -> class list -> class list | |
| 24768 | 68 | val uncheck_typs: Proof.context -> typ list -> typ list | 
| 69 | val uncheck_terms: Proof.context -> term list -> term list | |
| 24263 | 70 | val read_sort: Proof.context -> string -> sort | 
| 71 | val read_typ: Proof.context -> string -> typ | |
| 24488 | 72 | val read_term: Proof.context -> string -> term | 
| 73 | val read_prop: Proof.context -> string -> term | |
| 24263 | 74 | val read_terms: Proof.context -> string list -> term list | 
| 75 | val read_props: Proof.context -> string list -> term list | |
| 24709 | 76 | val read_sort_global: theory -> string -> sort | 
| 77 | val read_typ_global: theory -> string -> typ | |
| 78 | val read_term_global: theory -> string -> term | |
| 79 | val read_prop_global: theory -> string -> term | |
| 24923 | 80 | val pretty_term: Proof.context -> term -> Pretty.T | 
| 81 | val pretty_typ: Proof.context -> typ -> Pretty.T | |
| 82 | val pretty_sort: Proof.context -> sort -> Pretty.T | |
| 83 | val pretty_classrel: Proof.context -> class list -> Pretty.T | |
| 84 | val pretty_arity: Proof.context -> arity -> Pretty.T | |
| 85 | val string_of_term: Proof.context -> term -> string | |
| 86 | val string_of_typ: Proof.context -> typ -> string | |
| 87 | val string_of_sort: Proof.context -> sort -> string | |
| 88 | val string_of_classrel: Proof.context -> class list -> string | |
| 89 | val string_of_arity: Proof.context -> arity -> string | |
| 26951 
030e4a818b39
moved global pretty/string_of functions from Sign to Syntax;
 wenzelm parents: 
26704diff
changeset | 90 | val is_pretty_global: Proof.context -> bool | 
| 
030e4a818b39
moved global pretty/string_of functions from Sign to Syntax;
 wenzelm parents: 
26704diff
changeset | 91 | val set_pretty_global: bool -> Proof.context -> Proof.context | 
| 
030e4a818b39
moved global pretty/string_of functions from Sign to Syntax;
 wenzelm parents: 
26704diff
changeset | 92 | val init_pretty_global: theory -> Proof.context | 
| 42383 
0ae4ad40d7b5
simplified pretty printing context, which is only required for certain kernel operations;
 wenzelm parents: 
42382diff
changeset | 93 | val init_pretty: Context.pretty -> Proof.context | 
| 26951 
030e4a818b39
moved global pretty/string_of functions from Sign to Syntax;
 wenzelm parents: 
26704diff
changeset | 94 | val pretty_term_global: theory -> term -> Pretty.T | 
| 
030e4a818b39
moved global pretty/string_of functions from Sign to Syntax;
 wenzelm parents: 
26704diff
changeset | 95 | val pretty_typ_global: theory -> typ -> Pretty.T | 
| 
030e4a818b39
moved global pretty/string_of functions from Sign to Syntax;
 wenzelm parents: 
26704diff
changeset | 96 | val pretty_sort_global: theory -> sort -> Pretty.T | 
| 
030e4a818b39
moved global pretty/string_of functions from Sign to Syntax;
 wenzelm parents: 
26704diff
changeset | 97 | val string_of_term_global: theory -> term -> string | 
| 
030e4a818b39
moved global pretty/string_of functions from Sign to Syntax;
 wenzelm parents: 
26704diff
changeset | 98 | val string_of_typ_global: theory -> typ -> string | 
| 
030e4a818b39
moved global pretty/string_of functions from Sign to Syntax;
 wenzelm parents: 
26704diff
changeset | 99 | val string_of_sort_global: theory -> sort -> string | 
| 39135 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 100 | type syntax | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 101 | val eq_syntax: syntax * syntax -> bool | 
| 44802 
65c397cc44ec
explicit join_syntax ensures command transaction integrity of 'theory';
 wenzelm parents: 
44113diff
changeset | 102 | val join_syntax: syntax -> unit | 
| 42298 
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
 wenzelm parents: 
42297diff
changeset | 103 | val lookup_const: syntax -> string -> string option | 
| 39135 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 104 | val is_keyword: syntax -> string -> bool | 
| 42251 | 105 | val tokenize: syntax -> bool -> Symbol_Pos.T list -> Lexicon.token list | 
| 106 | val parse: Proof.context -> syntax -> string -> Lexicon.token list -> Parser.parsetree list | |
| 42253 | 107 | val parse_ast_translation: syntax -> string -> (Proof.context -> Ast.ast list -> Ast.ast) option | 
| 42255 | 108 | val parse_rules: syntax -> string -> (Ast.ast * Ast.ast) list | 
| 42253 | 109 | val parse_translation: syntax -> string -> (Proof.context -> term list -> term) option | 
| 42254 | 110 | val print_translation: syntax -> string -> | 
| 111 | Proof.context -> typ -> term list -> term (*exception Match*) | |
| 42255 | 112 | val print_rules: syntax -> string -> (Ast.ast * Ast.ast) list | 
| 42254 | 113 | val print_ast_translation: syntax -> string -> | 
| 114 | Proof.context -> Ast.ast list -> Ast.ast (*exception Match*) | |
| 42255 | 115 | val prtabs: syntax -> Printer.prtabs | 
| 39135 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 116 | type mode | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 117 | val mode_default: mode | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 118 | val mode_input: mode | 
| 42294 | 119 | val empty_syntax: syntax | 
| 39135 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 120 | val merge_syntaxes: syntax -> syntax -> syntax | 
| 42294 | 121 | val token_markers: string list | 
| 39135 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 122 | val basic_nonterms: string list | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 123 | val print_gram: syntax -> unit | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 124 | val print_trans: syntax -> unit | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 125 | val print_syntax: syntax -> unit | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 126 | val guess_infix: syntax -> string -> mixfix option | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 127 | datatype 'a trrule = | 
| 42204 | 128 | Parse_Rule of 'a * 'a | | 
| 129 | Print_Rule of 'a * 'a | | |
| 130 | Parse_Print_Rule of 'a * 'a | |
| 39135 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 131 |   val map_trrule: ('a -> 'b) -> 'a trrule -> 'b trrule
 | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 132 | val update_trfuns: | 
| 42224 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 wenzelm parents: 
42223diff
changeset | 133 | (string * ((Ast.ast list -> Ast.ast) * stamp)) list * | 
| 39135 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 134 | (string * ((term list -> term) * stamp)) list * | 
| 42247 
12fe41a92cd5
typed_print_translation: discontinued show_sorts argument;
 wenzelm parents: 
42245diff
changeset | 135 | (string * ((typ -> term list -> term) * stamp)) list * | 
| 42224 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 wenzelm parents: 
42223diff
changeset | 136 | (string * ((Ast.ast list -> Ast.ast) * stamp)) list -> syntax -> syntax | 
| 39135 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 137 | val update_advanced_trfuns: | 
| 42224 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 wenzelm parents: 
42223diff
changeset | 138 | (string * ((Proof.context -> Ast.ast list -> Ast.ast) * stamp)) list * | 
| 39135 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 139 | (string * ((Proof.context -> term list -> term) * stamp)) list * | 
| 42247 
12fe41a92cd5
typed_print_translation: discontinued show_sorts argument;
 wenzelm parents: 
42245diff
changeset | 140 | (string * ((Proof.context -> typ -> term list -> term) * stamp)) list * | 
| 42224 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 wenzelm parents: 
42223diff
changeset | 141 | (string * ((Proof.context -> Ast.ast list -> Ast.ast) * stamp)) list -> syntax -> syntax | 
| 39135 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 142 | val update_type_gram: bool -> mode -> (string * typ * mixfix) list -> syntax -> syntax | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 143 | val update_const_gram: bool -> (string -> bool) -> | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 144 | mode -> (string * typ * mixfix) list -> syntax -> syntax | 
| 42224 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 wenzelm parents: 
42223diff
changeset | 145 | val update_trrules: Ast.ast trrule list -> syntax -> syntax | 
| 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 wenzelm parents: 
42223diff
changeset | 146 | val remove_trrules: Ast.ast trrule list -> syntax -> syntax | 
| 42476 | 147 | val const: string -> term | 
| 148 | val free: string -> term | |
| 149 | val var: indexname -> term | |
| 2383 | 150 | end; | 
| 0 | 151 | |
| 12094 | 152 | structure Syntax: SYNTAX = | 
| 0 | 153 | struct | 
| 154 | ||
| 42288 
2074b31650e6
discontinued special treatment of structure Syntax_Ext (formerly Syn_Ext);
 wenzelm parents: 
42287diff
changeset | 155 | |
| 39135 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 156 | (** inner syntax operations **) | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 157 | |
| 42382 
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
 wenzelm parents: 
42360diff
changeset | 158 | (* back-patched operations *) | 
| 
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
 wenzelm parents: 
42360diff
changeset | 159 | |
| 
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
 wenzelm parents: 
42360diff
changeset | 160 | type operations = | 
| 
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
 wenzelm parents: 
42360diff
changeset | 161 |  {parse_sort: Proof.context -> string -> sort,
 | 
| 
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
 wenzelm parents: 
42360diff
changeset | 162 | parse_typ: Proof.context -> string -> typ, | 
| 
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
 wenzelm parents: 
42360diff
changeset | 163 | parse_term: Proof.context -> string -> term, | 
| 
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
 wenzelm parents: 
42360diff
changeset | 164 | parse_prop: Proof.context -> string -> term, | 
| 
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
 wenzelm parents: 
42360diff
changeset | 165 | unparse_sort: Proof.context -> sort -> Pretty.T, | 
| 
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
 wenzelm parents: 
42360diff
changeset | 166 | unparse_typ: Proof.context -> typ -> Pretty.T, | 
| 
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
 wenzelm parents: 
42360diff
changeset | 167 | unparse_term: Proof.context -> term -> Pretty.T, | 
| 
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
 wenzelm parents: 
42360diff
changeset | 168 | check_typs: Proof.context -> typ list -> typ list, | 
| 
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
 wenzelm parents: 
42360diff
changeset | 169 | check_terms: Proof.context -> term list -> term list, | 
| 
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
 wenzelm parents: 
42360diff
changeset | 170 | check_props: Proof.context -> term list -> term list, | 
| 
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
 wenzelm parents: 
42360diff
changeset | 171 | uncheck_typs: Proof.context -> typ list -> typ list, | 
| 
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
 wenzelm parents: 
42360diff
changeset | 172 | uncheck_terms: Proof.context -> term list -> term list}; | 
| 
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
 wenzelm parents: 
42360diff
changeset | 173 | |
| 
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
 wenzelm parents: 
42360diff
changeset | 174 | val operations: operations Single_Assignment.var = Single_Assignment.var "Syntax.operations"; | 
| 
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
 wenzelm parents: 
42360diff
changeset | 175 | fun install_operations ops = Single_Assignment.assign operations ops; | 
| 
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
 wenzelm parents: 
42360diff
changeset | 176 | |
| 
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
 wenzelm parents: 
42360diff
changeset | 177 | fun operation which ctxt x = | 
| 
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
 wenzelm parents: 
42360diff
changeset | 178 | (case Single_Assignment.peek operations of | 
| 
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
 wenzelm parents: 
42360diff
changeset | 179 | NONE => raise Fail "Inner syntax operations not installed" | 
| 
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
 wenzelm parents: 
42360diff
changeset | 180 | | SOME ops => which ops ctxt x); | 
| 
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
 wenzelm parents: 
42360diff
changeset | 181 | |
| 
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
 wenzelm parents: 
42360diff
changeset | 182 | |
| 42268 | 183 | (* configuration options *) | 
| 184 | ||
| 42293 
6cca0343ea48
renamed sprop "prop#" to "prop'" -- proper identifier;
 wenzelm parents: 
42290diff
changeset | 185 | val root = Config.string (Config.declare "syntax_root" (K (Config.String "any"))); | 
| 42288 
2074b31650e6
discontinued special treatment of structure Syntax_Ext (formerly Syn_Ext);
 wenzelm parents: 
42287diff
changeset | 186 | |
| 42268 | 187 | val positions_raw = Config.declare "syntax_positions" (fn _ => Config.Bool true); | 
| 188 | val positions = Config.bool positions_raw; | |
| 189 | ||
| 190 | val ambiguity_enabled = | |
| 191 | Config.bool (Config.declare "syntax_ambiguity_enabled" (fn _ => Config.Bool true)); | |
| 192 | ||
| 193 | val ambiguity_level_raw = Config.declare "syntax_ambiguity_level" (fn _ => Config.Int 1); | |
| 194 | val ambiguity_level = Config.int ambiguity_level_raw; | |
| 195 | ||
| 196 | val ambiguity_limit = | |
| 197 | Config.int (Config.declare "syntax_ambiguity_limit" (fn _ => Config.Int 10)); | |
| 198 | ||
| 44069 | 199 | val ambiguity_warnings_raw = | 
| 200 | Config.declare "syntax_ambiguity_warnings" (fn _ => Config.Bool true); | |
| 201 | val ambiguity_warnings = | |
| 202 | Config.bool ambiguity_warnings_raw; | |
| 203 | ||
| 42268 | 204 | |
| 43731 
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
 wenzelm parents: 
43558diff
changeset | 205 | (* outer syntax token -- with optional YXML content *) | 
| 39135 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 206 | |
| 43731 
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
 wenzelm parents: 
43558diff
changeset | 207 | fun explode_token tree = | 
| 39135 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 208 | let | 
| 39555 
ccb223a4d49c
added XML.content_of convenience -- cover XML.body, which is the general situation;
 wenzelm parents: 
39510diff
changeset | 209 | val text = XML.content_of [tree]; | 
| 39135 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 210 | val pos = | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 211 | (case tree of | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 212 | XML.Elem ((name, props), _) => | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 213 | if name = Markup.tokenN then Position.of_properties props | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 214 | else Position.none | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 215 | | XML.Text _ => Position.none); | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 216 | in (Symbol_Pos.explode (text, pos), pos) end; | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 217 | |
| 43731 
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
 wenzelm parents: 
43558diff
changeset | 218 | fun read_token str = explode_token (YXML.parse str handle Fail msg => error msg); | 
| 
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
 wenzelm parents: 
43558diff
changeset | 219 | |
| 
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
 wenzelm parents: 
43558diff
changeset | 220 | fun parse_token ctxt decode markup parse str = | 
| 
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
 wenzelm parents: 
43558diff
changeset | 221 | let | 
| 
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
 wenzelm parents: 
43558diff
changeset | 222 | fun parse_tree tree = | 
| 
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
 wenzelm parents: 
43558diff
changeset | 223 | let | 
| 
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
 wenzelm parents: 
43558diff
changeset | 224 | val (syms, pos) = explode_token tree; | 
| 
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
 wenzelm parents: 
43558diff
changeset | 225 | val _ = Context_Position.report ctxt pos markup; | 
| 
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
 wenzelm parents: 
43558diff
changeset | 226 | in parse (syms, pos) end; | 
| 
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
 wenzelm parents: 
43558diff
changeset | 227 | in | 
| 
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
 wenzelm parents: 
43558diff
changeset | 228 | (case YXML.parse_body str handle Fail msg => error msg of | 
| 
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
 wenzelm parents: 
43558diff
changeset | 229 | body as [tree as XML.Elem ((name, _), _)] => | 
| 
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
 wenzelm parents: 
43558diff
changeset | 230 | if name = Markup.tokenN then parse_tree tree else decode body | 
| 
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
 wenzelm parents: 
43558diff
changeset | 231 | | [tree as XML.Text _] => parse_tree tree | 
| 
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
 wenzelm parents: 
43558diff
changeset | 232 | | body => decode body) | 
| 
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
 wenzelm parents: 
43558diff
changeset | 233 | end; | 
| 
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
 wenzelm parents: 
43558diff
changeset | 234 | |
| 39135 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 235 | |
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 236 | (* (un)parsing *) | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 237 | |
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 238 | val parse_sort = operation #parse_sort; | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 239 | val parse_typ = operation #parse_typ; | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 240 | val parse_term = operation #parse_term; | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 241 | val parse_prop = operation #parse_prop; | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 242 | val unparse_sort = operation #unparse_sort; | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 243 | val unparse_typ = operation #unparse_typ; | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 244 | val unparse_term = operation #unparse_term; | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 245 | |
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 246 | |
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 247 | (* context-sensitive (un)checking *) | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 248 | |
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 249 | type key = int * bool; | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 250 | |
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 251 | structure Checks = Generic_Data | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 252 | ( | 
| 42402 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42383diff
changeset | 253 |   type 'a check = 'a list -> Proof.context -> ('a list * Proof.context) option;
 | 
| 39135 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 254 | type T = | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 255 | ((key * ((string * typ check) * stamp) list) list * | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 256 | (key * ((string * term check) * stamp) list) list); | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 257 | val empty = ([], []); | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 258 | val extend = I; | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 259 | fun merge ((typ_checks1, term_checks1), (typ_checks2, term_checks2)) : T = | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 260 | (AList.join (op =) (K (Library.merge (eq_snd (op =)))) (typ_checks1, typ_checks2), | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 261 | AList.join (op =) (K (Library.merge (eq_snd (op =)))) (term_checks1, term_checks2)); | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 262 | ); | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 263 | |
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 264 | fun print_checks ctxt = | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 265 | let | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 266 | fun split_checks checks = | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 267 | List.partition (fn ((_, un), _) => not un) checks | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 268 | |> pairself (map (fn ((i, _), fs) => (i, map (fst o fst) fs)) | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 269 | #> sort (int_ord o pairself fst)); | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 270 | fun pretty_checks kind checks = | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 271 | checks |> map (fn (i, names) => Pretty.block | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 272 | [Pretty.str (kind ^ " (stage " ^ signed_string_of_int i ^ "):"), | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 273 | Pretty.brk 1, Pretty.strs names]); | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 274 | |
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 275 | val (typs, terms) = Checks.get (Context.Proof ctxt); | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 276 | val (typ_checks, typ_unchecks) = split_checks typs; | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 277 | val (term_checks, term_unchecks) = split_checks terms; | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 278 | in | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 279 | pretty_checks "typ_checks" typ_checks @ | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 280 | pretty_checks "term_checks" term_checks @ | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 281 | pretty_checks "typ_unchecks" typ_unchecks @ | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 282 | pretty_checks "term_unchecks" term_unchecks | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 283 | end |> Pretty.chunks |> Pretty.writeln; | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 284 | |
| 42402 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42383diff
changeset | 285 | |
| 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42383diff
changeset | 286 | local | 
| 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42383diff
changeset | 287 | |
| 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42383diff
changeset | 288 | fun context_check which (key: key) name f = | 
| 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42383diff
changeset | 289 | Checks.map (which (AList.map_default op = (key, []) (cons ((name, f), stamp ())))); | 
| 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42383diff
changeset | 290 | |
| 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42383diff
changeset | 291 | fun simple_check eq f xs ctxt = | 
| 42403 
38b29c9fc742
slightly more special eq_list/eq_set, with shortcut involving pointer_eq;
 wenzelm parents: 
42402diff
changeset | 292 | let val xs' = f ctxt xs | 
| 
38b29c9fc742
slightly more special eq_list/eq_set, with shortcut involving pointer_eq;
 wenzelm parents: 
42402diff
changeset | 293 | in if eq_list eq (xs, xs') then NONE else SOME (xs', ctxt) end; | 
| 42402 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42383diff
changeset | 294 | |
| 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42383diff
changeset | 295 | in | 
| 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42383diff
changeset | 296 | |
| 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42383diff
changeset | 297 | fun context_typ_check stage = context_check apfst (stage, false); | 
| 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42383diff
changeset | 298 | fun context_term_check stage = context_check apsnd (stage, false); | 
| 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42383diff
changeset | 299 | fun context_typ_uncheck stage = context_check apfst (stage, true); | 
| 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42383diff
changeset | 300 | fun context_term_uncheck stage = context_check apsnd (stage, true); | 
| 39135 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 301 | |
| 42402 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42383diff
changeset | 302 | fun add_typ_check key name f = context_typ_check key name (simple_check (op =) f); | 
| 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42383diff
changeset | 303 | fun add_term_check key name f = context_term_check key name (simple_check (op aconv) f); | 
| 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42383diff
changeset | 304 | fun add_typ_uncheck key name f = context_typ_uncheck key name (simple_check (op =) f); | 
| 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42383diff
changeset | 305 | fun add_term_uncheck key name f = context_term_uncheck key name (simple_check (op aconv) f); | 
| 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42383diff
changeset | 306 | |
| 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42383diff
changeset | 307 | end; | 
| 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42383diff
changeset | 308 | |
| 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42383diff
changeset | 309 | |
| 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42383diff
changeset | 310 | local | 
| 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42383diff
changeset | 311 | |
| 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42383diff
changeset | 312 | fun check_stage fs = perhaps_loop (perhaps_apply (map uncurry fs)); | 
| 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42383diff
changeset | 313 | fun check_all fs = perhaps_apply (map check_stage fs); | 
| 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42383diff
changeset | 314 | |
| 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42383diff
changeset | 315 | fun check which uncheck ctxt0 xs0 = | 
| 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42383diff
changeset | 316 | let | 
| 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42383diff
changeset | 317 | val funs = which (Checks.get (Context.Proof ctxt0)) | 
| 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42383diff
changeset | 318 | |> map_filter (fn ((i, u), fs) => if uncheck = u then SOME (i, map (snd o fst) fs) else NONE) | 
| 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42383diff
changeset | 319 | |> Library.sort (int_ord o pairself fst) |> map snd | 
| 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42383diff
changeset | 320 | |> not uncheck ? map rev; | 
| 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42383diff
changeset | 321 | in #1 (perhaps (check_all funs) (xs0, ctxt0)) end; | 
| 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42383diff
changeset | 322 | |
| 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42383diff
changeset | 323 | in | 
| 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42383diff
changeset | 324 | |
| 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42383diff
changeset | 325 | val typ_check = check fst false; | 
| 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42383diff
changeset | 326 | val term_check = check snd false; | 
| 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42383diff
changeset | 327 | val typ_uncheck = check fst true; | 
| 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42383diff
changeset | 328 | val term_uncheck = check snd true; | 
| 42382 
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
 wenzelm parents: 
42360diff
changeset | 329 | |
| 
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
 wenzelm parents: 
42360diff
changeset | 330 | val check_typs = operation #check_typs; | 
| 
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
 wenzelm parents: 
42360diff
changeset | 331 | val check_terms = operation #check_terms; | 
| 
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
 wenzelm parents: 
42360diff
changeset | 332 | val check_props = operation #check_props; | 
| 39135 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 333 | |
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 334 | val check_typ = singleton o check_typs; | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 335 | val check_term = singleton o check_terms; | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 336 | val check_prop = singleton o check_props; | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 337 | |
| 42382 
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
 wenzelm parents: 
42360diff
changeset | 338 | val uncheck_typs = operation #uncheck_typs; | 
| 
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
 wenzelm parents: 
42360diff
changeset | 339 | val uncheck_terms = operation #uncheck_terms; | 
| 42402 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42383diff
changeset | 340 | |
| 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42383diff
changeset | 341 | end; | 
| 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42383diff
changeset | 342 | |
| 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42383diff
changeset | 343 | |
| 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42383diff
changeset | 344 | (* derived operations for algebra of sorts *) | 
| 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42383diff
changeset | 345 | |
| 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42383diff
changeset | 346 | local | 
| 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42383diff
changeset | 347 | |
| 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42383diff
changeset | 348 | fun map_sort f S = | 
| 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42383diff
changeset | 349 |   (case f (TFree ("", S)) of
 | 
| 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42383diff
changeset | 350 |     TFree ("", S') => S'
 | 
| 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42383diff
changeset | 351 |   | _ => raise TYPE ("map_sort", [TFree ("", S)], []));
 | 
| 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42383diff
changeset | 352 | |
| 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42383diff
changeset | 353 | in | 
| 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42383diff
changeset | 354 | |
| 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42383diff
changeset | 355 | val check_sort = map_sort o check_typ; | 
| 39135 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 356 | val uncheck_sort = map_sort o singleton o uncheck_typs; | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 357 | |
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 358 | end; | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 359 | |
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 360 | |
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 361 | val uncheck_classrel = map o singleton o uncheck_sort; | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 362 | |
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 363 | fun unparse_classrel ctxt cs = Pretty.block (flat | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 364 | (separate [Pretty.str " <", Pretty.brk 1] (map (single o unparse_sort ctxt o single) cs))); | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 365 | |
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 366 | fun uncheck_arity ctxt (a, Ss, S) = | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 367 | let | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 368 | val T = Type (a, replicate (length Ss) dummyT); | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 369 | val a' = | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 370 | (case singleton (uncheck_typs ctxt) T of | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 371 | Type (a', _) => a' | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 372 |       | T => raise TYPE ("uncheck_arity", [T], []));
 | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 373 | val Ss' = map (uncheck_sort ctxt) Ss; | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 374 | val S' = uncheck_sort ctxt S; | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 375 | in (a', Ss', S') end; | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 376 | |
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 377 | fun unparse_arity ctxt (a, Ss, S) = | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 378 | let | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 379 | val prtT = unparse_typ ctxt (Type (a, [])); | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 380 | val dom = | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 381 | if null Ss then [] | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 382 |       else [Pretty.list "(" ")" (map (unparse_sort ctxt) Ss), Pretty.brk 1];
 | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 383 | in Pretty.block ([prtT, Pretty.str " ::", Pretty.brk 1] @ dom @ [unparse_sort ctxt S]) end; | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 384 | |
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 385 | |
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 386 | (* read = parse + check *) | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 387 | |
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 388 | fun read_sort ctxt = parse_sort ctxt #> check_sort ctxt; | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 389 | fun read_typ ctxt = parse_typ ctxt #> singleton (check_typs ctxt); | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 390 | |
| 43558 | 391 | fun read_terms ctxt = Par_List.map_name "Syntax.read_terms" (parse_term ctxt) #> check_terms ctxt; | 
| 392 | fun read_props ctxt = Par_List.map_name "Syntax.read_props" (parse_prop ctxt) #> check_props ctxt; | |
| 39135 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 393 | |
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 394 | val read_term = singleton o read_terms; | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 395 | val read_prop = singleton o read_props; | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 396 | |
| 42360 | 397 | val read_sort_global = read_sort o Proof_Context.init_global; | 
| 398 | val read_typ_global = read_typ o Proof_Context.init_global; | |
| 399 | val read_term_global = read_term o Proof_Context.init_global; | |
| 400 | val read_prop_global = read_prop o Proof_Context.init_global; | |
| 39135 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 401 | |
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 402 | |
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 403 | (* pretty = uncheck + unparse *) | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 404 | |
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 405 | fun pretty_term ctxt = singleton (uncheck_terms ctxt) #> unparse_term ctxt; | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 406 | fun pretty_typ ctxt = singleton (uncheck_typs ctxt) #> unparse_typ ctxt; | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 407 | fun pretty_sort ctxt = uncheck_sort ctxt #> unparse_sort ctxt; | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 408 | fun pretty_classrel ctxt = uncheck_classrel ctxt #> unparse_classrel ctxt; | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 409 | fun pretty_arity ctxt = uncheck_arity ctxt #> unparse_arity ctxt; | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 410 | |
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 411 | val string_of_term = Pretty.string_of oo pretty_term; | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 412 | val string_of_typ = Pretty.string_of oo pretty_typ; | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 413 | val string_of_sort = Pretty.string_of oo pretty_sort; | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 414 | val string_of_classrel = Pretty.string_of oo pretty_classrel; | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 415 | val string_of_arity = Pretty.string_of oo pretty_arity; | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 416 | |
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 417 | |
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 418 | (* global pretty printing *) | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 419 | |
| 39508 
dfacdb01e1ec
simplified some internal flags using Config.T instead of full-blown Proof_Data;
 wenzelm parents: 
39507diff
changeset | 420 | val pretty_global = Config.bool (Config.declare "Syntax.pretty_global" (K (Config.Bool false))); | 
| 
dfacdb01e1ec
simplified some internal flags using Config.T instead of full-blown Proof_Data;
 wenzelm parents: 
39507diff
changeset | 421 | fun is_pretty_global ctxt = Config.get ctxt pretty_global; | 
| 
dfacdb01e1ec
simplified some internal flags using Config.T instead of full-blown Proof_Data;
 wenzelm parents: 
39507diff
changeset | 422 | val set_pretty_global = Config.put pretty_global; | 
| 42360 | 423 | val init_pretty_global = set_pretty_global true o Proof_Context.init_global; | 
| 42383 
0ae4ad40d7b5
simplified pretty printing context, which is only required for certain kernel operations;
 wenzelm parents: 
42382diff
changeset | 424 | val init_pretty = Context.pretty_context init_pretty_global; | 
| 39135 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 425 | |
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 426 | val pretty_term_global = pretty_term o init_pretty_global; | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 427 | val pretty_typ_global = pretty_typ o init_pretty_global; | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 428 | val pretty_sort_global = pretty_sort o init_pretty_global; | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 429 | |
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 430 | val string_of_term_global = string_of_term o init_pretty_global; | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 431 | val string_of_typ_global = string_of_typ o init_pretty_global; | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 432 | val string_of_sort_global = string_of_sort o init_pretty_global; | 
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 433 | |
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 434 | |
| 
6f5f64542405
structure Syntax: define "interfaces" before actual implementations;
 wenzelm parents: 
39126diff
changeset | 435 | |
| 237 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
175diff
changeset | 436 | (** tables of translation functions **) | 
| 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
175diff
changeset | 437 | |
| 5692 | 438 | (* parse (ast) translations *) | 
| 439 | ||
| 23655 
d2d1138e0ddc
replaced exception TableFun/GraphFun.DUPS by TableFun/GraphFun.DUP;
 wenzelm parents: 
23631diff
changeset | 440 | fun err_dup_trfun name c = | 
| 
d2d1138e0ddc
replaced exception TableFun/GraphFun.DUPS by TableFun/GraphFun.DUP;
 wenzelm parents: 
23631diff
changeset | 441 |   error ("More than one " ^ name ^ " for " ^ quote c);
 | 
| 237 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
175diff
changeset | 442 | |
| 42253 | 443 | fun lookup_tr tab c = Option.map fst (Symtab.lookup tab c); | 
| 444 | ||
| 42288 
2074b31650e6
discontinued special treatment of structure Syntax_Ext (formerly Syn_Ext);
 wenzelm parents: 
42287diff
changeset | 445 | fun remove_trtab trfuns = fold (Symtab.remove Syntax_Ext.eq_trfun) trfuns; | 
| 21536 
f119c730f509
extend_trtab: allow identical trfuns to be overwritten;
 wenzelm parents: 
20784diff
changeset | 446 | |
| 29004 | 447 | fun update_trtab name trfuns tab = fold Symtab.update_new trfuns (remove_trtab trfuns tab) | 
| 23655 
d2d1138e0ddc
replaced exception TableFun/GraphFun.DUPS by TableFun/GraphFun.DUP;
 wenzelm parents: 
23631diff
changeset | 448 | handle Symtab.DUP c => err_dup_trfun name c; | 
| 237 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
175diff
changeset | 449 | |
| 42288 
2074b31650e6
discontinued special treatment of structure Syntax_Ext (formerly Syn_Ext);
 wenzelm parents: 
42287diff
changeset | 450 | fun merge_trtabs name tab1 tab2 = Symtab.merge Syntax_Ext.eq_trfun (tab1, tab2) | 
| 23655 
d2d1138e0ddc
replaced exception TableFun/GraphFun.DUPS by TableFun/GraphFun.DUP;
 wenzelm parents: 
23631diff
changeset | 451 | handle Symtab.DUP c => err_dup_trfun name c; | 
| 237 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
175diff
changeset | 452 | |
| 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
175diff
changeset | 453 | |
| 5692 | 454 | (* print (ast) translations *) | 
| 455 | ||
| 42254 | 456 | fun apply_tr' tab c ctxt T args = | 
| 457 | let | |
| 458 | val fns = map fst (Symtab.lookup_list tab c); | |
| 459 | fun app_first [] = raise Match | |
| 460 | | app_first (f :: fs) = f ctxt T args handle Match => app_first fs; | |
| 461 | in app_first fns end; | |
| 462 | ||
| 463 | fun apply_ast_tr' tab c ctxt args = | |
| 464 | let | |
| 465 | val fns = map fst (Symtab.lookup_list tab c); | |
| 466 | fun app_first [] = raise Match | |
| 467 | | app_first (f :: fs) = f ctxt args handle Match => app_first fs; | |
| 468 | in app_first fns end; | |
| 42253 | 469 | |
| 42288 
2074b31650e6
discontinued special treatment of structure Syntax_Ext (formerly Syn_Ext);
 wenzelm parents: 
42287diff
changeset | 470 | fun update_tr'tab trfuns = fold_rev (Symtab.update_list Syntax_Ext.eq_trfun) trfuns; | 
| 
2074b31650e6
discontinued special treatment of structure Syntax_Ext (formerly Syn_Ext);
 wenzelm parents: 
42287diff
changeset | 471 | fun remove_tr'tab trfuns = fold (Symtab.remove_list Syntax_Ext.eq_trfun) trfuns; | 
| 
2074b31650e6
discontinued special treatment of structure Syntax_Ext (formerly Syn_Ext);
 wenzelm parents: 
42287diff
changeset | 472 | fun merge_tr'tabs tab1 tab2 = Symtab.merge_list Syntax_Ext.eq_trfun (tab1, tab2); | 
| 5692 | 473 | |
| 474 | ||
| 237 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
175diff
changeset | 475 | |
| 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
175diff
changeset | 476 | (** tables of translation rules **) | 
| 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
175diff
changeset | 477 | |
| 5692 | 478 | type ruletab = (Ast.ast * Ast.ast) list Symtab.table; | 
| 237 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
175diff
changeset | 479 | |
| 19482 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 wenzelm parents: 
19375diff
changeset | 480 | fun dest_ruletab tab = maps snd (Symtab.dest tab); | 
| 237 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
175diff
changeset | 481 | |
| 25394 
db25c98f32e1
syntax operations: turned extend'' into update'' (absorb duplicates);
 wenzelm parents: 
25387diff
changeset | 482 | val update_ruletab = fold_rev (fn r => Symtab.update_list (op =) (Ast.head_of_rule r, r)); | 
| 18931 | 483 | val remove_ruletab = fold (fn r => Symtab.remove_list (op =) (Ast.head_of_rule r, r)); | 
| 484 | fun merge_ruletabs tab1 tab2 = Symtab.merge_list (op =) (tab1, tab2); | |
| 0 | 485 | |
| 486 | ||
| 487 | ||
| 488 | (** datatype syntax **) | |
| 489 | ||
| 42820 
3daff3cc2214
future merge of grammars, to improve parallel performance;
 wenzelm parents: 
42476diff
changeset | 490 | fun future_gram deps e = | 
| 
3daff3cc2214
future merge of grammars, to improve parallel performance;
 wenzelm parents: 
42476diff
changeset | 491 | singleton | 
| 
3daff3cc2214
future merge of grammars, to improve parallel performance;
 wenzelm parents: 
42476diff
changeset | 492 |     (Future.cond_forks {name = "Syntax.gram", group = NONE,
 | 
| 44113 | 493 | deps = map Future.task_of deps, pri = 0, interrupts = true}) e; | 
| 42820 
3daff3cc2214
future merge of grammars, to improve parallel performance;
 wenzelm parents: 
42476diff
changeset | 494 | |
| 237 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
175diff
changeset | 495 | datatype syntax = | 
| 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
175diff
changeset | 496 |   Syntax of {
 | 
| 42288 
2074b31650e6
discontinued special treatment of structure Syntax_Ext (formerly Syn_Ext);
 wenzelm parents: 
42287diff
changeset | 497 | input: Syntax_Ext.xprod list, | 
| 4703 | 498 | lexicon: Scan.lexicon, | 
| 42820 
3daff3cc2214
future merge of grammars, to improve parallel performance;
 wenzelm parents: 
42476diff
changeset | 499 | gram: Parser.gram future, | 
| 42298 
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
 wenzelm parents: 
42297diff
changeset | 500 | consts: string Symtab.table, | 
| 2913 | 501 | prmodes: string list, | 
| 21772 | 502 | parse_ast_trtab: ((Proof.context -> Ast.ast list -> Ast.ast) * stamp) Symtab.table, | 
| 237 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
175diff
changeset | 503 | parse_ruletab: ruletab, | 
| 21772 | 504 | parse_trtab: ((Proof.context -> term list -> term) * stamp) Symtab.table, | 
| 42247 
12fe41a92cd5
typed_print_translation: discontinued show_sorts argument;
 wenzelm parents: 
42245diff
changeset | 505 | print_trtab: ((Proof.context -> typ -> term list -> term) * stamp) list Symtab.table, | 
| 237 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
175diff
changeset | 506 | print_ruletab: ruletab, | 
| 21772 | 507 | print_ast_trtab: ((Proof.context -> Ast.ast list -> Ast.ast) * stamp) list Symtab.table, | 
| 17079 | 508 | prtabs: Printer.prtabs} * stamp; | 
| 0 | 509 | |
| 17079 | 510 | fun eq_syntax (Syntax (_, s1), Syntax (_, s2)) = s1 = s2; | 
| 511 | ||
| 44802 
65c397cc44ec
explicit join_syntax ensures command transaction integrity of 'theory';
 wenzelm parents: 
44113diff
changeset | 512 | fun join_syntax (Syntax ({gram, ...}, _)) = ignore (Future.join gram);
 | 
| 
65c397cc44ec
explicit join_syntax ensures command transaction integrity of 'theory';
 wenzelm parents: 
44113diff
changeset | 513 | |
| 42298 
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
 wenzelm parents: 
42297diff
changeset | 514 | fun lookup_const (Syntax ({consts, ...}, _)) = Symtab.lookup consts;
 | 
| 17079 | 515 | fun is_keyword (Syntax ({lexicon, ...}, _)) = Scan.is_literal lexicon o Symbol.explode;
 | 
| 42251 | 516 | fun tokenize (Syntax ({lexicon, ...}, _)) = Lexicon.tokenize lexicon;
 | 
| 42820 
3daff3cc2214
future merge of grammars, to improve parallel performance;
 wenzelm parents: 
42476diff
changeset | 517 | fun parse ctxt (Syntax ({gram, ...}, _)) = Parser.parse ctxt (Future.join gram);
 | 
| 14687 | 518 | |
| 42253 | 519 | fun parse_ast_translation (Syntax ({parse_ast_trtab, ...}, _)) = lookup_tr parse_ast_trtab;
 | 
| 520 | fun parse_translation (Syntax ({parse_trtab, ...}, _)) = lookup_tr parse_trtab;
 | |
| 42255 | 521 | fun parse_rules (Syntax ({parse_ruletab, ...}, _)) = Symtab.lookup_list parse_ruletab;
 | 
| 42254 | 522 | fun print_translation (Syntax ({print_trtab, ...}, _)) = apply_tr' print_trtab;
 | 
| 42255 | 523 | fun print_rules (Syntax ({print_ruletab, ...}, _)) = Symtab.lookup_list print_ruletab;
 | 
| 42254 | 524 | fun print_ast_translation (Syntax ({print_ast_trtab, ...}, _)) = apply_ast_tr' print_ast_trtab;
 | 
| 42253 | 525 | |
| 42255 | 526 | fun prtabs (Syntax ({prtabs, ...}, _)) = prtabs;
 | 
| 527 | ||
| 20784 | 528 | type mode = string * bool; | 
| 24970 | 529 | val mode_default = ("", true);
 | 
| 37146 
f652333bbf8e
renamed structure PrintMode to Print_Mode, keeping the old name as legacy alias for some time;
 wenzelm parents: 
36739diff
changeset | 530 | val mode_input = (Print_Mode.input, true); | 
| 15755 
50ac97ebe7d8
expect translations functions to be stamped already;
 wenzelm parents: 
15574diff
changeset | 531 | |
| 18 | 532 | |
| 237 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
175diff
changeset | 533 | (* empty_syntax *) | 
| 18 | 534 | |
| 17079 | 535 | val empty_syntax = Syntax | 
| 536 |   ({input = [],
 | |
| 4703 | 537 | lexicon = Scan.empty_lexicon, | 
| 42820 
3daff3cc2214
future merge of grammars, to improve parallel performance;
 wenzelm parents: 
42476diff
changeset | 538 | gram = Future.value Parser.empty_gram, | 
| 42298 
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
 wenzelm parents: 
42297diff
changeset | 539 | consts = Symtab.empty, | 
| 2913 | 540 | prmodes = [], | 
| 5692 | 541 | parse_ast_trtab = Symtab.empty, | 
| 542 | parse_ruletab = Symtab.empty, | |
| 543 | parse_trtab = Symtab.empty, | |
| 544 | print_trtab = Symtab.empty, | |
| 545 | print_ruletab = Symtab.empty, | |
| 546 | print_ast_trtab = Symtab.empty, | |
| 17079 | 547 | prtabs = Printer.empty_prtabs}, stamp ()); | 
| 167 | 548 | |
| 549 | ||
| 25394 
db25c98f32e1
syntax operations: turned extend'' into update'' (absorb duplicates);
 wenzelm parents: 
25387diff
changeset | 550 | (* update_syntax *) | 
| 167 | 551 | |
| 42298 
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
 wenzelm parents: 
42297diff
changeset | 552 | fun update_const (c, b) tab = | 
| 
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
 wenzelm parents: 
42297diff
changeset | 553 | if c = "" orelse (b = "" andalso (Lexicon.is_marked c orelse Symtab.defined tab c)) | 
| 
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
 wenzelm parents: 
42297diff
changeset | 554 | then tab | 
| 
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
 wenzelm parents: 
42297diff
changeset | 555 | else Symtab.update (c, b) tab; | 
| 
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
 wenzelm parents: 
42297diff
changeset | 556 | |
| 25394 
db25c98f32e1
syntax operations: turned extend'' into update'' (absorb duplicates);
 wenzelm parents: 
25387diff
changeset | 557 | fun update_syntax (mode, inout) syn_ext (Syntax (tabs, _)) = | 
| 167 | 558 | let | 
| 42268 | 559 |     val {input, lexicon, gram, consts = consts1, prmodes, parse_ast_trtab, parse_ruletab,
 | 
| 560 | parse_trtab, print_trtab, print_ruletab, print_ast_trtab, prtabs} = tabs; | |
| 42288 
2074b31650e6
discontinued special treatment of structure Syntax_Ext (formerly Syn_Ext);
 wenzelm parents: 
42287diff
changeset | 561 |     val Syntax_Ext.Syn_Ext {xprods, consts = consts2, parse_ast_translation,
 | 
| 42268 | 562 | parse_rules, parse_translation, print_translation, print_rules, | 
| 563 | print_ast_translation} = syn_ext; | |
| 36208 
74c5e6e3c1d3
update_syntax: add new productions only once, to allow repeated local notation, for example;
 wenzelm parents: 
35668diff
changeset | 564 | val new_xprods = | 
| 
74c5e6e3c1d3
update_syntax: add new productions only once, to allow repeated local notation, for example;
 wenzelm parents: 
35668diff
changeset | 565 | if inout then distinct (op =) (filter_out (member (op =) input) xprods) else []; | 
| 19546 
00d5c7c7ce07
extend/remove_syntax: observe inout flag for translations, too;
 wenzelm parents: 
19482diff
changeset | 566 | fun if_inout xs = if inout then xs else []; | 
| 167 | 567 | in | 
| 17079 | 568 | Syntax | 
| 36208 
74c5e6e3c1d3
update_syntax: add new productions only once, to allow repeated local notation, for example;
 wenzelm parents: 
35668diff
changeset | 569 |     ({input = new_xprods @ input,
 | 
| 42288 
2074b31650e6
discontinued special treatment of structure Syntax_Ext (formerly Syn_Ext);
 wenzelm parents: 
42287diff
changeset | 570 | lexicon = fold Scan.extend_lexicon (Syntax_Ext.delims_of new_xprods) lexicon, | 
| 42820 
3daff3cc2214
future merge of grammars, to improve parallel performance;
 wenzelm parents: 
42476diff
changeset | 571 | gram = Future.value (Parser.extend_gram new_xprods (Future.join gram)), | 
| 42298 
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
 wenzelm parents: 
42297diff
changeset | 572 | consts = fold update_const consts2 consts1, | 
| 42268 | 573 | prmodes = insert (op =) mode prmodes, | 
| 167 | 574 | parse_ast_trtab = | 
| 25394 
db25c98f32e1
syntax operations: turned extend'' into update'' (absorb duplicates);
 wenzelm parents: 
25387diff
changeset | 575 | update_trtab "parse ast translation" (if_inout parse_ast_translation) parse_ast_trtab, | 
| 
db25c98f32e1
syntax operations: turned extend'' into update'' (absorb duplicates);
 wenzelm parents: 
25387diff
changeset | 576 | parse_ruletab = update_ruletab (if_inout parse_rules) parse_ruletab, | 
| 
db25c98f32e1
syntax operations: turned extend'' into update'' (absorb duplicates);
 wenzelm parents: 
25387diff
changeset | 577 | parse_trtab = update_trtab "parse translation" (if_inout parse_translation) parse_trtab, | 
| 
db25c98f32e1
syntax operations: turned extend'' into update'' (absorb duplicates);
 wenzelm parents: 
25387diff
changeset | 578 | print_trtab = update_tr'tab print_translation print_trtab, | 
| 
db25c98f32e1
syntax operations: turned extend'' into update'' (absorb duplicates);
 wenzelm parents: 
25387diff
changeset | 579 | print_ruletab = update_ruletab print_rules print_ruletab, | 
| 
db25c98f32e1
syntax operations: turned extend'' into update'' (absorb duplicates);
 wenzelm parents: 
25387diff
changeset | 580 | print_ast_trtab = update_tr'tab print_ast_translation print_ast_trtab, | 
| 
db25c98f32e1
syntax operations: turned extend'' into update'' (absorb duplicates);
 wenzelm parents: 
25387diff
changeset | 581 | prtabs = Printer.update_prtabs mode xprods prtabs}, stamp ()) | 
| 18 | 582 | end; | 
| 583 | ||
| 15755 
50ac97ebe7d8
expect translations functions to be stamped already;
 wenzelm parents: 
15574diff
changeset | 584 | |
| 
50ac97ebe7d8
expect translations functions to be stamped already;
 wenzelm parents: 
15574diff
changeset | 585 | (* remove_syntax *) | 
| 
50ac97ebe7d8
expect translations functions to be stamped already;
 wenzelm parents: 
15574diff
changeset | 586 | |
| 17079 | 587 | fun remove_syntax (mode, inout) syn_ext (Syntax (tabs, _)) = | 
| 15755 
50ac97ebe7d8
expect translations functions to be stamped already;
 wenzelm parents: 
15574diff
changeset | 588 | let | 
| 42288 
2074b31650e6
discontinued special treatment of structure Syntax_Ext (formerly Syn_Ext);
 wenzelm parents: 
42287diff
changeset | 589 |     val Syntax_Ext.Syn_Ext {xprods, consts = _, parse_ast_translation, parse_rules,
 | 
| 42268 | 590 | parse_translation, print_translation, print_rules, print_ast_translation} = syn_ext; | 
| 591 |     val {input, lexicon, gram, consts, prmodes, parse_ast_trtab, parse_ruletab,
 | |
| 592 | parse_trtab, print_trtab, print_ruletab, print_ast_trtab, prtabs} = tabs; | |
| 19300 | 593 | val input' = if inout then subtract (op =) xprods input else input; | 
| 25394 
db25c98f32e1
syntax operations: turned extend'' into update'' (absorb duplicates);
 wenzelm parents: 
25387diff
changeset | 594 | val changed = length input <> length input'; | 
| 19546 
00d5c7c7ce07
extend/remove_syntax: observe inout flag for translations, too;
 wenzelm parents: 
19482diff
changeset | 595 | fun if_inout xs = if inout then xs else []; | 
| 15755 
50ac97ebe7d8
expect translations functions to be stamped already;
 wenzelm parents: 
15574diff
changeset | 596 | in | 
| 17079 | 597 | Syntax | 
| 598 |     ({input = input',
 | |
| 42288 
2074b31650e6
discontinued special treatment of structure Syntax_Ext (formerly Syn_Ext);
 wenzelm parents: 
42287diff
changeset | 599 | lexicon = if changed then Scan.make_lexicon (Syntax_Ext.delims_of input') else lexicon, | 
| 42820 
3daff3cc2214
future merge of grammars, to improve parallel performance;
 wenzelm parents: 
42476diff
changeset | 600 | gram = if changed then Future.value (Parser.extend_gram input' Parser.empty_gram) else gram, | 
| 15755 
50ac97ebe7d8
expect translations functions to be stamped already;
 wenzelm parents: 
15574diff
changeset | 601 | consts = consts, | 
| 
50ac97ebe7d8
expect translations functions to be stamped already;
 wenzelm parents: 
15574diff
changeset | 602 | prmodes = prmodes, | 
| 19546 
00d5c7c7ce07
extend/remove_syntax: observe inout flag for translations, too;
 wenzelm parents: 
19482diff
changeset | 603 | parse_ast_trtab = remove_trtab (if_inout parse_ast_translation) parse_ast_trtab, | 
| 
00d5c7c7ce07
extend/remove_syntax: observe inout flag for translations, too;
 wenzelm parents: 
19482diff
changeset | 604 | parse_ruletab = remove_ruletab (if_inout parse_rules) parse_ruletab, | 
| 
00d5c7c7ce07
extend/remove_syntax: observe inout flag for translations, too;
 wenzelm parents: 
19482diff
changeset | 605 | parse_trtab = remove_trtab (if_inout parse_translation) parse_trtab, | 
| 15755 
50ac97ebe7d8
expect translations functions to be stamped already;
 wenzelm parents: 
15574diff
changeset | 606 | print_trtab = remove_tr'tab print_translation print_trtab, | 
| 
50ac97ebe7d8
expect translations functions to be stamped already;
 wenzelm parents: 
15574diff
changeset | 607 | print_ruletab = remove_ruletab print_rules print_ruletab, | 
| 
50ac97ebe7d8
expect translations functions to be stamped already;
 wenzelm parents: 
15574diff
changeset | 608 | print_ast_trtab = remove_tr'tab print_ast_translation print_ast_trtab, | 
| 17079 | 609 | prtabs = Printer.remove_prtabs mode xprods prtabs}, stamp ()) | 
| 15755 
50ac97ebe7d8
expect translations functions to be stamped already;
 wenzelm parents: 
15574diff
changeset | 610 | end; | 
| 14904 
7d8dc92fcb7f
removed separate logtypes field of syntax; removed test_read, simple_str_of_sort, simple_string_of_typ; provide default_mode;
 wenzelm parents: 
14798diff
changeset | 611 | |
| 18 | 612 | |
| 237 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
175diff
changeset | 613 | (* merge_syntaxes *) | 
| 0 | 614 | |
| 17079 | 615 | fun merge_syntaxes (Syntax (tabs1, _)) (Syntax (tabs2, _)) = | 
| 0 | 616 | let | 
| 15755 
50ac97ebe7d8
expect translations functions to be stamped already;
 wenzelm parents: 
15574diff
changeset | 617 |     val {input = input1, lexicon = lexicon1, gram = gram1, consts = consts1,
 | 
| 42268 | 618 | prmodes = prmodes1, parse_ast_trtab = parse_ast_trtab1, parse_ruletab = parse_ruletab1, | 
| 619 | parse_trtab = parse_trtab1, print_trtab = print_trtab1, print_ruletab = print_ruletab1, | |
| 620 | print_ast_trtab = print_ast_trtab1, prtabs = prtabs1} = tabs1; | |
| 237 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
175diff
changeset | 621 | |
| 15755 
50ac97ebe7d8
expect translations functions to be stamped already;
 wenzelm parents: 
15574diff
changeset | 622 |     val {input = input2, lexicon = lexicon2, gram = gram2, consts = consts2,
 | 
| 42268 | 623 | prmodes = prmodes2, parse_ast_trtab = parse_ast_trtab2, parse_ruletab = parse_ruletab2, | 
| 624 | parse_trtab = parse_trtab2, print_trtab = print_trtab2, print_ruletab = print_ruletab2, | |
| 625 | print_ast_trtab = print_ast_trtab2, prtabs = prtabs2} = tabs2; | |
| 0 | 626 | in | 
| 17079 | 627 | Syntax | 
| 18921 | 628 |     ({input = Library.merge (op =) (input1, input2),
 | 
| 27768 | 629 | lexicon = Scan.merge_lexicons (lexicon1, lexicon2), | 
| 42820 
3daff3cc2214
future merge of grammars, to improve parallel performance;
 wenzelm parents: 
42476diff
changeset | 630 | gram = future_gram [gram1, gram2] (fn () => | 
| 
3daff3cc2214
future merge of grammars, to improve parallel performance;
 wenzelm parents: 
42476diff
changeset | 631 | Parser.merge_gram (Future.join gram1, Future.join gram2)), | 
| 42298 
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
 wenzelm parents: 
42297diff
changeset | 632 | consts = Symtab.merge (K true) (consts1, consts2), | 
| 18921 | 633 | prmodes = Library.merge (op =) (prmodes1, prmodes2), | 
| 237 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
175diff
changeset | 634 | parse_ast_trtab = | 
| 15755 
50ac97ebe7d8
expect translations functions to be stamped already;
 wenzelm parents: 
15574diff
changeset | 635 | merge_trtabs "parse ast translation" parse_ast_trtab1 parse_ast_trtab2, | 
| 237 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
175diff
changeset | 636 | parse_ruletab = merge_ruletabs parse_ruletab1 parse_ruletab2, | 
| 15755 
50ac97ebe7d8
expect translations functions to be stamped already;
 wenzelm parents: 
15574diff
changeset | 637 | parse_trtab = merge_trtabs "parse translation" parse_trtab1 parse_trtab2, | 
| 5692 | 638 | print_trtab = merge_tr'tabs print_trtab1 print_trtab2, | 
| 237 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
175diff
changeset | 639 | print_ruletab = merge_ruletabs print_ruletab1 print_ruletab2, | 
| 5692 | 640 | print_ast_trtab = merge_tr'tabs print_ast_trtab1 print_ast_trtab2, | 
| 17079 | 641 | prtabs = Printer.merge_prtabs prtabs1 prtabs2}, stamp ()) | 
| 0 | 642 | end; | 
| 643 | ||
| 644 | ||
| 18720 | 645 | (* basic syntax *) | 
| 260 | 646 | |
| 42294 | 647 | val token_markers = | 
| 42408 | 648 | ["_tfree", "_tvar", "_free", "_bound", "_loose", "_var", "_numeral", "_inner_string"]; | 
| 260 | 649 | |
| 18720 | 650 | val basic_nonterms = | 
| 42293 
6cca0343ea48
renamed sprop "prop#" to "prop'" -- proper identifier;
 wenzelm parents: 
42290diff
changeset | 651 | (Lexicon.terminals @ ["logic", "type", "types", "sort", "classes", | 
| 
6cca0343ea48
renamed sprop "prop#" to "prop'" -- proper identifier;
 wenzelm parents: 
42290diff
changeset | 652 | "args", "cargs", "pttrn", "pttrns", "idt", "idts", "aprop", "asms", | 
| 
6cca0343ea48
renamed sprop "prop#" to "prop'" -- proper identifier;
 wenzelm parents: 
42290diff
changeset | 653 | "any", "prop'", "num_const", "float_const", "index", "struct", | 
| 
6cca0343ea48
renamed sprop "prop#" to "prop'" -- proper identifier;
 wenzelm parents: 
42290diff
changeset | 654 | "id_position", "longid_position", "type_name", "class_name"]); | 
| 18720 | 655 | |
| 0 | 656 | |
| 4887 | 657 | |
| 15759 | 658 | (** print syntax **) | 
| 659 | ||
| 660 | local | |
| 0 | 661 | |
| 260 | 662 | fun pretty_strs_qs name strs = | 
| 28840 | 663 | Pretty.strs (name :: map quote (sort_strings strs)); | 
| 0 | 664 | |
| 17079 | 665 | fun pretty_gram (Syntax (tabs, _)) = | 
| 0 | 666 | let | 
| 32784 | 667 |     val {lexicon, prmodes, gram, ...} = tabs;
 | 
| 28375 | 668 | val prmodes' = sort_strings (filter_out (fn s => s = "") prmodes); | 
| 0 | 669 | in | 
| 8720 | 670 | [pretty_strs_qs "lexicon:" (Scan.dest_lexicon lexicon), | 
| 42820 
3daff3cc2214
future merge of grammars, to improve parallel performance;
 wenzelm parents: 
42476diff
changeset | 671 | Pretty.big_list "prods:" (Parser.pretty_gram (Future.join gram)), | 
| 8720 | 672 | pretty_strs_qs "print modes:" prmodes'] | 
| 0 | 673 | end; | 
| 674 | ||
| 17079 | 675 | fun pretty_trans (Syntax (tabs, _)) = | 
| 0 | 676 | let | 
| 42298 
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
 wenzelm parents: 
42297diff
changeset | 677 | fun pretty_tab name tab = | 
| 
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
 wenzelm parents: 
42297diff
changeset | 678 | pretty_strs_qs name (sort_strings (Symtab.keys tab)); | 
| 0 | 679 | |
| 260 | 680 | fun pretty_ruletab name tab = | 
| 5692 | 681 | Pretty.big_list name (map Ast.pretty_rule (dest_ruletab tab)); | 
| 0 | 682 | |
| 237 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
175diff
changeset | 683 |     val {consts, parse_ast_trtab, parse_ruletab, parse_trtab, print_trtab,
 | 
| 42268 | 684 | print_ruletab, print_ast_trtab, ...} = tabs; | 
| 0 | 685 | in | 
| 42298 
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
 wenzelm parents: 
42297diff
changeset | 686 | [pretty_tab "consts:" consts, | 
| 
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
 wenzelm parents: 
42297diff
changeset | 687 | pretty_tab "parse_ast_translation:" parse_ast_trtab, | 
| 
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
 wenzelm parents: 
42297diff
changeset | 688 | pretty_ruletab "parse_rules:" parse_ruletab, | 
| 
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
 wenzelm parents: 
42297diff
changeset | 689 | pretty_tab "parse_translation:" parse_trtab, | 
| 
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
 wenzelm parents: 
42297diff
changeset | 690 | pretty_tab "print_translation:" print_trtab, | 
| 
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
 wenzelm parents: 
42297diff
changeset | 691 | pretty_ruletab "print_rules:" print_ruletab, | 
| 
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
 wenzelm parents: 
42297diff
changeset | 692 | pretty_tab "print_ast_translation:" print_ast_trtab] | 
| 0 | 693 | end; | 
| 694 | ||
| 15759 | 695 | in | 
| 0 | 696 | |
| 15759 | 697 | fun print_gram syn = Pretty.writeln (Pretty.chunks (pretty_gram syn)); | 
| 698 | fun print_trans syn = Pretty.writeln (Pretty.chunks (pretty_trans syn)); | |
| 699 | fun print_syntax syn = Pretty.writeln (Pretty.chunks (pretty_gram syn @ pretty_trans syn)); | |
| 0 | 700 | |
| 15759 | 701 | end; | 
| 0 | 702 | |
| 703 | ||
| 26951 
030e4a818b39
moved global pretty/string_of functions from Sign to Syntax;
 wenzelm parents: 
26704diff
changeset | 704 | (* reconstructing infixes -- educated guessing *) | 
| 
030e4a818b39
moved global pretty/string_of functions from Sign to Syntax;
 wenzelm parents: 
26704diff
changeset | 705 | |
| 
030e4a818b39
moved global pretty/string_of functions from Sign to Syntax;
 wenzelm parents: 
26704diff
changeset | 706 | fun guess_infix (Syntax ({gram, ...}, _)) c =
 | 
| 42820 
3daff3cc2214
future merge of grammars, to improve parallel performance;
 wenzelm parents: 
42476diff
changeset | 707 | (case Parser.guess_infix_lr (Future.join gram) c of | 
| 26951 
030e4a818b39
moved global pretty/string_of functions from Sign to Syntax;
 wenzelm parents: 
26704diff
changeset | 708 | SOME (s, l, r, j) => SOME | 
| 35130 | 709 | (if l then Mixfix.Infixl (s, j) | 
| 710 | else if r then Mixfix.Infixr (s, j) | |
| 711 | else Mixfix.Infix (s, j)) | |
| 26951 
030e4a818b39
moved global pretty/string_of functions from Sign to Syntax;
 wenzelm parents: 
26704diff
changeset | 712 | | NONE => NONE); | 
| 
030e4a818b39
moved global pretty/string_of functions from Sign to Syntax;
 wenzelm parents: 
26704diff
changeset | 713 | |
| 
030e4a818b39
moved global pretty/string_of functions from Sign to Syntax;
 wenzelm parents: 
26704diff
changeset | 714 | |
| 0 | 715 | |
| 1158 | 716 | (** prepare translation rules **) | 
| 717 | ||
| 42204 | 718 | (* rules *) | 
| 719 | ||
| 1158 | 720 | datatype 'a trrule = | 
| 42204 | 721 | Parse_Rule of 'a * 'a | | 
| 722 | Print_Rule of 'a * 'a | | |
| 723 | Parse_Print_Rule of 'a * 'a; | |
| 888 
3a1de9454d13
improved read_xrules: patterns no longer read twice;
 wenzelm parents: 
882diff
changeset | 724 | |
| 42204 | 725 | fun map_trrule f (Parse_Rule (x, y)) = Parse_Rule (f x, f y) | 
| 726 | | map_trrule f (Print_Rule (x, y)) = Print_Rule (f x, f y) | |
| 727 | | map_trrule f (Parse_Print_Rule (x, y)) = Parse_Print_Rule (f x, f y); | |
| 1158 | 728 | |
| 42204 | 729 | fun parse_rule (Parse_Rule pats) = SOME pats | 
| 730 | | parse_rule (Print_Rule _) = NONE | |
| 731 | | parse_rule (Parse_Print_Rule pats) = SOME pats; | |
| 1158 | 732 | |
| 42204 | 733 | fun print_rule (Parse_Rule _) = NONE | 
| 734 | | print_rule (Print_Rule pats) = SOME (swap pats) | |
| 735 | | print_rule (Parse_Print_Rule pats) = SOME (swap pats); | |
| 1158 | 736 | |
| 737 | ||
| 42204 | 738 | (* check_rules *) | 
| 739 | ||
| 19262 | 740 | local | 
| 741 | ||
| 42048 
afd11ca8e018
support for encoded positions (for id_position, longid_position) as pseudo type-constraints -- still inactive;
 wenzelm parents: 
42044diff
changeset | 742 | fun check_rule rule = | 
| 5692 | 743 | (case Ast.rule_error rule of | 
| 15531 | 744 | SOME msg => | 
| 1158 | 745 |       error ("Error in syntax translation rule: " ^ msg ^ "\n" ^
 | 
| 42048 
afd11ca8e018
support for encoded positions (for id_position, longid_position) as pseudo type-constraints -- still inactive;
 wenzelm parents: 
42044diff
changeset | 746 | Pretty.string_of (Ast.pretty_rule rule)) | 
| 15531 | 747 | | NONE => rule); | 
| 888 
3a1de9454d13
improved read_xrules: patterns no longer read twice;
 wenzelm parents: 
882diff
changeset | 748 | |
| 42204 | 749 | in | 
| 750 | ||
| 751 | fun check_rules rules = | |
| 752 | (map check_rule (map_filter parse_rule rules), | |
| 753 | map check_rule (map_filter print_rule rules)); | |
| 754 | ||
| 755 | end; | |
| 756 | ||
| 757 | ||
| 18 | 758 | |
| 19262 | 759 | (** modify syntax **) | 
| 383 
fcea89074e4c
added incremental extension functions: extend_log_types, extend_type_gram,
 wenzelm parents: 
330diff
changeset | 760 | |
| 25394 
db25c98f32e1
syntax operations: turned extend'' into update'' (absorb duplicates);
 wenzelm parents: 
25387diff
changeset | 761 | fun ext_syntax f decls = update_syntax mode_default (f decls); | 
| 383 
fcea89074e4c
added incremental extension functions: extend_log_types, extend_type_gram,
 wenzelm parents: 
330diff
changeset | 762 | |
| 42288 
2074b31650e6
discontinued special treatment of structure Syntax_Ext (formerly Syn_Ext);
 wenzelm parents: 
42287diff
changeset | 763 | val update_trfuns = ext_syntax Syntax_Ext.syn_ext_trfuns; | 
| 
2074b31650e6
discontinued special treatment of structure Syntax_Ext (formerly Syn_Ext);
 wenzelm parents: 
42287diff
changeset | 764 | val update_advanced_trfuns = ext_syntax Syntax_Ext.syn_ext_advanced_trfuns; | 
| 5692 | 765 | |
| 35412 
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
 wenzelm parents: 
35394diff
changeset | 766 | fun update_type_gram add prmode decls = | 
| 
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
 wenzelm parents: 
35394diff
changeset | 767 | (if add then update_syntax else remove_syntax) prmode (Mixfix.syn_ext_types decls); | 
| 25387 | 768 | |
| 35412 
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
 wenzelm parents: 
35394diff
changeset | 769 | fun update_const_gram add is_logtype prmode decls = | 
| 
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
 wenzelm parents: 
35394diff
changeset | 770 | (if add then update_syntax else remove_syntax) prmode (Mixfix.syn_ext_consts is_logtype decls); | 
| 15755 
50ac97ebe7d8
expect translations functions to be stamped already;
 wenzelm parents: 
15574diff
changeset | 771 | |
| 42288 
2074b31650e6
discontinued special treatment of structure Syntax_Ext (formerly Syn_Ext);
 wenzelm parents: 
42287diff
changeset | 772 | val update_trrules = ext_syntax Syntax_Ext.syn_ext_rules o check_rules; | 
| 
2074b31650e6
discontinued special treatment of structure Syntax_Ext (formerly Syn_Ext);
 wenzelm parents: 
42287diff
changeset | 773 | val remove_trrules = remove_syntax mode_default o Syntax_Ext.syn_ext_rules o check_rules; | 
| 5692 | 774 | |
| 42476 | 775 | |
| 776 | open Lexicon.Syntax; | |
| 777 | ||
| 0 | 778 | end; | 
| 5692 | 779 |